diff --git a/src/HOL/Computational_Algebra/Polynomial.thy b/src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy +++ b/src/HOL/Computational_Algebra/Polynomial.thy @@ -1,4717 +1,4717 @@ (* Title: HOL/Computational_Algebra/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial imports Complex_Main "HOL-Library.More_List" "HOL-Library.Infinite_Set" Factorial_Ring begin subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) where "x ## xs = (if xs = [] \ x = 0 then [] else x # xs)" lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" by (simp add: cCons_def) lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys" by (simp add: cCons_def) lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys" by (simp add: cCons_def) lemma cCons_not_0_eq [simp]: "x \ 0 \ x ## xs = x # xs" by (simp add: cCons_def) lemma strip_while_not_0_Cons_eq [simp]: "strip_while (\x. x = 0) (x # xs) = x ## strip_while (\x. x = 0) xs" proof (cases "x = 0") case False then show ?thesis by simp next case True show ?thesis proof (induct xs rule: rev_induct) case Nil with True show ?case by simp next case (snoc y ys) then show ?case by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) qed qed lemma tl_cCons [simp]: "tl (x ## xs) = xs" by (simp add: cCons_def) subsection \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" morphisms coeff Abs_poly by (auto intro!: ALL_MOST) setup_lifting type_definition_poly lemma poly_eq_iff: "p = q \ (\n. coeff p n = coeff q n)" by (simp add: coeff_inject [symmetric] fun_eq_iff) lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" by (simp add: poly_eq_iff) lemma MOST_coeff_eq_0: "\\<^sub>\ n. coeff p n = 0" using coeff [of p] by simp subsection \Degree of a polynomial\ definition degree :: "'a::zero poly \ nat" where "degree p = (LEAST n. \i>n. coeff p i = 0)" lemma coeff_eq_0: assumes "degree p < n" shows "coeff p n = 0" proof - have "\n. \i>n. coeff p i = 0" using MOST_coeff_eq_0 by (simp add: MOST_nat) then have "\i>degree p. coeff p i = 0" unfolding degree_def by (rule LeastI_ex) with assms show ?thesis by simp qed lemma le_degree: "coeff p n \ 0 \ n \ degree p" by (erule contrapos_np, rule coeff_eq_0, simp) lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" unfolding degree_def by (erule Least_le) lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" unfolding degree_def by (drule not_less_Least, simp) subsection \The zero polynomial\ instantiation poly :: (zero) zero begin lift_definition zero_poly :: "'a poly" is "\_. 0" by (rule MOST_I) simp instance .. end lemma coeff_0 [simp]: "coeff 0 n = 0" by transfer rule lemma degree_0 [simp]: "degree 0 = 0" by (rule order_antisym [OF degree_le le0]) simp lemma leading_coeff_neq_0: assumes "p \ 0" shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from \p \ 0\ obtain n where "coeff p n \ 0" by (auto simp add: poly_eq_iff) then have "n \ degree p" by (rule le_degree) with \coeff p n \ 0\ and \degree p = 0\ show "coeff p (degree p) \ 0" by simp next case (Suc n) from \degree p = Suc n\ have "n < degree p" by simp then have "\i>n. coeff p i \ 0" by (rule less_degree_imp) then obtain i where "n < i" and "coeff p i \ 0" by blast from \degree p = Suc n\ and \n < i\ have "degree p \ i" by simp also from \coeff p i \ 0\ have "i \ degree p" by (rule le_degree) finally have "degree p = i" . with \coeff p i \ 0\ show "coeff p (degree p) \ 0" by simp qed lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0") (simp_all add: leading_coeff_neq_0) lemma eq_zero_or_degree_less: assumes "degree p \ n" and "coeff p n = 0" shows "p = 0 \ degree p < n" proof (cases n) case 0 with \degree p \ n\ and \coeff p n = 0\ have "coeff p (degree p) = 0" by simp then have "p = 0" by simp then show ?thesis .. next case (Suc m) from \degree p \ n\ have "\i>n. coeff p i = 0" by (simp add: coeff_eq_0) with \coeff p n = 0\ have "\i\n. coeff p i = 0" by (simp add: le_less) with \n = Suc m\ have "\i>m. coeff p i = 0" by (simp add: less_eq_Suc_le) then have "degree p \ m" by (rule degree_le) with \n = Suc m\ have "degree p < n" by (simp add: less_Suc_eq_le) then show ?thesis .. qed lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" using eq_zero_or_degree_less by fastforce subsection \List-style constructor for polynomials\ lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" is "\a p. case_nat a (coeff p)" by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) lemmas coeff_pCons = pCons.rep_eq lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" by transfer simp lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" by (simp add: coeff_pCons) lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) lemma degree_pCons_eq: "p \ 0 \ degree (pCons a p) = Suc (degree p)" apply (rule order_antisym [OF degree_pCons_le]) apply (rule le_degree, simp) done lemma degree_pCons_0: "degree (pCons a 0) = 0" apply (rule order_antisym [OF _ le0]) apply (rule degree_le, simp add: coeff_pCons split: nat.split) done lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" apply (cases "p = 0", simp_all) apply (rule order_antisym [OF _ le0]) apply (rule degree_le, simp add: coeff_pCons split: nat.split) apply (rule order_antisym [OF degree_pCons_le]) apply (rule le_degree, simp) done lemma pCons_0_0 [simp]: "pCons 0 0 = 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \ a = b \ p = q" proof safe assume "pCons a p = pCons b q" then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp then show "a = b" by simp next assume "pCons a p = pCons b q" then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n by simp then show "p = q" by (simp add: poly_eq_iff) qed lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" using pCons_eq_iff [of a p 0 0] by simp lemma pCons_cases [cases type: poly]: obtains (pCons) a q where "p = pCons a q" proof show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" by transfer (simp_all add: MOST_inj[where f=Suc and P="\n. p n = 0" for p] fun_eq_iff Abs_poly_inverse split: nat.split) qed lemma pCons_induct [case_names 0 pCons, induct type: poly]: assumes zero: "P 0" assumes pCons: "\a p. a \ 0 \ p \ 0 \ P p \ P (pCons a p)" shows "P p" proof (induct p rule: measure_induct_rule [where f=degree]) case (less p) obtain a q where "p = pCons a q" by (rule pCons_cases) have "P q" proof (cases "q = 0") case True then show "P q" by (simp add: zero) next case False then have "degree (pCons a q) = Suc (degree q)" by (rule degree_pCons_eq) with \p = pCons a q\ have "degree q < degree p" by simp then show "P q" by (rule less.hyps) qed have "P (pCons a q)" proof (cases "a \ 0 \ q \ 0") case True with \P q\ show ?thesis by (auto intro: pCons) next case False with zero show ?thesis by simp qed with \p = pCons a q\ show ?case by simp qed lemma degree_eq_zeroE: fixes p :: "'a::zero poly" assumes "degree p = 0" obtains a where "p = pCons a 0" proof - obtain a q where p: "p = pCons a q" by (cases p) with assms have "q = 0" by (cases "q = 0") simp_all with p have "p = pCons a 0" by simp then show thesis .. qed subsection \Quickcheck generator for polynomials\ quickcheck_generator poly constructors: "0 :: _ poly", pCons subsection \List-style syntax for polynomials\ syntax "_poly" :: "args \ 'a poly" ("[:(_):]") translations "[:x, xs:]" \ "CONST pCons x [:xs:]" "[:x:]" \ "CONST pCons x 0" "[:x:]" \ "CONST pCons x (_constrain 0 t)" subsection \Representation of polynomials by lists of coefficients\ primrec Poly :: "'a::zero list \ 'a poly" where [code_post]: "Poly [] = 0" | [code_post]: "Poly (a # as) = pCons a (Poly as)" lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0" by (induct n) simp_all lemma Poly_eq_0: "Poly as = 0 \ (\n. as = replicate n 0)" by (induct as) (auto simp add: Cons_replicate_eq) lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as" by (induct as) simp_all lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as" using Poly_append_replicate_zero [of as 1] by simp lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)" by (simp add: cCons_def) lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \ Poly (rev (tl as)) = Poly (rev as)" by (cases as) simp_all lemma degree_Poly: "degree (Poly xs) \ length xs" by (induct xs) simp_all lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs" by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits) definition coeffs :: "'a poly \ 'a::zero list" where "coeffs p = (if p = 0 then [] else map (\i. coeff p i) [0 ..< Suc (degree p)])" lemma coeffs_eq_Nil [simp]: "coeffs p = [] \ p = 0" by (simp add: coeffs_def) lemma not_0_coeffs_not_Nil: "p \ 0 \ coeffs p \ []" by simp lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []" by simp lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p" proof - have *: "\m\set ms. m > 0 \ map (case_nat x f) ms = map f (map (\n. n - 1) ms)" for ms :: "nat list" and f :: "nat \ 'a" and x :: "'a" by (induct ms) (auto split: nat.split) show ?thesis by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) qed lemma length_coeffs: "p \ 0 \ length (coeffs p) = degree p + 1" by (simp add: coeffs_def) lemma coeffs_nth: "p \ 0 \ n \ degree p \ coeffs p ! n = coeff p n" by (auto simp: coeffs_def simp del: upt_Suc) lemma coeff_in_coeffs: "p \ 0 \ n \ degree p \ coeff p n \ set (coeffs p)" using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) lemma not_0_cCons_eq [simp]: "p \ 0 \ a ## coeffs p = a # coeffs p" by (simp add: cCons_def) lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p" by (induct p) auto lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as" proof (induct as) case Nil then show ?case by simp next case (Cons a as) from replicate_length_same [of as 0] have "(\n. as \ replicate n 0) \ (\a\set as. a \ 0)" by (auto dest: sym [of _ as]) with Cons show ?case by auto qed lemma no_trailing_coeffs [simp]: "no_trailing (HOL.eq 0) (coeffs p)" by (induct p) auto lemma strip_while_coeffs [simp]: "strip_while (HOL.eq 0) (coeffs p) = coeffs p" by simp lemma coeffs_eq_iff: "p = q \ coeffs p = coeffs q" (is "?P \ ?Q") proof assume ?P then show ?Q by simp next assume ?Q then have "Poly (coeffs p) = Poly (coeffs q)" by simp then show ?P by simp qed lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p" by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) lemma [code]: "coeff p = nth_default 0 (coeffs p)" by (simp add: nth_default_coeffs_eq) lemma coeffs_eqI: assumes coeff: "\n. coeff p n = nth_default 0 xs n" assumes zero: "no_trailing (HOL.eq 0) xs" shows "coeffs p = xs" proof - from coeff have "p = Poly xs" by (simp add: poly_eq_iff) with zero show ?thesis by simp qed lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1" by (simp add: coeffs_def) lemma length_coeffs_degree: "p \ 0 \ length (coeffs p) = Suc (degree p)" by (induct p) (auto simp: cCons_def) lemma [code abstract]: "coeffs 0 = []" by (fact coeffs_0_eq_Nil) lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p" by (fact coeffs_pCons_eq_cCons) lemma set_coeffs_subset_singleton_0_iff [simp]: "set (coeffs p) \ {0} \ p = 0" by (auto simp add: coeffs_def intro: classical) lemma set_coeffs_not_only_0 [simp]: "set (coeffs p) \ {0}" by (auto simp add: set_eq_subset) lemma forall_coeffs_conv: "(\n. P (coeff p n)) \ (\c \ set (coeffs p). P c)" if "P 0" using that by (auto simp add: coeffs_def) (metis atLeastLessThan_iff coeff_eq_0 not_less_iff_gr_or_eq zero_le) instantiation poly :: ("{zero, equal}") equal begin definition [code]: "HOL.equal (p::'a poly) q \ HOL.equal (coeffs p) (coeffs q)" instance by standard (simp add: equal equal_poly_def coeffs_eq_iff) end lemma [code nbe]: "HOL.equal (p :: _ poly) p \ True" by (fact equal_refl) definition is_zero :: "'a::zero poly \ bool" where [code]: "is_zero p \ List.null (coeffs p)" lemma is_zero_null [code_abbrev]: "is_zero p \ p = 0" by (simp add: is_zero_def null_def) subsubsection \Reconstructing the polynomial from the list\ \ \contributed by Sebastiaan J.C. Joosten and René Thiemann\ definition poly_of_list :: "'a::comm_monoid_add list \ 'a poly" where [simp]: "poly_of_list = Poly" lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as" by simp subsection \Fold combinator for polynomials\ definition fold_coeffs :: "('a::zero \ 'b \ 'b) \ 'a poly \ 'b \ 'b" where "fold_coeffs f p = foldr f (coeffs p)" lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def cCons_def fun_eq_iff) lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: "a \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_not_0_0_eq [simp]: "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: "'a::comm_semiring_0 poly \ 'a \ 'a" where "poly p = fold_coeffs (\a f x. a + x * f x) p (\x. 0)" \ \The Horner Schema\ lemma poly_0 [simp]: "poly 0 x = 0" by (simp add: poly_def) lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" by (cases "p = 0 \ a = 0") (auto simp add: poly_def) lemma poly_altdef: "poly p x = (\i\degree p. coeff p i * x ^ i)" for x :: "'a::{comm_semiring_0,semiring_1}" proof (induction p rule: pCons_induct) case 0 then show ?case by simp next case (pCons a p) show ?case proof (cases "p = 0") case True then show ?thesis by simp next case False let ?p' = "pCons a p" note poly_pCons[of a p x] also note pCons.IH also have "a + x * (\i\degree p. coeff p i * x ^ i) = coeff ?p' 0 * x^0 + (\i\degree p. coeff ?p' (Suc i) * x^Suc i)" by (simp add: field_simps sum_distrib_left coeff_pCons) also note sum.atMost_Suc_shift[symmetric] also note degree_pCons_eq[OF \p \ 0\, of a, symmetric] finally show ?thesis . qed qed lemma poly_0_coeff_0: "poly p 0 = coeff p 0" by (cases p) (auto simp: poly_altdef) subsection \Monomials\ lift_definition monom :: "'a \ nat \ 'a::zero poly" is "\a m n. if m = n then a else 0" by (simp add: MOST_iff_cofinite) lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)" by transfer rule lemma monom_0: "monom a 0 = pCons a 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_eq_0 [simp]: "monom 0 n = 0" by (rule poly_eqI) simp lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" apply (rule order_antisym [OF degree_monom_le]) apply (rule le_degree) apply simp done lemma coeffs_monom [code abstract]: "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" by (induct n) (simp_all add: monom_0 monom_Suc) lemma fold_coeffs_monom [simp]: "a \ 0 \ fold_coeffs f (monom a n) = f 0 ^^ n \ f a" by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) lemma poly_monom: "poly (monom a n) x = a * x ^ n" for a x :: "'a::comm_semiring_1" by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_def) lemma monom_eq_iff': "monom c n = monom d m \ c = d \ (c = 0 \ n = m)" by (auto simp: poly_eq_iff) lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) subsection \Leading coefficient\ abbreviation lead_coeff:: "'a::zero poly \ 'a" where "lead_coeff p \ coeff p (degree p)" lemma lead_coeff_pCons[simp]: "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" "p = 0 \ lead_coeff (pCons a p) = a" by auto lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" by (cases "c = 0") (simp_all add: degree_monom_eq) lemma last_coeffs_eq_coeff_degree: "last (coeffs p) = lead_coeff p" if "p \ 0" using that by (simp add: coeffs_def) subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add begin lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n + coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n + coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" by (simp add: plus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "(p + q) + r = p + (q + r)" by (simp add: poly_eq_iff add.assoc) show "p + q = q + p" by (simp add: poly_eq_iff add.commute) show "0 + p = p" by (simp add: poly_eq_iff) qed end instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n - coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n - coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" by (simp add: minus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "p + q - p = q" by (simp add: poly_eq_iff) show "p - q - r = p - (q + r)" by (simp add: poly_eq_iff diff_diff_eq) qed end instantiation poly :: (ab_group_add) ab_group_add begin lift_definition uminus_poly :: "'a poly \ 'a poly" is "\p n. - coeff p n" proof - fix p :: "'a poly" show "\\<^sub>\n. - coeff p n = 0" using MOST_coeff_eq_0 by simp qed lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" by (simp add: uminus_poly.rep_eq) instance proof fix p q :: "'a poly" show "- p + p = 0" by (simp add: poly_eq_iff) show "p - q = p + - q" by (simp add: poly_eq_iff) qed end lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" by (rule degree_le) (auto simp add: coeff_eq_0) lemma degree_add_le: "degree p \ n \ degree q \ n \ degree (p + q) \ n" by (auto intro: order_trans degree_add_le_max) lemma degree_add_less: "degree p < n \ degree q < n \ degree (p + q) < n" by (auto intro: le_less_trans degree_add_le_max) lemma degree_add_eq_right: "degree p < degree q \ degree (p + q) = degree q" apply (cases "q = 0") apply simp apply (rule order_antisym) apply (simp add: degree_add_le) apply (rule le_degree) apply (simp add: coeff_eq_0) done lemma degree_add_eq_left: "degree q < degree p \ degree (p + q) = degree p" using degree_add_eq_right [of q p] by (simp add: add.commute) lemma degree_minus [simp]: "degree (- p) = degree p" by (simp add: degree_def) lemma lead_coeff_add_le: "degree p < degree q \ lead_coeff (p + q) = lead_coeff q" by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) lemma lead_coeff_minus: "lead_coeff (- p) = - lead_coeff p" by (metis coeff_minus degree_minus) lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" for p q :: "'a::ab_group_add poly" using degree_add_le [where p=p and q="-q"] by simp lemma degree_diff_le: "degree p \ n \ degree q \ n \ degree (p - q) \ n" for p q :: "'a::ab_group_add poly" using degree_add_le [of p n "- q"] by simp lemma degree_diff_less: "degree p < n \ degree q < n \ degree (p - q) < n" for p q :: "'a::ab_group_add poly" using degree_add_less [of p n "- q"] by simp lemma add_monom: "monom a n + monom b n = monom (a + b) n" by (rule poly_eqI) simp lemma diff_monom: "monom a n - monom b n = monom (a - b) n" by (rule poly_eqI) simp lemma minus_monom: "- monom a n = monom (- a) n" by (rule poly_eqI) simp lemma coeff_sum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" by (induct A rule: infinite_finite_induct) simp_all lemma monom_sum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" by (rule poly_eqI) (simp add: coeff_sum) fun plus_coeffs :: "'a::comm_monoid_add list \ 'a list \ 'a list" where "plus_coeffs xs [] = xs" | "plus_coeffs [] ys = ys" | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" lemma coeffs_plus_eq_plus_coeffs [code abstract]: "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" proof - have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" for xs ys :: "'a list" and n proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) case (3 x xs y ys n) then show ?case by (cases n) (auto simp add: cCons_def) qed simp_all have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)" if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys" for xs ys :: "'a list" using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def) show ?thesis by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **) qed lemma coeffs_uminus [code abstract]: "coeffs (- p) = map uminus (coeffs p)" proof - have eq_0: "HOL.eq 0 \ uminus = HOL.eq (0::'a)" by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0) qed lemma [code]: "p - q = p + - q" for p q :: "'a::ab_group_add poly" by (fact diff_conv_add_uminus) lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" apply (induct p arbitrary: q) apply simp apply (case_tac q, simp, simp add: algebra_simps) done lemma poly_minus [simp]: "poly (- p) x = - poly p x" for x :: "'a::comm_ring" by (induct p) simp_all lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x" for x :: "'a::comm_ring" using poly_add [of p "- q" x] by simp lemma poly_sum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_sum_le: "finite S \ (\p. p \ S \ degree (f p) \ n) \ degree (sum f S) \ n" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert p S) then have "degree (sum f S) \ n" "degree (f p) \ n" by auto then show ?case unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le) qed lemma poly_as_sum_of_monoms': assumes "degree p \ n" shows "(\i\n. monom (coeff p i) i) = p" proof - have eq: "\i. {..n} \ {i} = (if i \ n then {i} else {})" by auto from assms show ?thesis by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq if_distrib[where f="\x. x * a" for a]) qed lemma poly_as_sum_of_monoms: "(\i\degree p. monom (coeff p i) i) = p" by (intro poly_as_sum_of_monoms' order_refl) lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" by (induct xs) (simp_all add: monom_0 monom_Suc) subsection \Multiplication by a constant, polynomial multiplication and the unit polynomial\ lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" is "\a p n. a * coeff p n" proof - fix a :: 'a and p :: "'a poly" show "\\<^sub>\ i. a * coeff p i = 0" using MOST_coeff_eq_0[of p] by eventually_elim simp qed lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" by (simp add: smult.rep_eq) lemma degree_smult_le: "degree (smult a p) \ degree p" by (rule degree_le) (simp add: coeff_eq_0) lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" by (rule poly_eqI) (simp add: mult.assoc) lemma smult_0_right [simp]: "smult a 0 = 0" by (rule poly_eqI) simp lemma smult_0_left [simp]: "smult 0 p = 0" by (rule poly_eqI) simp lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" by (rule poly_eqI) simp lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_add_left: "smult (a + b) p = smult a p + smult b p" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_minus_right [simp]: "smult a (- p) = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_minus_left [simp]: "smult (- a) p = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q" for a :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p" for a b :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemmas smult_distribs = smult_add_left smult_add_right smult_diff_left smult_diff_right lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma smult_monom: "smult a (monom b n) = monom (a * b) n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma smult_Poly: "smult c (Poly xs) = Poly (map ((*) c) xs)" by (auto simp: poly_eq_iff nth_default_def) lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (cases "a = 0") (simp_all add: degree_def) lemma smult_eq_0_iff [simp]: "smult a p = 0 \ a = 0 \ p = 0" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (simp add: poly_eq_iff) lemma coeffs_smult [code abstract]: "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have eq_0: "HOL.eq 0 \ times a = HOL.eq (0::'a)" if "a \ 0" using that by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0) qed lemma smult_eq_iff: fixes b :: "'a :: field" assumes "b \ 0" shows "smult a p = smult b q \ smult (a / b) p = q" (is "?lhs \ ?rhs") proof assume ?lhs also from assms have "smult (inverse b) \ = q" by simp finally show ?rhs by (simp add: field_simps) next assume ?rhs with assms show ?lhs by auto qed instantiation poly :: (comm_semiring_0) comm_semiring_0 begin definition "p * q = fold_coeffs (\a p. smult a q + pCons 0 p) p 0" lemma mult_poly_0_left: "(0::'a poly) * q = 0" by (simp add: times_poly_def) lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" by (cases "p = 0 \ a = 0") (auto simp add: times_poly_def) lemma mult_poly_0_right: "p * (0::'a poly) = 0" by (induct p) (simp_all add: mult_poly_0_left) lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" by (induct p) (simp_all add: mult_poly_0_left algebra_simps) lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" by (induct p) (simp_all add: mult_poly_0 smult_add_right) lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" by (induct q) (simp_all add: mult_poly_0 smult_add_right) lemma mult_poly_add_left: "(p + q) * r = p * r + q * r" for p q r :: "'a poly" by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps) instance proof fix p q r :: "'a poly" show 0: "0 * p = 0" by (rule mult_poly_0_left) show "p * 0 = 0" by (rule mult_poly_0_right) show "(p + q) * r = p * r + q * r" by (rule mult_poly_add_left) show "(p * q) * r = p * (q * r)" by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left) show "p * q = q * p" by (induct p) (simp_all add: mult_poly_0) qed end lemma coeff_mult_degree_sum: "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (induct p) (simp_all add: coeff_eq_0) instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors proof fix p q :: "'a poly" assume "p \ 0" and "q \ 0" have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (rule coeff_mult_degree_sum) also from \p \ 0\ \q \ 0\ have "coeff p (degree p) * coeff q (degree q) \ 0" by simp finally have "\n. coeff (p * q) n \ 0" .. then show "p * q \ 0" by (simp add: poly_eq_iff) qed instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. lemma coeff_mult: "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" proof (induct p arbitrary: n) case 0 show ?case by simp next case (pCons a p n) then show ?case by (cases n) (simp_all add: sum.atMost_Suc_shift del: sum.atMost_Suc) qed lemma degree_mult_le: "degree (p * q) \ degree p + degree q" apply (rule degree_le) apply (induct p) apply simp apply (simp add: coeff_eq_0 coeff_pCons split: nat.split) done lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) instantiation poly :: (comm_semiring_1) comm_semiring_1 begin lift_definition one_poly :: "'a poly" is "\n. of_bool (n = 0)" by (rule MOST_SucD) simp lemma coeff_1 [simp]: "coeff 1 n = of_bool (n = 0)" by (simp add: one_poly.rep_eq) lemma one_pCons: "1 = [:1:]" by (simp add: poly_eq_iff coeff_pCons split: nat.splits) lemma pCons_one: "[:1:] = 1" by (simp add: one_pCons) instance by standard (simp_all add: one_pCons) end lemma poly_1 [simp]: "poly 1 x = 1" by (simp add: one_pCons) lemma one_poly_eq_simps [simp]: "1 = [:1:] \ True" "[:1:] = 1 \ True" by (simp_all add: one_pCons) lemma degree_1 [simp]: "degree 1 = 0" by (simp add: one_pCons) lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]" by (simp add: one_pCons) lemma smult_one [simp]: "smult c 1 = [:c:]" by (simp add: one_pCons) lemma monom_eq_1 [simp]: "monom 1 0 = 1" by (simp add: monom_0 one_pCons) lemma monom_eq_1_iff: "monom c n = 1 \ c = 1 \ n = 0" using monom_eq_const_iff [of c n 1] by auto lemma monom_altdef: "monom c n = smult c ([:0, 1:] ^ n)" by (induct n) (simp_all add: monom_0 monom_Suc) instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors .. instance poly :: (comm_ring) comm_ring .. instance poly :: (comm_ring_1) comm_ring_1 .. instance poly :: (comm_ring_1) comm_semiring_1_cancel .. lemma degree_power_le: "degree (p ^ n) \ degree p * n" by (induct n) (auto intro: order_trans degree_mult_le) lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induct n) (simp_all add: coeff_mult) lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" by (induct p) (simp_all add: algebra_simps) lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" by (induct p) (simp_all add: algebra_simps) lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n" for p :: "'a::comm_semiring_1 poly" by (induct n) simp_all lemma poly_prod: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_prod_sum_le: "finite S \ degree (prod f S) \ sum (degree \ f) S" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert a S) show ?case unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] by (rule le_trans[OF degree_mult_le]) (use insert in auto) qed lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" by (induct xs) (simp_all add: coeff_mult) lemma coeff_monom_mult: "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" proof - have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" by (simp add: coeff_mult) also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" by (intro sum.cong) simp_all also have "\ = (if k < n then 0 else c * coeff p (k - n))" by simp finally show ?thesis . qed lemma monom_1_dvd_iff': "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" have "\\<^sub>\k. coeff p (k + n) = 0" by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, subst cofinite_eq_sequentially [symmetric]) transfer then have coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def by (subst poly.Abs_poly_inverse) simp_all have "p = monom 1 n * r" by (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero) then show "monom 1 n dvd p" by simp qed subsection \Mapping polynomials\ definition map_poly :: "('a :: zero \ 'b :: zero) \ 'a poly \ 'b poly" where "map_poly f p = Poly (map f (coeffs p))" lemma map_poly_0 [simp]: "map_poly f 0 = 0" by (simp add: map_poly_def) lemma map_poly_1: "map_poly f 1 = [:f 1:]" by (simp add: map_poly_def) lemma map_poly_1' [simp]: "f 1 = 1 \ map_poly f 1 = 1" by (simp add: map_poly_def one_pCons) lemma coeff_map_poly: assumes "f 0 = 0" shows "coeff (map_poly f p) n = f (coeff p n)" by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc) lemma coeffs_map_poly [code abstract]: "coeffs (map_poly f p) = strip_while ((=) 0) (map f (coeffs p))" by (simp add: map_poly_def) lemma coeffs_map_poly': assumes "\x. x \ 0 \ f x \ 0" shows "coeffs (map_poly f p) = map f (coeffs p)" using assms by (auto simp add: coeffs_map_poly strip_while_idem_iff last_coeffs_eq_coeff_degree no_trailing_unfold last_map) lemma set_coeffs_map_poly: "(\x. f x = 0 \ x = 0) \ set (coeffs (map_poly f p)) = f ` set (coeffs p)" by (simp add: coeffs_map_poly') lemma degree_map_poly: assumes "\x. x \ 0 \ f x \ 0" shows "degree (map_poly f p) = degree p" by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms) lemma map_poly_eq_0_iff: assumes "f 0 = 0" "\x. x \ set (coeffs p) \ x \ 0 \ f x \ 0" shows "map_poly f p = 0 \ p = 0" proof - have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n proof - have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms) also have "\ = 0 \ coeff p n = 0" proof (cases "n < length (coeffs p)") case True then have "coeff p n \ set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc) with assms show "f (coeff p n) = 0 \ coeff p n = 0" by auto next case False then show ?thesis by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def) qed finally show ?thesis . qed then show ?thesis by (auto simp: poly_eq_iff) qed lemma map_poly_smult: assumes "f 0 = 0""\c x. f (c * x) = f c * f x" shows "map_poly f (smult c p) = smult (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly) lemma map_poly_pCons: assumes "f 0 = 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits) lemma map_poly_map_poly: assumes "f 0 = 0" "g 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" by (intro poly_eqI) (simp add: coeff_map_poly assms) lemma map_poly_id [simp]: "map_poly id p = p" by (simp add: map_poly_def) lemma map_poly_id' [simp]: "map_poly (\x. x) p = p" by (simp add: map_poly_def) lemma map_poly_cong: assumes "(\x. x \ set (coeffs p) \ f x = g x)" shows "map_poly f p = map_poly g p" proof - from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all then show ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly) qed lemma map_poly_monom: "f 0 = 0 \ map_poly f (monom c n) = monom (f c) n" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma map_poly_idI: assumes "\x. x \ set (coeffs p) \ f x = x" shows "map_poly f p = p" using map_poly_cong[OF assms, of _ id] by simp lemma map_poly_idI': assumes "\x. x \ set (coeffs p) \ f x = x" shows "p = map_poly f p" using map_poly_cong[OF assms, of _ id] by simp lemma smult_conv_map_poly: "smult c p = map_poly (\x. c * x) p" by (intro poly_eqI) (simp_all add: coeff_map_poly) subsection \Conversions\ lemma of_nat_poly: "of_nat n = [:of_nat n:]" by (induct n) (simp_all add: one_pCons) lemma of_nat_monom: "of_nat n = monom (of_nat n) 0" by (simp add: of_nat_poly monom_0) lemma degree_of_nat [simp]: "degree (of_nat n) = 0" by (simp add: of_nat_poly) lemma lead_coeff_of_nat [simp]: "lead_coeff (of_nat n) = of_nat n" by (simp add: of_nat_poly) lemma of_int_poly: "of_int k = [:of_int k:]" by (simp only: of_int_of_nat of_nat_poly) simp lemma of_int_monom: "of_int k = monom (of_int k) 0" by (simp add: of_int_poly monom_0) lemma degree_of_int [simp]: "degree (of_int k) = 0" by (simp add: of_int_poly) lemma lead_coeff_of_int [simp]: "lead_coeff (of_int k) = of_int k" by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" proof - have "numeral n = of_nat (numeral n)" by simp also have "\ = [:of_nat (numeral n):]" by (simp add: of_nat_poly) finally show ?thesis by simp qed lemma numeral_monom: "numeral n = monom (numeral n) 0" by (simp add: numeral_poly monom_0) lemma degree_numeral [simp]: "degree (numeral n) = 0" by (simp add: numeral_poly) lemma lead_coeff_numeral [simp]: "lead_coeff (numeral n) = numeral n" by (simp add: numeral_poly) subsection \Lemmas about divisibility\ lemma dvd_smult: assumes "p dvd q" shows "p dvd smult a q" proof - from assms obtain k where "q = p * k" .. then have "smult a q = p * smult a k" by simp then show "p dvd smult a q" .. qed lemma dvd_smult_cancel: "p dvd smult a q \ a \ 0 \ p dvd q" for a :: "'a::field" by (drule dvd_smult [where a="inverse a"]) simp lemma dvd_smult_iff: "a \ 0 \ p dvd smult a q \ p dvd q" for a :: "'a::field" by (safe elim!: dvd_smult dvd_smult_cancel) lemma smult_dvd_cancel: assumes "smult a p dvd q" shows "p dvd q" proof - from assms obtain k where "q = smult a p * k" .. then have "q = p * smult a k" by simp then show "p dvd q" .. qed lemma smult_dvd: "p dvd q \ a \ 0 \ smult a p dvd q" for a :: "'a::field" by (rule smult_dvd_cancel [where a="inverse a"]) simp lemma smult_dvd_iff: "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" for a :: "'a::field" by (auto elim: smult_dvd smult_dvd_cancel) lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" proof - have "smult c p = [:c:] * p" by simp also have "\ dvd 1 \ c dvd 1 \ p dvd 1" proof safe assume *: "[:c:] * p dvd 1" then show "p dvd 1" by (rule dvd_mult_right) from * obtain q where q: "1 = [:c:] * p * q" by (rule dvdE) have "c dvd c * (coeff p 0 * coeff q 0)" by simp also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) also note q [symmetric] finally have "c dvd coeff 1 0" . then show "c dvd 1" by simp next assume "c dvd 1" "p dvd 1" from this(1) obtain d where "1 = c * d" by (rule dvdE) then have "1 = [:c:] * [:d:]" by (simp add: one_pCons ac_simps) then have "[:c:] dvd 1" by (rule dvdI) from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp qed finally show ?thesis . qed subsection \Polynomials form an integral domain\ instance poly :: (idom) idom .. instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0 by standard (auto simp add: of_nat_poly intro: injI) lemma degree_mult_eq: "p \ 0 \ q \ 0 \ degree (p * q) = degree p + degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum) lemma degree_mult_eq_0: "degree (p * q) = 0 \ p = 0 \ q = 0 \ (p \ 0 \ q \ 0 \ degree p = 0 \ degree q = 0)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (auto simp: degree_mult_eq) lemma degree_power_eq: "p \ 0 \ degree ((p :: 'a :: idom poly) ^ n) = n * degree p" by (induction n) (simp_all add: degree_mult_eq) lemma degree_mult_right_le: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "q \ 0" shows "degree p \ degree (p * q)" using assms by (cases "p = 0") (simp_all add: degree_mult_eq) lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) lemma dvd_imp_degree_le: "p dvd q \ q \ 0 \ degree p \ degree q" for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (erule dvdE, hypsubst, subst degree_mult_eq) auto lemma divides_degree: fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "degree p \ degree q \ q = 0" by (metis dvd_imp_degree_le assms) lemma const_poly_dvd_iff: fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" shows "[:c:] dvd p \ (\n. c dvd coeff p n)" proof (cases "c = 0 \ p = 0") case True then show ?thesis by (auto intro!: poly_eqI) next case False show ?thesis proof assume "[:c:] dvd p" then show "\n. c dvd coeff p n" by (auto elim!: dvdE simp: coeffs_def) next assume *: "\n. c dvd coeff p n" define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a have mydiv: "x = y * mydiv x y" if "y dvd x" for x y using that unfolding mydiv_def dvd_def by (rule someI_ex) define q where "q = Poly (map (\a. mydiv a c) (coeffs p))" from False * have "p = q * [:c:]" by (intro poly_eqI) (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth intro!: coeff_eq_0 mydiv) then show "[:c:] dvd p" by (simp only: dvd_triv_right) qed qed lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \ a dvd b" for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q" for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have "smult c p = [:c:] * p" by simp also have "lead_coeff \ = c * lead_coeff p" by (subst lead_coeff_mult) simp_all finally show ?thesis . qed lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by simp lemma lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: lead_coeff_mult) subsection \Polynomials form an ordered integral domain\ definition pos_poly :: "'a::linordered_semidom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" lemma pos_poly_pCons: "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" by (simp add: pos_poly_def) lemma not_pos_poly_0 [simp]: "\ pos_poly 0" by (simp add: pos_poly_def) lemma pos_poly_add: "pos_poly p \ pos_poly q \ pos_poly (p + q)" apply (induct p arbitrary: q) apply simp apply (case_tac q) apply (force simp add: pos_poly_pCons add_pos_pos) done lemma pos_poly_mult: "pos_poly p \ pos_poly q \ pos_poly (p * q)" unfolding pos_poly_def apply (subgoal_tac "p \ 0 \ q \ 0") apply (simp add: degree_mult_eq coeff_mult_degree_sum) apply auto done lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" for p :: "'a::linordered_idom poly" by (induct p) (auto simp: pos_poly_pCons) lemma pos_poly_coeffs [code]: "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) next assume ?lhs then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) then have "p \ 0" by auto with * show ?rhs by (simp add: last_coeffs_eq_coeff_degree) qed instantiation poly :: (linordered_idom) linordered_idom begin definition "x < y \ pos_poly (y - x)" definition "x \ y \ x = y \ pos_poly (y - x)" definition "\x::'a poly\ = (if x < 0 then - x else x)" definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" instance proof fix x y z :: "'a poly" show "x < y \ x \ y \ \ y \ x" unfolding less_eq_poly_def less_poly_def apply safe apply simp apply (drule (1) pos_poly_add) apply simp done show "x \ x" by (simp add: less_eq_poly_def) show "x \ y \ y \ z \ x \ z" unfolding less_eq_poly_def apply safe apply (drule (1) pos_poly_add) apply (simp add: algebra_simps) done show "x \ y \ y \ x \ x = y" unfolding less_eq_poly_def apply safe apply (drule (1) pos_poly_add) apply simp done show "x \ y \ z + x \ z + y" unfolding less_eq_poly_def apply safe apply (simp add: algebra_simps) done show "x \ y \ y \ x" unfolding less_eq_poly_def using pos_poly_total [of "x - y"] by auto show "x < y \ 0 < z \ z * x < z * y" by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) show "\x\ = (if x < 0 then - x else x)" by (rule abs_poly_def) show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" by (rule sgn_poly_def) qed end text \TODO: Simplification rules for comparisons\ subsection \Synthetic division and polynomial roots\ subsubsection \Synthetic division\ text \Synthetic division is simply division by the linear polynomial \<^term>\x - c\.\ definition synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" where "synthetic_divmod p c = fold_coeffs (\a (q, r). (pCons r q, a + c * r)) p (0, 0)" definition synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" where "synthetic_div p c = fst (synthetic_divmod p c)" lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)" by (simp add: synthetic_divmod_def) lemma synthetic_divmod_pCons [simp]: "synthetic_divmod (pCons a p) c = (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" by (cases "p = 0 \ a = 0") (auto simp add: synthetic_divmod_def) lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" by (simp add: synthetic_div_def) lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" by (induct p arbitrary: a) simp_all lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" by (induct p) (simp_all add: split_def) lemma synthetic_div_pCons [simp]: "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" by (simp add: synthetic_div_def split_def snd_synthetic_divmod) lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \ degree p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) then show ?case by (cases p) simp qed lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1" by (induct p) (simp_all add: synthetic_div_eq_0_iff) lemma synthetic_div_correct: "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" by (induct p) simp_all lemma synthetic_div_unique: "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" apply (induct p arbitrary: q r) apply simp apply (frule synthetic_div_unique_lemma) apply simp apply (case_tac q, force) done lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" for c :: "'a::comm_ring_1" using synthetic_div_correct [of p c] by (simp add: algebra_simps) subsubsection \Polynomial roots\ lemma poly_eq_0_iff_dvd: "poly p c = 0 \ [:- c, 1:] dvd p" (is "?lhs \ ?rhs") for c :: "'a::comm_ring_1" proof assume ?lhs with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp then show ?rhs .. next assume ?rhs then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) then show ?lhs by simp qed lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \ poly p (- c) = 0" for c :: "'a::comm_ring_1" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: "p \ 0 \ finite {x. poly p x = 0}" for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly" proof (induct n \ "degree p" arbitrary: p) case 0 then obtain a where "a \ 0" and "p = [:a:]" by (cases p) (simp split: if_splits) then show "finite {x. poly p x = 0}" by simp next case (Suc n) show "finite {x. poly p x = 0}" proof (cases "\x. poly p x = 0") case False then show "finite {x. poly p x = 0}" by simp next case True then obtain a where "poly p a = 0" .. then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) then obtain k where k: "p = [:-a, 1:] * k" .. with \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from this \k \ 0\ have "finite {x. poly k x = 0}" by (rule Suc.hyps) then have "finite (insert a {x. poly k x = 0})" by simp then show "finite {x. poly p x = 0}" by (simp add: k Collect_disj_eq del: mult_pCons_left) qed qed lemma poly_eq_poly_eq_iff: "poly p = poly q \ p = q" (is "?lhs \ ?rhs") for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly" proof assume ?rhs then show ?lhs by simp next assume ?lhs have "poly p = poly 0 \ p = 0" for p :: "'a poly" apply (cases "p = 0") apply simp_all apply (drule poly_roots_finite) apply (auto simp add: infinite_UNIV_char_0) done from \?lhs\ and this [of "p - q"] show ?rhs by auto qed lemma poly_all_0_iff_0: "(\x. poly p x = 0) \ p = 0" for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly" by (auto simp add: poly_eq_poly_eq_iff [symmetric]) subsubsection \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" for a :: "'a::comm_semiring_1" apply (induct n) apply simp_all apply (subst coeff_eq_0) apply (auto intro: le_less_trans degree_power_le) done lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" for a :: "'a::comm_semiring_1" apply (rule order_antisym) apply (rule ord_le_eq_trans [OF degree_power_le]) apply simp apply (rule le_degree) apply (simp add: coeff_linear_power) done lemma order_1: "[:-a, 1:] ^ order a p dvd p" apply (cases "p = 0") apply simp apply (cases "order a p") apply simp apply (subgoal_tac "nat < (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)") apply (drule not_less_Least) apply simp apply (fold order_def) apply simp done lemma order_2: "p \ 0 \ \ [:-a, 1:] ^ Suc (order a p) dvd p" unfolding order_def apply (rule LeastI_ex) apply (rule_tac x="degree p" in exI) apply (rule notI) apply (drule (1) dvd_imp_degree_le) apply (simp only: degree_linear_power) done lemma order: "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" by (rule conjI [OF order_1 order_2]) lemma order_degree: assumes p: "p \ 0" shows "order a p \ degree p" proof - have "order a p = degree ([:-a, 1:] ^ order a p)" by (simp only: degree_linear_power) also from order_1 p have "\ \ degree p" by (rule dvd_imp_degree_le) finally show ?thesis . qed lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" apply (cases "p = 0") apply simp_all apply (rule iffI) apply (metis order_2 not_gr0 poly_eq_0_iff_dvd power_0 power_Suc_0 power_one_right) unfolding poly_eq_0_iff_dvd apply (metis dvd_power dvd_trans order_1) done lemma order_0I: "poly p a \ 0 \ order a p = 0" by (subst (asm) order_root) auto lemma order_unique_lemma: fixes p :: "'a::idom poly" assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" shows "n = order a p" unfolding Polynomial.order_def apply (rule Least_equality [symmetric]) apply (fact assms) apply (rule classical) apply (erule notE) unfolding not_less_eq_eq using assms(1) apply (rule power_le_dvd) apply assumption done lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" proof - define i where "i = order a p" define j where "j = order a q" define t where "t = [:-a, 1:]" have t_dvd_iff: "\u. t dvd u \ poly u a = 0" by (simp add: t_def dvd_iff_poly_eq_0) assume "p * q \ 0" then show "order a (p * q) = i + j" apply clarsimp apply (drule order [where a=a and p=p, folded i_def t_def]) apply (drule order [where a=a and p=q, folded j_def t_def]) apply clarify apply (erule dvdE)+ apply (rule order_unique_lemma [symmetric], fold t_def) apply (simp_all add: power_add t_dvd_iff) done qed lemma order_smult: assumes "c \ 0" shows "order x (smult c p) = order x p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "smult c p = [:c:] * p" by simp also from assms False have "order x \ = order x [:c:] + order x p" by (subst order_mult) simp_all also have "order x [:c:] = 0" by (rule order_0I) (use assms in auto) finally show ?thesis by simp qed (* Next three lemmas contributed by Wenda Li *) lemma order_1_eq_0 [simp]:"order x 1 = 0" by (metis order_root poly_1 zero_neq_one) lemma order_uminus[simp]: "order x (-p) = order x p" by (metis neg_equal_0_iff_equal order_smult smult_1_left smult_minus_left) lemma order_power_n_n: "order a ([:-a,1:]^n)=n" proof (induct n) (*might be proved more concisely using nat_less_induct*) case 0 then show ?case by (metis order_root poly_1 power_0 zero_neq_one) next case (Suc n) have "order a ([:- a, 1:] ^ Suc n) = order a ([:- a, 1:] ^ n) + order a [:-a,1:]" by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) moreover have "order a [:-a,1:] = 1" unfolding order_def proof (rule Least_equality, rule notI) assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" then have "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:])" by (rule dvd_imp_degree_le) auto then show False by auto next fix y assume *: "\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" show "1 \ y" proof (rule ccontr) assume "\ 1 \ y" then have "y = 0" by auto then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto with * show False by auto qed qed ultimately show ?case using Suc by auto qed lemma order_0_monom [simp]: "c \ 0 \ order 0 (monom c n) = n" using order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) lemma dvd_imp_order_le: "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" by (auto simp: order_mult elim: dvdE) text \Now justify the standard squarefree decomposition, i.e. \f / gcd f f'\.\ lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" apply (cases "p = 0") apply auto apply (drule order_2 [where a=a and p=p]) apply (metis not_less_eq_eq power_le_dvd) apply (erule power_le_dvd [OF order_1]) done lemma order_decomp: assumes "p \ 0" shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" proof - from assms have *: "[:- a, 1:] ^ order a p dvd p" and **: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" .. with ** have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" by simp then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" by simp with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] have "\ [:- a, 1:] dvd q" by auto with q show ?thesis by blast qed lemma monom_1_dvd_iff: "p \ 0 \ monom 1 n dvd p \ n \ order 0 p" using order_divides[of 0 n p] by (simp add: monom_altdef) subsection \Additional induction rules on polynomials\ text \ An induction rule for induction over the roots of a polynomial with a certain property. (e.g. all positive roots) \ lemma poly_root_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "Q 0" and "\p. (\a. P a \ poly p a \ 0) \ Q p" and "\a p. P a \ Q p \ Q ([:a, -1:] * p)" shows "Q p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "p = 0") case True with assms(1) show ?thesis by simp next case False show ?thesis proof (cases "\a. P a \ poly p a = 0") case False then show ?thesis by (intro assms(2)) blast next case True then obtain a where a: "P a" "poly p a = 0" by blast then have "-[:-a, 1:] dvd p" by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp with False have "q \ 0" by auto have "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) (simp_all add: \q \ 0\) then have "Q q" by (intro less) simp with a(1) have "Q ([:a, -1:] * q)" by (rule assms(3)) with q show ?thesis by simp qed qed qed lemma dropWhile_replicate_append: "dropWhile ((=) a) (replicate n a @ ys) = dropWhile ((=) a) ys" by (induct n) simp_all lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) text \ An induction rule for simultaneous induction over two polynomials, prepending one coefficient in each step. \ lemma poly_induct2 [case_names 0 pCons]: assumes "P 0 0" "\a p b q. P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof - define n where "n = max (length (coeffs p)) (length (coeffs q))" define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" have "length xs = length ys" by (simp add: xs_def ys_def n_def) then have "P (Poly xs) (Poly ys)" by (induct rule: list_induct2) (simp_all add: assms) also have "Poly xs = p" by (simp add: xs_def Poly_append_replicate_0) also have "Poly ys = q" by (simp add: ys_def Poly_append_replicate_0) finally show ?thesis . qed subsection \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" notation pcompose (infixl "\\<^sub>p" 71) lemma pcompose_0 [simp]: "pcompose 0 q = 0" by (simp add: pcompose_def) lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q" by (cases "p = 0 \ a = 0") (auto simp add: pcompose_def) lemma pcompose_1: "pcompose 1 p = 1" for p :: "'a::comm_semiring_1 poly" by (auto simp: one_pCons pcompose_pCons) lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" by (induct p) (simp_all add: pcompose_pCons) lemma degree_pcompose_le: "degree (pcompose p q) \ degree p * degree q" apply (induct p) apply simp apply (simp add: pcompose_pCons) apply clarify apply (rule degree_add_le) apply simp apply (rule order_trans [OF degree_mult_le]) apply simp done lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r" for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly" proof (induction p q rule: poly_induct2) case 0 then show ?case by simp next case (pCons a p b q) have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r" by (simp_all add: pcompose_pCons pCons.IH algebra_simps) also have "[:a + b:] = [:a:] + [:b:]" by simp also have "\ + r * pcompose p r + r * pcompose q r = pcompose (pCons a p) r + pcompose (pCons b q) r" by (simp only: pcompose_pCons add_ac) finally show ?case . qed lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r" for p r :: "'a::comm_ring poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r" for p q r :: "'a::comm_ring poly" using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)" for p r :: "'a::comm_semiring_0 poly" by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right) lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult) lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p" for p :: "'a::comm_semiring_1 poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_sum: "pcompose (sum f A) p = sum (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add) lemma pcompose_prod: "pcompose (prod f A) p = prod (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult) lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]" by (subst pcompose_pCons) simp lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]" by (induct p) (auto simp add: pcompose_pCons) lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 show ?thesis proof (cases "p = 0") case True then show ?thesis by auto next case False from prems have "degree q = 0 \ pcompose p q = 0" by (auto simp add: degree_mult_eq_0) moreover have False if "pcompose p q = 0" "degree q \ 0" proof - from pCons.hyps(2) that have "degree p = 0" by auto then obtain a1 where "p = [:a1:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) with \pcompose p q = 0\ \p \ 0\ show False by auto qed ultimately have "degree (pCons a p) * degree q = 0" by auto moreover have "degree (pcompose (pCons a p) q) = 0" proof - from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))" by simp also have "\ \ degree ([:a:] + q * pcompose p q)" by (rule degree_add_le_max) finally show ?thesis by (auto simp add: pcompose_pCons) qed ultimately show ?thesis by simp qed next case prems: 2 then have "p \ 0" "q \ 0" "pcompose p q \ 0" by auto from prems degree_add_eq_right [of "[:a:]"] have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)" by (auto simp: pcompose_pCons) with pCons.hyps(2) degree_mult_eq[OF \q\0\ \pcompose p q\0\] show ?thesis by auto qed qed lemma pcompose_eq_0: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "pcompose p q = 0" "degree q > 0" shows "p = 0" proof - from assms degree_pcompose [of p q] have "degree p = 0" by auto then obtain a where "p = [:a:]" by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) with assms(1) have "a = 0" by auto with \p = [:a:]\ show ?thesis by simp qed lemma lead_coeff_comp: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree q > 0" shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 then have "pcompose p q = 0" by (metis assms degree_0 degree_mult_eq_0 neq0_conv) with pcompose_eq_0[OF _ \degree q > 0\] have "p = 0" by simp then show ?thesis by auto next case prems: 2 then have "degree [:a:] < degree (q * pcompose p q)" by simp then have "lead_coeff ([:a:] + q * p \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" by (rule lead_coeff_add_le) then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" by (simp add: pcompose_pCons) also have "\ = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "\ = lead_coeff p * lead_coeff q ^ (degree p + 1)" by (auto simp: mult_ac) finally show ?thesis by auto qed qed subsection \Shifting polynomials\ definition poly_shift :: "nat \ 'a::zero poly \ 'a poly" where "poly_shift n p = Abs_poly (\i. coeff p (i + n))" lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)" by (auto simp add: nth_default_def add_ac) lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" by (auto simp add: nth_default_def add_ac) lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" proof - from MOST_coeff_eq_0[of p] obtain m where "\k>m. coeff p k = 0" by (auto simp: MOST_nat) then have "\k>m. coeff p (k + n) = 0" by auto then have "\\<^sub>\k. coeff p (k + n) = 0" by (auto simp: MOST_nat) then show ?thesis by (simp add: poly_shift_def poly.Abs_poly_inverse) qed lemma poly_shift_id [simp]: "poly_shift 0 = (\x. x)" by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift) lemma poly_shift_0 [simp]: "poly_shift n 0 = 0" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \ n then monom c (m - n) else 0)" by (auto simp add: poly_eq_iff coeff_poly_shift) lemma coeffs_shift_poly [code abstract]: "coeffs (poly_shift n p) = drop n (coeffs p)" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq) qed subsection \Truncating polynomials\ definition poly_cutoff where "poly_cutoff n p = Abs_poly (\k. if k < n then coeff p k else 0)" lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)" unfolding poly_cutoff_def by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n]) lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma coeffs_poly_cutoff [code abstract]: "coeffs (poly_cutoff n p) = strip_while ((=) 0) (take n (coeffs p))" proof (cases "strip_while ((=) 0) (take n (coeffs p)) = []") case True then have "coeff (poly_cutoff n p) k = 0" for k unfolding coeff_poly_cutoff by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth) then have "poly_cutoff n p = 0" by (simp add: poly_eq_iff) then show ?thesis by (subst True) simp_all next case False have "no_trailing ((=) 0) (strip_while ((=) 0) (take n (coeffs p)))" by simp with False have "last (strip_while ((=) 0) (take n (coeffs p))) \ 0" unfolding no_trailing_unfold by auto then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq) qed subsection \Reflecting polynomials\ definition reflect_poly :: "'a::zero poly \ 'a poly" where "reflect_poly p = Poly (rev (coeffs p))" lemma coeffs_reflect_poly [code abstract]: "coeffs (reflect_poly p) = rev (dropWhile ((=) 0) (coeffs p))" by (simp add: reflect_poly_def) lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0" by (simp add: reflect_poly_def) lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1" by (simp add: reflect_poly_def one_pCons) lemma coeff_reflect_poly: "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))" by (cases "p = 0") (auto simp add: reflect_poly_def nth_default_def rev_nth degree_eq_length_coeffs coeffs_nth not_less dest: le_imp_less_Suc) lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly) lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly poly_0_coeff_0) lemma reflect_poly_pCons': "p \ 0 \ reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))" by (intro poly_eqI) (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split) lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]" by (cases "a = 0") (simp_all add: reflect_poly_def) lemma poly_reflect_poly_nz: "x \ 0 \ poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)" for x :: "'a::field" by (induct rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" by (simp add: coeff_reflect_poly) lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" by (simp add: poly_0_coeff_0) lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \ 0 \ reflect_poly (reflect_poly p) = p" by (cases p rule: pCons_cases) (simp add: reflect_poly_def ) lemma degree_reflect_poly_le: "degree (reflect_poly p) \ degree p" by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) lemma reflect_poly_pCons: "a \ 0 \ reflect_poly (pCons a p) = Poly (rev (a # coeffs p))" by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly) lemma degree_reflect_poly_eq [simp]: "coeff p 0 \ 0 \ degree (reflect_poly p) = degree p" by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs) (* TODO: does this work with zero divisors as well? Probably not. *) lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (cases "p = 0 \ q = 0") case False then have [simp]: "p \ 0" "q \ 0" by auto show ?thesis proof (rule poly_eqI) show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i proof (cases "i \ degree (p * q)") case True define A where "A = {..i} \ {i - degree q..degree p}" define B where "B = {..degree p} \ {degree p - i..degree (p*q) - i}" let ?f = "\j. degree p - j" from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)" by (simp add: coeff_reflect_poly) also have "\ = (\j\degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" by (simp add: coeff_mult) also have "\ = (\j\B. coeff p j * coeff q (degree (p * q) - i - j))" by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0) also from True have "\ = (\j\A. coeff p (degree p - j) * coeff q (degree q - (i - j)))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) (auto simp: A_def B_def degree_mult_eq add_ac) also have "\ = (\j\i. if j \ {i - degree q..degree p} then coeff p (degree p - j) * coeff q (degree q - (i - j)) else 0)" by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def) also have "\ = coeff (reflect_poly p * reflect_poly q) i" by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong) finally show ?thesis . qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral) qed qed auto lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" using reflect_poly_mult[of "[:c:]" p] by simp lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\x. reflect_poly (f x)) A" for f :: "_ \ _::{comm_semiring_0,semiring_no_zero_divisors} poly" by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)" for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list" by (induct xs) (simp_all add: reflect_poly_mult) lemma reflect_poly_Poly_nz: "no_trailing (HOL.eq 0) xs \ reflect_poly (Poly xs) = Poly (rev xs)" by (simp add: reflect_poly_def) lemmas reflect_poly_simps = reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult reflect_poly_power reflect_poly_prod reflect_poly_prod_list subsection \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ 'a poly" where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" by (auto intro: pCons_cases) termination pderiv by (relation "measure degree") simp_all declare pderiv.simps[simp del] lemma pderiv_0 [simp]: "pderiv 0 = 0" using pderiv.simps [of 0 0] by simp lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" by (simp add: pderiv.simps) lemma pderiv_1 [simp]: "pderiv 1 = 0" by (simp add: one_pCons pderiv_pCons) lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" and pderiv_numeral [simp]: "pderiv (numeral m) = 0" by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" by (induct p arbitrary: n) (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \ 'a list \ 'a list" where "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" | "pderiv_coeffs_code f [] = []" definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \ 'a list" where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" (* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) lemma pderiv_coeffs_code: "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n" proof (induct xs arbitrary: f n) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 then show ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") (auto simp: cCons_def) next case n: (Suc m) show ?thesis proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") case False then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" by (auto simp: cCons_def n) also have "\ = (f + of_nat n) * nth_default 0 xs m" by (simp add: Cons n add_ac) finally show ?thesis by (simp add: n) next case True have empty: "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" for g proof (induct xs arbitrary: g m) case Nil then show ?case by simp next case (Cons x xs) from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \ x = 0" by (auto simp: cCons_def split: if_splits) note IH = Cons(1)[OF empty] from IH[of m] IH[of "m - 1"] g show ?case by (cases m) (auto simp: field_simps) qed from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" by (auto simp: cCons_def n) moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" by (simp add: n) (use empty[of "f+1"] in \auto simp: field_simps\) ultimately show ?thesis by simp qed qed qed lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) case (1 n) have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" apply (rule iffI) apply (cases p) apply simp apply (simp add: poly_eq_iff coeff_pderiv del: of_nat_Suc) apply (simp add: poly_eq_iff coeff_pderiv coeff_eq_0) done lemma degree_pderiv: "degree (pderiv p) = degree p - 1" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" apply (rule order_antisym [OF degree_le]) apply (simp add: coeff_pderiv coeff_eq_0) apply (cases "degree p") apply simp apply (rule le_degree) apply (simp add: coeff_pderiv del: of_nat_Suc) apply (metis degree_0 leading_coeff_0_iff nat.distinct(1)) done lemma not_dvd_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" assumes "degree p \ 0" shows "\ p dvd pderiv p" proof assume dvd: "p dvd pderiv p" then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto from dvd have le: "degree p \ degree (pderiv p)" by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) from assms and this [unfolded degree_pderiv] show False by auto qed lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \ degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" by (simp add: pderiv_pCons) lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" apply (induct n) apply simp apply (subst power_Suc) apply (subst pderiv_mult) apply (erule ssubst) apply (simp only: of_nat_Suc smult_add_left smult_1_left) apply (simp add: algebra_simps) done lemma pderiv_pcompose: "pderiv (pcompose p q) = pcompose (pderiv p) q * pderiv q" by (induction p rule: pCons_induct) (auto simp: pcompose_pCons pderiv_add pderiv_mult pderiv_pCons pcompose_add algebra_simps) lemma pderiv_prod: "pderiv (prod f (as)) = (\a\as. prod f (as - {a}) * pderiv (f a))" proof (induct as rule: infinite_finite_induct) case (insert a as) then have id: "prod f (insert a as) = f a * prod f as" "\g. sum g (insert a as) = g a + sum g as" "insert a as - {a} = as" by auto have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \ as" for b proof - from \a \ as\ that have *: "insert a as - {b} = insert a (as - {b})" by auto show ?thesis unfolding * by (subst prod.insert) (use insert in auto) qed then show ?case unfolding id pderiv_mult insert(3) sum_distrib_left by (auto simp add: ac_simps intro!: sum.cong) qed auto lemma DERIV_pow2: "DERIV (\x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" by (rule DERIV_cong, rule DERIV_pow) simp declare DERIV_pow2 [simp] DERIV_pow [simp] lemma DERIV_add_const: "DERIV f x :> D \ DERIV (\x. a + f x :: 'a::real_normed_field) x :> D" by (rule DERIV_cong, rule DERIV_add) auto lemma poly_DERIV [simp]: "DERIV (\x. poly p x) x :> poly (pderiv p) x" by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) lemma poly_isCont[simp]: fixes x::"'a::real_normed_field" shows "isCont (\x. poly p x) x" by (rule poly_DERIV [THEN DERIV_isCont]) lemma tendsto_poly [tendsto_intros]: "(f \ a) F \ ((\x. poly p (f x)) \ poly p a) F" for f :: "_ \ 'a::real_normed_field" by (rule isCont_tendsto_compose [OF poly_isCont]) lemma continuous_within_poly: "continuous (at z within s) (poly p)" for z :: "'a::{real_normed_field}" by (simp add: continuous_within tendsto_poly) lemma continuous_poly [continuous_intros]: "continuous F f \ continuous F (\x. poly p (f x))" for f :: "_ \ 'a::real_normed_field" unfolding continuous_def by (rule tendsto_poly) lemma continuous_on_poly [continuous_intros]: fixes p :: "'a :: {real_normed_field} poly" assumes "continuous_on A f" shows "continuous_on A (\x. poly p (f x))" by (metis DERIV_continuous_on assms continuous_on_compose2 poly_DERIV subset_UNIV) text \Consequences of the derivative theorem above.\ lemma poly_differentiable[simp]: "(\x. poly p x) differentiable (at x)" for x :: real by (simp add: real_differentiable_def) (blast intro: poly_DERIV) lemma poly_IVT_pos: "a < b \ poly p a < 0 \ 0 < poly p b \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using IVT_objl [of "poly p" a 0 b] by (auto simp add: order_le_less) lemma poly_IVT_neg: "a < b \ 0 < poly p a \ poly p b < 0 \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using poly_IVT_pos [where p = "- p"] by simp lemma poly_IVT: "a < b \ poly p a * poly p b < 0 \ \x>a. x < b \ poly p x = 0" for p :: "real poly" by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) lemma poly_MVT: "a < b \ \x. a < x \ x < b \ poly p b - poly p a = (b - a) * poly (pderiv p) x" for a b :: real using MVT [of a b "poly p"] apply simp by (metis (full_types) DERIV_continuous_on DERIV_unique has_field_derivative_at_within poly_DERIV) lemma poly_MVT': fixes a b :: real assumes "{min a b..max a b} \ A" shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) x" proof (cases a b rule: linorder_cases) case less from poly_MVT[OF less, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) next case greater from poly_MVT[OF greater, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) qed (use assms in auto) lemma poly_pinfty_gt_lc: fixes p :: "real poly" assumes "lead_coeff p > 0" shows "\n. \ x \ n. poly p x \ lead_coeff p" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) from this(1) consider "a \ 0" "p = 0" | "p \ 0" by auto then show ?case proof cases case 1 then show ?thesis by auto next case 2 with pCons obtain n1 where gte_lcoeff: "\x\n1. lead_coeff p \ poly p x" by auto from pCons(3) \p \ 0\ have gt_0: "lead_coeff p > 0" by auto define n where "n = max n1 (1 + \a\ / lead_coeff p)" have "lead_coeff (pCons a p) \ poly (pCons a p) x" if "n \ x" for x proof - from gte_lcoeff that have "lead_coeff p \ poly p x" by (auto simp: n_def) with gt_0 have "\a\ / lead_coeff p \ \a\ / poly p x" and "poly p x > 0" by (auto intro: frac_le) with \n \ x\[unfolded n_def] have "x \ 1 + \a\ / poly p x" by auto with \lead_coeff p \ poly p x\ \poly p x > 0\ \p \ 0\ show "lead_coeff (pCons a p) \ poly (pCons a p) x" by (auto simp: field_simps) qed then show ?thesis by blast qed qed lemma lemma_order_pderiv1: "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons) lemma lemma_order_pderiv: fixes p :: "'a :: field_char_0 poly" assumes n: "0 < n" and pd: "pderiv p \ 0" and pe: "p = [:- a, 1:] ^ n * q" and nd: "\ [:- a, 1:] dvd q" shows "n = Suc (order a (pderiv p))" proof - from assms have "pderiv ([:- a, 1:] ^ n * q) \ 0" by auto from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" by (cases n) auto have *: "k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" for k l by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" proof (rule order_unique_lemma) show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" apply (subst lemma_order_pderiv1) apply (rule dvd_add) apply (metis dvdI dvd_mult2 power_Suc2) apply (metis dvd_smult dvd_triv_right) done show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" apply (subst lemma_order_pderiv1) apply (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) done qed then show ?thesis by (metis \n = Suc n'\ pe) qed lemma order_pderiv: "pderiv p \ 0 \ order a p \ 0 \ order a p = Suc (order a (pderiv p))" for p :: "'a::field_char_0 poly" apply (cases "p = 0") apply simp apply (drule_tac a = a and p = p in order_decomp) using neq0_conv apply (blast intro: lemma_order_pderiv) done lemma poly_squarefree_decomp_order: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and p: "p = q * d" and p': "pderiv p = e * d" and d: "d = r * p + s * pderiv p" shows "order a q = (if order a p = 0 then 0 else 1)" proof (rule classical) assume 1: "\ ?thesis" from \pderiv p \ 0\ have "p \ 0" by auto with p have "order a p = order a q + order a d" by (simp add: order_mult) with 1 have "order a p \ 0" by (auto split: if_splits) from \pderiv p \ 0\ \pderiv p = e * d\ have "order a (pderiv p) = order a e + order a d" by (simp add: order_mult) from \pderiv p \ 0\ \order a p \ 0\ have "order a p = Suc (order a (pderiv p))" by (rule order_pderiv) from \p \ 0\ \p = q * d\ have "d \ 0" by simp have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" apply (simp add: d) apply (rule dvd_add) apply (rule dvd_mult) apply (simp add: order_divides \p \ 0\ \order a p = Suc (order a (pderiv p))\) apply (rule dvd_mult) apply (simp add: order_divides) done with \d \ 0\ have "order a (pderiv p) \ order a d" by (simp add: order_divides) show ?thesis using \order a p = order a q + order a d\ and \order a (pderiv p) = order a e + order a d\ and \order a p = Suc (order a (pderiv p))\ and \order a (pderiv p) \ order a d\ by auto qed lemma poly_squarefree_decomp_order2: "pderiv p \ 0 \ p = q * d \ pderiv p = e * d \ d = r * p + s * pderiv p \ \a. order a q = (if order a p = 0 then 0 else 1)" for p :: "'a::field_char_0 poly" by (blast intro: poly_squarefree_decomp_order) lemma order_pderiv2: "pderiv p \ 0 \ order a p \ 0 \ order a (pderiv p) = n \ order a p = Suc n" for p :: "'a::field_char_0 poly" by (auto dest: order_pderiv) definition rsquarefree :: "'a::idom poly \ bool" where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h:]" for p :: "'a::{semidom,semiring_char_0} poly" by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) lemma rsquarefree_roots: "rsquarefree p \ (\a. \ (poly p a = 0 \ poly (pderiv p) a = 0))" for p :: "'a::field_char_0 poly" apply (simp add: rsquarefree_def) apply (case_tac "p = 0") apply simp apply simp apply (case_tac "pderiv p = 0") apply simp apply (drule pderiv_iszero, clarsimp) apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree) apply (force simp add: order_root order_pderiv2) done lemma poly_squarefree_decomp: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and "p = q * d" and "pderiv p = e * d" and "d = r * p + s * pderiv p" shows "rsquarefree q \ (\a. poly q a = 0 \ poly p a = 0)" proof - from \pderiv p \ 0\ have "p \ 0" by auto with \p = q * d\ have "q \ 0" by simp from assms have "\a. order a q = (if order a p = 0 then 0 else 1)" by (rule poly_squarefree_decomp_order2) with \p \ 0\ \q \ 0\ show ?thesis by (simp add: rsquarefree_def order_root) qed subsection \Algebraic numbers\ text \ Algebraic numbers can be defined in two equivalent ways: all real numbers that are roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry uses the rational definition, but we need the integer definition. The equivalence is obvious since any rational polynomial can be multiplied with the LCM of its coefficients, yielding an integer polynomial with the same roots. \ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" lemma algebraicI: "(\i. coeff p i \ \) \ p \ 0 \ poly p x = 0 \ algebraic x" unfolding algebraic_def by blast lemma algebraicE: assumes "algebraic x" obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast lemma algebraic_altdef: "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" for p :: "'a::field_char_0 poly" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i by (subst (asm) bchoice_iff) blast define cs' where "cs' = map (quotient_of \ f) (coeffs p)" define d where "d = Lcm (set (map snd cs'))" define p' where "p' = smult (of_int d) p" have "coeff p' n \ \" for n proof (cases "n \ degree p") case True define c where "c = coeff p n" define a where "a = fst (quotient_of (f (coeff p n)))" define b where "b = snd (quotient_of (f (coeff p n)))" have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric]) also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" by (simp add: of_rat_mult of_rat_divide) also from nz True have "b \ snd ` set cs'" by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc) then have "b dvd (a * d)" by (simp add: d_def) then have "of_int (a * d) / of_int b \ (\ :: rat set)" by (rule of_int_divide_in_Ints) then have "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto finally show ?thesis . next case False then show ?thesis by (auto simp: p'_def not_le coeff_eq_0) qed moreover have "set (map snd cs') \ {0<..}" unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) then have "d \ 0" unfolding d_def by (induct cs') simp_all with nz have "p' \ 0" by (simp add: p'_def) moreover from root have "poly p' x = 0" by (simp add: p'_def) ultimately show "algebraic x" unfolding algebraic_def by blast next assume "algebraic x" then obtain p where p: "coeff p i \ \" "poly p x = 0" "p \ 0" for i by (force simp: algebraic_def) moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp ultimately show "\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0" by auto qed subsection \Division of polynomials\ subsubsection \Division in general\ instantiation poly :: (idom_divide) idom_divide begin fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly" where "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in if False \ a * lc = cr then \ \\False \\ is only because of problem in function-package\ divide_poly_main lc (q + mon) (r - mon * d) d (dr - 1) n else 0)" | "divide_poly_main lc q r d dr 0 = q" definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where "divide_poly f g = (if g = 0 then 0 else divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" lemma divide_poly_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" shows "q' = q + r" using assms(3-) proof (induct n arbitrary: q r dr) case (Suc n) let ?rr = "d * r" let ?a = "coeff ?rr dr" let ?qq = "?a div lc" define b where [simp]: "b = monom ?qq n" let ?rrr = "d * (r - b)" let ?qqq = "q + b" note res = Suc(3) from Suc(4) have dr: "dr = n + degree d" by auto from d have lc: "lc \ 0" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by simp next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have c2: "coeff (b * d) dr = lc * coeff b n" . have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) have c1: "coeff (d * r) dr = lc * coeff r n" proof (cases "degree r = n") case True with Suc(2) show ?thesis unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) next case False from dr Suc(2) have "degree r \ n" by auto (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) with False have r_n: "degree r < n" by auto then have right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) have "coeff (d * r) dr = coeff (d * r) (degree d + n)" by (simp add: dr ac_simps) also from r_n have "\ = 0" by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) finally show ?thesis by (simp only: right) qed have c0: "coeff ?rrr dr = 0" and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 unfolding b_def coeff_monom coeff_smult c1 using lc by auto from res[unfolded divide_poly_main.simps[of lc q] Let_def] id have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" by (simp del: divide_poly_main.simps add: field_simps) note IH = Suc(1)[OF _ res] from Suc(4) have dr: "dr = n + degree d" by auto from Suc(2) have deg_rr: "degree ?rr \ dr" by auto have deg_bd: "degree (b * d) \ dr" unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed from IH[OF deg_rrr this] show ?case by simp next case 0 show ?case proof (cases "r = 0") case True with 0 show ?thesis by auto next case False from d False have "degree (d * r) = degree d + degree r" by (subst degree_mult_eq) auto with 0 d show ?thesis by auto qed qed lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" proof (induct n arbitrary: r d dr) case 0 then show ?case by simp next case Suc show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def by (simp add: Suc del: divide_poly_main.simps) qed lemma divide_poly: assumes g: "g \ 0" shows "(f * g) div g = (f :: 'a poly)" proof - have len: "length (coeffs f) = Suc (degree f)" if "f \ 0" for f :: "'a poly" using that unfolding degree_eq_length_coeffs by auto have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) = (f * g) div g" by (simp add: divide_poly_def Let_def ac_simps) note main = divide_poly_main[OF g refl le_refl this] have "(f * g) div g = 0 + f" proof (rule main, goal_cases) case 1 show ?case proof (cases "f = 0") case True with g show ?thesis by (auto simp: degree_eq_length_coeffs) next case False with g have fg: "g * f \ 0" by auto show ?thesis unfolding len[OF fg] len[OF g] by auto qed qed then show ?thesis by simp qed lemma divide_poly_0: "f div 0 = 0" for f :: "'a poly" by (simp add: divide_poly_def Let_def divide_poly_main_0) instance by standard (auto simp: divide_poly divide_poly_0) end instance poly :: (idom_divide) algebraic_semidom .. lemma div_const_poly_conv_map_poly: assumes "[:c:] dvd p" shows "p div [:c:] = map_poly (\x. x div c) p" proof (cases "c = 0") case True then show ?thesis by (auto intro!: poly_eqI simp: coeff_map_poly) next case False from assms obtain q where p: "p = [:c:] * q" by (rule dvdE) moreover { have "smult c q = [:c:] * q" by simp also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (use False in auto) finally have "smult c q div [:c:] = q" . } ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) qed lemma is_unit_monom_0: fixes a :: "'a::field" assumes "a \ 0" shows "is_unit (monom a 0)" proof from assms show "1 = monom a 0 * monom (inverse a) 0" by (simp add: mult_monom) qed lemma is_unit_triv: "a \ 0 \ is_unit [:a:]" for a :: "'a::field" by (simp add: is_unit_monom_0 monom_0 [symmetric]) lemma is_unit_iff_degree: fixes p :: "'a::field poly" assumes "p \ 0" shows "is_unit p \ degree p = 0" (is "?lhs \ ?rhs") proof assume ?rhs then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) with assms show ?lhs by (simp add: is_unit_triv) next assume ?lhs then obtain q where "q \ 0" "p * q = 1" .. then have "degree (p * q) = degree 1" by simp with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" by (simp add: degree_mult_eq) then show ?rhs by simp qed lemma is_unit_pCons_iff: "is_unit (pCons a p) \ p = 0 \ a \ 0" for p :: "'a::field poly" by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) lemma is_unit_monom_trival: "is_unit p \ monom (coeff p (degree p)) 0 = p" for p :: "'a::field poly" by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) lemma is_unit_const_poly_iff: "[:c:] dvd 1 \ c dvd 1" for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (auto simp: one_pCons) lemma is_unit_polyE: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" proof - from assms obtain q where "1 = p * q" by (rule dvdE) then have "p \ 0" and "q \ 0" by auto from \1 = p * q\ have "degree 1 = degree (p * q)" by simp also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" by (simp add: degree_mult_eq) finally have "degree p = 0" by simp with degree_eq_zeroE obtain c where c: "p = [:c:]" . with \p dvd 1\ have "c dvd 1" by (simp add: is_unit_const_poly_iff) with c show thesis .. qed lemma is_unit_polyE': fixes p :: "'a::field poly" assumes "is_unit p" obtains a where "p = monom a 0" and "a \ 0" proof - obtain a q where "p = pCons a q" by (cases p) with assms have "p = [:a:]" and "a \ 0" by (simp_all add: is_unit_pCons_iff) with that show thesis by (simp add: monom_0) qed lemma is_unit_poly_iff: "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) subsubsection \Pseudo-Division\ text \This part is by René Thiemann and Akihisa Yamada.\ fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly \ 'a poly" where "pseudo_divmod_main lc q r d dr (Suc n) = (let rr = smult lc r; qq = coeff r dr; rrr = rr - monom qq n * d; qqq = smult lc q + monom qq n in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" | "pseudo_divmod_main lc q r d dr 0 = (q,r)" definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where "pseudo_divmod p q \ if q = 0 then (0, p) else pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" using assms(3-) proof (induct n arbitrary: q r dr) case 0 then show ?case by auto next case (Suc n) let ?rr = "smult lc r" let ?qq = "coeff r dr" define b where [simp]: "b = monom ?qq n" let ?rrr = "?rr - b * d" let ?qqq = "smult lc q + b" note res = Suc(3) from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" by (simp del: pseudo_divmod_main.simps) from Suc(4) have dr: "dr = n + degree d" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by auto next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have "coeff (b * d) dr = lc * coeff b n" . moreover have "coeff ?rr dr = lc * coeff r dr" by simp ultimately have c0: "coeff ?rrr dr = 0" by auto from Suc(4) have dr: "dr = n + degree d" by auto have deg_rr: "degree ?rr \ dr" using Suc(2) degree_smult_le dual_order.trans by blast have deg_bd: "degree (b * d) \ dr" unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" using degree_diff_le[OF deg_rr deg_bd] by auto with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp then have "\a. ?rrr = [:a:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) from this obtain a where rrr: "?rrr = [:a:]" by auto show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed note IH = Suc(1)[OF deg_rrr res this] show ?case proof (intro conjI) from IH show "r' = 0 \ degree r' < degree d" by blast show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" unfolding IH[THEN conjunct2,symmetric] by (simp add: field_simps smult_add_right) qed qed lemma pseudo_divmod: assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) and "r = 0 \ degree r < degree g" (is ?B) proof - from *[unfolded pseudo_divmod_def Let_def] have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] from g have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) note main' = main[OF this] then show "r = 0 \ degree r < degree g" by auto show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" by (subst main'[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, cases "f = 0"; cases "coeffs g", use g in auto) qed definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" lemma snd_pseudo_divmod_main: "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" by (induct n arbitrary: q q' lc r d dr) (simp_all add: Let_def) definition pseudo_mod :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where "pseudo_mod f g = snd (pseudo_divmod f g)" lemma pseudo_mod: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes g: "g \ 0" shows "\a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" define a where "a = ?cge" from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" by (cases "pseudo_divmod f g") auto from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" by (auto simp: a_def) show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id show "\a q. a \ 0 \ smult a f = g * q + r" by auto qed lemma fst_pseudo_divmod_main_as_divide_poly_main: assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" proof (induct n arbitrary: q r dr) case 0 then show ?case by simp next case (Suc n) note lc0 = leading_coeff_neq_0[OF d, folded lc] then have "pseudo_divmod_main lc q r d dr (Suc n) = pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" by (simp add: Let_def ac_simps) also have "fst \ = divide_poly_main lc (smult (lc^n) (smult lc q + monom (coeff r dr) n)) (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) d (dr - 1) n" by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) also have "\ = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" unfolding smult_monom smult_distribs mult_smult_left[symmetric] using lc0 by (simp add: Let_def ac_simps) finally show ?case . qed subsubsection \Division in polynomials over fields\ lemma pseudo_divmod_field: fixes g :: "'a::field poly" assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" shows "f = g * smult (1/c) q + smult (1/c) r" proof - from leading_coeff_neq_0[OF g] have c0: "c \ 0" by (auto simp: c_def) from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r" by auto also have "smult (1 / c) \ = g * smult (1 / c) q + smult (1 / c) r" by (simp add: smult_add_right) finally show ?thesis using c0 by auto qed lemma divide_poly_main_field: fixes d :: "'a::field poly" assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1 / lc)^n) q) (smult ((1 / lc)^n) r) d dr n)" unfolding lc by (subst fst_pseudo_divmod_main_as_divide_poly_main) (auto simp: d power_one_over) lemma divide_poly_field: fixes f g :: "'a::field poly" defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" shows "f div g = fst (pseudo_divmod f' g)" proof (cases "g = 0") case True show ?thesis unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) next case False from leading_coeff_neq_0[OF False] have "degree f' = degree f" by (auto simp: f'_def) then show ?thesis using length_coeffs_degree[of f'] length_coeffs_degree[of f] unfolding divide_poly_def pseudo_divmod_def Let_def divide_poly_main_field[OF False] length_coeffs_degree[OF False] f'_def by force qed instantiation poly :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom begin definition unit_factor_poly :: "'a poly \ 'a poly" where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" definition normalize_poly :: "'a poly \ 'a poly" where "normalize p = p div [:unit_factor (lead_coeff p):]" instance proof fix p :: "'a poly" show "unit_factor p * normalize p = p" proof (cases "p = 0") case True then show ?thesis by (simp add: unit_factor_poly_def normalize_poly_def) next case False then have "lead_coeff p \ 0" by simp then have *: "unit_factor (lead_coeff p) \ 0" using unit_factor_is_unit [of "lead_coeff p"] by auto then have "unit_factor (lead_coeff p) dvd 1" by (auto intro: unit_factor_is_unit) then have **: "unit_factor (lead_coeff p) dvd c" for c by (rule dvd_trans) simp have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c proof - from ** obtain b where "c = unit_factor (lead_coeff p) * b" .. with False * show ?thesis by simp qed have "p div [:unit_factor (lead_coeff p):] = map_poly (\c. c div unit_factor (lead_coeff p)) p" by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **) then show ?thesis by (simp add: normalize_poly_def unit_factor_poly_def smult_conv_map_poly map_poly_map_poly o_def ***) qed next fix p :: "'a poly" assume "is_unit p" then obtain c where p: "p = [:c:]" "c dvd 1" by (auto simp: is_unit_poly_iff) then show "unit_factor p = p" by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor) next fix p :: "'a poly" assume "p \ 0" then show "is_unit (unit_factor p)" by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit) next fix a b :: "'a poly" assume "is_unit a" thus "unit_factor (a * b) = a * unit_factor b" by (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult elim!: is_unit_polyE) qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) end instance poly :: ("{semidom_divide_unit_factor,idom_divide,normalization_semidom_multiplicative}") normalization_semidom_multiplicative by intro_classes (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult) lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\x. x div unit_factor (lead_coeff p)) p" proof - have "[:unit_factor (lead_coeff p):] dvd p" by (metis unit_factor_poly_def unit_factor_self) then show ?thesis by (simp add: normalize_poly_def div_const_poly_conv_map_poly) qed lemma coeff_normalize [simp]: "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)" by (simp add: normalize_poly_eq_map_poly coeff_map_poly) class field_unit_factor = field + unit_factor + assumes unit_factor_field [simp]: "unit_factor = id" begin subclass semidom_divide_unit_factor proof fix a assume "a \ 0" then have "1 = a * inverse a" by simp then have "a dvd 1" .. then show "unit_factor a dvd 1" by simp qed simp_all end lemma unit_factor_pCons: "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] else unit_factor p)" by (simp add: unit_factor_poly_def) lemma normalize_monom [simp]: "normalize (monom a n) = monom (normalize a) n" by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_eq_map_poly degree_monom_eq) lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]" by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" by (simp add: normalize_poly_eq_map_poly map_poly_pCons) lemma normalize_smult: fixes c :: "'a :: {normalization_semidom_multiplicative, idom_divide}" shows "normalize (smult c p) = smult (normalize c) (normalize p)" proof - have "smult c p = [:c:] * p" by simp also have "normalize \ = smult (normalize c) (normalize p)" by (subst normalize_mult) (simp add: normalize_const_poly) finally show ?thesis . qed inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y \ x = q * y + r \ eucl_rel_poly x y (q, r)" lemma eucl_rel_poly_iff: "eucl_rel_poly x y (q, r) \ x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" by (auto elim: eucl_rel_poly.cases intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) lemma eucl_rel_poly_0: "eucl_rel_poly 0 y (0, 0)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_pCons: assumes rel: "eucl_rel_poly x y (q, r)" assumes y: "y \ 0" assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" (is "eucl_rel_poly ?x y (?q, ?r)") proof - from assms have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" by (simp_all add: eucl_rel_poly_iff) from b x have "?x = ?q * y + ?r" by simp moreover have "?r = 0 \ degree ?r < degree y" proof (rule eq_zero_or_degree_less) show "degree ?r \ degree y" proof (rule degree_diff_le) from r show "degree (pCons a r) \ degree y" by auto show "degree (smult b y) \ degree y" by (rule degree_smult_le) qed from \y \ 0\ show "coeff ?r (degree y) = 0" by (simp add: b) qed ultimately show ?thesis unfolding eucl_rel_poly_iff using \y \ 0\ by simp qed lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" apply (cases "y = 0") apply (fast intro!: eucl_rel_poly_by_0) apply (induct x) apply (fast intro!: eucl_rel_poly_0) apply (fast intro!: eucl_rel_poly_pCons) done lemma eucl_rel_poly_unique: assumes 1: "eucl_rel_poly x y (q1, r1)" assumes 2: "eucl_rel_poly x y (q2, r2)" shows "q1 = q2 \ r1 = r2" proof (cases "y = 0") assume "y = 0" with assms show ?thesis by (simp add: eucl_rel_poly_iff) next assume [simp]: "y \ 0" from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" unfolding eucl_rel_poly_iff by simp_all from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" unfolding eucl_rel_poly_iff by simp_all from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" by (simp add: algebra_simps) from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" by (auto intro: degree_diff_less) show "q1 = q2 \ r1 = r2" proof (rule classical) assume "\ ?thesis" with q3 have "q1 \ q2" and "r1 \ r2" by auto with r3 have "degree (r2 - r1) < degree y" by simp also have "degree y \ degree (q1 - q2) + degree y" by simp also from \q1 \ q2\ have "\ = degree ((q1 - q2) * y)" by (simp add: degree_mult_eq) also from q3 have "\ = degree (r2 - r1)" by simp finally have "degree (r2 - r1) < degree (r2 - r1)" . then show ?thesis by simp qed qed lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] instantiation poly :: (field) semidom_modulo begin definition modulo_poly :: "'a poly \ 'a poly \ 'a poly" where mod_poly_def: "f mod g = (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)" instance proof fix x y :: "'a poly" show "x div y * y + x mod y = x" proof (cases "y = 0") case True then show ?thesis by (simp add: divide_poly_0 mod_poly_def) next case False then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y = (x div y, x mod y)" by (simp add: divide_poly_field mod_poly_def pseudo_mod_def) with False pseudo_divmod [OF False this] show ?thesis by (simp add: power_mult_distrib [symmetric] ac_simps) qed qed end lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)" unfolding eucl_rel_poly_iff proof show "x = x div y * y + x mod y" by (simp add: div_mult_mod_eq) show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" proof (cases "y = 0") case True then show ?thesis by auto next case False with pseudo_mod[OF this] show ?thesis by (simp add: mod_poly_def) qed qed lemma div_poly_eq: "eucl_rel_poly x y (q, r) \ x div y = q" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly]) lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \ x mod y = r" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly]) instance poly :: (field) idom_modulo .. lemma div_pCons_eq: "pCons a p div q = (if q = 0 then 0 else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: div_poly_eq) lemma mod_pCons_eq: "pCons a p mod q = (if q = 0 then pCons a p else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: mod_poly_eq) lemma div_mod_fold_coeffs: "(p div q, p mod q) = (if q = 0 then (0, p) else fold_coeffs (\a (s, r). let b = coeff (pCons a r) (degree q) / coeff q (degree q) in (pCons b s, pCons a r - smult b q)) p (0, 0))" by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def) lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" using degree_mod_less[of b a] by auto lemma div_poly_less: fixes x :: "'a::field poly" assumes "degree x < degree y" shows "x div y = 0" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x div y = 0" by (rule div_poly_eq) qed lemma mod_poly_less: assumes "degree x < degree y" shows "x mod y = x" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x mod y = x" by (rule mod_poly_eq) qed lemma eucl_rel_poly_smult_left: "eucl_rel_poly x y (q, r) \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" by (simp add: eucl_rel_poly_iff smult_add_right) lemma div_smult_left: "(smult a x) div y = smult a (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma poly_div_minus_left [simp]: "(- x) div y = - (x div y)" for x y :: "'a::field poly" using div_smult_left [of "- 1::'a"] by simp lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)" for x y :: "'a::field poly" using mod_smult_left [of "- 1::'a"] by simp lemma eucl_rel_poly_add_left: assumes "eucl_rel_poly x y (q, r)" assumes "eucl_rel_poly x' y (q', r')" shows "eucl_rel_poly (x + x') y (q + q', r + r')" using assms unfolding eucl_rel_poly_iff by (auto simp: algebra_simps degree_add_less) lemma poly_div_add_left: "(x + y) div z = x div z + y div z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule div_poly_eq) lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule mod_poly_eq) lemma poly_div_diff_left: "(x - y) div z = x div z - y div z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) lemma eucl_rel_poly_smult_right: "a \ 0 \ eucl_rel_poly x y (q, r) \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" by (simp add: eucl_rel_poly_iff) lemma div_smult_right: "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma poly_div_minus_right [simp]: "x div (- y) = - (x div y)" for x y :: "'a::field poly" using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y" for x y :: "'a::field poly" using mod_smult_right [of "- 1::'a"] by simp lemma eucl_rel_poly_mult: "eucl_rel_poly x y (q, r) \ eucl_rel_poly q z (q', r') \ eucl_rel_poly x (y * z) (q', y * r' + r)" apply (cases "z = 0", simp add: eucl_rel_poly_iff) apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff) apply (cases "r = 0") apply (cases "r' = 0") apply (simp add: eucl_rel_poly_iff) apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq) apply (cases "r' = 0") apply (simp add: eucl_rel_poly_iff degree_mult_eq) apply (simp add: eucl_rel_poly_iff field_simps) apply (simp add: degree_mult_eq degree_add_less) done lemma poly_div_mult_right: "x div (y * z) = (x div y) div z" for x y z :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y" for x y z :: "'a::field poly" by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma mod_pCons: fixes a :: "'a::field" and x y :: "'a::field poly" assumes y: "y \ 0" defines "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" shows "(pCons a x) mod y = pCons a (x mod y) - smult b y" unfolding b_def by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) subsubsection \List-based versions for fast implementation\ (* Subsection by: Sebastiaan Joosten René Thiemann Akihisa Yamada *) fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" | "minus_poly_rev_list xs [] = xs" | "minus_poly_rev_list [] (y # ys) = []" fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "pseudo_divmod_main_list lc q r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; qqq = cCons a (map ((*) lc) q); rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_divmod_main_list lc qqq rrr d n)" | "pseudo_divmod_main_list lc q r d 0 = (q, r)" fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ nat \ 'a list" where "pseudo_mod_main_list lc r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_mod_main_list lc rrr d n)" | "pseudo_mod_main_list lc r d 0 = r" fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "divmod_poly_one_main_list q r d (Suc n) = (let a = hd r; qqq = cCons a q; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in divmod_poly_one_main_list qqq rr d n)" | "divmod_poly_one_main_list q r d 0 = (q, r)" fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ nat \ 'a list" where "mod_poly_one_main_list r d (Suc n) = (let a = hd r; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in mod_poly_one_main_list rr d n)" | "mod_poly_one_main_list r d 0 = r" definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where "pseudo_divmod_list p q = (if q = [] then ([], p) else (let rq = rev q; (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in (qu, rev re)))" definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where "pseudo_mod_list p q = (if q = [] then p else (let rq = rev q; re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in rev re))" lemma minus_zero_does_nothing: "minus_poly_rev_list x (map ((*) 0) y) = x" for x :: "'a::ring list" by (induct x y rule: minus_poly_rev_list.induct) auto lemma length_minus_poly_rev_list [simp]: "length (minus_poly_rev_list xs ys) = length xs" by (induct xs ys rule: minus_poly_rev_list.induct) auto lemma if_0_minus_poly_rev_list: "(if a = 0 then x else minus_poly_rev_list x (map ((*) a) y)) = minus_poly_rev_list x (map ((*) a) y)" for a :: "'a::ring" by(cases "a = 0") (simp_all add: minus_zero_does_nothing) lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b" for a :: "'a::comm_semiring_1 list" by (induct a) (auto simp: monom_0 monom_Suc) lemma minus_poly_rev_list: "length p \ length q \ Poly (rev (minus_poly_rev_list (rev p) (rev q))) = Poly p - monom 1 (length p - length q) * Poly q" for p q :: "'a :: comm_ring_1 list" proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) case (1 x xs y ys) then have "length (rev q) \ length (rev p)" by simp from this[folded 1(2,3)] have ys_xs: "length ys \ length xs" by simp then have *: "Poly (rev (minus_poly_rev_list xs ys)) = Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" by (subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev]) auto have "Poly p - monom 1 (length p - length q) * Poly q = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" by simp also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" unfolding 1(2,3) by simp also from ys_xs have "\ = Poly (rev xs) + monom x (length xs) - (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" by (simp add: Poly_append distrib_left mult_monom smult_monom) also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" unfolding * diff_monom[symmetric] by simp finally show ?case by (simp add: 1(2,3)[symmetric] smult_monom Poly_append) qed auto lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" using smult_monom [of a _ n] by (metis mult_smult_left) lemma head_minus_poly_rev_list: "length d \ length r \ d \ [] \ hd (minus_poly_rev_list (map ((*) (last d)) r) (map ((*) (hd r)) (rev d))) = 0" for d r :: "'a::comm_ring list" proof (induct r) case Nil then show ?case by simp next case (Cons a rs) then show ?case by (cases "rev d") (simp_all add: ac_simps) qed lemma Poly_map: "Poly (map ((*) a) p) = smult a (Poly p)" proof (induct p) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases "Poly xs = 0") auto qed lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) lemma pseudo_divmod_main_list_invar: assumes leading_nonzero: "last d \ 0" and lc: "last d = lc" and "d \ []" and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q', rev r')" and "n = 1 + length r - length d" shows "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = (Poly q', Poly r')" using assms(4-) proof (induct n arbitrary: r q) case (Suc n) from Suc.prems have *: "\ Suc (length r) \ length d" by simp with \d \ []\ have "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce let ?a = "(hd (rev r))" let ?rr = "map ((*) lc) (rev r)" let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map ((*) ?a) (rev d))))" let ?qq = "cCons ?a (map ((*) lc) q)" from * Suc(3) have n: "n = (1 + length r - length d - 1)" by simp from * have rr_val:"(length ?rrr) = (length r - 1)" by auto with \r \ []\ * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" by auto from * have id: "Suc (length r) - length d = Suc (length r - length d)" by auto from Suc.prems * have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" by (simp add: Let_def if_0_minus_poly_rev_list id) with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" by auto from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)" using Suc_diff_le not_less_eq_eq by blast from Suc(3) \r \ []\ have n_ok : "n = 1 + (length ?rrr) - length d" by simp have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp have hd_rev: "coeff (Poly r) (length r - Suc 0) = hd (rev r)" using last_coeff_is_hd[OF \r \ []\] by simp show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def proof (rule cong[OF _ _ refl], goal_cases) case 1 show ?case by (simp add: monom_Suc hd_rev[symmetric] smult_monom Poly_map) next case 2 show ?case proof (subst Poly_on_rev_starting_with_0, goal_cases) show "hd (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))) = 0" by (fold lc, subst head_minus_poly_rev_list, insert * \d \ []\, auto) from * have "length d \ length r" by simp then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = Poly (rev (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))))" by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] minus_poly_rev_list) qed qed simp qed simp lemma pseudo_divmod_impl [code]: "pseudo_divmod f g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" for f g :: "'a::comm_ring_1 poly" proof (cases "g = 0") case False then have "last (coeffs g) \ 0" and "last (coeffs g) = lead_coeff g" and "coeffs g \ []" by (simp_all add: last_coeffs_eq_coeff_degree) moreover obtain q r where qr: "pseudo_divmod_main_list (last (coeffs g)) (rev []) (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, rev (rev r))" by force ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))" by (subst pseudo_divmod_main_list_invar [symmetric]) auto moreover have "pseudo_divmod_main_list (hd (rev (coeffs g))) [] (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" using qr hd_rev [OF \coeffs g \ []\] by simp ultimately show ?thesis by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def) next case True then show ?thesis by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def) qed lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q xs ys n) = pseudo_mod_main_list l xs ys n" by (induct n arbitrary: l q xs ys) (auto simp: Let_def) lemma pseudo_mod_impl[code]: "pseudo_mod f g = poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" proof - have snd_case: "\f g p. snd ((\(x,y). (f x, g y)) p) = g (snd p)" by auto show ?thesis unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def pseudo_mod_list_def Let_def by (simp add: snd_case pseudo_mod_main_list) qed subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ (p div q, p mod q) = (r, s)" by (metis eucl_rel_poly eucl_rel_poly_unique) lemma pdivmod_via_pseudo_divmod: "(f div g, f mod g) = (if g = 0 then (0, f) else let ilc = inverse (coeff g (degree g)); h = smult ilc g; (q,r) = pseudo_divmod f h in (smult ilc q, r))" (is "?l = ?r") proof (cases "g = 0") case True then show ?thesis by simp next case False define lc where "lc = inverse (coeff g (degree g))" define h where "h = smult lc g" from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" by (auto simp: h_def lc_def) then have h0: "h \ 0" by auto obtain q r where p: "pseudo_divmod f h = (q, r)" by force from False have id: "?r = (smult lc q, r)" by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p) from pseudo_divmod[OF h0 p, unfolded h1] have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto from f r h0 have "eucl_rel_poly f h (q, r)" by (auto simp: eucl_rel_poly_iff) then have "(f div h, f mod h) = (q, r)" by (simp add: pdivmod_pdivmodrel) with lc have "(f div g, f mod g) = (smult lc q, r)" by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc]) with id show ?thesis by auto qed lemma pdivmod_via_pseudo_divmod_list: "(f div g, f mod g) = (let cg = coeffs g in if cg = [] then (0, f) else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; (q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) ilc) q), poly_of_list (rev r)))" proof - note d = pdivmod_via_pseudo_divmod pseudo_divmod_impl pseudo_divmod_list_def show ?thesis proof (cases "g = 0") case True with d show ?thesis by auto next case False define ilc where "ilc = inverse (coeff g (degree g))" from False have ilc: "ilc \ 0" by (auto simp: ilc_def) with False have id: "g = 0 \ False" "coeffs g = [] \ False" "last (coeffs g) = coeff g (degree g)" "coeffs (smult ilc g) = [] \ False" by (auto simp: last_coeffs_eq_coeff_degree) have id2: "hd (rev (coeffs (smult ilc g))) = 1" by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" "rev (coeffs (smult ilc g)) = rev (map ((*) ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map ((*) ilc) (coeffs g))) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by force show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 unfolding id3 pair map_prod_def split by (auto simp: Poly_map) qed qed lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" proof (intro ext, goal_cases) case (1 q r d n) have *: "map ((*) 1) xs = xs" for xs :: "'a list" by (induct xs) auto show ?case by (induct n arbitrary: q r d) (auto simp: * Let_def) qed fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list \ nat \ 'a list" where "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" | "divide_poly_main_list lc q r d 0 = q" lemma divide_poly_main_list_simp [simp]: "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r; a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" by (simp add: Let_def minus_zero_does_nothing) declare divide_poly_main_list.simps(1)[simp del] definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where "divide_poly_list f g = (let cg = coeffs g in if cg = [] then g else let cf = coeffs f; cgr = rev cg in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" lemmas pdivmod_via_divmod_list = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" by (induct n arbitrary: q r d) (auto simp: Let_def) lemma mod_poly_code [code]: "f mod g = (let cg = coeffs g in if cg = [] then f else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) in poly_of_list (rev r))" (is "_ = ?rhs") proof - have "snd (f div g, f mod g) = ?rhs" unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil] by (auto split: prod.splits) then show ?thesis by simp qed definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where "div_field_poly_impl f g = (let cg = coeffs g in if cg = [] then 0 else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) in poly_of_list ((map ((*) ilc) q)))" text \We do not declare the following lemma as code equation, since then polynomial division on non-fields will no longer be executable. However, a code-unfold is possible, since \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ lemma div_field_poly_impl[code_unfold]: "(div) = div_field_poly_impl" proof (intro ext) fix f g :: "'a poly" have "fst (f div g, f mod g) = div_field_poly_impl f g" unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) then show "f div g = div_field_poly_impl f g" by simp qed lemma divide_poly_main_list: assumes lc0: "lc \ 0" and lc: "last d = lc" and d: "d \ []" and "n = (1 + length r - length d)" shows "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" using assms(4-) proof (induct "n" arbitrary: r q) case (Suc n) from Suc.prems have ifCond: "\ Suc (length r) \ length d" by simp with d have r: "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases) auto from d lc obtain dd where d: "d = dd @ [lc]" by (cases d rule: rev_cases) auto from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) from ifCond have len: "length dd \ length rr" by (simp add: r d) show ?case proof (cases "lcr div lc * lc = lcr") case False with r d show ?thesis unfolding Suc(2)[symmetric] by (auto simp add: Let_def nth_default_append) next case True with r d have id: "?thesis \ Poly (divide_poly_main_list lc (cCons (lcr div lc) q) (rev (rev (minus_poly_rev_list (rev rr) (rev (map ((*) (lcr div lc)) dd))))) (rev d) n) = divide_poly_main lc (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) (Poly r - monom (lcr div lc) n * Poly d) (Poly d) (length rr - 1) n" by (cases r rule: rev_cases; cases "d" rule: rev_cases) (auto simp add: Let_def rev_map nth_default_append) have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp show ?thesis unfolding id proof (subst Suc(1), simp add: n, subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) case 2 have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" by (simp add: mult_monom len True) then show ?case unfolding r d Poly_append n ring_distribs by (auto simp: Poly_map smult_monom smult_monom_mult) qed (auto simp: len monom_Suc smult_monom) qed qed simp lemma divide_poly_list[code]: "f div g = divide_poly_list f g" proof - note d = divide_poly_def divide_poly_list_def show ?thesis proof (cases "g = 0") case True show ?thesis by (auto simp: d True) next case False then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases) auto with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto from cg False have lcg: "coeff g (degree g) = lcg" using last_coeffs_eq_coeff_degree last_snoc by force with False have "lcg \ 0" by auto from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = g" by auto show ?thesis unfolding d cg Let_def id if_False poly_of_list_def by (subst divide_poly_main_list, insert False cg \lcg \ 0\) (auto simp: lcg ltp, simp add: degree_eq_length_coeffs) qed qed subsection \Primality and irreducibility in polynomial rings\ lemma prod_mset_const_poly: "(\x\#A. [:f x:]) = [:prod_mset (image_mset f A):]" by (induct A) (simp_all add: ac_simps) lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ irreducible c" proof assume A: "irreducible c" show "irreducible [:c:]" proof (rule irreducibleI) fix a b assume ab: "[:c:] = a * b" hence "degree [:c:] = degree (a * b)" by (simp only: ) also from A ab have "a \ 0" "b \ 0" by auto hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq) finally have "degree a = 0" "degree b = 0" by auto then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE) from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: ) hence "c = a' * b'" by (simp add: ab' mult_ac) from A and this have "a' dvd 1 \ b' dvd 1" by (rule irreducibleD) with ab' show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed (insert A, auto simp: irreducible_def is_unit_poly_iff) next assume A: "irreducible [:c:]" then have "c \ 0" and "\ c dvd 1" by (auto simp add: irreducible_def is_unit_const_poly_iff) then show "irreducible c" proof (rule irreducibleI) fix a b assume ab: "c = a * b" hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:a:] dvd 1 \ [:b:] dvd 1" by (rule irreducibleD) then show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed qed lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast - have B: "\i. \c dvd coeff b i \ i \ degree b" + have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i - unfolding m_def by (rule Greatest_le_nat[OF that B]) + unfolding m_def by (metis (mono_tags, lifting) B Greatest_le_nat that) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) subsection \Content and primitive part of a polynomial\ definition content :: "'a::semiring_gcd poly \ 'a" where "content p = gcd_list (coeffs p)" lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0" by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps) lemma content_0 [simp]: "content 0 = 0" by (simp add: content_def) lemma content_1 [simp]: "content 1 = 1" by (simp add: content_def) lemma content_const [simp]: "content [:c:] = normalize c" by (simp add: content_def cCons_def) lemma const_poly_dvd_iff_dvd_content: "[:c:] dvd p \ c dvd content p" for c :: "'a::semiring_gcd" proof (cases "p = 0") case True then show ?thesis by simp next case False have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) also have "\ \ (\a\set (coeffs p). c dvd a)" proof safe fix n :: nat assume "\a\set (coeffs p). c dvd a" then show "c dvd coeff p n" by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) also have "\ \ c dvd content p" by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff) finally show ?thesis . qed lemma content_dvd [simp]: "[:content p:] dvd p" by (subst const_poly_dvd_iff_dvd_content) simp_all lemma content_dvd_coeff [simp]: "content p dvd coeff p n" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (cases "n \ degree p") (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd) qed lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" by (simp add: content_def Gcd_fin_dvd) lemma normalize_content [simp]: "normalize (content p) = content p" by (simp add: content_def) lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" proof assume "is_unit (content p)" then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) then show "content p = 1" by simp qed auto lemma content_smult [simp]: fixes c :: "'a :: {normalization_semidom_multiplicative, semiring_gcd}" shows "content (smult c p) = normalize c * content p" by (simp add: content_def coeffs_smult Gcd_fin_mult normalize_mult) lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" by (auto simp: content_def simp: poly_eq_iff coeffs_def) definition primitive_part :: "'a :: semiring_gcd poly \ 'a poly" where "primitive_part p = map_poly (\x. x div content p) p" lemma primitive_part_0 [simp]: "primitive_part 0 = 0" by (simp add: primitive_part_def) lemma content_times_primitive_part [simp]: "smult (content p) (primitive_part p) = p" for p :: "'a :: semiring_gcd poly" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis unfolding primitive_part_def by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs intro: map_poly_idI) qed lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" proof (cases "p = 0") case True then show ?thesis by simp next case False then have "primitive_part p = map_poly (\x. x div content p) p" by (simp add: primitive_part_def) also from False have "\ = 0 \ p = 0" by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) finally show ?thesis using False by simp qed lemma content_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" assumes "p \ 0" shows "content (primitive_part p) = 1" proof - have "p = smult (content p) (primitive_part p)" by simp also have "content \ = content (primitive_part p) * content p" by (simp del: content_times_primitive_part add: ac_simps) finally have "1 * content p = content (primitive_part p) * content p" by simp then have "1 * content p div content p = content (primitive_part p) * content p div content p" by simp with assms show ?thesis by simp qed lemma content_decompose: obtains p' :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" where "p = smult (content p) p'" "content p' = 1" proof (cases "p = 0") case True then have "p = smult (content p) 1" "content 1 = 1" by simp_all then show ?thesis .. next case False then have "p = smult (content p) (primitive_part p)" "content (primitive_part p) = 1" by simp_all then show ?thesis .. qed lemma content_dvd_contentI [intro]: "p dvd q \ content p dvd content q" using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" by (simp add: primitive_part_def map_poly_pCons) lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" by (auto simp: primitive_part_def) lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "p = smult (content p) (primitive_part p)" by simp also from False have "degree \ = degree (primitive_part p)" by (subst degree_smult_eq) simp_all finally show ?thesis .. qed lemma smult_content_normalize_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd, idom_divide} poly" shows "smult (content p) (normalize (primitive_part p)) = normalize p" proof - have "smult (content p) (normalize (primitive_part p)) = normalize ([:content p:] * primitive_part p)" by (subst normalize_mult) (simp_all add: normalize_const_poly) also have "[:content p:] * primitive_part p = p" by simp finally show ?thesis . qed context begin private lemma content_1_mult: fixes f g :: "'a :: {semiring_gcd, factorial_semiring} poly" assumes "content f = 1" "content g = 1" shows "content (f * g) = 1" proof (cases "f * g = 0") case False from assms have "f \ 0" "g \ 0" by auto hence "f * g \ 0" by auto { assume "\is_unit (content (f * g))" with False have "\p. p dvd content (f * g) \ prime p" by (intro prime_divisor_exists) simp_all then obtain p where "p dvd content (f * g)" "prime p" by blast from \p dvd content (f * g)\ have "[:p:] dvd f * g" by (simp add: const_poly_dvd_iff_dvd_content) moreover from \prime p\ have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly) ultimately have "[:p:] dvd f \ [:p:] dvd g" by (simp add: prime_elem_dvd_mult_iff) with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content) with \prime p\ have False by simp } hence "is_unit (content (f * g))" by blast hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content) thus ?thesis by simp qed (insert assms, auto) lemma content_mult: fixes p q :: "'a :: {factorial_semiring, semiring_gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = content p * content q" proof (cases "p * q = 0") case False then have "p \ 0" and "q \ 0" by simp_all then have *: "content (primitive_part p * primitive_part q) = 1" by (auto intro: content_1_mult) have "p * q = smult (content p) (primitive_part p) * smult (content q) (primitive_part q)" by simp also have "\ = smult (content p * content q) (primitive_part p * primitive_part q)" by (metis mult.commute mult_smult_right smult_smult) with * show ?thesis by (simp add: normalize_mult) next case True then show ?thesis by auto qed end lemma primitive_part_mult: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (p * q) = primitive_part p * primitive_part q" proof - have "primitive_part (p * q) = p * q div [:content (p * q):]" by (simp add: primitive_part_def div_const_poly_conv_map_poly) also have "\ = (p div [:content p:]) * (q div [:content q:])" by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac) also have "\ = primitive_part p * primitive_part q" by (simp add: primitive_part_def div_const_poly_conv_map_poly) finally show ?thesis . qed lemma primitive_part_smult: fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)" proof - have "smult a p = [:a:] * p" by simp also have "primitive_part \ = smult (unit_factor a) (primitive_part p)" by (subst primitive_part_mult) simp_all finally show ?thesis . qed lemma primitive_part_dvd_primitive_partI [intro]: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "p dvd q \ primitive_part p dvd primitive_part q" by (auto elim!: dvdE simp: primitive_part_mult) lemma content_prod_mset: fixes A :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly multiset" shows "content (prod_mset A) = prod_mset (image_mset content A)" by (induction A) (simp_all add: content_mult mult_ac) lemma content_prod_eq_1_iff: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = 1 \ content p = 1 \ content q = 1" proof safe assume A: "content (p * q) = 1" { fix p q :: "'a poly" assume "content p * content q = 1" hence "1 = content p * content q" by simp hence "content p dvd 1" by (rule dvdI) hence "content p = 1" by simp } note B = this from A B[of p q] B [of q p] show "content p = 1" "content q = 1" by (simp_all add: content_mult mult_ac) qed (auto simp: content_mult) no_notation cCons (infixr "##" 65) end diff --git a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy --- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy +++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy @@ -1,869 +1,869 @@ (* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy Author: Manuel Eberl *) section \Polynomials, fractions and rings\ theory Polynomial_Factorial imports Complex_Main Polynomial Normalized_Fraction begin subsection \Lifting elements into the field of fractions\ definition to_fract :: "'a :: idom \ 'a fract" where "to_fract x = Fract x 1" \ \FIXME: more idiomatic name, abbreviation\ lemma to_fract_0 [simp]: "to_fract 0 = 0" by (simp add: to_fract_def eq_fract Zero_fract_def) lemma to_fract_1 [simp]: "to_fract 1 = 1" by (simp add: to_fract_def eq_fract One_fract_def) lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y" by (simp add: to_fract_def) lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y" by (simp add: to_fract_def) lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x" by (simp add: to_fract_def) lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y" by (simp add: to_fract_def) lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \ x = y" by (simp add: to_fract_def eq_fract) lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \ x = 0" by (simp add: to_fract_def Zero_fract_def eq_fract) lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \ 0" by transfer simp lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x" by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp) lemma to_fract_quot_of_fract: assumes "snd (quot_of_fract x) = 1" shows "to_fract (fst (quot_of_fract x)) = x" proof - have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp also note assms finally show ?thesis by (simp add: to_fract_def) qed lemma snd_quot_of_fract_Fract_whole: assumes "y dvd x" shows "snd (quot_of_fract (Fract x y)) = 1" using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd) lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b" by (simp add: to_fract_def) lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)" unfolding to_fract_def by transfer (simp add: normalize_quot_def) lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \ x = 0" by transfer simp lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1" unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all lemma coprime_quot_of_fract: "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))" by transfer (simp add: coprime_normalize_quot) lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1" using quot_of_fract_in_normalized_fracts[of x] by (simp add: normalized_fracts_def case_prod_unfold) lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \ normalize x = x" by (subst (2) normalize_mult_unit_factor [symmetric, of x]) (simp del: normalize_mult_unit_factor) lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)" by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) subsection \Lifting polynomial coefficients to the field of fractions\ abbreviation (input) fract_poly where "fract_poly \ map_poly to_fract" abbreviation (input) unfract_poly where "unfract_poly \ map_poly (fst \ quot_of_fract)" lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)" by (simp add: smult_conv_map_poly map_poly_map_poly o_def) lemma fract_poly_0 [simp]: "fract_poly 0 = 0" by (simp add: poly_eqI coeff_map_poly) lemma fract_poly_1 [simp]: "fract_poly 1 = 1" by (simp add: map_poly_pCons) lemma fract_poly_add [simp]: "fract_poly (p + q) = fract_poly p + fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma fract_poly_diff [simp]: "fract_poly (p - q) = fract_poly p - fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\x. to_fract (f x)) A" by (cases "finite A", induction A rule: finite_induct) simp_all lemma fract_poly_mult [simp]: "fract_poly (p * q) = fract_poly p * fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult) lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \ p = q" by (auto simp: poly_eq_iff coeff_map_poly) lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \ p = 0" using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff) lemma fract_poly_dvd: "p dvd q \ fract_poly p dvd fract_poly q" by (auto elim!: dvdE) lemma prod_mset_fract_poly: "(\x\#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))" by (induct A) (simp_all add: ac_simps) lemma is_unit_fract_poly_iff: "p dvd 1 \ fract_poly p dvd 1 \ content p = 1" proof safe assume A: "p dvd 1" with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)" by simp from A show "content p = 1" by (auto simp: is_unit_poly_iff normalize_1_iff) next assume A: "fract_poly p dvd 1" and B: "content p = 1" from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff) { fix n :: nat assume "n > 0" have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly) also note c also from \n > 0\ have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits) finally have "coeff p n = 0" by simp } hence "degree p \ 0" by (intro degree_le) simp_all with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE) qed lemma fract_poly_is_unit: "p dvd 1 \ fract_poly p dvd 1" using fract_poly_dvd[of p 1] by simp lemma fract_poly_smult_eqE: fixes c :: "'a :: {idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract" assumes "fract_poly p = smult c (fract_poly q)" obtains a b where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a" proof - define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)" have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)" by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms) hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff) hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff) moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b" by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric]) ultimately show ?thesis by (intro that[of a b]) qed subsection \Fractional content\ abbreviation (input) Lcm_coeff_denoms :: "'a :: {semiring_Gcd,idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract poly \ 'a" where "Lcm_coeff_denoms p \ Lcm (snd ` quot_of_fract ` set (coeffs p))" definition fract_content :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a fract" where "fract_content p = (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" definition primitive_part_fract :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a poly" where "primitive_part_fract p = primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))" lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0" by (simp add: primitive_part_fract_def) lemma fract_content_eq_0_iff [simp]: "fract_content p = 0 \ p = 0" unfolding fract_content_def Let_def Zero_fract_def by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff) lemma content_primitive_part_fract [simp]: fixes p :: "'a :: {semiring_gcd_mult_normalize, factorial_semiring, ring_gcd, semiring_Gcd,idom_divide} fract poly" shows "p \ 0 \ content (primitive_part_fract p) = 1" unfolding primitive_part_fract_def by (rule content_primitive_part) (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff) lemma content_times_primitive_part_fract: "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p" proof - define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)" have "fract_poly p' = map_poly (to_fract \ fst \ quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)" unfolding primitive_part_fract_def p'_def by (subst map_poly_map_poly) (simp_all add: o_assoc) also have "\ = smult (to_fract (Lcm_coeff_denoms p)) p" proof (intro map_poly_idI, unfold o_apply) fix c assume "c \ set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))" then obtain c' where c: "c' \ set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'" by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits) note c(2) also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))" by simp also have "to_fract (Lcm_coeff_denoms p) * \ = Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))" unfolding to_fract_def by (subst mult_fract) simp_all also have "snd (quot_of_fract \) = 1" by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto) finally show "to_fract (fst (quot_of_fract c)) = c" by (rule to_fract_quot_of_fract) qed also have "p' = smult (content p') (primitive_part p')" by (rule content_times_primitive_part [symmetric]) also have "primitive_part p' = primitive_part_fract p" by (simp add: primitive_part_fract_def p'_def) also have "fract_poly (smult (content p') (primitive_part_fract p)) = smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) = smult (to_fract (Lcm_coeff_denoms p)) p" . thus ?thesis by (subst (asm) smult_eq_iff) (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def) qed lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)" proof - have "Lcm_coeff_denoms (fract_poly p) = 1" by (auto simp: set_coeffs_map_poly) hence "fract_content (fract_poly p) = to_fract (content (map_poly (fst \ quot_of_fract \ to_fract) p))" by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff) also have "map_poly (fst \ quot_of_fract \ to_fract) p = p" by (intro map_poly_idI) simp_all finally show ?thesis . qed lemma content_decompose_fract: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} fract poly" obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1" proof (cases "p = 0") case True hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all thus ?thesis .. next case False thus ?thesis by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract]) qed subsection \More properties of content and primitive part\ lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast - have B: "\i. \c dvd coeff b i \ i \ degree b" + have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i - unfolding m_def by (rule Greatest_le_nat[OF that B]) + unfolding m_def by (blast intro!: Greatest_le_nat that B) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) lemma fract_poly_dvdD: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} poly" assumes "fract_poly p dvd fract_poly q" "content p = 1" shows "p dvd q" proof - from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE) from content_decompose_fract[of r] guess c r' . note r' = this from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp from fract_poly_smult_eqE[OF this] guess a b . note ab = this have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2)) hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4)) have "1 = gcd a (normalize b)" by (simp add: ab) also note eq' also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4)) finally have [simp]: "a = 1" by simp from eq ab have "q = p * ([:b:] * r')" by simp thus ?thesis by (rule dvdI) qed subsection \Polynomials over a field are a Euclidean ring\ context begin interpretation field_poly: normalization_euclidean_semiring_multiplicative where zero = "0 :: 'a :: field poly" and one = 1 and plus = plus and minus = minus and times = times and normalize = "\p. smult (inverse (lead_coeff p)) p" and unit_factor = "\p. [:lead_coeff p:]" and euclidean_size = "\p. if p = 0 then 0 else 2 ^ degree p" and divide = divide and modulo = modulo rewrites "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" and "comm_monoid_mult.prod_mset times 1 = prod_mset" and "comm_semiring_1.irreducible times 1 0 = irreducible" and "comm_semiring_1.prime_elem times 1 0 = prime_elem" proof - show "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" by (simp add: dvd_dict) show "comm_monoid_mult.prod_mset times 1 = prod_mset" by (simp add: prod_mset_dict) show "comm_semiring_1.irreducible times 1 0 = irreducible" by (simp add: irreducible_dict) show "comm_semiring_1.prime_elem times 1 0 = prime_elem" by (simp add: prime_elem_dict) show "class.normalization_euclidean_semiring_multiplicative divide plus minus (0 :: 'a poly) times 1 modulo (\p. if p = 0 then 0 else 2 ^ degree p) (\p. [:lead_coeff p:]) (\p. smult (inverse (lead_coeff p)) p)" proof (standard, fold dvd_dict) fix p :: "'a poly" show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p" by (cases "p = 0") simp_all next fix p :: "'a poly" assume "is_unit p" then show "[:lead_coeff p:] = p" by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps) next fix p :: "'a poly" assume "p \ 0" then show "is_unit [:lead_coeff p:]" by (simp add: is_unit_pCons_iff) next fix a b :: "'a poly" assume "is_unit a" thus "[:lead_coeff (a * b):] = a * [:lead_coeff b:]" by (auto elim!: is_unit_polyE) qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le) qed lemma field_poly_irreducible_imp_prime: "prime_elem p" if "irreducible p" for p :: "'a :: field poly" using that by (fact field_poly.irreducible_imp_prime_elem) find_theorems name:prod_mset_prime_factorization lemma field_poly_prod_mset_prime_factorization: "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p" if "p \ 0" for p :: "'a :: field poly" using that by (fact field_poly.prod_mset_prime_factorization) lemma field_poly_in_prime_factorization_imp_prime: "prime_elem p" if "p \# field_poly.prime_factorization x" for p :: "'a :: field poly" by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime) (fact that) subsection \Primality and irreducibility in polynomial rings\ lemma nonconst_poly_irreducible_iff: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "degree p \ 0" shows "irreducible p \ irreducible (fract_poly p) \ content p = 1" proof safe assume p: "irreducible p" from content_decompose[of p] guess p' . note p' = this hence "p = [:content p:] * p'" by simp from p this have "[:content p:] dvd 1 \ p' dvd 1" by (rule irreducibleD) moreover have "\p' dvd 1" proof assume "p' dvd 1" hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff) with assms show False by contradiction qed ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff) show "irreducible (map_poly to_fract p)" proof (rule irreducibleI) have "fract_poly p = 0 \ p = 0" by (intro map_poly_eq_0_iff) auto with assms show "map_poly to_fract p \ 0" by auto next show "\is_unit (fract_poly p)" proof assume "is_unit (map_poly to_fract p)" hence "degree (map_poly to_fract p) = 0" by (auto simp: is_unit_poly_iff) hence "degree p = 0" by (simp add: degree_map_poly) with assms show False by contradiction qed next fix q r assume qr: "fract_poly p = q * r" from content_decompose_fract[of q] guess cg q' . note q = this from content_decompose_fract[of r] guess cr r' . note r = this from qr q r p have nz: "cg \ 0" "cr \ 0" by auto from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))" by (simp add: q r) from fract_poly_smult_eqE[OF this] guess a b . note ab = this hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:) with ab(4) have a: "a = normalize b" by (simp add: content_mult q r) then have "normalize b = gcd a b" by simp with \coprime a b\ have "normalize b = 1" by simp then have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff) note eq also from ab(1) \a = 1\ have "cr * cg = to_fract b" by simp also have "smult \ (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult) from p and this have "([:b:] * q') dvd 1 \ r' dvd 1" by (rule irreducibleD) hence "q' dvd 1 \ r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left) hence "fract_poly q' dvd 1 \ fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit) with q r show "is_unit q \ is_unit r" by (auto simp add: is_unit_smult_iff dvd_field_iff nz) qed next assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1" show "irreducible p" proof (rule irreducibleI) from irred show "p \ 0" by auto next from irred show "\p dvd 1" by (auto simp: irreducible_def dest: fract_poly_is_unit) next fix q r assume qr: "p = q * r" hence "fract_poly p = fract_poly q * fract_poly r" by simp from irred and this have "fract_poly q dvd 1 \ fract_poly r dvd 1" by (rule irreducibleD) with primitive qr show "q dvd 1 \ r dvd 1" by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff) qed qed private lemma irreducible_imp_prime_poly: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "irreducible p" shows "prime_elem p" proof (cases "degree p = 0") case True with assms show ?thesis by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE) next case False from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1" by (simp_all add: nonconst_poly_irreducible_iff) from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime) show ?thesis proof (rule prime_elemI) fix q r assume "p dvd q * r" hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd) hence "fract_poly p dvd fract_poly q * fract_poly r" by simp from prime and this have "fract_poly p dvd fract_poly q \ fract_poly p dvd fract_poly r" by (rule prime_elem_dvd_multD) with primitive show "p dvd q \ p dvd r" by (auto dest: fract_poly_dvdD) qed (insert assms, auto simp: irreducible_def) qed lemma degree_primitive_part_fract [simp]: "degree (primitive_part_fract p) = degree p" proof - have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))" by (simp add: content_times_primitive_part_fract) also have "degree \ = degree (primitive_part_fract p)" by (auto simp: degree_map_poly) finally show ?thesis .. qed lemma irreducible_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" assumes "irreducible p" shows "irreducible (primitive_part_fract p)" proof - from assms have deg: "degree (primitive_part_fract p) \ 0" by (intro notI) (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff) hence [simp]: "p \ 0" by auto note \irreducible p\ also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" by (simp add: content_times_primitive_part_fract) also have "irreducible \ \ irreducible (fract_poly (primitive_part_fract p))" by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff) finally show ?thesis using deg by (simp add: nonconst_poly_irreducible_iff) qed lemma prime_elem_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" shows "irreducible p \ prime_elem (primitive_part_fract p)" by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract) lemma irreducible_linear_field_poly: fixes a b :: "'a::field" assumes "b \ 0" shows "irreducible [:a,b:]" proof (rule irreducibleI) fix p q assume pq: "[:a,b:] = p * q" also from pq assms have "degree \ = degree p + degree q" by (intro degree_mult_eq) auto finally have "degree p = 0 \ degree q = 0" using assms by auto with assms pq show "is_unit p \ is_unit q" by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE) qed (insert assms, auto simp: is_unit_poly_iff) lemma prime_elem_linear_field_poly: "(b :: 'a :: field) \ 0 \ prime_elem [:a,b:]" by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly) lemma irreducible_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ irreducible [:a,b:]" by (auto intro!: irreducible_linear_field_poly simp: nonconst_poly_irreducible_iff content_def map_poly_pCons) lemma prime_elem_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ prime_elem [:a,b:]" by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly) subsection \Prime factorisation of polynomials\ private lemma poly_prime_factorization_exists_content_1: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" "content p = 1" shows "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize p" proof - let ?P = "field_poly.prime_factorization (fract_poly p)" define c where "c = prod_mset (image_mset fract_content ?P)" define c' where "c' = c * to_fract (lead_coeff p)" define e where "e = prod_mset (image_mset primitive_part_fract ?P)" define A where "A = image_mset (normalize \ primitive_part_fract) ?P" have "content e = (\x\#field_poly.prime_factorization (map_poly to_fract p). content (primitive_part_fract x))" by (simp add: e_def content_prod_mset multiset.map_comp o_def) also have "image_mset (\x. content (primitive_part_fract x)) ?P = image_mset (\_. 1) ?P" by (intro image_mset_cong content_primitive_part_fract) auto finally have content_e: "content e = 1" by simp from \p \ 0\ have "fract_poly p = [:lead_coeff (fract_poly p):] * smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)" by simp also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]" by (simp add: monom_0 degree_map_poly coeff_map_poly) also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P" by (subst field_poly_prod_mset_prime_factorization) simp_all also have "\ = prod_mset (image_mset id ?P)" by simp also have "image_mset id ?P = image_mset (\x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P" by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract) also have "prod_mset \ = smult c (fract_poly e)" by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def) also have "[:to_fract (lead_coeff p):] * \ = smult c' (fract_poly e)" by (simp add: c'_def) finally have eq: "fract_poly p = smult c' (fract_poly e)" . also obtain b where b: "c' = to_fract b" "is_unit b" proof - from fract_poly_smult_eqE[OF eq] guess a b . note ab = this from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: ) with assms content_e have "a = normalize b" by (simp add: ab(4)) with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff) with ab ab' have "c' = to_fract b" by auto from this and \is_unit b\ show ?thesis by (rule that) qed hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp finally have "p = smult b e" by (simp only: fract_poly_eq_iff) hence "p = [:b:] * e" by simp with b have "normalize p = normalize e" by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff) also have "normalize e = prod_mset A" by (simp add: multiset.map_comp e_def A_def normalize_prod_mset) finally have "prod_mset A = normalize p" .. have "prime_elem p" if "p \# A" for p using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible dest!: field_poly_in_prime_factorization_imp_prime ) from this and \prod_mset A = normalize p\ show ?thesis by (intro exI[of _ A]) blast qed lemma poly_prime_factorization_exists: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" shows "\A. (\p. p \# A \ prime_elem p) \ normalize (prod_mset A) = normalize p" proof - define B where "B = image_mset (\x. [:x:]) (prime_factorization (content p))" have "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize (primitive_part p)" by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all) then guess A by (elim exE conjE) note A = this have "normalize (prod_mset (A + B)) = normalize (prod_mset A * normalize (prod_mset B))" by simp also from assms have "normalize (prod_mset B) = normalize [:content p:]" by (simp add: prod_mset_const_poly normalize_const_poly prod_mset_prime_factorization_weak B_def) also have "prod_mset A = normalize (primitive_part p)" using A by simp finally have "normalize (prod_mset (A + B)) = normalize (primitive_part p * [:content p:])" by simp moreover have "\p. p \# B \ prime_elem p" by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime) ultimately show ?thesis using A by (intro exI[of _ "A + B"]) (auto) qed end subsection \Typeclass instances\ instance poly :: ("{factorial_ring_gcd,semiring_gcd_mult_normalize}") factorial_semiring by standard (rule poly_prime_factorization_exists) instantiation poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") factorial_ring_gcd begin definition gcd_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "gcd_poly = gcd_factorial" definition lcm_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "lcm_poly = lcm_factorial" definition Gcd_poly :: "'a poly set \ 'a poly" where [code del]: "Gcd_poly = Gcd_factorial" definition Lcm_poly :: "'a poly set \ 'a poly" where [code del]: "Lcm_poly = Lcm_factorial" instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def) end instance poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") semiring_gcd_mult_normalize .. instantiation poly :: ("{field,factorial_ring_gcd,semiring_gcd_mult_normalize}") "{unique_euclidean_ring, normalization_euclidean_semiring}" begin definition euclidean_size_poly :: "'a poly \ nat" where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" definition division_segment_poly :: "'a poly \ 'a poly" where [simp]: "division_segment_poly p = 1" instance proof show "(q * p + r) div p = q" if "p \ 0" and "euclidean_size r < euclidean_size p" for q p r :: "'a poly" proof - from \p \ 0\ eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)" by (simp add: eucl_rel_poly_iff distrib_right) then have "(r + q * p) div p = q + r div p" by (rule div_poly_eq) with that show ?thesis by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps) qed qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq power_add intro!: degree_mod_less' split: if_splits) end instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd, semiring_gcd_mult_normalize}") euclidean_ring_gcd by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard subsection \Polynomial GCD\ lemma gcd_poly_decompose: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "gcd p q = smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" proof (rule sym, rule gcdI) have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p" by simp next have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q" by simp next fix d assume "d dvd p" "d dvd q" hence "[:content d:] * primitive_part d dvd [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)" by (intro mult_dvd_mono) auto thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" by simp qed (auto simp: normalize_smult) lemma gcd_poly_pseudo_mod: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes nz: "q \ 0" and prim: "content p = 1" "content q = 1" shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))" proof - define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)" define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]" have [simp]: "primitive_part a = unit_factor a" by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0) from nz have [simp]: "a \ 0" by (auto simp: a_def) have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def) have "gcd (q * r + s) q = gcd q s" using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac) with pseudo_divmod(1)[OF nz rs] have "gcd (p * a) q = gcd q s" by (simp add: a_def) also from prim have "gcd (p * a) q = gcd p q" by (subst gcd_poly_decompose) (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim simp del: mult_pCons_right ) also from prim have "gcd q s = gcd q (primitive_part s)" by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim) also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def) finally show ?thesis . qed lemma degree_pseudo_mod_less: assumes "q \ 0" "pseudo_mod p q \ 0" shows "degree (pseudo_mod p q) < degree q" using pseudo_mod(2)[of q p] assms by auto function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code_aux p q = (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" by auto termination by (relation "measure ((\p. if p = 0 then 0 else Suc (degree p)) \ snd)") (auto simp: degree_pseudo_mod_less) declare gcd_poly_code_aux.simps [simp del] lemma gcd_poly_code_aux_correct: assumes "content p = 1" "q = 0 \ content q = 1" shows "gcd_poly_code_aux p q = gcd p q" using assms proof (induction p q rule: gcd_poly_code_aux.induct) case (1 p q) show ?case proof (cases "q = 0") case True thus ?thesis by (subst gcd_poly_code_aux.simps) auto next case False hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))" by (subst gcd_poly_code_aux.simps) simp_all also from "1.prems" False have "primitive_part (pseudo_mod p q) = 0 \ content (primitive_part (pseudo_mod p q)) = 1" by (cases "pseudo_mod p q = 0") auto with "1.prems" False have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = gcd q (primitive_part (pseudo_mod p q))" by (intro 1) simp_all also from "1.prems" False have "\ = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto finally show ?thesis . qed qed definition gcd_poly_code :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code p q = (if p = 0 then normalize q else if q = 0 then normalize p else smult (gcd (content p) (content q)) (gcd_poly_code_aux (primitive_part p) (primitive_part q)))" lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q" by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric]) lemma lcm_poly_code [code]: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "lcm p q = normalize (p * q div gcd p q)" by (fact lcm_gcd) lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] text \Example: @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval} \ end diff --git a/src/HOL/List.thy b/src/HOL/List.thy --- a/src/HOL/List.thy +++ b/src/HOL/List.thy @@ -1,7682 +1,7681 @@ (* Title: HOL/List.thy Author: Tobias Nipkow; proofs tidied by LCP *) section \The datatype of finite lists\ theory List imports Sledgehammer Code_Numeral Lifting_Set begin datatype (set: 'a) list = Nil ("[]") | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) for map: map rel: list_all2 pred: list_all where "tl [] = []" datatype_compat list lemma [case_names Nil Cons, cases type: list]: \ \for backward compatibility -- names of variables differ\ "(y = [] \ P) \ (\a list. y = a # list \ P) \ P" by (rule list.exhaust) lemma [case_names Nil Cons, induct type: list]: \ \for backward compatibility -- names of variables differ\ "P [] \ (\a list. P list \ P (a # list)) \ P list" by (rule list.induct) text \Compatibility:\ setup \Sign.mandatory_path "list"\ lemmas inducts = list.induct lemmas recs = list.rec lemmas cases = list.case setup \Sign.parent_path\ lemmas set_simps = list.set (* legacy *) syntax \ \list Enumeration\ "_list" :: "args => 'a list" ("[(_)]") translations "[x, xs]" == "x#[xs]" "[x]" == "x#[]" subsection \Basic list processing functions\ primrec (nonexhaustive) last :: "'a list \ 'a" where "last (x # xs) = (if xs = [] then x else last xs)" primrec butlast :: "'a list \ 'a list" where "butlast [] = []" | "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)" lemma set_rec: "set xs = rec_list {} (\x _. insert x) xs" by (induct xs) auto definition coset :: "'a list \ 'a set" where [simp]: "coset xs = - set xs" primrec append :: "'a list \ 'a list \ 'a list" (infixr "@" 65) where append_Nil: "[] @ ys = ys" | append_Cons: "(x#xs) @ ys = x # xs @ ys" primrec rev :: "'a list \ 'a list" where "rev [] = []" | "rev (x # xs) = rev xs @ [x]" primrec filter:: "('a \ bool) \ 'a list \ 'a list" where "filter P [] = []" | "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" text \Special input syntax for filter:\ syntax (ASCII) "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])") syntax "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_\_ ./ _])") translations "[x<-xs . P]" \ "CONST filter (\x. P) xs" primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where fold_Nil: "fold f [] = id" | fold_Cons: "fold f (x # xs) = fold f xs \ f x" primrec foldr :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where foldr_Nil: "foldr f [] = id" | foldr_Cons: "foldr f (x # xs) = f x \ foldr f xs" primrec foldl :: "('b \ 'a \ 'b) \ 'b \ 'a list \ 'b" where foldl_Nil: "foldl f a [] = a" | foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs" primrec concat:: "'a list list \ 'a list" where "concat [] = []" | "concat (x # xs) = x @ concat xs" primrec drop:: "nat \ 'a list \ 'a list" where drop_Nil: "drop n [] = []" | drop_Cons: "drop n (x # xs) = (case n of 0 \ x # xs | Suc m \ drop m xs)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec take:: "nat \ 'a list \ 'a list" where take_Nil:"take n [] = []" | take_Cons: "take n (x # xs) = (case n of 0 \ [] | Suc m \ x # take m xs)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where nth_Cons: "(x # xs) ! n = (case n of 0 \ x | Suc k \ xs ! k)" \ \Warning: simpset does not contain this definition, but separate theorems for \n = 0\ and \n = Suc k\\ primrec list_update :: "'a list \ nat \ 'a \ 'a list" where "list_update [] i v = []" | "list_update (x # xs) i v = (case i of 0 \ v # xs | Suc j \ x # list_update xs j v)" nonterminal lupdbinds and lupdbind syntax "_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)") "" :: "lupdbind => lupdbinds" ("_") "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _") "_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [1000,0] 900) translations "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs" "xs[i:=x]" == "CONST list_update xs i x" primrec takeWhile :: "('a \ bool) \ 'a list \ 'a list" where "takeWhile P [] = []" | "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])" primrec dropWhile :: "('a \ bool) \ 'a list \ 'a list" where "dropWhile P [] = []" | "dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)" primrec zip :: "'a list \ 'b list \ ('a \ 'b) list" where "zip xs [] = []" | zip_Cons: "zip xs (y # ys) = (case xs of [] \ [] | z # zs \ (z, y) # zip zs ys)" \ \Warning: simpset does not contain this definition, but separate theorems for \xs = []\ and \xs = z # zs\\ abbreviation map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where "map2 f xs ys \ map (\(x,y). f x y) (zip xs ys)" primrec product :: "'a list \ 'b list \ ('a \ 'b) list" where "product [] _ = []" | "product (x#xs) ys = map (Pair x) ys @ product xs ys" hide_const (open) product primrec product_lists :: "'a list list \ 'a list list" where "product_lists [] = [[]]" | "product_lists (xs # xss) = concat (map (\x. map (Cons x) (product_lists xss)) xs)" primrec upt :: "nat \ nat \ nat list" ("(1[_.. j then [i.. 'a list \ 'a list" where "insert x xs = (if x \ set xs then xs else x # xs)" definition union :: "'a list \ 'a list \ 'a list" where "union = fold insert" hide_const (open) insert union hide_fact (open) insert_def union_def primrec find :: "('a \ bool) \ 'a list \ 'a option" where "find _ [] = None" | "find P (x#xs) = (if P x then Some x else find P xs)" text \In the context of multisets, \count_list\ is equivalent to \<^term>\count \ mset\ and it it advisable to use the latter.\ primrec count_list :: "'a list \ 'a \ nat" where "count_list [] y = 0" | "count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)" definition "extract" :: "('a \ bool) \ 'a list \ ('a list * 'a * 'a list) option" where "extract P xs = (case dropWhile (Not \ P) xs of [] \ None | y#ys \ Some(takeWhile (Not \ P) xs, y, ys))" hide_const (open) "extract" primrec those :: "'a option list \ 'a list option" where "those [] = Some []" | "those (x # xs) = (case x of None \ None | Some y \ map_option (Cons y) (those xs))" primrec remove1 :: "'a \ 'a list \ 'a list" where "remove1 x [] = []" | "remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)" primrec removeAll :: "'a \ 'a list \ 'a list" where "removeAll x [] = []" | "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)" primrec distinct :: "'a list \ bool" where "distinct [] \ True" | "distinct (x # xs) \ x \ set xs \ distinct xs" primrec remdups :: "'a list \ 'a list" where "remdups [] = []" | "remdups (x # xs) = (if x \ set xs then remdups xs else x # remdups xs)" fun remdups_adj :: "'a list \ 'a list" where "remdups_adj [] = []" | "remdups_adj [x] = [x]" | "remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))" primrec replicate :: "nat \ 'a \ 'a list" where replicate_0: "replicate 0 x = []" | replicate_Suc: "replicate (Suc n) x = x # replicate n x" text \ Function \size\ is overloaded for all datatypes. Users may refer to the list version as \length\.\ abbreviation length :: "'a list \ nat" where "length \ size" definition enumerate :: "nat \ 'a list \ (nat \ 'a) list" where enumerate_eq_zip: "enumerate n xs = zip [n.. 'a list" where "rotate1 [] = []" | "rotate1 (x # xs) = xs @ [x]" definition rotate :: "nat \ 'a list \ 'a list" where "rotate n = rotate1 ^^ n" definition nths :: "'a list => nat set => 'a list" where "nths xs A = map fst (filter (\p. snd p \ A) (zip xs [0.. 'a list list" where "subseqs [] = [[]]" | "subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)" primrec n_lists :: "nat \ 'a list \ 'a list list" where "n_lists 0 xs = [[]]" | "n_lists (Suc n) xs = concat (map (\ys. map (\y. y # ys) xs) (n_lists n xs))" hide_const (open) n_lists function splice :: "'a list \ 'a list \ 'a list" where "splice [] ys = ys" | "splice (x#xs) ys = x # splice ys xs" by pat_completeness auto termination by(relation "measure(\(xs,ys). size xs + size ys)") auto function shuffles where "shuffles [] ys = {ys}" | "shuffles xs [] = {xs}" | "shuffles (x # xs) (y # ys) = (#) x ` shuffles xs (y # ys) \ (#) y ` shuffles (x # xs) ys" by pat_completeness simp_all termination by lexicographic_order text\Use only if you cannot use \<^const>\Min\ instead:\ fun min_list :: "'a::ord list \ 'a" where "min_list (x # xs) = (case xs of [] \ x | _ \ min x (min_list xs))" text\Returns first minimum:\ fun arg_min_list :: "('a \ ('b::linorder)) \ 'a list \ 'a" where "arg_min_list f [x] = x" | "arg_min_list f (x#y#zs) = (let m = arg_min_list f (y#zs) in if f x \ f m then x else m)" text\ \begin{figure}[htbp] \fbox{ \begin{tabular}{l} @{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\ @{lemma "length [a,b,c] = 3" by simp}\\ @{lemma "set [a,b,c] = {a,b,c}" by simp}\\ @{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\ @{lemma "rev [a,b,c] = [c,b,a]" by simp}\\ @{lemma "hd [a,b,c,d] = a" by simp}\\ @{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\ @{lemma "last [a,b,c,d] = d" by simp}\\ @{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\ @{lemma[source] "filter (\n::nat. n<2) [0,2,1] = [0,1]" by simp}\\ @{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\ @{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\ @{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\ @{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\ @{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\ @{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\ @{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\ @{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\ @{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\ @{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\ @{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\ @{lemma "shuffles [a,b] [c,d] = {[a,b,c,d],[a,c,b,d],[a,c,d,b],[c,a,b,d],[c,a,d,b],[c,d,a,b]}" by (simp add: insert_commute)}\\ @{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\ @{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\ @{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\ @{lemma "drop 6 [a,b,c,d] = []" by simp}\\ @{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\ @{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\ @{lemma "distinct [2,0,1::nat]" by simp}\\ @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\ @{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\ @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\ @{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\ @{lemma "List.union [2,3,4] [0::int,1,2] = [4,3,0,1,2]" by (simp add: List.insert_def List.union_def)}\\ @{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\ @{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\ @{lemma "count_list [0,1,0,2::int] 0 = 2" by (simp)}\\ @{lemma "List.extract (%i::int. i>0) [0,0] = None" by(simp add: extract_def)}\\ @{lemma "List.extract (%i::int. i>0) [0,1,0,2] = Some([0], 1, [0,2])" by(simp add: extract_def)}\\ @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\ @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\ @{lemma "nth [a,b,c,d] 2 = c" by simp}\\ @{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\ @{lemma "nths [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:nths_def)}\\ @{lemma "subseqs [a,b] = [[a, b], [a], [b], []]" by simp}\\ @{lemma "List.n_lists 2 [a,b,c] = [[a, a], [b, a], [c, a], [a, b], [b, b], [c, b], [a, c], [b, c], [c, c]]" by (simp add: eval_nat_numeral)}\\ @{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\ @{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\ @{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\ @{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\ @{lemma "min_list [3,1,-2::int] = -2" by (simp)}\\ @{lemma "arg_min_list (\i. i*i) [3,-1,1,-2::int] = -1" by (simp)} \end{tabular}} \caption{Characteristic examples} \label{fig:Characteristic} \end{figure} Figure~\ref{fig:Characteristic} shows characteristic examples that should give an intuitive understanding of the above functions. \ text\The following simple sort functions are intended for proofs, not for efficient implementations.\ text \A sorted predicate w.r.t. a relation:\ fun sorted_wrt :: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted_wrt P [] = True" | "sorted_wrt P (x # ys) = ((\y \ set ys. P x y) \ sorted_wrt P ys)" (* FIXME: define sorted in terms of sorted_wrt *) text \A class-based sorted predicate:\ context linorder begin fun sorted :: "'a list \ bool" where "sorted [] = True" | "sorted (x # ys) = ((\y \ set ys. x \ y) \ sorted ys)" fun strict_sorted :: "'a list \ bool" where "strict_sorted [] = True" | "strict_sorted (x # ys) = ((\y \ List.set ys. x < y) \ strict_sorted ys)" lemma sorted_sorted_wrt: "sorted = sorted_wrt (\)" proof (rule ext) fix xs show "sorted xs = sorted_wrt (\) xs" by(induction xs rule: sorted.induct) auto qed lemma strict_sorted_sorted_wrt: "strict_sorted = sorted_wrt (<)" proof (rule ext) fix xs show "strict_sorted xs = sorted_wrt (<) xs" by(induction xs rule: strict_sorted.induct) auto qed primrec insort_key :: "('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_key f x [] = [x]" | "insort_key f x (y#ys) = (if f x \ f y then (x#y#ys) else y#(insort_key f x ys))" definition sort_key :: "('b \ 'a) \ 'b list \ 'b list" where "sort_key f xs = foldr (insort_key f) xs []" definition insort_insert_key :: "('b \ 'a) \ 'b \ 'b list \ 'b list" where "insort_insert_key f x xs = (if f x \ f ` set xs then xs else insort_key f x xs)" abbreviation "sort \ sort_key (\x. x)" abbreviation "insort \ insort_key (\x. x)" abbreviation "insort_insert \ insort_insert_key (\x. x)" definition stable_sort_key :: "(('b \ 'a) \ 'b list \ 'b list) \ bool" where "stable_sort_key sk = (\f xs k. filter (\y. f y = k) (sk f xs) = filter (\y. f y = k) xs)" end subsubsection \List comprehension\ text\Input syntax for Haskell-like list comprehension notation. Typical example: \[(x,y). x \ xs, y \ ys, x \ y]\, the list of all pairs of distinct elements from \xs\ and \ys\. The syntax is as in Haskell, except that \|\ becomes a dot (like in Isabelle's set comprehension): \[e. x \ xs, \]\ rather than \verb![e| x <- xs, ...]!. The qualifiers after the dot are \begin{description} \item[generators] \p \ xs\, where \p\ is a pattern and \xs\ an expression of list type, or \item[guards] \b\, where \b\ is a boolean expression. %\item[local bindings] @ {text"let x = e"}. \end{description} Just like in Haskell, list comprehension is just a shorthand. To avoid misunderstandings, the translation into desugared form is not reversed upon output. Note that the translation of \[e. x \ xs]\ is optmized to \<^term>\map (%x. e) xs\. It is easy to write short list comprehensions which stand for complex expressions. During proofs, they may become unreadable (and mangled). In such cases it can be advisable to introduce separate definitions for the list comprehensions in question.\ nonterminal lc_qual and lc_quals syntax "_listcompr" :: "'a \ lc_qual \ lc_quals \ 'a list" ("[_ . __") "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ \ _") "_lc_test" :: "bool \ lc_qual" ("_") (*"_lc_let" :: "letbinds => lc_qual" ("let _")*) "_lc_end" :: "lc_quals" ("]") "_lc_quals" :: "lc_qual \ lc_quals \ lc_quals" (", __") syntax (ASCII) "_lc_gen" :: "'a \ 'a list \ lc_qual" ("_ <- _") parse_translation \ let val NilC = Syntax.const \<^const_syntax>\Nil\; val ConsC = Syntax.const \<^const_syntax>\Cons\; val mapC = Syntax.const \<^const_syntax>\map\; val concatC = Syntax.const \<^const_syntax>\concat\; val IfC = Syntax.const \<^const_syntax>\If\; val dummyC = Syntax.const \<^const_syntax>\Pure.dummy_pattern\ fun single x = ConsC $ x $ NilC; fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *) let (* FIXME proper name context!? *) val x = Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT); val e = if opti then single e else e; val case1 = Syntax.const \<^syntax_const>\_case1\ $ p $ e; val case2 = Syntax.const \<^syntax_const>\_case1\ $ dummyC $ NilC; val cs = Syntax.const \<^syntax_const>\_case2\ $ case1 $ case2; in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end; fun pair_pat_tr (x as Free _) e = Syntax_Trans.abs_tr [x, e] | pair_pat_tr (_ $ p1 $ p2) e = Syntax.const \<^const_syntax>\case_prod\ $ pair_pat_tr p1 (pair_pat_tr p2 e) | pair_pat_tr dummy e = Syntax_Trans.abs_tr [Syntax.const "_idtdummy", e] fun pair_pat ctxt (Const (\<^const_syntax>\Pair\,_) $ s $ t) = pair_pat ctxt s andalso pair_pat ctxt t | pair_pat ctxt (Free (s,_)) = let val thy = Proof_Context.theory_of ctxt; val s' = Proof_Context.intern_const ctxt s; in not (Sign.declared_const thy s') end | pair_pat _ t = (t = dummyC); fun abs_tr ctxt p e opti = let val p = Term_Position.strip_positions p in if pair_pat ctxt p then (pair_pat_tr p e, true) else (pat_tr ctxt p e opti, false) end fun lc_tr ctxt [e, Const (\<^syntax_const>\_lc_test\, _) $ b, qs] = let val res = (case qs of Const (\<^syntax_const>\_lc_end\, _) => single e | Const (\<^syntax_const>\_lc_quals\, _) $ q $ qs => lc_tr ctxt [e, q, qs]); in IfC $ b $ res $ NilC end | lc_tr ctxt [e, Const (\<^syntax_const>\_lc_gen\, _) $ p $ es, Const(\<^syntax_const>\_lc_end\, _)] = (case abs_tr ctxt p e true of (f, true) => mapC $ f $ es | (f, false) => concatC $ (mapC $ f $ es)) | lc_tr ctxt [e, Const (\<^syntax_const>\_lc_gen\, _) $ p $ es, Const (\<^syntax_const>\_lc_quals\, _) $ q $ qs] = let val e' = lc_tr ctxt [e, q, qs]; in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end; in [(\<^syntax_const>\_listcompr\, lc_tr)] end \ ML_val \ let val read = Syntax.read_term \<^context> o Syntax.implode_input; fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote (#1 (Input.source_content s1)) ^ Position.here_list [Input.pos_of s1, Input.pos_of s2]); in check \[(x,y,z). b]\ \if b then [(x, y, z)] else []\; check \[(x,y,z). (x,_,y)\xs]\ \map (\(x,_,y). (x, y, z)) xs\; check \[e x y. (x,_)\xs, y\ys]\ \concat (map (\(x,_). map (\y. e x y) ys) xs)\; check \[(x,y,z). xb]\ \if x < a then if b < x then [(x, y, z)] else [] else []\; check \[(x,y,z). x\xs, x>b]\ \concat (map (\x. if b < x then [(x, y, z)] else []) xs)\; check \[(x,y,z). xxs]\ \if x < a then map (\x. (x, y, z)) xs else []\; check \[(x,y). Cons True x \ xs]\ \concat (map (\xa. case xa of [] \ [] | True # x \ [(x, y)] | False # x \ []) xs)\; check \[(x,y,z). Cons x [] \ xs]\ \concat (map (\xa. case xa of [] \ [] | [x] \ [(x, y, z)] | x # aa # lista \ []) xs)\; check \[(x,y,z). xb, x=d]\ \if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []\; check \[(x,y,z). xb, y\ys]\ \if x < a then if b < x then map (\y. (x, y, z)) ys else [] else []\; check \[(x,y,z). xxs,y>b]\ \if x < a then concat (map (\(_,x). if b < y then [(x, y, z)] else []) xs) else []\; check \[(x,y,z). xxs, y\ys]\ \if x < a then concat (map (\x. map (\y. (x, y, z)) ys) xs) else []\; check \[(x,y,z). x\xs, x>b, y \concat (map (\x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)\; check \[(x,y,z). x\xs, x>b, y\ys]\ \concat (map (\x. if b < x then map (\y. (x, y, z)) ys else []) xs)\; check \[(x,y,z). x\xs, (y,_)\ys,y>x]\ \concat (map (\x. concat (map (\(y,_). if x < y then [(x, y, z)] else []) ys)) xs)\; check \[(x,y,z). x\xs, y\ys,z\zs]\ \concat (map (\x. concat (map (\y. map (\z. (x, y, z)) zs) ys)) xs)\ end; \ ML \ (* Simproc for rewriting list comprehensions applied to List.set to set comprehension. *) signature LIST_TO_SET_COMPREHENSION = sig val simproc : Proof.context -> cterm -> thm option end structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION = struct (* conversion *) fun all_exists_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Ex\, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct | _ => cv ctxt ct) fun all_but_last_exists_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Ex\, _) $ Abs (_, _, Const (\<^const_name>\Ex\, _) $ _) => Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct | _ => cv ctxt ct) fun Collect_conv cv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Collect\, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct | _ => raise CTERM ("Collect_conv", [ct])) fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th) fun conjunct_assoc_conv ct = Conv.try_conv (rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct fun right_hand_set_comprehension_conv conv ctxt = HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (all_exists_conv conv o #2) ctxt)) (* term abstraction of list comprehension patterns *) datatype termlets = If | Case of typ * int local val set_Nil_I = @{lemma "set [] = {x. False}" by (simp add: empty_def [symmetric])} val set_singleton = @{lemma "set [a] = {x. x = a}" by simp} val inst_Collect_mem_eq = @{lemma "set A = {x. x \ set A}" by simp} val del_refl_eq = @{lemma "(t = t \ P) \ P" by simp} fun mk_set T = Const (\<^const_name>\set\, HOLogic.listT T --> HOLogic.mk_setT T) fun dest_set (Const (\<^const_name>\set\, _) $ xs) = xs fun dest_singleton_list (Const (\<^const_name>\Cons\, _) $ t $ (Const (\<^const_name>\Nil\, _))) = t | dest_singleton_list t = raise TERM ("dest_singleton_list", [t]) (*We check that one case returns a singleton list and all other cases return [], and return the index of the one singleton list case.*) fun possible_index_of_singleton_case cases = let fun check (i, case_t) s = (case strip_abs_body case_t of (Const (\<^const_name>\Nil\, _)) => s | _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE)) in fold_index check cases (SOME NONE) |> the_default NONE end (*returns condition continuing term option*) fun dest_if (Const (\<^const_name>\If\, _) $ cond $ then_t $ Const (\<^const_name>\Nil\, _)) = SOME (cond, then_t) | dest_if _ = NONE (*returns (case_expr type index chosen_case constr_name) option*) fun dest_case ctxt case_term = let val (case_const, args) = strip_comb case_term in (case try dest_Const case_const of SOME (c, T) => (case Ctr_Sugar.ctr_sugar_of_case ctxt c of SOME {ctrs, ...} => (case possible_index_of_singleton_case (fst (split_last args)) of SOME i => let val constr_names = map (fst o dest_Const) ctrs val (Ts, _) = strip_type T val T' = List.last Ts in SOME (List.last args, T', i, nth args i, nth constr_names i) end | NONE => NONE) | NONE => NONE) | NONE => NONE) end fun tac ctxt [] = resolve_tac ctxt [set_singleton] 1 ORELSE resolve_tac ctxt [inst_Collect_mem_eq] 1 | tac ctxt (If :: cont) = Splitter.split_tac ctxt @{thms if_split} 1 THEN resolve_tac ctxt @{thms conjI} 1 THEN resolve_tac ctxt @{thms impI} 1 THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv then_conv rewr_conv' @{lemma "(True \ P) = P" by simp})) ctxt') 1) ctxt 1 THEN tac ctxt cont THEN resolve_tac ctxt @{thms impI} 1 THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv then_conv rewr_conv' @{lemma "(False \ P) = False" by simp})) ctxt') 1) ctxt 1 THEN resolve_tac ctxt [set_Nil_I] 1 | tac ctxt (Case (T, i) :: cont) = let val SOME {injects, distincts, case_thms, split, ...} = Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T)) in (* do case distinction *) Splitter.split_tac ctxt [split] 1 THEN EVERY (map_index (fn (i', _) => (if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac) THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1) THEN resolve_tac ctxt @{thms impI} 1 THEN (if i' = i then (* continue recursively *) Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K ((HOLogic.conj_conv (HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects)))) Conv.all_conv) then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq)) then_conv conjunct_assoc_conv)) ctxt' then_conv (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => Conv.repeat_conv (all_but_last_exists_conv (K (rewr_conv' @{lemma "(\x. x = t \ P x) = P t" by simp})) ctxt'')) ctxt')))) 1) ctxt 1 THEN tac ctxt cont else Subgoal.FOCUS (fn {prems, context = ctxt', ...} => CONVERSION (right_hand_set_comprehension_conv (K (HOLogic.conj_conv ((HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems))) then_conv (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts))) Conv.all_conv then_conv (rewr_conv' @{lemma "(False \ P) = False" by simp}))) ctxt' then_conv HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') => Conv.repeat_conv (Conv.bottom_conv (K (rewr_conv' @{lemma "(\x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1 THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms) end in fun simproc ctxt redex = let fun make_inner_eqs bound_vs Tis eqs t = (case dest_case ctxt t of SOME (x, T, i, cont, constr_name) => let val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont) val x' = incr_boundvars (length vs) x val eqs' = map (incr_boundvars (length vs)) eqs val constr_t = list_comb (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0)) val constr_eq = Const (\<^const_name>\HOL.eq\, T --> T --> \<^typ>\bool\) $ constr_t $ x' in make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body end | NONE => (case dest_if t of SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont | NONE => if null eqs then NONE (*no rewriting, nothing to be done*) else let val Type (\<^type_name>\list\, [rT]) = fastype_of1 (map snd bound_vs, t) val pat_eq = (case try dest_singleton_list t of SOME t' => Const (\<^const_name>\HOL.eq\, rT --> rT --> \<^typ>\bool\) $ Bound (length bound_vs) $ t' | NONE => Const (\<^const_name>\Set.member\, rT --> HOLogic.mk_setT rT --> \<^typ>\bool\) $ Bound (length bound_vs) $ (mk_set rT $ t)) val reverse_bounds = curry subst_bounds ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)]) val eqs' = map reverse_bounds eqs val pat_eq' = reverse_bounds pat_eq val inner_t = fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t) (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq') val lhs = Thm.term_of redex val rhs = HOLogic.mk_Collect ("x", rT, inner_t) val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) in SOME ((Goal.prove ctxt [] [] rewrite_rule_t (fn {context = ctxt', ...} => tac ctxt' (rev Tis))) RS @{thm eq_reflection}) end)) in make_inner_eqs [] [] [] (dest_set (Thm.term_of redex)) end end end \ simproc_setup list_to_set_comprehension ("set xs") = \K List_to_Set_Comprehension.simproc\ code_datatype set coset hide_const (open) coset subsubsection \\<^const>\Nil\ and \<^const>\Cons\\ lemma not_Cons_self [simp]: "xs \ x # xs" by (induct xs) auto lemma not_Cons_self2 [simp]: "x # xs \ xs" by (rule not_Cons_self [symmetric]) lemma neq_Nil_conv: "(xs \ []) = (\y ys. xs = y # ys)" by (induct xs) auto lemma tl_Nil: "tl xs = [] \ xs = [] \ (\x. xs = [x])" by (cases xs) auto lemma Nil_tl: "[] = tl xs \ xs = [] \ (\x. xs = [x])" by (cases xs) auto lemma length_induct: "(\xs. \ys. length ys < length xs \ P ys \ P xs) \ P xs" by (fact measure_induct) lemma induct_list012: "\P []; \x. P [x]; \x y zs. \ P zs; P (y # zs) \ \ P (x # y # zs)\ \ P xs" by induction_schema (pat_completeness, lexicographic_order) lemma list_nonempty_induct [consumes 1, case_names single cons]: "\ xs \ []; \x. P [x]; \x xs. xs \ [] \ P xs \ P (x # xs)\ \ P xs" by(induction xs rule: induct_list012) auto lemma inj_split_Cons: "inj_on (\(xs, n). n#xs) X" by (auto intro!: inj_onI) lemma inj_on_Cons1 [simp]: "inj_on ((#) x) A" by(simp add: inj_on_def) subsubsection \\<^const>\length\\ text \ Needs to come before \@\ because of theorem \append_eq_append_conv\. \ lemma length_append [simp]: "length (xs @ ys) = length xs + length ys" by (induct xs) auto lemma length_map [simp]: "length (map f xs) = length xs" by (induct xs) auto lemma length_rev [simp]: "length (rev xs) = length xs" by (induct xs) auto lemma length_tl [simp]: "length (tl xs) = length xs - 1" by (cases xs) auto lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])" by (induct xs) auto lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \ [])" by (induct xs) auto lemma length_pos_if_in_set: "x \ set xs \ length xs > 0" by auto lemma length_Suc_conv: "(length xs = Suc n) = (\y ys. xs = y # ys \ length ys = n)" by (induct xs) auto lemma Suc_length_conv: "(Suc n = length xs) = (\y ys. xs = y # ys \ length ys = n)" by (induct xs; simp; blast) lemma Suc_le_length_iff: "(Suc n \ length xs) = (\x ys. xs = x # ys \ n \ length ys)" by (metis Suc_le_D[of n] Suc_le_mono[of n] Suc_length_conv[of _ xs]) -lemma impossible_Cons: "length xs \ length ys ==> xs = x # ys = False" +lemma impossible_Cons: "length xs \ length ys \ xs = x # ys = False" by (induct xs) auto lemma list_induct2 [consumes 1, case_names Nil Cons]: "length xs = length ys \ P [] [] \ (\x xs y ys. length xs = length ys \ P xs ys \ P (x#xs) (y#ys)) \ P xs ys" proof (induct xs arbitrary: ys) case (Cons x xs ys) then show ?case by (cases ys) simp_all qed simp lemma list_induct3 [consumes 2, case_names Nil Cons]: "length xs = length ys \ length ys = length zs \ P [] [] [] \ (\x xs y ys z zs. length xs = length ys \ length ys = length zs \ P xs ys zs \ P (x#xs) (y#ys) (z#zs)) \ P xs ys zs" proof (induct xs arbitrary: ys zs) case Nil then show ?case by simp next case (Cons x xs ys zs) then show ?case by (cases ys, simp_all) (cases zs, simp_all) qed lemma list_induct4 [consumes 3, case_names Nil Cons]: "length xs = length ys \ length ys = length zs \ length zs = length ws \ P [] [] [] [] \ (\x xs y ys z zs w ws. length xs = length ys \ length ys = length zs \ length zs = length ws \ P xs ys zs ws \ P (x#xs) (y#ys) (z#zs) (w#ws)) \ P xs ys zs ws" proof (induct xs arbitrary: ys zs ws) case Nil then show ?case by simp next case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all) qed lemma list_induct2': "\ P [] []; \x xs. P (x#xs) []; \y ys. P [] (y#ys); \x xs y ys. P xs ys \ P (x#xs) (y#ys) \ \ P xs ys" by (induct xs arbitrary: ys) (case_tac x, auto)+ lemma list_all2_iff: "list_all2 P xs ys \ length xs = length ys \ (\(x, y) \ set (zip xs ys). P x y)" by (induct xs ys rule: list_induct2') auto lemma neq_if_length_neq: "length xs \ length ys \ (xs = ys) == False" by (rule Eq_FalseI) auto simproc_setup list_neq ("(xs::'a list) = ys") = \ (* Reduces xs=ys to False if xs and ys cannot be of the same length. This is the case if the atomic sublists of one are a submultiset of those of the other list and there are fewer Cons's in one than the other. *) let fun len (Const(\<^const_name>\Nil\,_)) acc = acc | len (Const(\<^const_name>\Cons\,_) $ _ $ xs) (ts,n) = len xs (ts,n+1) | len (Const(\<^const_name>\append\,_) $ xs $ ys) acc = len xs (len ys acc) | len (Const(\<^const_name>\rev\,_) $ xs) acc = len xs acc | len (Const(\<^const_name>\map\,_) $ _ $ xs) acc = len xs acc | len t (ts,n) = (t::ts,n); val ss = simpset_of \<^context>; fun list_neq ctxt ct = let val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct; val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0); fun prove_neq() = let val Type(_,listT::_) = eqT; val size = HOLogic.size_const listT; val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs); val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len); val thm = Goal.prove ctxt [] [] neq_len (K (simp_tac (put_simpset ss ctxt) 1)); in SOME (thm RS @{thm neq_if_length_neq}) end in if m < n andalso submultiset (op aconv) (ls,rs) orelse n < m andalso submultiset (op aconv) (rs,ls) then prove_neq() else NONE end; in K list_neq end \ subsubsection \\@\ -- append\ global_interpretation append: monoid append Nil proof fix xs ys zs :: "'a list" show "(xs @ ys) @ zs = xs @ (ys @ zs)" by (induct xs) simp_all show "xs @ [] = xs" by (induct xs) simp_all qed simp lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)" by (fact append.assoc) lemma append_Nil2: "xs @ [] = xs" by (fact append.right_neutral) lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \ ys = [])" by (induct xs) auto lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \ ys = [])" by (induct xs) auto lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])" by (induct xs) auto lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])" by (induct xs) auto lemma append_eq_append_conv [simp]: "length xs = length ys \ length us = length vs - ==> (xs@us = ys@vs) = (xs=ys \ us=vs)" + \ (xs@us = ys@vs) = (xs=ys \ us=vs)" by (induct xs arbitrary: ys; case_tac ys; force) lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) = (\us. xs = zs @ us \ us @ ys = ts \ xs @ us = zs \ ys = us @ ts)" proof (induct xs arbitrary: ys zs ts) case (Cons x xs) then show ?case by (cases zs) auto qed fastforce lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)" by simp lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \ x = y)" by simp lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)" by simp lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])" using append_same_eq [of _ _ "[]"] by auto lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])" using append_same_eq [of "[]"] by auto -lemma hd_Cons_tl: "xs \ [] ==> hd xs # tl xs = xs" +lemma hd_Cons_tl: "xs \ [] \ hd xs # tl xs = xs" by (fact list.collapse) lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)" by (induct xs) auto -lemma hd_append2 [simp]: "xs \ [] ==> hd (xs @ ys) = hd xs" +lemma hd_append2 [simp]: "xs \ [] \ hd (xs @ ys) = hd xs" by (simp add: hd_append split: list.split) lemma tl_append: "tl (xs @ ys) = (case xs of [] \ tl ys | z#zs \ zs @ ys)" by (simp split: list.split) -lemma tl_append2 [simp]: "xs \ [] ==> tl (xs @ ys) = tl xs @ ys" +lemma tl_append2 [simp]: "xs \ [] \ tl (xs @ ys) = tl xs @ ys" by (simp add: tl_append split: list.split) lemma Cons_eq_append_conv: "x#xs = ys@zs = (ys = [] \ x#xs = zs \ (\ys'. x#ys' = ys \ xs = ys'@zs))" by(cases ys) auto lemma append_eq_Cons_conv: "(ys@zs = x#xs) = (ys = [] \ zs = x#xs \ (\ys'. ys = x#ys' \ ys'@zs = xs))" by(cases ys) auto lemma longest_common_prefix: "\ps xs' ys'. xs = ps @ xs' \ ys = ps @ ys' \ (xs' = [] \ ys' = [] \ hd xs' \ hd ys')" by (induct xs ys rule: list_induct2') (blast, blast, blast, metis (no_types, hide_lams) append_Cons append_Nil list.sel(1)) text \Trivial rules for solving \@\-equations automatically.\ -lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys" -by simp - -lemma Cons_eq_appendI: -"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs" -by (drule sym) simp - -lemma append_eq_appendI: -"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us" -by (drule sym) simp +lemma eq_Nil_appendI: "xs = ys \ xs = [] @ ys" + by simp + +lemma Cons_eq_appendI: "\x # xs1 = ys; xs = xs1 @ zs\ \ x # xs = ys @ zs" + by auto + +lemma append_eq_appendI: "\xs @ xs1 = zs; ys = xs1 @ us\ \ xs @ ys = zs @ us" + by auto text \ Simplification procedure for all list equalities. Currently only tries to rearrange \@\ to see if - both lists end in a singleton list, - or both lists end in the same list. \ simproc_setup list_eq ("(xs::'a list) = ys") = \ let fun last (cons as Const (\<^const_name>\Cons\, _) $ _ $ xs) = (case xs of Const (\<^const_name>\Nil\, _) => cons | _ => last xs) | last (Const(\<^const_name>\append\,_) $ _ $ ys) = last ys | last t = t; fun list1 (Const(\<^const_name>\Cons\,_) $ _ $ Const(\<^const_name>\Nil\,_)) = true | list1 _ = false; fun butlast ((cons as Const(\<^const_name>\Cons\,_) $ x) $ xs) = (case xs of Const (\<^const_name>\Nil\, _) => xs | _ => cons $ butlast xs) | butlast ((app as Const (\<^const_name>\append\, _) $ xs) $ ys) = app $ butlast ys | butlast xs = Const(\<^const_name>\Nil\, fastype_of xs); val rearr_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]); fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) = let val lastl = last lhs and lastr = last rhs; fun rearr conv = let val lhs1 = butlast lhs and rhs1 = butlast rhs; val Type(_,listT::_) = eqT val appT = [listT,listT] ---> listT val app = Const(\<^const_name>\append\,appT) val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2)); val thm = Goal.prove ctxt [] [] eq (K (simp_tac (put_simpset rearr_ss ctxt) 1)); in SOME ((conv RS (thm RS trans)) RS eq_reflection) end; in if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv} else if lastl aconv lastr then rearr @{thm append_same_eq} else NONE end; in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end \ subsubsection \\<^const>\map\\ lemma hd_map: "xs \ [] \ hd (map f xs) = f (hd xs)" by (cases xs) simp_all lemma map_tl: "map f (tl xs) = tl (map f xs)" by (cases xs) simp_all -lemma map_ext: "(\x. x \ set xs \ f x = g x) ==> map f xs = map g xs" +lemma map_ext: "(\x. x \ set xs \ f x = g x) \ map f xs = map g xs" by (induct xs) simp_all lemma map_ident [simp]: "map (\x. x) = (\xs. xs)" by (rule ext, induct_tac xs) auto lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys" by (induct xs) auto lemma map_map [simp]: "map f (map g xs) = map (f \ g) xs" by (induct xs) auto lemma map_comp_map[simp]: "((map f) \ (map g)) = map(f \ g)" by (rule ext) simp lemma rev_map: "rev (map f xs) = map f (rev xs)" by (induct xs) auto lemma map_eq_conv[simp]: "(map f xs = map g xs) = (\x \ set xs. f x = g x)" by (induct xs) auto lemma map_cong [fundef_cong]: "xs = ys \ (\x. x \ set ys \ f x = g x) \ map f xs = map g ys" by simp lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])" by (cases xs) auto lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])" by (cases xs) auto lemma map_eq_Cons_conv: "(map f xs = y#ys) = (\z zs. xs = z#zs \ f z = y \ map f zs = ys)" by (cases xs) auto lemma Cons_eq_map_conv: "(x#xs = map f ys) = (\z zs. ys = z#zs \ x = f z \ xs = map f zs)" by (cases ys) auto lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1] lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1] declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] lemma ex_map_conv: "(\xs. ys = map f xs) = (\y \ set ys. \x. y = f x)" by(induct ys, auto simp add: Cons_eq_map_conv) lemma map_eq_imp_length_eq: assumes "map f xs = map g ys" shows "length xs = length ys" using assms proof (induct ys arbitrary: xs) case Nil then show ?case by simp next case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto from Cons xs have "map f zs = map g ys" by simp with Cons have "length zs = length ys" by blast with xs show ?case by simp qed lemma map_inj_on: assumes map: "map f xs = map f ys" and inj: "inj_on f (set xs Un set ys)" shows "xs = ys" using map_eq_imp_length_eq [OF map] assms proof (induct rule: list_induct2) case (Cons x xs y ys) then show ?case by (auto intro: sym) qed auto lemma inj_on_map_eq_map: "inj_on f (set xs Un set ys) \ (map f xs = map f ys) = (xs = ys)" by(blast dest:map_inj_on) lemma map_injective: - "map f xs = map f ys ==> inj f ==> xs = ys" + "map f xs = map f ys \ inj f \ xs = ys" by (induct ys arbitrary: xs) (auto dest!:injD) lemma inj_map_eq_map[simp]: "inj f \ (map f xs = map f ys) = (xs = ys)" by(blast dest:map_injective) -lemma inj_mapI: "inj f ==> inj (map f)" +lemma inj_mapI: "inj f \ inj (map f)" by (iprover dest: map_injective injD intro: inj_onI) -lemma inj_mapD: "inj (map f) ==> inj f" +lemma inj_mapD: "inj (map f) \ inj f" by (metis (no_types, hide_lams) injI list.inject list.simps(9) the_inv_f_f) lemma inj_map[iff]: "inj (map f) = inj f" by (blast dest: inj_mapD intro: inj_mapI) lemma inj_on_mapI: "inj_on f (\(set ` A)) \ inj_on (map f) A" by (blast intro:inj_onI dest:inj_onD map_inj_on) lemma map_idI: "(\x. x \ set xs \ f x = x) \ map f xs = xs" by (induct xs, auto) lemma map_fun_upd [simp]: "y \ set xs \ map (f(y:=v)) xs = map f xs" by (induct xs) auto lemma map_fst_zip[simp]: "length xs = length ys \ map fst (zip xs ys) = xs" by (induct rule:list_induct2, simp_all) lemma map_snd_zip[simp]: "length xs = length ys \ map snd (zip xs ys) = ys" by (induct rule:list_induct2, simp_all) lemma map_fst_zip_take: "map fst (zip xs ys) = take (min (length xs) (length ys)) xs" by (induct xs ys rule: list_induct2') simp_all lemma map_snd_zip_take: "map snd (zip xs ys) = take (min (length xs) (length ys)) ys" by (induct xs ys rule: list_induct2') simp_all lemma map2_map_map: "map2 h (map f xs) (map g xs) = map (\x. h (f x) (g x)) xs" by (induction xs) (auto) functor map: map by (simp_all add: id_def) declare map.id [simp] subsubsection \\<^const>\rev\\ lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs" by (induct xs) auto lemma rev_rev_ident [simp]: "rev (rev xs) = xs" by (induct xs) auto lemma rev_swap: "(rev xs = ys) = (xs = rev ys)" by auto lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])" by (induct xs) auto lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])" by (induct xs) auto lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])" by (cases xs) auto lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])" by (cases xs) auto lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by force next case Cons then show ?case by (cases ys) auto qed lemma inj_on_rev[iff]: "inj_on rev A" by(simp add:inj_on_def) lemma rev_induct [case_names Nil snoc]: - "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs" -apply(simplesubst rev_rev_ident[symmetric]) -apply(rule_tac list = "rev xs" in list.induct, simp_all) -done + assumes "P []" and "\x xs. P xs \ P (xs @ [x])" + shows "P xs" +proof - + have "P (rev (rev xs))" + by (rule_tac list = "rev xs" in list.induct, simp_all add: assms) + then show ?thesis by simp +qed lemma rev_exhaust [case_names Nil snoc]: - "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P" + "(xs = [] \ P) \(\ys y. xs = ys @ [y] \ P) \ P" by (induct xs rule: rev_induct) auto lemmas rev_cases = rev_exhaust lemma rev_nonempty_induct [consumes 1, case_names single snoc]: assumes "xs \ []" and single: "\x. P [x]" and snoc': "\x xs. xs \ [] \ P xs \ P (xs@[x])" shows "P xs" using \xs \ []\ proof (induct xs rule: rev_induct) case (snoc x xs) then show ?case proof (cases xs) case Nil thus ?thesis by (simp add: single) next case Cons with snoc show ?thesis by (fastforce intro!: snoc') qed qed simp lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])" by(rule rev_cases[of xs]) auto subsubsection \\<^const>\set\\ declare list.set[code_post] \ \pretty output\ lemma finite_set [iff]: "finite (set xs)" by (induct xs) auto lemma set_append [simp]: "set (xs @ ys) = (set xs \ set ys)" by (induct xs) auto lemma hd_in_set[simp]: "xs \ [] \ hd xs \ set xs" by(cases xs) auto lemma set_subset_Cons: "set xs \ set (x # xs)" by auto lemma set_ConsD: "y \ set (x # xs) \ y=x \ y \ set xs" by auto lemma set_empty [iff]: "(set xs = {}) = (xs = [])" by (induct xs) auto lemma set_empty2[iff]: "({} = set xs) = (xs = [])" by(induct xs) auto lemma set_rev [simp]: "set (rev xs) = set xs" by (induct xs) auto lemma set_map [simp]: "set (map f xs) = f`(set xs)" by (induct xs) auto lemma set_filter [simp]: "set (filter P xs) = {x. x \ set xs \ P x}" by (induct xs) auto lemma set_upt [simp]: "set[i.. set xs \ \ys zs. xs = ys @ x # zs" proof (induct xs) case Nil thus ?case by simp next case Cons thus ?case by (auto intro: Cons_eq_appendI) qed lemma in_set_conv_decomp: "x \ set xs \ (\ys zs. xs = ys @ x # zs)" by (auto elim: split_list) lemma split_list_first: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set ys" proof (induct xs) case Nil thus ?case by simp next case (Cons a xs) show ?case proof cases assume "x = a" thus ?case using Cons by fastforce next assume "x \ a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI) qed qed lemma in_set_conv_decomp_first: "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)" by (auto dest!: split_list_first) lemma split_list_last: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs" proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc a xs) show ?case proof cases assume "x = a" thus ?case using snoc by (auto intro!: exI) next assume "x \ a" thus ?case using snoc by fastforce qed qed lemma in_set_conv_decomp_last: "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set zs)" by (auto dest!: split_list_last) lemma split_list_prop: "\x \ set xs. P x \ \ys x zs. xs = ys @ x # zs \ P x" proof (induct xs) case Nil thus ?case by simp next case Cons thus ?case by(simp add:Bex_def)(metis append_Cons append.simps(1)) qed lemma split_list_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" using split_list_prop [OF assms] by blast lemma split_list_first_prop: "\x \ set xs. P x \ \ys x zs. xs = ys@x#zs \ P x \ (\y \ set ys. \ P y)" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) show ?case proof cases assume "P x" hence "x # xs = [] @ x # xs \ P x \ (\y\set []. \ P y)" by simp thus ?thesis by fast next assume "\ P x" hence "\x\set xs. P x" using Cons(2) by simp thus ?thesis using \\ P x\ Cons(1) by (metis append_Cons set_ConsD) qed qed lemma split_list_first_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\y \ set ys. \ P y" using split_list_first_prop [OF assms] by blast lemma split_list_first_prop_iff: "(\x \ set xs. P x) \ (\ys x zs. xs = ys@x#zs \ P x \ (\y \ set ys. \ P y))" by (rule, erule split_list_first_prop) auto lemma split_list_last_prop: "\x \ set xs. P x \ \ys x zs. xs = ys@x#zs \ P x \ (\z \ set zs. \ P z)" proof(induct xs rule:rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof cases assume "P x" thus ?thesis by (auto intro!: exI) next assume "\ P x" hence "\x\set xs. P x" using snoc(2) by simp thus ?thesis using \\ P x\ snoc(1) by fastforce qed qed lemma split_list_last_propE: assumes "\x \ set xs. P x" obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\z \ set zs. \ P z" using split_list_last_prop [OF assms] by blast lemma split_list_last_prop_iff: "(\x \ set xs. P x) \ (\ys x zs. xs = ys@x#zs \ P x \ (\z \ set zs. \ P z))" by rule (erule split_list_last_prop, auto) lemma finite_list: "finite A \ \xs. set xs = A" by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2)) lemma card_length: "card (set xs) \ length xs" by (induct xs) (auto simp add: card_insert_if) lemma set_minus_filter_out: "set xs - {y} = set (filter (\x. \ (x = y)) xs)" by (induct xs) auto lemma append_Cons_eq_iff: "\ x \ set xs; x \ set ys \ \ xs @ x # ys = xs' @ x # ys' \ (xs = xs' \ ys = ys')" by(auto simp: append_eq_Cons_conv Cons_eq_append_conv append_eq_append_conv2) subsubsection \\<^const>\filter\\ lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys" by (induct xs) auto lemma rev_filter: "rev (filter P xs) = filter P (rev xs)" by (induct xs) simp_all lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\x. Q x \ P x) xs" by (induct xs) auto lemma length_filter_le [simp]: "length (filter P xs) \ length xs" by (induct xs) (auto simp add: le_SucI) lemma sum_length_filter_compl: "length(filter P xs) + length(filter (\x. \P x) xs) = length xs" by(induct xs) simp_all -lemma filter_True [simp]: "\x \ set xs. P x ==> filter P xs = xs" +lemma filter_True [simp]: "\x \ set xs. P x \ filter P xs = xs" by (induct xs) auto -lemma filter_False [simp]: "\x \ set xs. \ P x ==> filter P xs = []" +lemma filter_False [simp]: "\x \ set xs. \ P x \ filter P xs = []" by (induct xs) auto lemma filter_empty_conv: "(filter P xs = []) = (\x\set xs. \ P x)" by (induct xs) simp_all lemma filter_id_conv: "(filter P xs = xs) = (\x\set xs. P x)" proof (induct xs) case (Cons x xs) then show ?case using length_filter_le by (simp add: impossible_Cons) qed auto lemma filter_map: "filter P (map f xs) = map f (filter (P \ f) xs)" by (induct xs) simp_all lemma length_filter_map[simp]: "length (filter P (map f xs)) = length(filter (P \ f) xs)" by (simp add:filter_map) lemma filter_is_subset [simp]: "set (filter P xs) \ set xs" by auto lemma length_filter_less: "\ x \ set xs; \ P x \ \ length(filter P xs) < length xs" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case using Suc_le_eq by fastforce qed lemma length_filter_conv_card: "length(filter p xs) = card{i. i < length xs \ p(xs!i)}" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) let ?S = "{i. i < length xs \ p(xs!i)}" have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite) show ?case (is "?l = card ?S'") proof (cases) assume "p x" hence eq: "?S' = insert 0 (Suc ` ?S)" by(auto simp: image_def split:nat.split dest:gr0_implies_Suc) have "length (filter p (x # xs)) = Suc(card ?S)" using Cons \p x\ by simp also have "\ = Suc(card(Suc ` ?S))" using fin by (simp add: card_image) also have "\ = card ?S'" using eq fin by (simp add:card_insert_if) finally show ?thesis . next assume "\ p x" hence eq: "?S' = Suc ` ?S" by(auto simp add: image_def split:nat.split elim:lessE) have "length (filter p (x # xs)) = card ?S" using Cons \\ p x\ by simp also have "\ = card(Suc ` ?S)" using fin by (simp add: card_image) also have "\ = card ?S'" using eq fin by (simp add:card_insert_if) finally show ?thesis . qed qed lemma Cons_eq_filterD: "x#xs = filter P ys \ \us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs" (is "_ \ \us vs. ?P ys us vs") proof(induct ys) case Nil thus ?case by simp next case (Cons y ys) show ?case (is "\x. ?Q x") proof cases assume Py: "P y" show ?thesis proof cases assume "x = y" with Py Cons.prems have "?Q []" by simp then show ?thesis .. next assume "x \ y" with Py Cons.prems show ?thesis by simp qed next assume "\ P y" with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce then have "?Q (y#us)" by simp then show ?thesis .. qed qed lemma filter_eq_ConsD: "filter P ys = x#xs \ \us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs" by(rule Cons_eq_filterD) simp lemma filter_eq_Cons_iff: "(filter P ys = x#xs) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" by(auto dest:filter_eq_ConsD) lemma Cons_eq_filter_iff: "(x#xs = filter P ys) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" by(auto dest:Cons_eq_filterD) lemma inj_on_filter_key_eq: assumes "inj_on f (insert y (set xs))" shows "filter (\x. f y = f x) xs = filter (HOL.eq y) xs" using assms by (induct xs) auto lemma filter_cong[fundef_cong]: "xs = ys \ (\x. x \ set ys \ P x = Q x) \ filter P xs = filter Q ys" by (induct ys arbitrary: xs) auto subsubsection \List partitioning\ primrec partition :: "('a \ bool) \'a list \ 'a list \ 'a list" where "partition P [] = ([], [])" | "partition P (x # xs) = (let (yes, no) = partition P xs in if P x then (x # yes, no) else (yes, x # no))" lemma partition_filter1: "fst (partition P xs) = filter P xs" by (induct xs) (auto simp add: Let_def split_def) lemma partition_filter2: "snd (partition P xs) = filter (Not \ P) xs" by (induct xs) (auto simp add: Let_def split_def) lemma partition_P: assumes "partition P xs = (yes, no)" shows "(\p \ set yes. P p) \ (\p \ set no. \ P p)" proof - from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" by simp_all then show ?thesis by (simp_all add: partition_filter1 partition_filter2) qed lemma partition_set: assumes "partition P xs = (yes, no)" shows "set yes \ set no = set xs" proof - from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)" by simp_all then show ?thesis by (auto simp add: partition_filter1 partition_filter2) qed lemma partition_filter_conv[simp]: "partition f xs = (filter f xs,filter (Not \ f) xs)" unfolding partition_filter2[symmetric] unfolding partition_filter1[symmetric] by simp declare partition.simps[simp del] subsubsection \\<^const>\concat\\ lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys" by (induct xs) auto lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\xs \ set xss. xs = [])" by (induct xss) auto lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\xs \ set xss. xs = [])" by (induct xss) auto lemma set_concat [simp]: "set (concat xs) = (\x\set xs. set x)" by (induct xs) auto lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs" by (induct xs) auto lemma map_concat: "map f (concat xs) = concat (map (map f) xs)" by (induct xs) auto lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)" by (induct xs) auto lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))" by (induct xs) auto -lemma concat_eq_concat_iff: "\(x, y) \ set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)" +lemma concat_eq_concat_iff: "\(x, y) \ set (zip xs ys). length x = length y \ length xs = length ys \ (concat xs = concat ys) = (xs = ys)" proof (induct xs arbitrary: ys) case (Cons x xs ys) thus ?case by (cases ys) auto qed (auto) -lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \(x, y) \ set (zip xs ys). length x = length y ==> xs = ys" +lemma concat_injective: "concat xs = concat ys \ length xs = length ys \ \(x, y) \ set (zip xs ys). length x = length y \ xs = ys" by (simp add: concat_eq_concat_iff) lemma concat_eq_appendD: assumes "concat xss = ys @ zs" "xss \ []" shows "\xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \ ys = concat xss1 @ xs \ zs = xs' @ concat xss2" using assms proof(induction xss arbitrary: ys) case (Cons xs xss) from Cons.prems consider us where "xs @ us = ys" "concat xss = us @ zs" | us where "xs = ys @ us" "us @ concat xss = zs" by(auto simp add: append_eq_append_conv2) then show ?case proof cases case 1 then show ?thesis using Cons.IH[OF 1(2)] by(cases xss)(auto intro: exI[where x="[]"], metis append.assoc append_Cons concat.simps(2)) qed(auto intro: exI[where x="[]"]) qed simp lemma concat_eq_append_conv: "concat xss = ys @ zs \ (if xss = [] then ys = [] \ zs = [] else \xss1 xs xs' xss2. xss = xss1 @ (xs @ xs') # xss2 \ ys = concat xss1 @ xs \ zs = xs' @ concat xss2)" by(auto dest: concat_eq_appendD) subsubsection \\<^const>\nth\\ lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x" by auto lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n" by auto declare nth.simps [simp del] lemma nth_Cons_pos[simp]: "0 < n \ (x#xs) ! n = xs ! (n - 1)" by(auto simp: Nat.gr0_conv_Suc) lemma nth_append: "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))" proof (induct xs arbitrary: n) case (Cons x xs) then show ?case using less_Suc_eq_0_disj by auto qed simp lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x" by (induct xs) auto lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n" by (induct xs) auto -lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)" +lemma nth_map [simp]: "n < length xs \ (map f xs)!n = f(xs!n)" proof (induct xs arbitrary: n) case (Cons x xs) then show ?case using less_Suc_eq_0_disj by auto qed simp lemma nth_tl: "n < length (tl xs) \ tl xs ! n = xs ! Suc n" by (induction xs) auto lemma hd_conv_nth: "xs \ [] \ hd xs = xs!0" by(cases xs) simp_all lemma list_eq_iff_nth_eq: "(xs = ys) = (length xs = length ys \ (\i ?R" by force show "?R \ ?L" using less_Suc_eq_0_disj by auto qed with Cons show ?case by simp qed simp lemma in_set_conv_nth: "(x \ set xs) = (\i < length xs. xs!i = x)" by(auto simp:set_conv_nth) lemma nth_equal_first_eq: assumes "x \ set xs" assumes "n \ length xs" shows "(x # xs) ! n = x \ n = 0" (is "?lhs \ ?rhs") proof assume ?lhs show ?rhs proof (rule ccontr) assume "n \ 0" then have "n > 0" by simp with \?lhs\ have "xs ! (n - 1) = x" by simp moreover from \n > 0\ \n \ length xs\ have "n - 1 < length xs" by simp ultimately have "\ix \ set xs\ in_set_conv_nth [of x xs] show False by simp qed next assume ?rhs then show ?lhs by simp qed lemma nth_non_equal_first_eq: assumes "x \ y" shows "(x # xs) ! n = y \ xs ! (n - 1) = y \ n > 0" (is "?lhs \ ?rhs") proof assume "?lhs" with assms have "n > 0" by (cases n) simp_all with \?lhs\ show ?rhs by simp next assume "?rhs" then show "?lhs" by simp qed lemma list_ball_nth: "\n < length xs; \x \ set xs. P x\ \ P(xs!n)" by (auto simp add: set_conv_nth) lemma nth_mem [simp]: "n < length xs \ xs!n \ set xs" by (auto simp add: set_conv_nth) lemma all_nth_imp_all_set: "\\i < length xs. P(xs!i); x \ set xs\ \ P x" by (auto simp add: set_conv_nth) lemma all_set_conv_all_nth: "(\x \ set xs. P x) = (\i. i < length xs \ P (xs ! i))" by (auto simp add: set_conv_nth) lemma rev_nth: "n < size xs \ rev xs ! n = xs ! (length xs - Suc n)" proof (induct xs arbitrary: n) case Nil thus ?case by simp next case (Cons x xs) hence n: "n < Suc (length xs)" by simp moreover { assume "n < length xs" with n obtain n' where n': "length xs - n = Suc n'" by (cases "length xs - n", auto) moreover from n' have "length xs - Suc n = n'" by simp ultimately have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp } ultimately show ?case by (clarsimp simp add: Cons nth_append) qed lemma Skolem_list_nth: "(\ix. P i x) = (\xs. size xs = k \ (\ixs. ?P k xs)") proof(induct k) case 0 show ?case by simp next case (Suc k) show ?case (is "?L = ?R" is "_ = (\xs. ?P' xs)") proof assume "?R" thus "?L" using Suc by auto next assume "?L" with Suc obtain x xs where "?P k xs \ P k x" by (metis less_Suc_eq) hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq) thus "?R" .. qed qed subsubsection \\<^const>\list_update\\ lemma length_list_update [simp]: "length(xs[i:=x]) = length xs" by (induct xs arbitrary: i) (auto split: nat.split) lemma nth_list_update: - "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)" + "i < length xs\ (xs[i:=x])!j = (if i = j then x else xs!j)" by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) -lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x" +lemma nth_list_update_eq [simp]: "i < length xs \ (xs[i:=x])!i = x" by (simp add: nth_list_update) -lemma nth_list_update_neq [simp]: "i \ j ==> xs[i:=x]!j = xs!j" +lemma nth_list_update_neq [simp]: "i \ j \ xs[i:=x]!j = xs!j" by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) lemma list_update_id[simp]: "xs[i := xs!i] = xs" by (induct xs arbitrary: i) (simp_all split:nat.splits) lemma list_update_beyond[simp]: "length xs \ i \ xs[i:=x] = xs" proof (induct xs arbitrary: i) case (Cons x xs i) then show ?case by (metis leD length_list_update list_eq_iff_nth_eq nth_list_update_neq) qed simp lemma list_update_nonempty[simp]: "xs[k:=x] = [] \ xs=[]" by (simp only: length_0_conv[symmetric] length_list_update) lemma list_update_same_conv: - "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)" + "i < length xs \ (xs[i := x] = xs) = (xs!i = x)" by (induct xs arbitrary: i) (auto split: nat.split) lemma list_update_append1: "i < size xs \ (xs @ ys)[i:=x] = xs[i:=x] @ ys" by (induct xs arbitrary: i)(auto split:nat.split) lemma list_update_append: "(xs @ ys) [n:= x] = (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))" by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_length [simp]: "(xs @ x # ys)[length xs := y] = (xs @ y # ys)" by (induct xs, auto) lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]" by(induct xs arbitrary: k)(auto split:nat.splits) lemma rev_update: "k < length xs \ rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]" by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits) lemma update_zip: "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])" by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split) lemma set_update_subset_insert: "set(xs[i:=x]) \ insert x (set xs)" by (induct xs arbitrary: i) (auto split: nat.split) lemma set_update_subsetI: "\set xs \ A; x \ A\ \ set(xs[i := x]) \ A" by (blast dest!: set_update_subset_insert [THEN subsetD]) lemma set_update_memI: "n < length xs \ x \ set (xs[n := x])" by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_overwrite[simp]: "xs [i := x, i := y] = xs [i := y]" by (induct xs arbitrary: i) (simp_all split: nat.split) lemma list_update_swap: "i \ i' \ xs [i := x, i' := x'] = xs [i' := x', i := x]" by (induct xs arbitrary: i i') (simp_all split: nat.split) lemma list_update_code [code]: "[][i := y] = []" "(x # xs)[0 := y] = y # xs" "(x # xs)[Suc i := y] = x # xs[i := y]" by simp_all subsubsection \\<^const>\last\ and \<^const>\butlast\\ lemma last_snoc [simp]: "last (xs @ [x]) = x" by (induct xs) auto lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs" by (induct xs) auto lemma last_ConsL: "xs = [] \ last(x#xs) = x" by simp lemma last_ConsR: "xs \ [] \ last(x#xs) = last xs" by simp lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)" by (induct xs) (auto) lemma last_appendL[simp]: "ys = [] \ last(xs @ ys) = last xs" by(simp add:last_append) lemma last_appendR[simp]: "ys \ [] \ last(xs @ ys) = last ys" by(simp add:last_append) lemma last_tl: "xs = [] \ tl xs \ [] \last (tl xs) = last xs" by (induct xs) simp_all lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)" by (induct xs) simp_all lemma hd_rev: "xs \ [] \ hd(rev xs) = last xs" by(rule rev_exhaust[of xs]) simp_all lemma last_rev: "xs \ [] \ last(rev xs) = hd xs" by(cases xs) simp_all lemma last_in_set[simp]: "as \ [] \ last as \ set as" by (induct as) auto lemma length_butlast [simp]: "length (butlast xs) = length xs - 1" by (induct xs rule: rev_induct) auto lemma butlast_append: "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)" by (induct xs arbitrary: ys) auto lemma append_butlast_last_id [simp]: "xs \ [] \ butlast xs @ [last xs] = xs" by (induct xs) auto lemma in_set_butlastD: "x \ set (butlast xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) lemma in_set_butlast_appendI: "x \ set (butlast xs) \ x \ set (butlast ys) \ x \ set (butlast (xs @ ys))" by (auto dest: in_set_butlastD simp add: butlast_append) lemma last_drop[simp]: "n < length xs \ last (drop n xs) = last xs" by (induct xs arbitrary: n)(auto split:nat.split) lemma nth_butlast: assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n" proof (cases xs) case (Cons y ys) moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n" by (simp add: nth_append) ultimately show ?thesis using append_butlast_last_id by simp qed simp lemma last_conv_nth: "xs\[] \ last xs = xs!(length xs - 1)" by(induct xs)(auto simp:neq_Nil_conv) lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" by (induction xs rule: induct_list012) simp_all lemma last_list_update: "xs \ [] \ last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)" by (auto simp: last_conv_nth) lemma butlast_list_update: "butlast(xs[k:=x]) = (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])" by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits) lemma last_map: "xs \ [] \ last (map f xs) = f (last xs)" by (cases xs rule: rev_cases) simp_all lemma map_butlast: "map f (butlast xs) = butlast (map f xs)" by (induct xs) simp_all lemma snoc_eq_iff_butlast: "xs @ [x] = ys \ (ys \ [] \ butlast ys = xs \ last ys = x)" by fastforce corollary longest_common_suffix: "\ss xs' ys'. xs = xs' @ ss \ ys = ys' @ ss \ (xs' = [] \ ys' = [] \ last xs' \ last ys')" using longest_common_prefix[of "rev xs" "rev ys"] unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) lemma butlast_rev [simp]: "butlast (rev xs) = rev (tl xs)" by (cases xs) simp_all subsubsection \\<^const>\take\ and \<^const>\drop\\ lemma take_0: "take 0 xs = []" by (induct xs) auto lemma drop_0: "drop 0 xs = xs" by (induct xs) auto lemma take0[simp]: "take 0 = (\xs. [])" by(rule ext) (rule take_0) lemma drop0[simp]: "drop 0 = (\x. x)" by(rule ext) (rule drop_0) lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs" by simp lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs" by simp declare take_Cons [simp del] and drop_Cons [simp del] lemma take_Suc: "xs \ [] \ take (Suc n) xs = hd xs # take n (tl xs)" by(clarsimp simp add:neq_Nil_conv) lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)" by(cases xs, simp_all) lemma hd_take[simp]: "j > 0 \ hd (take j xs) = hd xs" by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc) lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)" by (induct xs arbitrary: n) simp_all lemma drop_tl: "drop n (tl xs) = tl(drop n xs)" by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split) lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)" by (cases n, simp, cases xs, auto) lemma tl_drop: "tl (drop n xs) = drop n (tl xs)" by (simp only: drop_tl) lemma nth_via_drop: "drop n xs = y#ys \ xs!n = y" by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits) lemma take_Suc_conv_app_nth: "i < length xs \ take (Suc i) xs = take i xs @ [xs!i]" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma Cons_nth_drop_Suc: "i < length xs \ (xs!i) # (drop (Suc i) xs) = drop i xs" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma length_take [simp]: "length (take n xs) = min (length xs) n" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma length_drop [simp]: "length (drop n xs) = (length xs - n)" by (induct n arbitrary: xs) (auto, case_tac xs, auto) -lemma take_all [simp]: "length xs \ n ==> take n xs = xs" +lemma take_all [simp]: "length xs \ n \ take n xs = xs" by (induct n arbitrary: xs) (auto, case_tac xs, auto) -lemma drop_all [simp]: "length xs \ n ==> drop n xs = []" +lemma drop_all [simp]: "length xs \ n \ drop n xs = []" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_append [simp]: "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma drop_append [simp]: "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys" by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_take [simp]: "take n (take m xs) = take (min n m) xs" proof (induct m arbitrary: xs n) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases n) simp_all qed lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs" proof (induct m arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)" proof (induct m arbitrary: xs n) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases n) simp_all qed lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)" by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split) lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \ xs = [])" by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs \ n)" by (induct xs arbitrary: n) (auto simp: drop_Cons split:nat.split) lemma take_map: "take n (map f xs) = map f (take n xs)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma drop_map: "drop n (map f xs) = map f (drop n xs)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)" by (cases "length xs < n") (auto simp: rev_take) lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)" by (cases "length xs < n") (auto simp: rev_drop) -lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i" +lemma nth_take [simp]: "i < n \ (take n xs)!i = xs!i" proof (induct xs arbitrary: i n) case Nil then show ?case by simp next case Cons then show ?case by (cases n; cases i) simp_all qed lemma nth_drop [simp]: - "n \ length xs ==> (drop n xs)!i = xs!(n + i)" + "n \ length xs \ (drop n xs)!i = xs!(n + i)" proof (induct n arbitrary: xs) case 0 then show ?case by simp next case Suc then show ?case by (cases xs) simp_all qed lemma butlast_take: - "n \ length xs ==> butlast (take n xs) = take (n - 1) xs" + "n \ length xs \ butlast (take n xs) = take (n - 1) xs" by (simp add: butlast_conv_take min.absorb1 min.absorb2) lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)" by (simp add: butlast_conv_take drop_take ac_simps) -lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs" +lemma take_butlast: "n < length xs \ take n (butlast xs) = take n xs" by (simp add: butlast_conv_take min.absorb1) lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" by (simp add: butlast_conv_take drop_take ac_simps) lemma hd_drop_conv_nth: "n < length xs \ hd(drop n xs) = xs!n" by(simp add: hd_conv_nth) lemma set_take_subset_set_take: "m \ n \ set(take m xs) \ set(take n xs)" proof (induct xs arbitrary: m n) case (Cons x xs m n) then show ?case by (cases n) (auto simp: take_Cons) qed simp lemma set_take_subset: "set(take n xs) \ set xs" by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split) lemma set_drop_subset: "set(drop n xs) \ set xs" by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split) lemma set_drop_subset_set_drop: "m \ n \ set(drop m xs) \ set(drop n xs)" proof (induct xs arbitrary: m n) case (Cons x xs m n) then show ?case by (clarsimp simp: drop_Cons split: nat.split) (metis set_drop_subset subset_iff) qed simp lemma in_set_takeD: "x \ set(take n xs) \ x \ set xs" using set_take_subset by fast lemma in_set_dropD: "x \ set(drop n xs) \ x \ set xs" using set_drop_subset by fast lemma append_eq_conv_conj: "(xs @ ys = zs) = (xs = take (length xs) zs \ ys = drop (length xs) zs)" proof (induct xs arbitrary: zs) case (Cons x xs zs) then show ?case by (cases zs, auto) qed auto lemma map_eq_append_conv: "map f xs = ys @ zs \ (\us vs. xs = us @ vs \ ys = map f us \ zs = map f vs)" proof - have "map f xs \ ys @ zs \ map f xs \ ys @ zs \ map f xs \ ys @ zs \ map f xs = ys @ zs \ (\bs bsa. xs = bs @ bsa \ ys = map f bs \ zs = map f bsa)" by (metis append_eq_conv_conj append_take_drop_id drop_map take_map) then show ?thesis using map_append by blast qed lemma append_eq_map_conv: "ys @ zs = map f xs \ (\us vs. xs = us @ vs \ ys = map f us \ zs = map f vs)" by (metis map_eq_append_conv) lemma take_add: "take (i+j) xs = take i xs @ take j (drop i xs)" proof (induct xs arbitrary: i) case (Cons x xs i) then show ?case by (cases i, auto) qed auto lemma append_eq_append_conv_if: "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) = (if size xs\<^sub>1 \ size ys\<^sub>1 then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \ xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2 else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \ drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)" proof (induct xs\<^sub>1 arbitrary: ys\<^sub>1) case (Cons a xs\<^sub>1 ys\<^sub>1) then show ?case by (cases ys\<^sub>1, auto) qed auto lemma take_hd_drop: "n < length xs \ take n xs @ [hd (drop n xs)] = take (Suc n) xs" by (induct xs arbitrary: n) (simp_all add:drop_Cons split:nat.split) lemma id_take_nth_drop: "i < length xs \ xs = take i xs @ xs!i # drop (Suc i) xs" proof - assume si: "i < length xs" hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto moreover from si have "take (Suc i) xs = take i xs @ [xs!i]" using take_Suc_conv_app_nth by blast ultimately show ?thesis by auto qed lemma take_update_cancel[simp]: "n \ m \ take n (xs[m := y]) = take n xs" by(simp add: list_eq_iff_nth_eq) lemma drop_update_cancel[simp]: "n < m \ drop m (xs[n := x]) = drop m xs" by(simp add: list_eq_iff_nth_eq) lemma upd_conv_take_nth_drop: "i < length xs \ xs[i:=a] = take i xs @ a # drop (Suc i) xs" proof - assume i: "i < length xs" have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]" by(rule arg_cong[OF id_take_nth_drop[OF i]]) also have "\ = take i xs @ a # drop (Suc i) xs" using i by (simp add: list_update_append) finally show ?thesis . qed lemma take_update_swap: "take m (xs[n := x]) = (take m xs)[n := x]" proof (cases "n \ length xs") case False then show ?thesis by (simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc split: nat.split) qed auto lemma drop_update_swap: assumes "m \ n" shows "drop m (xs[n := x]) = (drop m xs)[n-m := x]" proof (cases "n \ length xs") case False with assms show ?thesis by (simp add: upd_conv_take_nth_drop drop_take) qed auto lemma nth_image: "l \ size xs \ nth xs ` {0..\<^const>\takeWhile\ and \<^const>\dropWhile\\ lemma length_takeWhile_le: "length (takeWhile P xs) \ length xs" by (induct xs) auto lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs" by (induct xs) auto lemma takeWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ takeWhile P (xs @ ys) = takeWhile P xs" by (induct xs) auto lemma takeWhile_append2 [simp]: "(\x. x \ set xs \ P x) \ takeWhile P (xs @ ys) = xs @ takeWhile P ys" by (induct xs) auto lemma takeWhile_tail: "\ P x \ takeWhile P (xs @ (x#l)) = takeWhile P xs" by (induct xs) auto lemma takeWhile_nth: "j < length (takeWhile P xs) \ takeWhile P xs ! j = xs ! j" by (metis nth_append takeWhile_dropWhile_id) lemma dropWhile_nth: "j < length (dropWhile P xs) \ dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))" by (metis add.commute nth_append_length_plus takeWhile_dropWhile_id) lemma length_dropWhile_le: "length (dropWhile P xs) \ length xs" by (induct xs) auto lemma dropWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ dropWhile P (xs @ ys) = (dropWhile P xs)@ys" by (induct xs) auto lemma dropWhile_append2 [simp]: - "(\x. x \ set xs \ P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys" + "(\x. x \ set xs \ P(x)) \ dropWhile P (xs @ ys) = dropWhile P ys" by (induct xs) auto lemma dropWhile_append3: "\ P y \dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys" by (induct xs) auto lemma dropWhile_last: "x \ set xs \ \ P x \ last (dropWhile P xs) = last xs" by (auto simp add: dropWhile_append3 in_set_conv_decomp) lemma set_dropWhileD: "x \ set (dropWhile P xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) lemma set_takeWhileD: "x \ set (takeWhile P xs) \ x \ set xs \ P x" by (induct xs) (auto split: if_split_asm) lemma takeWhile_eq_all_conv[simp]: "(takeWhile P xs = xs) = (\x \ set xs. P x)" by(induct xs, auto) lemma dropWhile_eq_Nil_conv[simp]: "(dropWhile P xs = []) = (\x \ set xs. P x)" by(induct xs, auto) lemma dropWhile_eq_Cons_conv: "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \ \ P y)" by(induct xs, auto) -lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)" +lemma distinct_takeWhile[simp]: "distinct xs \ distinct (takeWhile P xs)" by (induct xs) (auto dest: set_takeWhileD) -lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)" +lemma distinct_dropWhile[simp]: "distinct xs \ distinct (dropWhile P xs)" by (induct xs) auto lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \ f) xs)" by (induct xs) auto lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \ f) xs)" by (induct xs) auto lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs" by (induct xs) auto lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs" by (induct xs) auto lemma hd_dropWhile: "dropWhile P xs \ [] \ \ P (hd (dropWhile P xs))" by (induct xs) auto lemma takeWhile_eq_filter: assumes "\ x. x \ set (dropWhile P xs) \ \ P x" shows "takeWhile P xs = filter P xs" proof - have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)" by simp have B: "filter P (dropWhile P xs) = []" unfolding filter_empty_conv using assms by blast have "filter P xs = takeWhile P xs" unfolding A filter_append B by (auto simp add: filter_id_conv dest: set_takeWhileD) thus ?thesis .. qed lemma takeWhile_eq_take_P_nth: "\ \ i. \ i < n ; i < length xs \ \ P (xs ! i) ; n < length xs \ \ P (xs ! n) \ \ takeWhile P xs = take n xs" proof (induct xs arbitrary: n) case Nil thus ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 with Cons show ?thesis by simp next case [simp]: (Suc n') have "P x" using Cons.prems(1)[of 0] by simp moreover have "takeWhile P xs = take n' xs" proof (rule Cons.hyps) fix i assume "i < n'" "i < length xs" thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp next assume "n' < length xs" thus "\ P (xs ! n')" using Cons by auto qed ultimately show ?thesis by simp qed qed lemma nth_length_takeWhile: "length (takeWhile P xs) < length xs \ \ P (xs ! length (takeWhile P xs))" by (induct xs) auto lemma length_takeWhile_less_P_nth: assumes all: "\ i. i < j \ P (xs ! i)" and "j \ length xs" shows "j \ length (takeWhile P xs)" proof (rule classical) assume "\ ?thesis" hence "length (takeWhile P xs) < length xs" using assms by simp thus ?thesis using all \\ ?thesis\ nth_length_takeWhile[of P xs] by auto qed lemma takeWhile_neq_rev: "\distinct xs; x \ set xs\ \ takeWhile (\y. y \ x) (rev xs) = rev (tl (dropWhile (\y. y \ x) xs))" by(induct xs) (auto simp: takeWhile_tail[where l="[]"]) lemma dropWhile_neq_rev: "\distinct xs; x \ set xs\ \ dropWhile (\y. y \ x) (rev xs) = x # rev (takeWhile (\y. y \ x) xs)" proof (induct xs) case (Cons a xs) then show ?case by(auto, subst dropWhile_append2, auto) qed simp lemma takeWhile_not_last: "distinct xs \ takeWhile (\y. y \ last xs) xs = butlast xs" by(induction xs rule: induct_list012) auto lemma takeWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ takeWhile P l = takeWhile Q k" by (induct k arbitrary: l) (simp_all) lemma dropWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ dropWhile P l = dropWhile Q k" by (induct k arbitrary: l, simp_all) lemma takeWhile_idem [simp]: "takeWhile P (takeWhile P xs) = takeWhile P xs" by (induct xs) auto lemma dropWhile_idem [simp]: "dropWhile P (dropWhile P xs) = dropWhile P xs" by (induct xs) auto subsubsection \\<^const>\zip\\ lemma zip_Nil [simp]: "zip [] ys = []" by (induct ys) auto lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys" by simp declare zip_Cons [simp del] lemma [code]: "zip [] ys = []" "zip xs [] = []" "zip (x # xs) (y # ys) = (x, y) # zip xs ys" by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+ lemma zip_Cons1: "zip (x#xs) ys = (case ys of [] \ [] | y#ys \ (x,y)#zip xs ys)" by(auto split:list.split) lemma length_zip [simp]: "length (zip xs ys) = min (length xs) (length ys)" by (induct xs ys rule:list_induct2') auto lemma zip_obtain_same_length: assumes "\zs ws n. length zs = length ws \ n = min (length xs) (length ys) \ zs = take n xs \ ws = take n ys \ P (zip zs ws)" shows "P (zip xs ys)" proof - let ?n = "min (length xs) (length ys)" have "P (zip (take ?n xs) (take ?n ys))" by (rule assms) simp_all moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases ys) simp_all qed ultimately show ?thesis by simp qed lemma zip_append1: "zip (xs @ ys) zs = zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)" by (induct xs zs rule:list_induct2') auto lemma zip_append2: "zip xs (ys @ zs) = zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs" by (induct xs ys rule:list_induct2') auto lemma zip_append [simp]: - "[| length xs = length us |] ==> + "\length xs = length us\ \ zip (xs@ys) (us@vs) = zip xs us @ zip ys vs" by (simp add: zip_append1) lemma zip_rev: - "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)" + "length xs = length ys \ zip (rev xs) (rev ys) = rev (zip xs ys)" by (induct rule:list_induct2, simp_all) lemma zip_map_map: "zip (map f xs) (map g ys) = map (\ (x, y). (f x, g y)) (zip xs ys)" proof (induct xs arbitrary: ys) case (Cons x xs) note Cons_x_xs = Cons.hyps show ?case proof (cases ys) case (Cons y ys') show ?thesis unfolding Cons using Cons_x_xs by simp qed simp qed simp lemma zip_map1: "zip (map f xs) ys = map (\(x, y). (f x, y)) (zip xs ys)" using zip_map_map[of f xs "\x. x" ys] by simp lemma zip_map2: "zip xs (map f ys) = map (\(x, y). (x, f y)) (zip xs ys)" using zip_map_map[of "\x. x" xs f ys] by simp lemma map_zip_map: "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)" by (auto simp: zip_map1) lemma map_zip_map2: "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)" by (auto simp: zip_map2) text\Courtesy of Andreas Lochbihler:\ lemma zip_same_conv_map: "zip xs xs = map (\x. (x, x)) xs" by(induct xs) auto lemma nth_zip [simp]: - "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)" + "\i < length xs; i < length ys\ \ (zip xs ys)!i = (xs!i, ys!i)" proof (induct ys arbitrary: i xs) case (Cons y ys) then show ?case by (cases xs) (simp_all add: nth.simps split: nat.split) qed auto lemma set_zip: "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}" by(simp add: set_conv_nth cong: rev_conj_cong) lemma zip_same: "((a,b) \ set (zip xs xs)) = (a \ set xs \ a = b)" by(induct xs) auto lemma zip_update: "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]" by (simp add: update_zip) lemma zip_replicate [simp]: "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)" proof (induct i arbitrary: j) case (Suc i) then show ?case by (cases j, auto) qed auto lemma zip_replicate1: "zip (replicate n x) ys = map (Pair x) (take n ys)" by(induction ys arbitrary: n)(case_tac [2] n, simp_all) lemma take_zip: "take n (zip xs ys) = zip (take n xs) (take n ys)" proof (induct n arbitrary: xs ys) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases ys) simp_all qed lemma drop_zip: "drop n (zip xs ys) = zip (drop n xs) (drop n ys)" proof (induct n arbitrary: xs ys) case 0 then show ?case by simp next case Suc then show ?case by (cases xs; cases ys) simp_all qed lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \ fst) (zip xs ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \ snd) (zip xs ys)" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma set_zip_leftD: "(x,y)\ set (zip xs ys) \ x \ set xs" by (induct xs ys rule:list_induct2') auto lemma set_zip_rightD: "(x,y)\ set (zip xs ys) \ y \ set ys" by (induct xs ys rule:list_induct2') auto lemma in_set_zipE: "(x,y) \ set(zip xs ys) \ (\ x \ set xs; y \ set ys \ \ R) \ R" by(blast dest: set_zip_leftD set_zip_rightD) lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs" by (induct zs) simp_all lemma zip_eq_conv: "length xs = length ys \ zip xs ys = zs \ map fst zs = xs \ map snd zs = ys" by (auto simp add: zip_map_fst_snd) lemma in_set_zip: "p \ set (zip xs ys) \ (\n. xs ! n = fst p \ ys ! n = snd p \ n < length xs \ n < length ys)" by (cases p) (auto simp add: set_zip) lemma in_set_impl_in_set_zip1: assumes "length xs = length ys" assumes "x \ set xs" obtains y where "(x, y) \ set (zip xs ys)" proof - from assms have "x \ set (map fst (zip xs ys))" by simp from this that show ?thesis by fastforce qed lemma in_set_impl_in_set_zip2: assumes "length xs = length ys" assumes "y \ set ys" obtains x where "(x, y) \ set (zip xs ys)" proof - from assms have "y \ set (map snd (zip xs ys))" by simp from this that show ?thesis by fastforce qed lemma zip_eq_Nil_iff: "zip xs ys = [] \ xs = [] \ ys = []" by (cases xs; cases ys) simp_all lemma zip_eq_ConsE: assumes "zip xs ys = xy # xys" obtains x xs' y ys' where "xs = x # xs'" and "ys = y # ys'" and "xy = (x, y)" and "xys = zip xs' ys'" proof - from assms have "xs \ []" and "ys \ []" using zip_eq_Nil_iff [of xs ys] by simp_all then obtain x xs' y ys' where xs: "xs = x # xs'" and ys: "ys = y # ys'" by (cases xs; cases ys) auto with assms have "xy = (x, y)" and "xys = zip xs' ys'" by simp_all with xs ys show ?thesis .. qed lemma semilattice_map2: "semilattice (map2 (\<^bold>*))" if "semilattice (\<^bold>*)" for f (infixl "\<^bold>*" 70) proof - from that interpret semilattice f . show ?thesis proof show "map2 (\<^bold>*) (map2 (\<^bold>*) xs ys) zs = map2 (\<^bold>*) xs (map2 (\<^bold>*) ys zs)" for xs ys zs :: "'a list" proof (induction "zip xs (zip ys zs)" arbitrary: xs ys zs) case Nil from Nil [symmetric] show ?case by (auto simp add: zip_eq_Nil_iff) next case (Cons xyz xyzs) from Cons.hyps(2) [symmetric] show ?case by (rule zip_eq_ConsE) (erule zip_eq_ConsE, auto intro: Cons.hyps(1) simp add: ac_simps) qed show "map2 (\<^bold>*) xs ys = map2 (\<^bold>*) ys xs" for xs ys :: "'a list" proof (induction "zip xs ys" arbitrary: xs ys) case Nil then show ?case by (auto simp add: zip_eq_Nil_iff dest: sym) next case (Cons xy xys) from Cons.hyps(2) [symmetric] show ?case by (rule zip_eq_ConsE) (auto intro: Cons.hyps(1) simp add: ac_simps) qed show "map2 (\<^bold>*) xs xs = xs" for xs :: "'a list" by (induction xs) simp_all qed qed lemma pair_list_eqI: assumes "map fst xs = map fst ys" and "map snd xs = map snd ys" shows "xs = ys" proof - from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq) from this assms show ?thesis by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI) qed lemma hd_zip: \hd (zip xs ys) = (hd xs, hd ys)\ if \xs \ []\ and \ys \ []\ using that by (cases xs; cases ys) simp_all lemma last_zip: \last (zip xs ys) = (last xs, last ys)\ if \xs \ []\ and \ys \ []\ and \length xs = length ys\ using that by (cases xs rule: rev_cases; cases ys rule: rev_cases) simp_all subsubsection \\<^const>\list_all2\\ lemma list_all2_lengthD [intro?]: - "list_all2 P xs ys ==> length xs = length ys" + "list_all2 P xs ys \ length xs = length ys" by (simp add: list_all2_iff) lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])" by (simp add: list_all2_iff) lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])" by (simp add: list_all2_iff) lemma list_all2_Cons [iff, code]: "list_all2 P (x # xs) (y # ys) = (P x y \ list_all2 P xs ys)" by (auto simp add: list_all2_iff) lemma list_all2_Cons1: "list_all2 P (x # xs) ys = (\z zs. ys = z # zs \ P x z \ list_all2 P xs zs)" by (cases ys) auto lemma list_all2_Cons2: "list_all2 P xs (y # ys) = (\z zs. xs = z # zs \ P z y \ list_all2 P zs ys)" by (cases xs) auto lemma list_all2_induct [consumes 1, case_names Nil Cons, induct set: list_all2]: assumes P: "list_all2 P xs ys" assumes Nil: "R [] []" assumes Cons: "\x xs y ys. \P x y; list_all2 P xs ys; R xs ys\ \ R (x # xs) (y # ys)" shows "R xs ys" using P by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons) lemma list_all2_rev [iff]: "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys" by (simp add: list_all2_iff zip_rev cong: conj_cong) lemma list_all2_rev1: "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)" by (subst list_all2_rev [symmetric]) simp lemma list_all2_append1: "list_all2 P (xs @ ys) zs = (\us vs. zs = us @ vs \ length us = length xs \ length vs = length ys \ list_all2 P xs us \ list_all2 P ys vs)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule_tac x = "take (length xs) zs" in exI) apply (rule_tac x = "drop (length xs) zs" in exI) apply (force split: nat_diff_split simp add: list_all2_iff zip_append1) done next assume ?rhs then show ?lhs by (auto simp add: list_all2_iff) qed lemma list_all2_append2: "list_all2 P xs (ys @ zs) = (\us vs. xs = us @ vs \ length us = length ys \ length vs = length zs \ list_all2 P us ys \ list_all2 P vs zs)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule_tac x = "take (length ys) xs" in exI) apply (rule_tac x = "drop (length ys) xs" in exI) apply (force split: nat_diff_split simp add: list_all2_iff zip_append2) done next assume ?rhs then show ?lhs by (auto simp add: list_all2_iff) qed lemma list_all2_append: "length xs = length ys \ list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \ list_all2 P us vs)" by (induct rule:list_induct2, simp_all) lemma list_all2_appendI [intro?, trans]: "\ list_all2 P a b; list_all2 P c d \ \ list_all2 P (a@c) (b@d)" by (simp add: list_all2_append list_all2_lengthD) lemma list_all2_conv_all_nth: "list_all2 P xs ys = (length xs = length ys \ (\i < length xs. P (xs!i) (ys!i)))" by (force simp add: list_all2_iff set_zip) lemma list_all2_trans: - assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c" - shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs" + assumes tr: "!!a b c. P1 a b \ P2 b c \ P3 a c" + shows "!!bs cs. list_all2 P1 as bs \ list_all2 P2 bs cs \ list_all2 P3 as cs" (is "!!bs cs. PROP ?Q as bs cs") proof (induct as) fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs" show "!!cs. PROP ?Q (x # xs) bs cs" proof (induct bs) fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs" show "PROP ?Q (x # xs) (y # ys) cs" by (induct cs) (auto intro: tr I1 I2) qed simp qed simp lemma list_all2_all_nthI [intro?]: "length a = length b \ (\n. n < length a \ P (a!n) (b!n)) \ list_all2 P a b" by (simp add: list_all2_conv_all_nth) lemma list_all2I: "\x \ set (zip a b). case_prod P x \ length a = length b \ list_all2 P a b" by (simp add: list_all2_iff) lemma list_all2_nthD: "\ list_all2 P xs ys; p < size xs \ \ P (xs!p) (ys!p)" by (simp add: list_all2_conv_all_nth) lemma list_all2_nthD2: "\list_all2 P xs ys; p < size ys\ \ P (xs!p) (ys!p)" by (frule list_all2_lengthD) (auto intro: list_all2_nthD) lemma list_all2_map1: "list_all2 P (map f as) bs = list_all2 (\x y. P (f x) y) as bs" by (simp add: list_all2_conv_all_nth) lemma list_all2_map2: "list_all2 P as (map f bs) = list_all2 (\x y. P x (f y)) as bs" by (auto simp add: list_all2_conv_all_nth) lemma list_all2_refl [intro?]: "(\x. P x x) \ list_all2 P xs xs" by (simp add: list_all2_conv_all_nth) lemma list_all2_update_cong: "\ list_all2 P xs ys; P x y \ \ list_all2 P (xs[i:=x]) (ys[i:=y])" by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update) lemma list_all2_takeI [simp,intro?]: "list_all2 P xs ys \ list_all2 P (take n xs) (take n ys)" proof (induct xs arbitrary: n ys) case (Cons x xs) then show ?case by (cases n) (auto simp: list_all2_Cons1) qed auto lemma list_all2_dropI [simp,intro?]: "list_all2 P xs ys \ list_all2 P (drop n xs) (drop n ys)" proof (induct xs arbitrary: n ys) case (Cons x xs) then show ?case by (cases n) (auto simp: list_all2_Cons1) qed auto lemma list_all2_mono [intro?]: "list_all2 P xs ys \ (\xs ys. P xs ys \ Q xs ys) \ list_all2 Q xs ys" by (rule list.rel_mono_strong) lemma list_all2_eq: "xs = ys \ list_all2 (=) xs ys" by (induct xs ys rule: list_induct2') auto lemma list_eq_iff_zip_eq: "xs = ys \ length xs = length ys \ (\(x,y) \ set (zip xs ys). x = y)" by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong) lemma list_all2_same: "list_all2 P xs xs \ (\x\set xs. P x x)" by(auto simp add: list_all2_conv_all_nth set_conv_nth) lemma zip_assoc: "zip xs (zip ys zs) = map (\((x, y), z). (x, y, z)) (zip (zip xs ys) zs)" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_commute: "zip xs ys = map (\(x, y). (y, x)) (zip ys xs)" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_left_commute: "zip xs (zip ys zs) = map (\(y, (x, z)). (x, y, z)) (zip ys (zip xs zs))" by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_replicate2: "zip xs (replicate n y) = map (\x. (x, y)) (take n xs)" by(subst zip_commute)(simp add: zip_replicate1) subsubsection \\<^const>\List.product\ and \<^const>\product_lists\\ lemma product_concat_map: "List.product xs ys = concat (map (\x. map (\y. (x,y)) ys) xs)" by(induction xs) (simp)+ lemma set_product[simp]: "set (List.product xs ys) = set xs \ set ys" by (induct xs) auto lemma length_product [simp]: "length (List.product xs ys) = length xs * length ys" by (induct xs) simp_all lemma product_nth: assumes "n < length xs * length ys" shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))" using assms proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs n) then have "length ys > 0" by auto with Cons show ?case by (auto simp add: nth_append not_less le_mod_geq le_div_geq) qed lemma in_set_product_lists_length: "xs \ set (product_lists xss) \ length xs = length xss" by (induct xss arbitrary: xs) auto lemma product_lists_set: "set (product_lists xss) = {xs. list_all2 (\x ys. x \ set ys) xs xss}" (is "?L = Collect ?R") proof (intro equalityI subsetI, unfold mem_Collect_eq) fix xs assume "xs \ ?L" then have "length xs = length xss" by (rule in_set_product_lists_length) from this \xs \ ?L\ show "?R xs" by (induct xs xss rule: list_induct2) auto next fix xs assume "?R xs" then show "xs \ ?L" by induct auto qed subsubsection \\<^const>\fold\ with natural argument order\ lemma fold_simps [code]: \ \eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\ "fold f [] s = s" "fold f (x # xs) s = fold f xs (f x s)" by simp_all lemma fold_remove1_split: "\ \x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x; x \ set xs \ \ fold f xs = fold f (remove1 x xs) \ f x" by (induct xs) (auto simp add: comp_assoc) lemma fold_cong [fundef_cong]: "a = b \ xs = ys \ (\x. x \ set xs \ f x = g x) \ fold f xs a = fold g ys b" by (induct ys arbitrary: a b xs) simp_all lemma fold_id: "(\x. x \ set xs \ f x = id) \ fold f xs = id" by (induct xs) simp_all lemma fold_commute: "(\x. x \ set xs \ h \ g x = f x \ h) \ h \ fold g xs = fold f xs \ h" by (induct xs) (simp_all add: fun_eq_iff) lemma fold_commute_apply: assumes "\x. x \ set xs \ h \ g x = f x \ h" shows "h (fold g xs s) = fold f xs (h s)" proof - from assms have "h \ fold g xs = fold f xs \ h" by (rule fold_commute) then show ?thesis by (simp add: fun_eq_iff) qed lemma fold_invariant: "\ \x. x \ set xs \ Q x; P s; \x s. Q x \ P s \ P (f x s) \ \ P (fold f xs s)" by (induct xs arbitrary: s) simp_all lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \ fold f xs" by (induct xs) simp_all lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \ f) xs" by (induct xs) simp_all lemma fold_filter: "fold f (filter P xs) = fold (\x. if P x then f x else id) xs" by (induct xs) simp_all lemma fold_rev: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ fold f (rev xs) = fold f xs" by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) lemma fold_Cons_rev: "fold Cons xs = append (rev xs)" by (induct xs) simp_all lemma rev_conv_fold [code]: "rev xs = fold Cons xs []" by (simp add: fold_Cons_rev) lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))" by (induct xss) simp_all text \\<^const>\Finite_Set.fold\ and \<^const>\fold\\ lemma (in comp_fun_commute) fold_set_fold_remdups: "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb) lemma (in comp_fun_idem) fold_set_fold: "Finite_Set.fold f y (set xs) = fold f xs y" by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm) lemma union_set_fold [code]: "set xs \ A = fold Set.insert xs A" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by (simp add: union_fold_insert fold_set_fold) qed lemma union_coset_filter [code]: "List.coset xs \ A = List.coset (List.filter (\x. x \ A) xs)" by auto lemma minus_set_fold [code]: "A - set xs = fold Set.remove xs A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) show ?thesis by (simp add: minus_fold_remove [of _ A] fold_set_fold) qed lemma minus_coset_filter [code]: "A - List.coset xs = set (List.filter (\x. x \ A) xs)" by auto lemma inter_set_filter [code]: "A \ set xs = set (List.filter (\x. x \ A) xs)" by auto lemma inter_coset_fold [code]: "A \ List.coset xs = fold Set.remove xs A" by (simp add: Diff_eq [symmetric] minus_set_fold) lemma (in semilattice_set) set_eq_fold [code]: "F (set (x # xs)) = fold f xs x" proof - interpret comp_fun_idem f by standard (simp_all add: fun_eq_iff left_commute) show ?thesis by (simp add: eq_fold fold_set_fold) qed lemma (in complete_lattice) Inf_set_fold: "Inf (set xs) = fold inf xs top" proof - interpret comp_fun_idem "inf :: 'a \ 'a \ 'a" by (fact comp_fun_idem_inf) show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute) qed declare Inf_set_fold [where 'a = "'a set", code] lemma (in complete_lattice) Sup_set_fold: "Sup (set xs) = fold sup xs bot" proof - interpret comp_fun_idem "sup :: 'a \ 'a \ 'a" by (fact comp_fun_idem_sup) show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute) qed declare Sup_set_fold [where 'a = "'a set", code] lemma (in complete_lattice) INF_set_fold: "\(f ` set xs) = fold (inf \ f) xs top" using Inf_set_fold [of "map f xs"] by (simp add: fold_map) lemma (in complete_lattice) SUP_set_fold: "\(f ` set xs) = fold (sup \ f) xs bot" using Sup_set_fold [of "map f xs"] by (simp add: fold_map) subsubsection \Fold variants: \<^const>\foldr\ and \<^const>\foldl\\ text \Correspondence\ lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)" by (induct xs) simp_all lemma foldl_conv_fold: "foldl f s xs = fold (\x s. f s x) xs s" by (induct xs arbitrary: s) simp_all lemma foldr_conv_foldl: \ \The ``Third Duality Theorem'' in Bird \& Wadler:\ "foldr f xs a = foldl (\x y. f y x) a (rev xs)" by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldl_conv_foldr: "foldl f a xs = foldr (\x y. f y x) (rev xs) a" by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldr_fold: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ foldr f xs = fold f xs" unfolding foldr_conv_fold by (rule fold_rev) lemma foldr_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f x a = g x a) \ foldr f l a = foldr g k b" by (auto simp add: foldr_conv_fold intro!: fold_cong) lemma foldl_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f a x = g a x) \ foldl f a l = foldl g b k" by (auto simp add: foldl_conv_fold intro!: fold_cong) lemma foldr_append [simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)" by (simp add: foldr_conv_fold) lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys" by (simp add: foldl_conv_fold) lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \ f) xs a" by (simp add: foldr_conv_fold fold_map rev_map) lemma foldr_filter: "foldr f (filter P xs) = foldr (\x. if P x then f x else id) xs" by (simp add: foldr_conv_fold rev_filter fold_filter) lemma foldl_map [code_unfold]: "foldl g a (map f xs) = foldl (\a x. g a (f x)) a xs" by (simp add: foldl_conv_fold fold_map comp_def) lemma concat_conv_foldr [code]: "concat xss = foldr append xss []" by (simp add: fold_append_concat_rev foldr_conv_fold) subsubsection \\<^const>\upt\\ lemma upt_rec[code]: "[i.. \simp does not terminate!\ by (induct j) auto lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n -lemma upt_conv_Nil [simp]: "j \ i ==> [i.. i \ [i.. j \ i)" by(induct j)simp_all lemma upt_eq_Cons_conv: "([i.. i = x \ [i+1.. j ==> [i..<(Suc j)] = [i.. j \ [i..<(Suc j)] = [i.. \Only needed if \upt_Suc\ is deleted from the simpset.\ by simp -lemma upt_conv_Cons: "i < j ==> [i.. [i.. \no precondition\ "m # n # ns = [m.. n # ns = [Suc m.. [i.. [i.. \LOOPS as a simprule, since \j \ j\.\ by (induct k) auto lemma length_upt [simp]: "length [i.. [i.. [i.. hd[i.. last[i.. n ==> take m [i.. n \ take m [i..i. i + n) [0.. (map f [m.. (map f [m..n. n - Suc 0) [Suc m..i. f (Suc i)) [0 ..< n]" by (induct n arbitrary: f) auto lemma nth_take_lemma: "k \ length xs \ k \ length ys \ (\i. i < k \ xs!i = ys!i) \ take k xs = take k ys" proof (induct k arbitrary: xs ys) case (Suc k) then show ?case apply (simp add: less_Suc_eq_0_disj) by (simp add: Suc.prems(3) take_Suc_conv_app_nth) qed simp lemma nth_equalityI: "\length xs = length ys; \i. i < length xs \ xs!i = ys!i\ \ xs = ys" by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all lemma map_nth: "map (\i. xs ! i) [0.. (\x y. \P x y; Q y x\ \ x = y); list_all2 P xs ys; list_all2 Q ys xs \ \ xs = ys" by (simp add: list_all2_conv_all_nth nth_equalityI) -lemma take_equalityI: "(\i. take i xs = take i ys) ==> xs = ys" +lemma take_equalityI: "(\i. take i xs = take i ys) \ xs = ys" \ \The famous take-lemma.\ by (metis length_take min.commute order_refl take_all) lemma take_Cons': "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)" by (cases n) simp_all lemma drop_Cons': "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)" by (cases n) simp_all lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))" by (cases n) simp_all lemma take_Cons_numeral [simp]: "take (numeral v) (x # xs) = x # take (numeral v - 1) xs" by (simp add: take_Cons') lemma drop_Cons_numeral [simp]: "drop (numeral v) (x # xs) = drop (numeral v - 1) xs" by (simp add: drop_Cons') lemma nth_Cons_numeral [simp]: "(x # xs) ! numeral v = xs ! (numeral v - 1)" by (simp add: nth_Cons') subsubsection \\upto\: interval-list on \<^typ>\int\\ function upto :: "int \ int \ int list" ("(1[_../_])") where "upto i j = (if i \ j then i # [i+1..j] else [])" by auto termination by(relation "measure(%(i::int,j). nat(j - i + 1))") auto declare upto.simps[simp del] lemmas upto_rec_numeral [simp] = upto.simps[of "numeral m" "numeral n"] upto.simps[of "numeral m" "- numeral n"] upto.simps[of "- numeral m" "numeral n"] upto.simps[of "- numeral m" "- numeral n"] for m n lemma upto_empty[simp]: "j < i \ [i..j] = []" by(simp add: upto.simps) lemma upto_single[simp]: "[i..i] = [i]" by(simp add: upto.simps) lemma upto_Nil[simp]: "[i..j] = [] \ j < i" by (simp add: upto.simps) lemma upto_Nil2[simp]: "[] = [i..j] \ j < i" by (simp add: upto.simps) lemma upto_rec1: "i \ j \ [i..j] = i#[i+1..j]" by(simp add: upto.simps) lemma upto_rec2: "i \ j \ [i..j] = [i..j - 1]@[j]" proof(induct "nat(j-i)" arbitrary: i j) case 0 thus ?case by(simp add: upto.simps) next case (Suc n) hence "n = nat (j - (i + 1))" "i < j" by linarith+ from this(2) Suc.hyps(1)[OF this(1)] Suc(2,3) upto_rec1 show ?case by simp qed lemma length_upto[simp]: "length [i..j] = nat(j - i + 1)" by(induction i j rule: upto.induct) (auto simp: upto.simps) lemma set_upto[simp]: "set[i..j] = {i..j}" proof(induct i j rule:upto.induct) case (1 i j) from this show ?case unfolding upto.simps[of i j] by auto qed lemma nth_upto[simp]: "i + int k \ j \ [i..j] ! k = i + int k" - apply(induction i j arbitrary: k rule: upto.induct) -apply(subst upto_rec1) -apply(auto simp add: nth_Cons') -done +proof(induction i j arbitrary: k rule: upto.induct) + case (1 i j) + then show ?case + by (auto simp add: upto_rec1 [of i j] nth_Cons') +qed lemma upto_split1: "i \ j \ j \ k \ [i..k] = [i..j-1] @ [j..k]" proof (induction j rule: int_ge_induct) case base thus ?case by (simp add: upto_rec1) next case step thus ?case using upto_rec1 upto_rec2 by simp qed lemma upto_split2: "i \ j \ j \ k \ [i..k] = [i..j] @ [j+1..k]" using upto_rec1 upto_rec2 upto_split1 by auto lemma upto_split3: "\ i \ j; j \ k \ \ [i..k] = [i..j-1] @ j # [j+1..k]" using upto_rec1 upto_split1 by auto text\Tail recursive version for code generation:\ definition upto_aux :: "int \ int \ int list \ int list" where "upto_aux i j js = [i..j] @ js" lemma upto_aux_rec [code]: "upto_aux i j js = (if j\<^const>\distinct\ and \<^const>\remdups\ and \<^const>\remdups_adj\\ lemma distinct_tl: "distinct xs \ distinct (tl xs)" by (cases xs) simp_all lemma distinct_append [simp]: "distinct (xs @ ys) = (distinct xs \ distinct ys \ set xs \ set ys = {})" by (induct xs) auto lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs" by(induct xs) auto lemma set_remdups [simp]: "set (remdups xs) = set xs" by (induct xs) (auto simp add: insert_absorb) lemma distinct_remdups [iff]: "distinct (remdups xs)" by (induct xs) auto -lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs" +lemma distinct_remdups_id: "distinct xs \ remdups xs = xs" by (induct xs, auto) lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \ distinct xs" by (metis distinct_remdups distinct_remdups_id) lemma finite_distinct_list: "finite A \ \xs. set xs = A \ distinct xs" by (metis distinct_remdups finite_list set_remdups) lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])" by (induct x, auto) lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])" by (induct x, auto) lemma length_remdups_leq[iff]: "length(remdups xs) \ length xs" by (induct xs) auto lemma length_remdups_eq[iff]: "(length (remdups xs) = length xs) = (remdups xs = xs)" proof (induct xs) case (Cons a xs) then show ?case by simp (metis Suc_n_not_le_n impossible_Cons length_remdups_leq) qed auto lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)" by (induct xs) auto lemma distinct_map: "distinct(map f xs) = (distinct xs \ inj_on f (set xs))" by (induct xs) auto lemma distinct_map_filter: "distinct (map f xs) \ distinct (map f (filter P xs))" by (induct xs) auto -lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)" +lemma distinct_filter [simp]: "distinct xs \ distinct (filter P xs)" by (induct xs) auto lemma distinct_upt[simp]: "distinct[i.. distinct (take i xs)" proof (induct xs arbitrary: i) case (Cons a xs) then show ?case by (metis Cons.prems append_take_drop_id distinct_append) qed auto lemma distinct_drop[simp]: "distinct xs \ distinct (drop i xs)" proof (induct xs arbitrary: i) case (Cons a xs) then show ?case by (metis Cons.prems append_take_drop_id distinct_append) qed auto lemma distinct_list_update: assumes d: "distinct xs" and a: "a \ set xs - {xs!i}" shows "distinct (xs[i:=a])" proof (cases "i < length xs") case True with a have anot: "a \ set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}" by simp (metis in_set_dropD in_set_takeD) show ?thesis proof (cases "a = xs!i") case True with d show ?thesis by auto next case False + have "set (take i xs) \ set (drop (Suc i) xs) = {}" + by (metis True d disjoint_insert(1) distinct_append id_take_nth_drop list.set(2)) then show ?thesis - using d anot \i < length xs\ - apply (simp add: upd_conv_take_nth_drop) - by (metis disjoint_insert(1) distinct_append id_take_nth_drop set_simps(2)) + using d False anot \i < length xs\ by (simp add: upd_conv_take_nth_drop) qed next case False with d show ?thesis by auto qed lemma distinct_concat: "\ distinct xs; \ ys. ys \ set xs \ distinct ys; \ ys zs. \ ys \ set xs ; zs \ set xs ; ys \ zs \ \ set ys \ set zs = {} \ \ distinct (concat xs)" by (induct xs) auto text \It is best to avoid this indexed version of distinct, but sometimes it is useful.\ lemma distinct_conv_nth: "distinct xs = (\i < size xs. \j < size xs. i \ j \ xs!i \ xs!j)" proof (induct xs) case (Cons x xs) show ?case apply (auto simp add: Cons nth_Cons split: nat.split_asm) apply (metis Suc_less_eq2 in_set_conv_nth less_not_refl zero_less_Suc)+ done qed auto lemma nth_eq_iff_index_eq: "\ distinct xs; i < length xs; j < length xs \ \ (xs!i = xs!j) = (i = j)" by(auto simp: distinct_conv_nth) lemma distinct_Ex1: "distinct xs \ x \ set xs \ (\!i. i < length xs \ xs ! i = x)" by (auto simp: in_set_conv_nth nth_eq_iff_index_eq) lemma inj_on_nth: "distinct xs \ \i \ I. i < length xs \ inj_on (nth xs) I" by (rule inj_onI) (simp add: nth_eq_iff_index_eq) lemma bij_betw_nth: assumes "distinct xs" "A = {.. distinct xs; n < length xs \ \ set(xs[n := x]) = insert x (set xs - {xs!n})" by(auto simp: set_eq_iff in_set_conv_nth nth_list_update nth_eq_iff_index_eq) -lemma distinct_swap[simp]: "\ i < size xs; j < size xs \ \ +lemma distinct_swap[simp]: "\ i < size xs; j < size xs\ \ distinct(xs[i := xs!j, j := xs!i]) = distinct xs" apply (simp add: distinct_conv_nth nth_list_update) -apply safe -apply metis+ -done + apply (safe; metis) + done lemma set_swap[simp]: "\ i < size xs; j < size xs \ \ set(xs[i := xs!j, j := xs!i]) = set xs" by(simp add: set_conv_nth nth_list_update) metis -lemma distinct_card: "distinct xs ==> card (set xs) = size xs" +lemma distinct_card: "distinct xs \ card (set xs) = size xs" by (induct xs) auto -lemma card_distinct: "card (set xs) = size xs ==> distinct xs" +lemma card_distinct: "card (set xs) = size xs \ distinct xs" proof (induct xs) case (Cons x xs) show ?case proof (cases "x \ set xs") case False with Cons show ?thesis by simp next case True with Cons.prems have "card (set xs) = Suc (length xs)" by (simp add: card_insert_if split: if_split_asm) moreover have "card (set xs) \ length xs" by (rule card_length) ultimately have False by simp thus ?thesis .. qed qed simp lemma distinct_length_filter: "distinct xs \ length (filter P xs) = card ({x. P x} Int set xs)" by (induct xs) (auto) lemma not_distinct_decomp: "\ distinct ws \ \xs ys zs y. ws = xs@[y]@ys@[y]@zs" proof (induct n == "length ws" arbitrary:ws) case (Suc n ws) then show ?case using length_Suc_conv [of ws n] apply (auto simp: eq_commute) apply (metis append_Nil in_set_conv_decomp_first) by (metis append_Cons) qed simp lemma not_distinct_conv_prefix: defines "dec as xs y ys \ y \ set xs \ distinct xs \ as = xs @ y # ys" shows "\distinct as \ (\xs y ys. dec as xs y ys)" (is "?L = ?R") proof assume "?L" then show "?R" proof (induct "length as" arbitrary: as rule: less_induct) case less obtain xs ys zs y where decomp: "as = (xs @ y # ys) @ y # zs" using not_distinct_decomp[OF less.prems] by auto show ?case proof (cases "distinct (xs @ y # ys)") case True with decomp have "dec as (xs @ y # ys) y zs" by (simp add: dec_def) then show ?thesis by blast next case False with less decomp obtain xs' y' ys' where "dec (xs @ y # ys) xs' y' ys'" by atomize_elim auto with decomp have "dec as xs' y' (ys' @ y # zs)" by (simp add: dec_def) then show ?thesis by blast qed qed qed (auto simp: dec_def) lemma distinct_product: "distinct xs \ distinct ys \ distinct (List.product xs ys)" by (induct xs) (auto intro: inj_onI simp add: distinct_map) lemma distinct_product_lists: assumes "\xs \ set xss. distinct xs" shows "distinct (product_lists xss)" using assms proof (induction xss) case (Cons xs xss) note * = this then show ?case proof (cases "product_lists xss") case Nil then show ?thesis by (induct xs) simp_all next case (Cons ps pss) with * show ?thesis by (auto intro!: inj_onI distinct_concat simp add: distinct_map) qed qed simp lemma length_remdups_concat: "length (remdups (concat xss)) = card (\xs\set xss. set xs)" by (simp add: distinct_card [symmetric]) lemma remdups_append2: "remdups (xs @ remdups ys) = remdups (xs @ ys)" by(induction xs) auto lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)" proof - have xs: "concat[xs] = xs" by simp from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp qed lemma remdups_remdups: "remdups (remdups xs) = remdups xs" by (induct xs) simp_all lemma distinct_butlast: assumes "distinct xs" shows "distinct (butlast xs)" proof (cases "xs = []") case False from \xs \ []\ obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto with \distinct xs\ show ?thesis by simp qed (auto) lemma remdups_map_remdups: "remdups (map f (remdups xs)) = remdups (map f xs)" by (induct xs) simp_all lemma distinct_zipI1: assumes "distinct xs" shows "distinct (zip xs ys)" proof (rule zip_obtain_same_length) fix xs' :: "'a list" and ys' :: "'b list" and n assume "length xs' = length ys'" assume "xs' = take n xs" with assms have "distinct xs'" by simp with \length xs' = length ys'\ show "distinct (zip xs' ys')" by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE) qed lemma distinct_zipI2: assumes "distinct ys" shows "distinct (zip xs ys)" proof (rule zip_obtain_same_length) fix xs' :: "'b list" and ys' :: "'a list" and n assume "length xs' = length ys'" assume "ys' = take n ys" with assms have "distinct ys'" by simp with \length xs' = length ys'\ show "distinct (zip xs' ys')" by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE) qed lemma set_take_disj_set_drop_if_distinct: "distinct vs \ i \ j \ set (take i vs) \ set (drop j vs) = {}" by (auto simp: in_set_conv_nth distinct_conv_nth) (* The next two lemmas help Sledgehammer. *) lemma distinct_singleton: "distinct [x]" by simp lemma distinct_length_2_or_more: "distinct (a # b # xs) \ (a \ b \ distinct (a # xs) \ distinct (b # xs))" by force lemma remdups_adj_altdef: "(remdups_adj xs = ys) \ (\f::nat => nat. mono f \ f ` {0 ..< size xs} = {0 ..< size ys} \ (\i < size xs. xs!i = ys!(f i)) \ (\i. i + 1 < size xs \ (xs!i = xs!(i+1) \ f i = f(i+1))))" (is "?L \ (\f. ?p f xs ys)") proof assume ?L then show "\f. ?p f xs ys" proof (induct xs arbitrary: ys rule: remdups_adj.induct) case (1 ys) thus ?case by (intro exI[of _ id]) (auto simp: mono_def) next case (2 x ys) thus ?case by (intro exI[of _ id]) (auto simp: mono_def) next case (3 x1 x2 xs ys) let ?xs = "x1 # x2 # xs" let ?cond = "x1 = x2" define zs where "zs = remdups_adj (x2 # xs)" from 3(1-2)[of zs] obtain f where p: "?p f (x2 # xs) zs" unfolding zs_def by (cases ?cond) auto then have f0: "f 0 = 0" by (intro mono_image_least[where f=f]) blast+ from p have mono: "mono f" and f_xs_zs: "f ` {0.. []" unfolding zs_def by (induct xs) auto let ?Succ = "if ?cond then id else Suc" let ?x1 = "if ?cond then id else Cons x1" let ?f = "\ i. if i = 0 then 0 else ?Succ (f (i - 1))" have ys: "ys = ?x1 zs" unfolding ys by (cases ?cond, auto) have mono: "mono ?f" using \mono f\ unfolding mono_def by auto show ?case unfolding ys proof (intro exI[of _ ?f] conjI allI impI) show "mono ?f" by fact next fix i assume i: "i < length ?xs" with p show "?xs ! i = ?x1 zs ! (?f i)" using zs0 by auto next fix i assume i: "i + 1 < length ?xs" with p show "(?xs ! i = ?xs ! (i + 1)) = (?f i = ?f (i + 1))" by (cases i) (auto simp: f0) next have id: "{0 ..< length (?x1 zs)} = insert 0 (?Succ ` {0 ..< length zs})" using zsne by (cases ?cond, auto) { fix i assume "i < Suc (length xs)" hence "Suc i \ {0.. Collect ((<) 0)" by auto from imageI[OF this, of "\i. ?Succ (f (i - Suc 0))"] have "?Succ (f i) \ (\i. ?Succ (f (i - Suc 0))) ` ({0.. Collect ((<) 0))" by auto } then show "?f ` {0 ..< length ?xs} = {0 ..< length (?x1 zs)}" unfolding id f_xs_zs[symmetric] by auto qed qed next assume "\ f. ?p f xs ys" then show ?L proof (induct xs arbitrary: ys rule: remdups_adj.induct) case 1 then show ?case by auto next case (2 x) then obtain f where f_img: "f ` {0 ..< size [x]} = {0 ..< size ys}" and f_nth: "\i. i < size [x] \ [x]!i = ys!(f i)" by blast have "length ys = card (f ` {0 ..< size [x]})" using f_img by auto then have *: "length ys = 1" by auto then have "f 0 = 0" using f_img by auto with * show ?case using f_nth by (cases ys) auto next case (3 x1 x2 xs) from "3.prems" obtain f where f_mono: "mono f" and f_img: "f ` {0..i. i < length (x1 # x2 # xs) \ (x1 # x2 # xs) ! i = ys ! f i" "\i. i + 1 < length (x1 # x2 #xs) \ ((x1 # x2 # xs) ! i = (x1 # x2 # xs) ! (i + 1)) = (f i = f (i + 1))" by blast show ?case proof cases assume "x1 = x2" let ?f' = "f \ Suc" have "remdups_adj (x1 # xs) = ys" proof (intro "3.hyps" exI conjI impI allI) show "mono ?f'" using f_mono by (simp add: mono_iff_le_Suc) next have "?f' ` {0 ..< length (x1 # xs)} = f ` {Suc 0 ..< length (x1 # x2 # xs)}" - apply safe - apply fastforce - subgoal for \ x by (cases x) auto - done + using less_Suc_eq_0_disj by auto also have "\ = f ` {0 ..< length (x1 # x2 # xs)}" proof - have "f 0 = f (Suc 0)" using \x1 = x2\ f_nth[of 0] by simp then show ?thesis - apply safe - apply fastforce - subgoal for \ x by (cases x) auto - done + using less_Suc_eq_0_disj by auto qed also have "\ = {0 ..< length ys}" by fact finally show "?f' ` {0 ..< length (x1 # xs)} = {0 ..< length ys}" . qed (insert f_nth[of "Suc i" for i], auto simp: \x1 = x2\) then show ?thesis using \x1 = x2\ by simp next assume "x1 \ x2" - have "2 \ length ys" + have two: "Suc (Suc 0) \ length ys" proof - have "2 = card {f 0, f 1}" using \x1 \ x2\ f_nth[of 0] by auto also have "\ \ card (f ` {0..< length (x1 # x2 # xs)})" by (rule card_mono) auto finally show ?thesis using f_img by simp qed have "f 0 = 0" using f_mono f_img by (rule mono_image_least) simp have "f (Suc 0) = Suc 0" proof (rule ccontr) assume "f (Suc 0) \ Suc 0" then have "Suc 0 < f (Suc 0)" using f_nth[of 0] \x1 \ x2\ \f 0 = 0\ by auto then have "\i. Suc 0 < f (Suc i)" using f_mono by (meson Suc_le_mono le0 less_le_trans monoD) then have "Suc 0 \ f i" for i using \f 0 = 0\ by (cases i) fastforce+ then have "Suc 0 \ f ` {0 ..< length (x1 # x2 # xs)}" by auto - then show False using f_img \2 \ length ys\ by auto + then show False using f_img two by auto qed obtain ys' where "ys = x1 # x2 # ys'" - using \2 \ length ys\ f_nth[of 0] f_nth[of 1] - apply (cases ys) - apply (rename_tac [2] ys', case_tac [2] ys') - apply (auto simp: \f 0 = 0\ \f (Suc 0) = Suc 0\) - done + using two f_nth[of 0] f_nth[of 1] + by (auto simp: Suc_le_length_iff \f 0 = 0\ \f (Suc 0) = Suc 0\) + + have Suc0_le_f_Suc: "Suc 0 \ f (Suc i)" for i + by (metis Suc_le_mono \f (Suc 0) = Suc 0\ f_mono le0 mono_def) define f' where "f' x = f (Suc x) - 1" for x - - { fix i - have "Suc 0 \ f (Suc 0)" using f_nth[of 0] \x1 \ x2\ \f 0 = 0\ by auto - also have "\ \ f (Suc i)" using f_mono by (rule monoD) arith - finally have "Suc 0 \ f (Suc i)" . - } note Suc0_le_f_Suc = this - - { fix i have "f (Suc i) = Suc (f' i)" + have f_Suc: "f (Suc i) = Suc (f' i)" for i using Suc0_le_f_Suc[of i] by (auto simp: f'_def) - } note f_Suc = this have "remdups_adj (x2 # xs) = (x2 # ys')" proof (intro "3.hyps" exI conjI impI allI) show "mono f'" using Suc0_le_f_Suc f_mono by (auto simp: f'_def mono_iff_le_Suc le_diff_iff) next have "f' ` {0 ..< length (x2 # xs)} = (\x. f x - 1) ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: f'_def \f 0 = 0\ \f (Suc 0) = Suc 0\ image_def Bex_def less_Suc_eq_0_disj) also have "\ = (\x. x - 1) ` f ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: image_comp) also have "\ = (\x. x - 1) ` {0 ..< length ys}" by (simp only: f_img) also have "\ = {0 ..< length (x2 # ys')}" using \ys = _\ by (fastforce intro: rev_image_eqI) finally show "f' ` {0 ..< length (x2 # xs)} = {0 ..< length (x2 # ys')}" . qed (insert f_nth[of "Suc i" for i] \x1 \ x2\, auto simp add: f_Suc \ys = _\) then show ?case using \ys = _\ \x1 \ x2\ by simp qed qed qed lemma hd_remdups_adj[simp]: "hd (remdups_adj xs) = hd xs" -by (induction xs rule: remdups_adj.induct) simp_all + by (induction xs rule: remdups_adj.induct) simp_all lemma remdups_adj_Cons: "remdups_adj (x # xs) = (case remdups_adj xs of [] \ [x] | y # xs \ if x = y then y # xs else x # y # xs)" -by (induct xs arbitrary: x) (auto split: list.splits) + by (induct xs arbitrary: x) (auto split: list.splits) lemma remdups_adj_append_two: "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])" -by (induct xs rule: remdups_adj.induct, simp_all) + by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_adjacent: "Suc i < length (remdups_adj xs) \ remdups_adj xs ! i \ remdups_adj xs ! Suc i" proof (induction xs arbitrary: i rule: remdups_adj.induct) case (3 x y xs i) thus ?case by (cases i, cases "x = y") (simp, auto simp: hd_conv_nth[symmetric]) qed simp_all lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)" by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two) lemma remdups_adj_length[simp]: "length (remdups_adj xs) \ length xs" by (induct xs rule: remdups_adj.induct, auto) lemma remdups_adj_length_ge1[simp]: "xs \ [] \ length (remdups_adj xs) \ Suc 0" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_Nil_iff[simp]: "remdups_adj xs = [] \ xs = []" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_set[simp]: "set (remdups_adj xs) = set xs" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_Cons_alt[simp]: "x # tl (remdups_adj (x # xs)) = remdups_adj (x # xs)" by (induct xs rule: remdups_adj.induct, auto) lemma remdups_adj_distinct: "distinct xs \ remdups_adj xs = xs" by (induct xs rule: remdups_adj.induct, simp_all) lemma remdups_adj_append: "remdups_adj (xs\<^sub>1 @ x # xs\<^sub>2) = remdups_adj (xs\<^sub>1 @ [x]) @ tl (remdups_adj (x # xs\<^sub>2))" by (induct xs\<^sub>1 rule: remdups_adj.induct, simp_all) lemma remdups_adj_singleton: "remdups_adj xs = [x] \ xs = replicate (length xs) x" by (induct xs rule: remdups_adj.induct, auto split: if_split_asm) lemma remdups_adj_map_injective: assumes "inj f" shows "remdups_adj (map f xs) = map f (remdups_adj xs)" by (induct xs rule: remdups_adj.induct) (auto simp add: injD[OF assms]) lemma remdups_adj_replicate: "remdups_adj (replicate n x) = (if n = 0 then [] else [x])" by (induction n) (auto simp: remdups_adj_Cons) lemma remdups_upt [simp]: "remdups [m.. n") case False then show ?thesis by simp next case True then obtain q where "n = m + q" by (auto simp add: le_iff_add) moreover have "remdups [m..\<^const>\insert\\ lemma in_set_insert [simp]: "x \ set xs \ List.insert x xs = xs" by (simp add: List.insert_def) lemma not_in_set_insert [simp]: "x \ set xs \ List.insert x xs = x # xs" by (simp add: List.insert_def) lemma insert_Nil [simp]: "List.insert x [] = [x]" by simp lemma set_insert [simp]: "set (List.insert x xs) = insert x (set xs)" by (auto simp add: List.insert_def) lemma distinct_insert [simp]: "distinct (List.insert x xs) = distinct xs" by (simp add: List.insert_def) lemma insert_remdups: "List.insert x (remdups xs) = remdups (List.insert x xs)" by (simp add: List.insert_def) subsubsection \\<^const>\List.union\\ text\This is all one should need to know about union:\ lemma set_union[simp]: "set (List.union xs ys) = set xs \ set ys" unfolding List.union_def by(induct xs arbitrary: ys) simp_all lemma distinct_union[simp]: "distinct(List.union xs ys) = distinct ys" unfolding List.union_def by(induct xs arbitrary: ys) simp_all subsubsection \\<^const>\List.find\\ lemma find_None_iff: "List.find P xs = None \ \ (\x. x \ set xs \ P x)" proof (induction xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (fastforce split: if_splits) qed lemma find_Some_iff: "List.find P xs = Some x \ (\i x = xs!i \ (\j P (xs!j)))" proof (induction xs) case Nil thus ?case by simp next case (Cons x xs) thus ?case apply(auto simp: nth_Cons' split: if_splits) using diff_Suc_1[unfolded One_nat_def] less_Suc_eq_0_disj by fastforce qed lemma find_cong[fundef_cong]: assumes "xs = ys" and "\x. x \ set ys \ P x = Q x" shows "List.find P xs = List.find Q ys" proof (cases "List.find P xs") case None thus ?thesis by (metis find_None_iff assms) next case (Some x) hence "List.find Q ys = Some x" using assms by (auto simp add: find_Some_iff) thus ?thesis using Some by auto qed lemma find_dropWhile: "List.find P xs = (case dropWhile (Not \ P) xs of [] \ None | x # _ \ Some x)" by (induct xs) simp_all subsubsection \\<^const>\count_list\\ lemma count_notin[simp]: "x \ set xs \ count_list xs x = 0" by (induction xs) auto lemma count_le_length: "count_list xs x \ length xs" by (induction xs) auto lemma sum_count_set: "set xs \ X \ finite X \ sum (count_list xs) X = length xs" proof (induction xs arbitrary: X) case (Cons x xs) then show ?case - apply (auto simp: sum.If_cases sum.remove) - by (metis (no_types) Cons.IH Cons.prems(2) diff_eq sum.remove) + using sum.remove [of X x "count_list xs"] + by (auto simp: sum.If_cases simp flip: diff_eq) qed simp subsubsection \\<^const>\List.extract\\ lemma extract_None_iff: "List.extract P xs = None \ \ (\ x\set xs. P x)" by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits) (metis in_set_conv_decomp) lemma extract_SomeE: "List.extract P xs = Some (ys, y, zs) \ xs = ys @ y # zs \ P y \ \ (\ y \ set ys. P y)" by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits) lemma extract_Some_iff: "List.extract P xs = Some (ys, y, zs) \ xs = ys @ y # zs \ P y \ \ (\ y \ set ys. P y)" by(auto simp: extract_def dropWhile_eq_Cons_conv dest: set_takeWhileD split: list.splits) lemma extract_Nil_code[code]: "List.extract P [] = None" by(simp add: extract_def) lemma extract_Cons_code[code]: "List.extract P (x # xs) = (if P x then Some ([], x, xs) else (case List.extract P xs of None \ None | Some (ys, y, zs) \ Some (x#ys, y, zs)))" by(auto simp add: extract_def comp_def split: list.splits) (metis dropWhile_eq_Nil_conv list.distinct(1)) subsubsection \\<^const>\remove1\\ lemma remove1_append: "remove1 x (xs @ ys) = (if x \ set xs then remove1 x xs @ ys else xs @ remove1 x ys)" by (induct xs) auto lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)" by (induct zs) auto lemma in_set_remove1[simp]: "a \ b \ a \ set(remove1 b xs) = (a \ set xs)" by (induct xs) auto lemma set_remove1_subset: "set(remove1 x xs) \ set xs" by (induct xs) auto lemma set_remove1_eq [simp]: "distinct xs \ set(remove1 x xs) = set xs - {x}" by (induct xs) auto lemma length_remove1: "length(remove1 x xs) = (if x \ set xs then length xs - 1 else length xs)" by (induct xs) (auto dest!:length_pos_if_in_set) lemma remove1_filter_not[simp]: "\ P x \ remove1 x (filter P xs) = filter P xs" by(induct xs) auto lemma filter_remove1: "filter Q (remove1 x xs) = remove1 x (filter Q xs)" by (induct xs) auto lemma notin_set_remove1[simp]: "x \ set xs \ x \ set(remove1 y xs)" by(insert set_remove1_subset) fast lemma distinct_remove1[simp]: "distinct xs \ distinct(remove1 x xs)" by (induct xs) simp_all lemma remove1_remdups: "distinct xs \ remove1 x (remdups xs) = remdups (remove1 x xs)" by (induct xs) simp_all lemma remove1_idem: "x \ set xs \ remove1 x xs = xs" by (induct xs) simp_all subsubsection \\<^const>\removeAll\\ lemma removeAll_filter_not_eq: "removeAll x = filter (\y. x \ y)" proof fix xs show "removeAll x xs = filter (\y. x \ y) xs" by (induct xs) auto qed lemma removeAll_append[simp]: "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys" by (induct xs) auto lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}" by (induct xs) auto lemma removeAll_id[simp]: "x \ set xs \ removeAll x xs = xs" by (induct xs) auto (* Needs count:: 'a \ 'a list \ nat lemma length_removeAll: "length(removeAll x xs) = length xs - count x xs" *) lemma removeAll_filter_not[simp]: "\ P x \ removeAll x (filter P xs) = filter P xs" by(induct xs) auto lemma distinct_removeAll: "distinct xs \ distinct (removeAll x xs)" by (simp add: removeAll_filter_not_eq) lemma distinct_remove1_removeAll: "distinct xs \ remove1 x xs = removeAll x xs" by (induct xs) simp_all lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \ map f (removeAll x xs) = removeAll (f x) (map f xs)" by (induct xs) (simp_all add:inj_on_def) lemma map_removeAll_inj: "inj f \ map f (removeAll x xs) = removeAll (f x) (map f xs)" by (rule map_removeAll_inj_on, erule subset_inj_on, rule subset_UNIV) lemma length_removeAll_less_eq [simp]: "length (removeAll x xs) \ length xs" by (simp add: removeAll_filter_not_eq) lemma length_removeAll_less [termination_simp]: "x \ set xs \ length (removeAll x xs) < length xs" by (auto dest: length_filter_less simp add: removeAll_filter_not_eq) subsubsection \\<^const>\replicate\\ lemma length_replicate [simp]: "length (replicate n x) = n" by (induct n) auto lemma replicate_eqI: assumes "length xs = n" and "\y. y \ set xs \ y = x" shows "xs = replicate n x" using assms proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases n) simp_all qed lemma Ex_list_of_length: "\xs. length xs = n" by (rule exI[of _ "replicate n undefined"]) simp lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)" by (induct n) auto lemma map_replicate_const: "map (\ x. k) lst = replicate (length lst) k" by (induct lst) auto lemma replicate_app_Cons_same: "(replicate n x) @ (x # xs) = x # replicate n x @ xs" by (induct n) auto lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x" by (induct n) (auto simp: replicate_app_Cons_same) lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x" by (induct n) auto text\Courtesy of Matthias Daum:\ lemma append_replicate_commute: "replicate n x @ replicate k x = replicate k x @ replicate n x" by (metis add.commute replicate_add) text\Courtesy of Andreas Lochbihler:\ lemma filter_replicate: "filter P (replicate n x) = (if P x then replicate n x else [])" by(induct n) auto -lemma hd_replicate [simp]: "n \ 0 ==> hd (replicate n x) = x" +lemma hd_replicate [simp]: "n \ 0 \ hd (replicate n x) = x" by (induct n) auto lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x" by (induct n) auto -lemma last_replicate [simp]: "n \ 0 ==> last (replicate n x) = x" +lemma last_replicate [simp]: "n \ 0 \ last (replicate n x) = x" by (atomize (full), induct n) auto -lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x" +lemma nth_replicate[simp]: "i < n \ (replicate n x)!i = x" by (induct n arbitrary: i)(auto simp: nth_Cons split: nat.split) text\Courtesy of Matthias Daum (2 lemmas):\ lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x" proof (cases "k \ i") case True then show ?thesis by (simp add: min_def) next case False then have "replicate k x = replicate i x @ replicate (k - i) x" by (simp add: replicate_add [symmetric]) then show ?thesis by (simp add: min_def) qed lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x" proof (induct k arbitrary: i) case (Suc k) then show ?case by (simp add: drop_Cons') qed simp lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}" by (induct n) auto -lemma set_replicate [simp]: "n \ 0 ==> set (replicate n x) = {x}" +lemma set_replicate [simp]: "n \ 0 \ set (replicate n x) = {x}" by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc) lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})" by auto lemma in_set_replicate[simp]: "(x \ set (replicate n y)) = (x = y \ n \ 0)" by (simp add: set_replicate_conv_if) lemma Ball_set_replicate[simp]: "(\x \ set(replicate n a). P x) = (P a \ n=0)" by(simp add: set_replicate_conv_if) lemma Bex_set_replicate[simp]: "(\x \ set(replicate n a). P x) = (P a \ n\0)" by(simp add: set_replicate_conv_if) lemma replicate_append_same: "replicate i x @ [x] = x # replicate i x" by (induct i) simp_all lemma map_replicate_trivial: "map (\i. x) [0.. n=0" by (induct n) auto lemma empty_replicate[simp]: "([] = replicate n x) \ n=0" by (induct n) auto lemma replicate_eq_replicate[simp]: "(replicate m x = replicate n y) \ (m=n \ (m\0 \ x=y))" proof (induct m arbitrary: n) case (Suc m n) then show ?case by (induct n) auto qed simp lemma replicate_length_filter: "replicate (length (filter (\y. x = y) xs)) x = filter (\y. x = y) xs" by (induct xs) auto lemma comm_append_are_replicate: "\ xs \ []; ys \ []; xs @ ys = ys @ xs \ \ \m n zs. concat (replicate m zs) = xs \ concat (replicate n zs) = ys" proof (induction "length (xs @ ys)" arbitrary: xs ys rule: less_induct) case less define xs' ys' where "xs' = (if (length xs \ length ys) then xs else ys)" and "ys' = (if (length xs \ length ys) then ys else xs)" then have prems': "length xs' \ length ys'" "xs' @ ys' = ys' @ xs'" and "xs' \ []" and len: "length (xs @ ys) = length (xs' @ ys')" using less by (auto intro: less.hyps) from prems' obtain ws where "ys' = xs' @ ws" by (auto simp: append_eq_append_conv2) have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ys'" proof (cases "ws = []") case True then have "concat (replicate 1 xs') = xs'" and "concat (replicate 1 xs') = ys'" using \ys' = xs' @ ws\ by auto then show ?thesis by blast next case False from \ys' = xs' @ ws\ and \xs' @ ys' = ys' @ xs'\ have "xs' @ ws = ws @ xs'" by simp then have "\m n zs. concat (replicate m zs) = xs' \ concat (replicate n zs) = ws" using False and \xs' \ []\ and \ys' = xs' @ ws\ and len by (intro less.hyps) auto then obtain m n zs where *: "concat (replicate m zs) = xs'" and "concat (replicate n zs) = ws" by blast then have "concat (replicate (m + n) zs) = ys'" using \ys' = xs' @ ws\ by (simp add: replicate_add) with * show ?thesis by blast qed then show ?case using xs'_def ys'_def by meson qed lemma comm_append_is_replicate: fixes xs ys :: "'a list" assumes "xs \ []" "ys \ []" assumes "xs @ ys = ys @ xs" shows "\n zs. n > 1 \ concat (replicate n zs) = xs @ ys" proof - obtain m n zs where "concat (replicate m zs) = xs" and "concat (replicate n zs) = ys" using comm_append_are_replicate[of xs ys, OF assms] by blast then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys" using \xs \ []\ and \ys \ []\ by (auto simp: replicate_add) then show ?thesis by blast qed lemma Cons_replicate_eq: "x # xs = replicate n y \ x = y \ n > 0 \ xs = replicate (n - 1) x" by (induct n) auto lemma replicate_length_same: "(\y\set xs. y = x) \ replicate (length xs) x = xs" by (induct xs) simp_all lemma foldr_replicate [simp]: "foldr f (replicate n x) = f x ^^ n" by (induct n) (simp_all) lemma fold_replicate [simp]: "fold f (replicate n x) = f x ^^ n" by (subst foldr_fold [symmetric]) simp_all subsubsection \\<^const>\enumerate\\ lemma enumerate_simps [simp, code]: "enumerate n [] = []" "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs" by (simp_all add: enumerate_eq_zip upt_rec) lemma length_enumerate [simp]: "length (enumerate n xs) = length xs" by (simp add: enumerate_eq_zip) lemma map_fst_enumerate [simp]: "map fst (enumerate n xs) = [n.. set (enumerate n xs) \ n \ fst p \ fst p < length xs + n \ nth xs (fst p - n) = snd p" proof - { fix m assume "n \ m" moreover assume "m < length xs + n" ultimately have "[n.. xs ! (m - n) = xs ! (m - n) \ m - n < length xs" by auto then have "\q. [n.. xs ! q = xs ! (m - n) \ q < length xs" .. } then show ?thesis by (cases p) (auto simp add: enumerate_eq_zip in_set_zip) qed lemma nth_enumerate_eq: "m < length xs \ enumerate n xs ! m = (n + m, xs ! m)" by (simp add: enumerate_eq_zip) lemma enumerate_replicate_eq: "enumerate n (replicate m a) = map (\q. (q, a)) [n..k. (k, f k)) [n.. m") (simp_all add: zip_map2 zip_same_conv_map enumerate_eq_zip) subsubsection \\<^const>\rotate1\ and \<^const>\rotate\\ lemma rotate0[simp]: "rotate 0 = id" by(simp add:rotate_def) lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)" by(simp add:rotate_def) lemma rotate_add: "rotate (m+n) = rotate m \ rotate n" by(simp add:rotate_def funpow_add) lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs" by(simp add:rotate_add) lemma rotate1_map: "rotate1 (map f xs) = map f (rotate1 xs)" by(cases xs) simp_all lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)" by(simp add:rotate_def funpow_swap1) lemma rotate1_length01[simp]: "length xs \ 1 \ rotate1 xs = xs" by(cases xs) simp_all lemma rotate_length01[simp]: "length xs \ 1 \ rotate n xs = xs" by (induct n) (simp_all add:rotate_def) lemma rotate1_hd_tl: "xs \ [] \ rotate1 xs = tl xs @ [hd xs]" by (cases xs) simp_all lemma rotate_drop_take: "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs" proof (induct n) case (Suc n) show ?case proof (cases "xs = []") case False then show ?thesis proof (cases "n mod length xs = 0") case True then show ?thesis - apply (simp add: mod_Suc) - by (simp add: False Suc.hyps drop_Suc rotate1_hd_tl take_Suc) + by (auto simp add: mod_Suc False Suc.hyps drop_Suc rotate1_hd_tl take_Suc Suc_length_conv) next case False with \xs \ []\ Suc show ?thesis by (simp add: rotate_def mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric] take_hd_drop linorder_not_le) qed qed simp qed simp lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs" -by(simp add:rotate_drop_take) + by(simp add:rotate_drop_take) lemma rotate_id[simp]: "n mod length xs = 0 \ rotate n xs = xs" -by(simp add:rotate_drop_take) + by(simp add:rotate_drop_take) lemma length_rotate1[simp]: "length(rotate1 xs) = length xs" -by (cases xs) simp_all + by (cases xs) simp_all lemma length_rotate[simp]: "length(rotate n xs) = length xs" -by (induct n arbitrary: xs) (simp_all add:rotate_def) + by (induct n arbitrary: xs) (simp_all add:rotate_def) lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs" -by (cases xs) auto + by (cases xs) auto lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs" -by (induct n) (simp_all add:rotate_def) + by (induct n) (simp_all add:rotate_def) lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)" -by(simp add:rotate_drop_take take_map drop_map) + by(simp add:rotate_drop_take take_map drop_map) lemma set_rotate1[simp]: "set(rotate1 xs) = set xs" -by (cases xs) auto + by (cases xs) auto lemma set_rotate[simp]: "set(rotate n xs) = set xs" -by (induct n) (simp_all add:rotate_def) + by (induct n) (simp_all add:rotate_def) lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])" -by (cases xs) auto + by (cases xs) auto lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])" -by (induct n) (simp_all add:rotate_def) + by (induct n) (simp_all add:rotate_def) lemma rotate_rev: "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)" proof (cases "length xs = 0 \ n mod length xs = 0") case False then show ?thesis by(simp add:rotate_drop_take rev_drop rev_take) qed force lemma hd_rotate_conv_nth: assumes "xs \ []" shows "hd(rotate n xs) = xs!(n mod length xs)" proof - have "n mod length xs < length xs" using assms by simp then show ?thesis by (metis drop_eq_Nil hd_append2 hd_drop_conv_nth leD rotate_drop_take) qed lemma rotate_append: "rotate (length l) (l @ q) = q @ l" by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap) subsubsection \\<^const>\nths\ --- a generalization of \<^const>\nth\ to sets\ lemma nths_empty [simp]: "nths xs {} = []" -by (auto simp add: nths_def) + by (auto simp add: nths_def) lemma nths_nil [simp]: "nths [] A = []" -by (auto simp add: nths_def) + by (auto simp add: nths_def) lemma nths_all: "\i < length xs. i \ I \ nths xs I = xs" apply (simp add: nths_def) apply (subst filter_True) - apply (clarsimp simp: in_set_zip subset_iff) -apply simp + apply (auto simp: in_set_zip subset_iff) done lemma length_nths: "length (nths xs I) = card{i. i < length xs \ i \ I}" -by(simp add: nths_def length_filter_conv_card cong:conj_cong) + by(simp add: nths_def length_filter_conv_card cong:conj_cong) lemma nths_shift_lemma_Suc: "map fst (filter (\p. P(Suc(snd p))) (zip xs is)) = map fst (filter (\p. P(snd p)) (zip xs (map Suc is)))" proof (induct xs arbitrary: "is") case (Cons x xs "is") show ?case by (cases "is") (auto simp add: Cons.hyps) qed simp lemma nths_shift_lemma: "map fst (filter (\p. snd p \ A) (zip xs [i..p. snd p + i \ A) (zip xs [0.. A}" unfolding nths_def proof (induct l' rule: rev_induct) case (snoc x xs) then show ?case by (simp add: upt_add_eq_append[of 0] nths_shift_lemma add.commute) qed auto lemma nths_Cons: "nths (x # l) A = (if 0 \ A then [x] else []) @ nths l {j. Suc j \ A}" proof (induct l rule: rev_induct) case (snoc x xs) then show ?case by (simp flip: append_Cons add: nths_append) qed (auto simp: nths_def) lemma nths_map: "nths (map f xs) I = map f (nths xs I)" -by(induction xs arbitrary: I) (simp_all add: nths_Cons) + by(induction xs arbitrary: I) (simp_all add: nths_Cons) lemma set_nths: "set(nths xs I) = {xs!i|i. i i \ I}" -by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc) + by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc) lemma set_nths_subset: "set(nths xs I) \ set xs" -by(auto simp add:set_nths) + by(auto simp add:set_nths) lemma notin_set_nthsI[simp]: "x \ set xs \ x \ set(nths xs I)" -by(auto simp add:set_nths) + by(auto simp add:set_nths) lemma in_set_nthsD: "x \ set(nths xs I) \ x \ set xs" -by(auto simp add:set_nths) + by(auto simp add:set_nths) lemma nths_singleton [simp]: "nths [x] A = (if 0 \ A then [x] else [])" -by (simp add: nths_Cons) + by (simp add: nths_Cons) lemma distinct_nthsI[simp]: "distinct xs \ distinct (nths xs I)" -by (induct xs arbitrary: I) (auto simp: nths_Cons) + by (induct xs arbitrary: I) (auto simp: nths_Cons) lemma nths_upt_eq_take [simp]: "nths l {.. A. \j \ B. card {i' \ A. i' < i} = j}" -apply(induction xs arbitrary: A B) -apply(auto simp add: nths_Cons card_less_Suc card_less_Suc2) -done + by (induction xs arbitrary: A B) (auto simp add: nths_Cons card_less_Suc card_less_Suc2) lemma drop_eq_nths: "drop n xs = nths xs {i. i \ n}" -apply(induction xs arbitrary: n) -apply (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl]) -done + by (induction xs arbitrary: n) (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl]) lemma nths_drop: "nths (drop n xs) I = nths xs ((+) n ` I)" by(force simp: drop_eq_nths nths_nths simp flip: atLeastLessThan_iff intro: arg_cong2[where f=nths, OF refl]) lemma filter_eq_nths: "filter P xs = nths xs {i. i P(xs!i)}" -by(induction xs) (auto simp: nths_Cons) + by(induction xs) (auto simp: nths_Cons) lemma filter_in_nths: "distinct xs \ filter (%x. x \ set (nths xs s)) xs = nths xs s" proof (induct xs arbitrary: s) case Nil thus ?case by simp next case (Cons a xs) then have "\x. x \ set xs \ x \ a" by auto with Cons show ?case by(simp add: nths_Cons cong:filter_cong) qed subsubsection \\<^const>\subseqs\ and \<^const>\List.n_lists\\ lemma length_subseqs: "length (subseqs xs) = 2 ^ length xs" by (induct xs) (simp_all add: Let_def) lemma subseqs_powset: "set ` set (subseqs xs) = Pow (set xs)" proof - have aux: "\x A. set ` Cons x ` A = insert x ` set ` A" by (auto simp add: image_def) have "set (map set (subseqs xs)) = Pow (set xs)" by (induct xs) (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map) then show ?thesis by simp qed lemma distinct_set_subseqs: assumes "distinct xs" shows "distinct (map set (subseqs xs))" proof (rule card_distinct) have "finite (set xs)" .. then have "card (Pow (set xs)) = 2 ^ card (set xs)" by (rule card_Pow) with assms distinct_card [of xs] have "card (Pow (set xs)) = 2 ^ length xs" by simp then show "card (set (map set (subseqs xs))) = length (map set (subseqs xs))" by (simp add: subseqs_powset length_subseqs) qed lemma n_lists_Nil [simp]: "List.n_lists n [] = (if n = 0 then [[]] else [])" by (induct n) simp_all lemma length_n_lists_elem: "ys \ set (List.n_lists n xs) \ length ys = n" by (induct n arbitrary: ys) auto lemma set_n_lists: "set (List.n_lists n xs) = {ys. length ys = n \ set ys \ set xs}" proof (rule set_eqI) fix ys :: "'a list" show "ys \ set (List.n_lists n xs) \ ys \ {ys. length ys = n \ set ys \ set xs}" proof - have "ys \ set (List.n_lists n xs) \ length ys = n" by (induct n arbitrary: ys) auto moreover have "\x. ys \ set (List.n_lists n xs) \ x \ set ys \ x \ set xs" by (induct n arbitrary: ys) auto moreover have "set ys \ set xs \ ys \ set (List.n_lists (length ys) xs)" by (induct ys) auto ultimately show ?thesis by auto qed qed lemma subseqs_refl: "xs \ set (subseqs xs)" by (induct xs) (simp_all add: Let_def) lemma subset_subseqs: "X \ set xs \ X \ set ` set (subseqs xs)" unfolding subseqs_powset by simp lemma Cons_in_subseqsD: "y # ys \ set (subseqs xs) \ ys \ set (subseqs xs)" by (induct xs) (auto simp: Let_def) lemma subseqs_distinctD: "\ ys \ set (subseqs xs); distinct xs \ \ distinct ys" proof (induct xs arbitrary: ys) case (Cons x xs ys) then show ?case by (auto simp: Let_def) (metis Pow_iff contra_subsetD image_eqI subseqs_powset) qed simp subsubsection \\<^const>\splice\\ lemma splice_Nil2 [simp]: "splice xs [] = xs" -by (cases xs) simp_all + by (cases xs) simp_all lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys" -by (induct xs ys rule: splice.induct) auto + by (induct xs ys rule: splice.induct) auto lemma split_Nil_iff[simp]: "splice xs ys = [] \ xs = [] \ ys = []" -by (induct xs ys rule: splice.induct) auto + by (induct xs ys rule: splice.induct) auto lemma splice_replicate[simp]: "splice (replicate m x) (replicate n x) = replicate (m+n) x" -apply(induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct) -apply (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc) -done +proof (induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct) + case (2 x xs) + then show ?case + by (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc) +qed auto subsubsection \\<^const>\shuffles\\ lemma shuffles_commutes: "shuffles xs ys = shuffles ys xs" by (induction xs ys rule: shuffles.induct) (simp_all add: Un_commute) lemma Nil_in_shuffles[simp]: "[] \ shuffles xs ys \ xs = [] \ ys = []" by (induct xs ys rule: shuffles.induct) auto lemma shufflesE: "zs \ shuffles xs ys \ (zs = xs \ ys = [] \ P) \ (zs = ys \ xs = [] \ P) \ (\x xs' z zs'. xs = x # xs' \ zs = z # zs' \ x = z \ zs' \ shuffles xs' ys \ P) \ (\y ys' z zs'. ys = y # ys' \ zs = z # zs' \ y = z \ zs' \ shuffles xs ys' \ P) \ P" by (induct xs ys rule: shuffles.induct) auto lemma Cons_in_shuffles_iff: "z # zs \ shuffles xs ys \ (xs \ [] \ hd xs = z \ zs \ shuffles (tl xs) ys \ ys \ [] \ hd ys = z \ zs \ shuffles xs (tl ys))" by (induct xs ys rule: shuffles.induct) auto lemma splice_in_shuffles [simp, intro]: "splice xs ys \ shuffles xs ys" -by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffles_iff shuffles_commutes) + by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffles_iff shuffles_commutes) lemma Nil_in_shufflesI: "xs = [] \ ys = [] \ [] \ shuffles xs ys" by simp lemma Cons_in_shuffles_leftI: "zs \ shuffles xs ys \ z # zs \ shuffles (z # xs) ys" by (cases ys) auto lemma Cons_in_shuffles_rightI: "zs \ shuffles xs ys \ z # zs \ shuffles xs (z # ys)" by (cases xs) auto lemma finite_shuffles [simp, intro]: "finite (shuffles xs ys)" by (induction xs ys rule: shuffles.induct) simp_all lemma length_shuffles: "zs \ shuffles xs ys \ length zs = length xs + length ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma set_shuffles: "zs \ shuffles xs ys \ set zs = set xs \ set ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma distinct_disjoint_shuffles: assumes "distinct xs" "distinct ys" "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "distinct zs" using assms proof (induction xs ys arbitrary: zs rule: shuffles.induct) case (3 x xs y ys) show ?case proof (cases zs) case (Cons z zs') with "3.prems" and "3.IH"[of zs'] show ?thesis by (force dest: set_shuffles) qed simp_all qed simp_all lemma Cons_shuffles_subset1: "(#) x ` shuffles xs ys \ shuffles (x # xs) ys" by (cases ys) auto lemma Cons_shuffles_subset2: "(#) y ` shuffles xs ys \ shuffles xs (y # ys)" by (cases xs) auto lemma filter_shuffles: "filter P ` shuffles xs ys = shuffles (filter P xs) (filter P ys)" proof - have *: "filter P ` (#) x ` A = (if P x then (#) x ` filter P ` A else filter P ` A)" for x A by (auto simp: image_image) show ?thesis by (induction xs ys rule: shuffles.induct) (simp_all split: if_splits add: image_Un * Un_absorb1 Un_absorb2 Cons_shuffles_subset1 Cons_shuffles_subset2) qed lemma filter_shuffles_disjoint1: assumes "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "filter (\x. x \ set xs) zs = xs" (is "filter ?P _ = _") and "filter (\x. x \ set xs) zs = ys" (is "filter ?Q _ = _") using assms proof - from assms have "filter ?P zs \ filter ?P ` shuffles xs ys" by blast also have "filter ?P ` shuffles xs ys = shuffles (filter ?P xs) (filter ?P ys)" by (rule filter_shuffles) also have "filter ?P xs = xs" by (rule filter_True) simp_all also have "filter ?P ys = []" by (rule filter_False) (insert assms(1), auto) also have "shuffles xs [] = {xs}" by simp finally show "filter ?P zs = xs" by simp next from assms have "filter ?Q zs \ filter ?Q ` shuffles xs ys" by blast also have "filter ?Q ` shuffles xs ys = shuffles (filter ?Q xs) (filter ?Q ys)" by (rule filter_shuffles) also have "filter ?Q ys = ys" by (rule filter_True) (insert assms(1), auto) also have "filter ?Q xs = []" by (rule filter_False) (insert assms(1), auto) also have "shuffles [] ys = {ys}" by simp finally show "filter ?Q zs = ys" by simp qed lemma filter_shuffles_disjoint2: assumes "set xs \ set ys = {}" "zs \ shuffles xs ys" shows "filter (\x. x \ set ys) zs = ys" "filter (\x. x \ set ys) zs = xs" using filter_shuffles_disjoint1[of ys xs zs] assms by (simp_all add: shuffles_commutes Int_commute) lemma partition_in_shuffles: "xs \ shuffles (filter P xs) (filter (\x. \P x) xs)" proof (induction xs) case (Cons x xs) show ?case proof (cases "P x") case True hence "x # xs \ (#) x ` shuffles (filter P xs) (filter (\x. \P x) xs)" by (intro imageI Cons.IH) also have "\ \ shuffles (filter P (x # xs)) (filter (\x. \P x) (x # xs))" by (simp add: True Cons_shuffles_subset1) finally show ?thesis . next case False hence "x # xs \ (#) x ` shuffles (filter P xs) (filter (\x. \P x) xs)" by (intro imageI Cons.IH) also have "\ \ shuffles (filter P (x # xs)) (filter (\x. \P x) (x # xs))" by (simp add: False Cons_shuffles_subset2) finally show ?thesis . qed qed auto lemma inv_image_partition: assumes "\x. x \ set xs \ P x" "\y. y \ set ys \ \P y" shows "partition P -` {(xs, ys)} = shuffles xs ys" proof (intro equalityI subsetI) fix zs assume zs: "zs \ shuffles xs ys" hence [simp]: "set zs = set xs \ set ys" by (rule set_shuffles) from assms have "filter P zs = filter (\x. x \ set xs) zs" "filter (\x. \P x) zs = filter (\x. x \ set ys) zs" by (intro filter_cong refl; force)+ moreover from assms have "set xs \ set ys = {}" by auto ultimately show "zs \ partition P -` {(xs, ys)}" using zs by (simp add: o_def filter_shuffles_disjoint1 filter_shuffles_disjoint2) next fix zs assume "zs \ partition P -` {(xs, ys)}" thus "zs \ shuffles xs ys" using partition_in_shuffles[of zs] by (auto simp: o_def) qed subsubsection \Transpose\ function transpose where "transpose [] = []" | "transpose ([] # xss) = transpose xss" | "transpose ((x#xs) # xss) = (x # [h. (h#t) \ xss]) # transpose (xs # [t. (h#t) \ xss])" by pat_completeness auto lemma transpose_aux_filter_head: "concat (map (case_list [] (\h t. [h])) xss) = map (\xs. hd xs) (filter (\ys. ys \ []) xss)" by (induct xss) (auto split: list.split) lemma transpose_aux_filter_tail: "concat (map (case_list [] (\h t. [t])) xss) = map (\xs. tl xs) (filter (\ys. ys \ []) xss)" by (induct xss) (auto split: list.split) lemma transpose_aux_max: "max (Suc (length xs)) (foldr (\xs. max (length xs)) xss 0) = Suc (max (length xs) (foldr (\x. max (length x - Suc 0)) (filter (\ys. ys \ []) xss) 0))" (is "max _ ?foldB = Suc (max _ ?foldA)") proof (cases "(filter (\ys. ys \ []) xss) = []") case True hence "foldr (\xs. max (length xs)) xss 0 = 0" proof (induct xss) case (Cons x xs) then have "x = []" by (cases x) auto with Cons show ?case by auto qed simp thus ?thesis using True by simp next case False have foldA: "?foldA = foldr (\x. max (length x)) (filter (\ys. ys \ []) xss) 0 - 1" by (induct xss) auto have foldB: "?foldB = foldr (\x. max (length x)) (filter (\ys. ys \ []) xss) 0" by (induct xss) auto have "0 < ?foldB" proof - from False obtain z zs where zs: "(filter (\ys. ys \ []) xss) = z#zs" by (auto simp: neq_Nil_conv) hence "z \ set (filter (\ys. ys \ []) xss)" by auto hence "z \ []" by auto thus ?thesis unfolding foldB zs by (auto simp: max_def intro: less_le_trans) qed thus ?thesis unfolding foldA foldB max_Suc_Suc[symmetric] by simp qed termination transpose by (relation "measure (\xs. foldr (\xs. max (length xs)) xs 0 + length xs)") (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le) lemma transpose_empty: "(transpose xs = []) \ (\x \ set xs. x = [])" by (induct rule: transpose.induct) simp_all lemma length_transpose: fixes xs :: "'a list list" shows "length (transpose xs) = foldr (\xs. max (length xs)) xs 0" by (induct rule: transpose.induct) (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max max_Suc_Suc[symmetric] simp del: max_Suc_Suc) lemma nth_transpose: fixes xs :: "'a list list" assumes "i < length (transpose xs)" shows "transpose xs ! i = map (\xs. xs ! i) (filter (\ys. i < length ys) xs)" using assms proof (induct arbitrary: i rule: transpose.induct) case (3 x xs xss) define XS where "XS = (x # xs) # xss" hence [simp]: "XS \ []" by auto thus ?case proof (cases i) case 0 thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth) next case (Suc j) have *: "\xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp have **: "\xss. (x#xs) # filter (\ys. ys \ []) xss = filter (\ys. ys \ []) ((x#xs)#xss)" by simp { fix x have "Suc j < length x \ x \ [] \ j < length x - Suc 0" by (cases x) simp_all } note *** = this have j_less: "j < length (transpose (xs # concat (map (case_list [] (\h t. [t])) xss)))" using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc) show ?thesis unfolding transpose.simps \i = Suc j\ nth_Cons_Suc "3.hyps"[OF j_less] apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric]) by (simp add: nth_tl) qed qed simp_all lemma transpose_map_map: "transpose (map (map f) xs) = map (map f) (transpose xs)" proof (rule nth_equalityI) have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)" by (simp add: length_transpose foldr_map comp_def) show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp fix i assume "i < length (transpose (map (map f) xs))" thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i" by (simp add: nth_transpose filter_map comp_def) qed subsubsection \\<^const>\min\ and \<^const>\arg_min\\ lemma min_list_Min: "xs \ [] \ min_list xs = Min (set xs)" -by (induction xs rule: induct_list012)(auto) + by (induction xs rule: induct_list012)(auto) lemma f_arg_min_list_f: "xs \ [] \ f (arg_min_list f xs) = Min (f ` (set xs))" -by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym) + by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym) lemma arg_min_list_in: "xs \ [] \ arg_min_list f xs \ set xs" -by(induction xs rule: induct_list012) (auto simp: Let_def) + by(induction xs rule: induct_list012) (auto simp: Let_def) subsubsection \(In)finiteness\ lemma finite_maxlen: "finite (M::'a list set) \ \n. \s\M. size s < n" proof (induct rule: finite.induct) case emptyI show ?case by simp next case (insertI M xs) then obtain n where "\s\M. length s < n" by blast hence "\s\insert xs M. size s < max n (size xs) + 1" by auto thus ?case .. qed lemma lists_length_Suc_eq: "{xs. set xs \ A \ length xs = Suc n} = (\(xs, n). n#xs) ` ({xs. set xs \ A \ length xs = n} \ A)" by (auto simp: length_Suc_conv) lemma assumes "finite A" shows finite_lists_length_eq: "finite {xs. set xs \ A \ length xs = n}" and card_lists_length_eq: "card {xs. set xs \ A \ length xs = n} = (card A)^n" using \finite A\ by (induct n) (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong) lemma finite_lists_length_le: assumes "finite A" shows "finite {xs. set xs \ A \ length xs \ n}" (is "finite ?S") proof- have "?S = (\n\{0..n}. {xs. set xs \ A \ length xs = n})" by auto thus ?thesis by (auto intro!: finite_lists_length_eq[OF \finite A\] simp only:) qed lemma card_lists_length_le: assumes "finite A" shows "card {xs. set xs \ A \ length xs \ n} = (\i\n. card A^i)" proof - have "(\i\n. card A^i) = card (\i\n. {xs. set xs \ A \ length xs = i})" using \finite A\ by (subst card_UN_disjoint) (auto simp add: card_lists_length_eq finite_lists_length_eq) also have "(\i\n. {xs. set xs \ A \ length xs = i}) = {xs. set xs \ A \ length xs \ n}" by auto finally show ?thesis by simp qed lemma finite_lists_distinct_length_eq [intro]: assumes "finite A" shows "finite {xs. length xs = n \ distinct xs \ set xs \ A}" (is "finite ?S") proof - have "finite {xs. set xs \ A \ length xs = n}" using \finite A\ by (rule finite_lists_length_eq) moreover have "?S \ {xs. set xs \ A \ length xs = n}" by auto ultimately show ?thesis using finite_subset by auto qed lemma card_lists_distinct_length_eq: assumes "finite A" "k \ card A" shows "card {xs. length xs = k \ distinct xs \ set xs \ A} = \{card A - k + 1 .. card A}" using assms proof (induct k) case 0 then have "{xs. length xs = 0 \ distinct xs \ set xs \ A} = {[]}" by auto then show ?case by simp next case (Suc k) let "?k_list" = "\k xs. length xs = k \ distinct xs \ set xs \ A" have inj_Cons: "\A. inj_on (\(xs, n). n # xs) A" by (rule inj_onI) auto from Suc have "k \ card A" by simp moreover note \finite A\ moreover have "finite {xs. ?k_list k xs}" by (rule finite_subset) (use finite_lists_length_eq[OF \finite A\, of k] in auto) moreover have "\i j. i \ j \ {i} \ (A - set i) \ {j} \ (A - set j) = {}" by auto moreover have "\i. i \ {xs. ?k_list k xs} \ card (A - set i) = card A - k" by (simp add: card_Diff_subset distinct_card) moreover have "{xs. ?k_list (Suc k) xs} = (\(xs, n). n#xs) ` \((\xs. {xs} \ (A - set xs)) ` {xs. ?k_list k xs})" by (auto simp: length_Suc_conv) moreover have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp then have "(card A - k) * \{Suc (card A - k)..card A} = \{Suc (card A - Suc k)..card A}" by (subst prod.insert[symmetric]) (simp add: atLeastAtMost_insertL)+ ultimately show ?case by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps) qed lemma card_lists_distinct_length_eq': assumes "k < card A" shows "card {xs. length xs = k \ distinct xs \ set xs \ A} = \{card A - k + 1 .. card A}" proof - from \k < card A\ have "finite A" and "k \ card A" using card_infinite by force+ from this show ?thesis by (rule card_lists_distinct_length_eq) qed lemma infinite_UNIV_listI: "\ finite(UNIV::'a list set)" by (metis UNIV_I finite_maxlen length_replicate less_irrefl) lemma same_length_different: assumes "xs \ ys" and "length xs = length ys" shows "\pre x xs' y ys'. x\y \ xs = pre @ [x] @ xs' \ ys = pre @ [y] @ ys'" using assms proof (induction xs arbitrary: ys) case Nil then show ?case by auto next case (Cons x xs) then obtain z zs where ys: "ys = Cons z zs" by (metis length_Suc_conv) show ?case proof (cases "x=z") case True then have "xs \ zs" "length xs = length zs" using Cons.prems ys by auto then obtain pre u xs' v ys' where "u\v" and xs: "xs = pre @ [u] @ xs'" and zs: "zs = pre @ [v] @ys'" using Cons.IH by meson then have "x # xs = (z#pre) @ [u] @ xs' \ ys = (z#pre) @ [v] @ ys'" by (simp add: True ys) with \u\v\ show ?thesis by blast next case False then have "x # xs = [] @ [x] @ xs \ ys = [] @ [z] @ zs" by (simp add: ys) then show ?thesis using False by blast qed qed subsection \Sorting\ subsubsection \\<^const>\sorted_wrt\\ text \Sometimes the second equation in the definition of \<^const>\sorted_wrt\ is too aggressive because it relates each list element to \emph{all} its successors. Then this equation should be removed and \sorted_wrt2_simps\ should be added instead.\ lemma sorted_wrt1: "sorted_wrt P [x] = True" by(simp) lemma sorted_wrt2: "transp P \ sorted_wrt P (x # y # zs) = (P x y \ sorted_wrt P (y # zs))" proof (induction zs arbitrary: x y) case (Cons z zs) then show ?case by simp (meson transpD)+ qed auto lemmas sorted_wrt2_simps = sorted_wrt1 sorted_wrt2 lemma sorted_wrt_true [simp]: "sorted_wrt (\_ _. True) xs" by (induction xs) simp_all lemma sorted_wrt_append: "sorted_wrt P (xs @ ys) \ sorted_wrt P xs \ sorted_wrt P ys \ (\x\set xs. \y\set ys. P x y)" by (induction xs) auto lemma sorted_wrt_map: "sorted_wrt R (map f xs) = sorted_wrt (\x y. R (f x) (f y)) xs" by (induction xs) simp_all lemma assumes "sorted_wrt f xs" shows sorted_wrt_take: "sorted_wrt f (take n xs)" and sorted_wrt_drop: "sorted_wrt f (drop n xs)" proof - from assms have "sorted_wrt f (take n xs @ drop n xs)" by simp thus "sorted_wrt f (take n xs)" and "sorted_wrt f (drop n xs)" unfolding sorted_wrt_append by simp_all qed lemma sorted_wrt_filter: "sorted_wrt f xs \ sorted_wrt f (filter P xs)" by (induction xs) auto lemma sorted_wrt_rev: "sorted_wrt P (rev xs) = sorted_wrt (\x y. P y x) xs" by (induction xs) (auto simp add: sorted_wrt_append) lemma sorted_wrt_mono_rel: "(\x y. \ x \ set xs; y \ set xs; P x y \ \ Q x y) \ sorted_wrt P xs \ sorted_wrt Q xs" by(induction xs)(auto) lemma sorted_wrt01: "length xs \ 1 \ sorted_wrt P xs" by(auto simp: le_Suc_eq length_Suc_conv) lemma sorted_wrt_iff_nth_less: "sorted_wrt P xs = (\i j. i < j \ j < length xs \ P (xs ! i) (xs ! j))" by (induction xs) (auto simp add: in_set_conv_nth Ball_def nth_Cons split: nat.split) lemma sorted_wrt_nth_less: "\ sorted_wrt P xs; i < j; j < length xs \ \ P (xs ! i) (xs ! j)" by(auto simp: sorted_wrt_iff_nth_less) lemma sorted_wrt_upt[simp]: "sorted_wrt (<) [m..Each element is greater or equal to its index:\ lemma sorted_wrt_less_idx: "sorted_wrt (<) ns \ i < length ns \ i \ ns!i" proof (induction ns arbitrary: i rule: rev_induct) case Nil thus ?case by simp next case snoc thus ?case by (auto simp: nth_append sorted_wrt_append) (metis less_antisym not_less nth_mem) qed subsubsection \\<^const>\sorted\\ context linorder begin text \Sometimes the second equation in the definition of \<^const>\sorted\ is too aggressive because it relates each list element to \emph{all} its successors. Then this equation should be removed and \sorted2_simps\ should be added instead. Executable code is one such use case.\ lemma sorted1: "sorted [x] = True" by simp lemma sorted2: "sorted (x # y # zs) = (x \ y \ sorted (y # zs))" by(induction zs) auto lemmas sorted2_simps = sorted1 sorted2 lemmas [code] = sorted.simps(1) sorted2_simps lemma sorted_append: "sorted (xs@ys) = (sorted xs \ sorted ys \ (\x \ set xs. \y \ set ys. x\y))" by (simp add: sorted_sorted_wrt sorted_wrt_append) lemma sorted_map: "sorted (map f xs) = sorted_wrt (\x y. f x \ f y) xs" by (simp add: sorted_sorted_wrt sorted_wrt_map) lemma sorted01: "length xs \ 1 \ sorted xs" by (simp add: sorted_sorted_wrt sorted_wrt01) lemma sorted_tl: "sorted xs \ sorted (tl xs)" by (cases xs) (simp_all) lemma sorted_iff_nth_mono_less: "sorted xs = (\i j. i < j \ j < length xs \ xs ! i \ xs ! j)" by (simp add: sorted_sorted_wrt sorted_wrt_iff_nth_less) lemma sorted_iff_nth_mono: "sorted xs = (\i j. i \ j \ j < length xs \ xs ! i \ xs ! j)" by (auto simp: sorted_iff_nth_mono_less nat_less_le) lemma sorted_nth_mono: "sorted xs \ i \ j \ j < length xs \ xs!i \ xs!j" by (auto simp: sorted_iff_nth_mono) lemma sorted_rev_nth_mono: "sorted (rev xs) \ i \ j \ j < length xs \ xs!j \ xs!i" using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"] rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"] by auto lemma sorted_map_remove1: "sorted (map f xs) \ sorted (map f (remove1 x xs))" by (induct xs) (auto) lemma sorted_remove1: "sorted xs \ sorted (remove1 a xs)" using sorted_map_remove1 [of "\x. x"] by simp lemma sorted_butlast: assumes "xs \ []" and "sorted xs" shows "sorted (butlast xs)" proof - from \xs \ []\ obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto with \sorted xs\ show ?thesis by (simp add: sorted_append) qed lemma sorted_replicate [simp]: "sorted(replicate n x)" by(induction n) (auto) lemma sorted_remdups[simp]: "sorted xs \ sorted (remdups xs)" by (induct xs) (auto) lemma sorted_remdups_adj[simp]: "sorted xs \ sorted (remdups_adj xs)" by (induct xs rule: remdups_adj.induct, simp_all split: if_split_asm) lemma sorted_nths: "sorted xs \ sorted (nths xs I)" by(induction xs arbitrary: I)(auto simp: nths_Cons) lemma sorted_distinct_set_unique: assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys" shows "xs = ys" proof - from assms have 1: "length xs = length ys" by (auto dest!: distinct_card) from assms show ?thesis proof(induct rule:list_induct2[OF 1]) case 1 show ?case by simp next case 2 thus ?case by simp (metis Diff_insert_absorb antisym insertE insert_iff) qed qed lemma map_sorted_distinct_set_unique: assumes "inj_on f (set xs \ set ys)" assumes "sorted (map f xs)" "distinct (map f xs)" "sorted (map f ys)" "distinct (map f ys)" assumes "set xs = set ys" shows "xs = ys" proof - from assms have "map f xs = map f ys" by (simp add: sorted_distinct_set_unique) with \inj_on f (set xs \ set ys)\ show "xs = ys" by (blast intro: map_inj_on) qed lemma assumes "sorted xs" shows sorted_take: "sorted (take n xs)" and sorted_drop: "sorted (drop n xs)" proof - from assms have "sorted (take n xs @ drop n xs)" by simp then show "sorted (take n xs)" and "sorted (drop n xs)" unfolding sorted_append by simp_all qed lemma sorted_dropWhile: "sorted xs \ sorted (dropWhile P xs)" by (auto dest: sorted_drop simp add: dropWhile_eq_drop) lemma sorted_takeWhile: "sorted xs \ sorted (takeWhile P xs)" by (subst takeWhile_eq_take) (auto dest: sorted_take) lemma sorted_filter: "sorted (map f xs) \ sorted (map f (filter P xs))" by (induct xs) simp_all lemma foldr_max_sorted: assumes "sorted (rev xs)" shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)" using assms proof (induct xs) case (Cons x xs) then have "sorted (rev xs)" using sorted_append by auto with Cons show ?case by (cases xs) (auto simp add: sorted_append max_def) qed simp lemma filter_equals_takeWhile_sorted_rev: assumes sorted: "sorted (rev (map f xs))" shows "filter (\x. t < f x) xs = takeWhile (\ x. t < f x) xs" (is "filter ?P xs = ?tW") proof (rule takeWhile_eq_filter[symmetric]) let "?dW" = "dropWhile ?P xs" fix x assume "x \ set ?dW" then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i" unfolding in_set_conv_nth by auto hence "length ?tW + i < length (?tW @ ?dW)" unfolding length_append by simp hence i': "length (map f ?tW) + i < length (map f xs)" by simp have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \ (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)" using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"] unfolding map_append[symmetric] by simp hence "f x \ f (?dW ! 0)" unfolding nth_append_length_plus nth_i using i preorder_class.le_less_trans[OF le0 i] by simp also have "... \ t" using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i] using hd_conv_nth[of "?dW"] by simp finally show "\ t < f x" by simp qed lemma sorted_map_same: "sorted (map f (filter (\x. f x = g xs) xs))" proof (induct xs arbitrary: g) case Nil then show ?case by simp next case (Cons x xs) then have "sorted (map f (filter (\y. f y = (\xs. f x) xs) xs))" . moreover from Cons have "sorted (map f (filter (\y. f y = (g \ Cons x) xs) xs))" . ultimately show ?case by simp_all qed lemma sorted_same: "sorted (filter (\x. x = g xs) xs)" using sorted_map_same [of "\x. x"] by simp end lemma sorted_upt[simp]: "sorted [m..Sorting functions\ text\Currently it is not shown that \<^const>\sort\ returns a permutation of its input because the nicest proof is via multisets, which are not part of Main. Alternatively one could define a function that counts the number of occurrences of an element in a list and use that instead of multisets to state the correctness property.\ context linorder begin lemma set_insort_key: "set (insort_key f x xs) = insert x (set xs)" by (induct xs) auto lemma length_insort [simp]: "length (insort_key f x xs) = Suc (length xs)" by (induct xs) simp_all lemma insort_key_left_comm: assumes "f x \ f y" shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)" by (induct xs) (auto simp add: assms dest: antisym) lemma insort_left_comm: "insort x (insort y xs) = insort y (insort x xs)" by (cases "x = y") (auto intro: insort_key_left_comm) lemma comp_fun_commute_insort: "comp_fun_commute insort" proof qed (simp add: insort_left_comm fun_eq_iff) lemma sort_key_simps [simp]: "sort_key f [] = []" "sort_key f (x#xs) = insort_key f x (sort_key f xs)" by (simp_all add: sort_key_def) lemma sort_key_conv_fold: assumes "inj_on f (set xs)" shows "sort_key f xs = fold (insort_key f) xs []" proof - have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" proof (rule fold_rev, rule ext) fix zs fix x y assume "x \ set xs" "y \ set xs" with assms have *: "f y = f x \ y = x" by (auto dest: inj_onD) have **: "x = y \ y = x" by auto show "(insort_key f y \ insort_key f x) zs = (insort_key f x \ insort_key f y) zs" by (induct zs) (auto intro: * simp add: **) qed then show ?thesis by (simp add: sort_key_def foldr_conv_fold) qed lemma sort_conv_fold: "sort xs = fold insort xs []" by (rule sort_key_conv_fold) simp lemma length_sort[simp]: "length (sort_key f xs) = length xs" by (induct xs, auto) lemma set_sort[simp]: "set(sort_key f xs) = set xs" by (induct xs) (simp_all add: set_insort_key) lemma distinct_insort: "distinct (insort_key f x xs) = (x \ set xs \ distinct xs)" by(induct xs)(auto simp: set_insort_key) lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs" by (induct xs) (simp_all add: distinct_insort) lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)" by (induct xs) (auto simp: set_insort_key) lemma sorted_insort: "sorted (insort x xs) = sorted xs" using sorted_insort_key [where f="\x. x"] by simp theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))" by (induct xs) (auto simp:sorted_insort_key) theorem sorted_sort [simp]: "sorted (sort xs)" using sorted_sort_key [where f="\x. x"] by simp lemma insort_not_Nil [simp]: "insort_key f a xs \ []" by (induction xs) simp_all lemma insort_is_Cons: "\x\set xs. f a \ f x \ insort_key f a xs = a # xs" by (cases xs) auto lemma sorted_sort_id: "sorted xs \ sort xs = xs" by (induct xs) (auto simp add: insort_is_Cons) lemma insort_key_remove1: assumes "a \ set xs" and "sorted (map f xs)" and "hd (filter (\x. f a = f x) xs) = a" shows "insort_key f a (remove1 a xs) = xs" using assms proof (induct xs) case (Cons x xs) then show ?case proof (cases "x = a") case False then have "f x \ f a" using Cons.prems by auto then have "f x < f a" using Cons.prems by auto with \f x \ f a\ show ?thesis using Cons by (auto simp: insort_is_Cons) qed (auto simp: insort_is_Cons) qed simp lemma insort_remove1: assumes "a \ set xs" and "sorted xs" shows "insort a (remove1 a xs) = xs" proof (rule insort_key_remove1) define n where "n = length (filter ((=) a) xs) - 1" from \a \ set xs\ show "a \ set xs" . from \sorted xs\ show "sorted (map (\x. x) xs)" by simp from \a \ set xs\ have "a \ set (filter ((=) a) xs)" by auto then have "set (filter ((=) a) xs) \ {}" by auto then have "filter ((=) a) xs \ []" by (auto simp only: set_empty) then have "length (filter ((=) a) xs) > 0" by simp then have n: "Suc n = length (filter ((=) a) xs)" by (simp add: n_def) moreover have "replicate (Suc n) a = a # replicate n a" by simp ultimately show "hd (filter ((=) a) xs) = a" by (simp add: replicate_length_filter) qed lemma finite_sorted_distinct_unique: assumes "finite A" shows "\!xs. set xs = A \ sorted xs \ distinct xs" proof - obtain xs where "distinct xs" "A = set xs" using finite_distinct_list [OF assms] by metis then show ?thesis by (rule_tac a="sort xs" in ex1I) (auto simp: sorted_distinct_set_unique) qed lemma insort_insert_key_triv: "f x \ f ` set xs \ insort_insert_key f x xs = xs" by (simp add: insort_insert_key_def) lemma insort_insert_triv: "x \ set xs \ insort_insert x xs = xs" using insort_insert_key_triv [of "\x. x"] by simp lemma insort_insert_insort_key: "f x \ f ` set xs \ insort_insert_key f x xs = insort_key f x xs" by (simp add: insort_insert_key_def) lemma insort_insert_insort: "x \ set xs \ insort_insert x xs = insort x xs" using insort_insert_insort_key [of "\x. x"] by simp lemma set_insort_insert: "set (insort_insert x xs) = insert x (set xs)" by (auto simp add: insort_insert_key_def set_insort_key) lemma distinct_insort_insert: assumes "distinct xs" shows "distinct (insort_insert_key f x xs)" using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort_key) lemma sorted_insort_insert_key: assumes "sorted (map f xs)" shows "sorted (map f (insort_insert_key f x xs))" using assms by (simp add: insort_insert_key_def sorted_insort_key) lemma sorted_insort_insert: assumes "sorted xs" shows "sorted (insort_insert x xs)" using assms sorted_insort_insert_key [of "\x. x"] by simp lemma filter_insort_triv: "\ P x \ filter P (insort_key f x xs) = filter P xs" by (induct xs) simp_all lemma filter_insort: "sorted (map f xs) \ P x \ filter P (insort_key f x xs) = insort_key f x (filter P xs)" by (induct xs) (auto, subst insort_is_Cons, auto) lemma filter_sort: "filter P (sort_key f xs) = sort_key f (filter P xs)" by (induct xs) (simp_all add: filter_insort_triv filter_insort) lemma remove1_insort [simp]: "remove1 x (insort x xs) = xs" by (induct xs) simp_all end lemma sort_upt [simp]: "sort [m.. \x \ set xs. P x \ List.find P xs = Some (Min {x\set xs. P x})" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases "P x") case True with Cons show ?thesis by (auto intro: Min_eqI [symmetric]) next case False then have "{y. (y = x \ y \ set xs) \ P y} = {y \ set xs. P y}" by auto with Cons False show ?thesis by (simp_all) qed qed lemma sorted_enumerate [simp]: "sorted (map fst (enumerate n xs))" by (simp add: enumerate_eq_zip) text \Stability of \<^const>\sort_key\:\ lemma sort_key_stable: "filter (\y. f y = k) (sort_key f xs) = filter (\y. f y = k) xs" by (induction xs) (auto simp: filter_insort insort_is_Cons filter_insort_triv) corollary stable_sort_key_sort_key: "stable_sort_key sort_key" by(simp add: stable_sort_key_def sort_key_stable) lemma sort_key_const: "sort_key (\x. c) xs = xs" by (metis (mono_tags) filter_True sort_key_stable) subsubsection \\<^const>\transpose\ on sorted lists\ lemma sorted_transpose[simp]: "sorted (rev (map length (transpose xs)))" by (auto simp: sorted_iff_nth_mono rev_nth nth_transpose length_filter_conv_card intro: card_mono) lemma transpose_max_length: "foldr (\xs. max (length xs)) (transpose xs) 0 = length (filter (\x. x \ []) xs)" (is "?L = ?R") proof (cases "transpose xs = []") case False have "?L = foldr max (map length (transpose xs)) 0" by (simp add: foldr_map comp_def) also have "... = length (transpose xs ! 0)" using False sorted_transpose by (simp add: foldr_max_sorted) finally show ?thesis using False by (simp add: nth_transpose) next case True hence "filter (\x. x \ []) xs = []" by (auto intro!: filter_False simp: transpose_empty) thus ?thesis by (simp add: transpose_empty True) qed lemma length_transpose_sorted: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))" proof (cases "xs = []") case False thus ?thesis using foldr_max_sorted[OF sorted] False unfolding length_transpose foldr_map comp_def by simp qed simp lemma nth_nth_transpose_sorted[simp]: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and i: "i < length (transpose xs)" and j: "j < length (filter (\ys. i < length ys) xs)" shows "transpose xs ! i ! j = xs ! j ! i" using j filter_equals_takeWhile_sorted_rev[OF sorted, of i] nth_transpose[OF i] nth_map[OF j] by (simp add: takeWhile_nth) lemma transpose_column_length: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and "i < length xs" shows "length (filter (\ys. i < length ys) (transpose xs)) = length (xs ! i)" proof - have "xs \ []" using \i < length xs\ by auto note filter_equals_takeWhile_sorted_rev[OF sorted, simp] { fix j assume "j \ i" note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this \i < length xs\] } note sortedE = this[consumes 1] have "{j. j < length (transpose xs) \ i < length (transpose xs ! j)} = {..< length (xs ! i)}" proof safe fix j assume "j < length (transpose xs)" and "i < length (transpose xs ! j)" with this(2) nth_transpose[OF this(1)] have "i < length (takeWhile (\ys. j < length ys) xs)" by simp from nth_mem[OF this] takeWhile_nth[OF this] show "j < length (xs ! i)" by (auto dest: set_takeWhileD) next fix j assume "j < length (xs ! i)" thus "j < length (transpose xs)" using foldr_max_sorted[OF sorted] \xs \ []\ sortedE[OF le0] by (auto simp: length_transpose comp_def foldr_map) have "Suc i \ length (takeWhile (\ys. j < length ys) xs)" using \i < length xs\ \j < length (xs ! i)\ less_Suc_eq_le by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE) with nth_transpose[OF \j < length (transpose xs)\] show "i < length (transpose xs ! j)" by simp qed thus ?thesis by (simp add: length_filter_conv_card) qed lemma transpose_column: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" and "i < length xs" shows "map (\ys. ys ! i) (filter (\ys. i < length ys) (transpose xs)) = xs ! i" (is "?R = _") proof (rule nth_equalityI) show length: "length ?R = length (xs ! i)" using transpose_column_length[OF assms] by simp fix j assume j: "j < length ?R" note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le] from j have j_less: "j < length (xs ! i)" using length by simp have i_less_tW: "Suc i \ length (takeWhile (\ys. Suc j \ length ys) xs)" proof (rule length_takeWhile_less_P_nth) show "Suc i \ length xs" using \i < length xs\ by simp fix k assume "k < Suc i" hence "k \ i" by auto with sorted_rev_nth_mono[OF sorted this] \i < length xs\ have "length (xs ! i) \ length (xs ! k)" by simp thus "Suc j \ length (xs ! k)" using j_less by simp qed have i_less_filter: "i < length (filter (\ys. j < length ys) xs) " unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j] using i_less_tW by (simp_all add: Suc_le_eq) from j show "?R ! j = xs ! i ! j" unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i] by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter]) qed lemma transpose_transpose: fixes xs :: "'a list list" assumes sorted: "sorted (rev (map length xs))" shows "transpose (transpose xs) = takeWhile (\x. x \ []) xs" (is "?L = ?R") proof - have len: "length ?L = length ?R" unfolding length_transpose transpose_max_length using filter_equals_takeWhile_sorted_rev[OF sorted, of 0] by simp { fix i assume "i < length ?R" with less_le_trans[OF _ length_takeWhile_le[of _ xs]] have "i < length xs" by simp } note * = this show ?thesis by (rule nth_equalityI) (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth) qed theorem transpose_rectangle: assumes "xs = [] \ n = 0" assumes rect: "\ i. i < length xs \ length (xs ! i) = n" shows "transpose xs = map (\ i. map (\ j. xs ! j ! i) [0..ys. i < length ys) xs = xs" using rect by (auto simp: in_set_conv_nth intro!: filter_True) } ultimately show "\i. i < length (transpose xs) \ ?trans ! i = ?map ! i" by (auto simp: nth_transpose intro: nth_equalityI) qed subsubsection \\sorted_list_of_set\\ text\This function maps (finite) linearly ordered sets to sorted lists. Warning: in most cases it is not a good idea to convert from sets to lists but one should convert in the other direction (via \<^const>\set\).\ context linorder begin definition sorted_list_of_set :: "'a set \ 'a list" where "sorted_list_of_set = folding.F insort []" sublocale sorted_list_of_set: folding insort Nil rewrites "folding.F insort [] = sorted_list_of_set" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show "folding insort" by standard (fact comp_fun_commute) show "folding.F insort [] = sorted_list_of_set" by (simp only: sorted_list_of_set_def) qed lemma sorted_list_of_set_empty: "sorted_list_of_set {} = []" -by (fact sorted_list_of_set.empty) + by (fact sorted_list_of_set.empty) lemma sorted_list_of_set_insert [simp]: "finite A \ sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))" -by (fact sorted_list_of_set.insert_remove) + by (fact sorted_list_of_set.insert_remove) lemma sorted_list_of_set_eq_Nil_iff [simp]: "finite A \ sorted_list_of_set A = [] \ A = {}" -by (auto simp: sorted_list_of_set.remove) + by (auto simp: sorted_list_of_set.remove) lemma set_sorted_list_of_set [simp]: "finite A \ set (sorted_list_of_set A) = A" -by(induct A rule: finite_induct) (simp_all add: set_insort_key) + by(induct A rule: finite_induct) (simp_all add: set_insort_key) lemma sorted_sorted_list_of_set [simp]: "sorted (sorted_list_of_set A)" proof (cases "finite A") case True thus ?thesis by(induction A) (simp_all add: sorted_insort) next case False thus ?thesis by simp qed lemma distinct_sorted_list_of_set [simp]: "distinct (sorted_list_of_set A)" proof (cases "finite A") case True thus ?thesis by(induction A) (simp_all add: distinct_insort) next case False thus ?thesis by simp qed lemmas sorted_list_of_set = set_sorted_list_of_set sorted_sorted_list_of_set distinct_sorted_list_of_set lemma sorted_list_of_set_sort_remdups [code]: "sorted_list_of_set (set xs) = sort (remdups xs)" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_set.eq_fold sort_conv_fold fold_set_fold_remdups) qed lemma sorted_list_of_set_remove: assumes "finite A" shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)" proof (cases "x \ A") case False with assms have "x \ set (sorted_list_of_set A)" by simp with False show ?thesis by (simp add: remove1_idem) next case True then obtain B where A: "A = insert x B" by (rule Set.set_insert) with assms show ?thesis by simp qed end lemma sorted_list_of_set_range [simp]: "sorted_list_of_set {m.. sorted l \ distinct l" by (induction l) (use less_le in auto) lemma strict_sorted_list_of_set [simp]: "strict_sorted (sorted_list_of_set A)" by (simp add: strict_sorted_iff) lemma finite_set_strict_sorted: fixes A :: "'a::linorder set" assumes "finite A" obtains l where "strict_sorted l" "List.set l = A" "length l = card A" by (metis assms distinct_card distinct_sorted_list_of_set set_sorted_list_of_set strict_sorted_list_of_set) subsubsection \\lists\: the list-forming operator over sets\ inductive_set lists :: "'a set => 'a list set" for A :: "'a set" where Nil [intro!, simp]: "[] \ lists A" | Cons [intro!, simp]: "\a \ A; l \ lists A\ \ a#l \ lists A" inductive_cases listsE [elim!]: "x#l \ lists A" inductive_cases listspE [elim!]: "listsp A (x # l)" inductive_simps listsp_simps[code]: "listsp A []" "listsp A (x # xs)" -lemma listsp_mono [mono]: "A \ B ==> listsp A \ listsp B" +lemma listsp_mono [mono]: "A \ B \ listsp A \ listsp B" by (rule predicate1I, erule listsp.induct, blast+) lemmas lists_mono = listsp_mono [to_set] lemma listsp_infI: - assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l + assumes l: "listsp A l" shows "listsp B l \ listsp (inf A B) l" using l by induct blast+ lemmas lists_IntI = listsp_infI [to_set] lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)" proof (rule mono_inf [where f=listsp, THEN order_antisym]) show "mono listsp" by (simp add: mono_def listsp_mono) show "inf (listsp A) (listsp B) \ listsp (inf A B)" by (blast intro!: listsp_infI) qed lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def] lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set] lemma Cons_in_lists_iff[simp]: "x#xs \ lists A \ x \ A \ xs \ lists A" by auto lemma append_in_listsp_conv [iff]: "(listsp A (xs @ ys)) = (listsp A xs \ listsp A ys)" by (induct xs) auto lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set] lemma in_listsp_conv_set: "(listsp A xs) = (\x \ set xs. A x)" \ \eliminate \listsp\ in favour of \set\\ by (induct xs) auto lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set] -lemma in_listspD [dest!]: "listsp A xs ==> \x\set xs. A x" +lemma in_listspD [dest!]: "listsp A xs \ \x\set xs. A x" by (rule in_listsp_conv_set [THEN iffD1]) lemmas in_listsD [dest!] = in_listspD [to_set] -lemma in_listspI [intro!]: "\x\set xs. A x ==> listsp A xs" +lemma in_listspI [intro!]: "\x\set xs. A x \ listsp A xs" by (rule in_listsp_conv_set [THEN iffD2]) lemmas in_listsI [intro!] = in_listspI [to_set] lemma lists_eq_set: "lists A = {xs. set xs \ A}" by auto lemma lists_empty [simp]: "lists {} = {[]}" by auto lemma lists_UNIV [simp]: "lists UNIV = UNIV" by auto lemma lists_image: "lists (f`A) = map f ` lists A" proof - { fix xs have "\x\set xs. x \ f ` A \ xs \ map f ` lists A" by (induct xs) (auto simp del: list.map simp add: list.map[symmetric] intro!: imageI) } then show ?thesis by auto qed subsubsection \Inductive definition for membership\ inductive ListMem :: "'a \ 'a list \ bool" where elem: "ListMem x (x # xs)" | insert: "ListMem x xs \ ListMem x (y # xs)" lemma ListMem_iff: "(ListMem x xs) = (x \ set xs)" proof show "ListMem x xs \ x \ set xs" by (induct set: ListMem) auto show "x \ set xs \ ListMem x xs" by (induct xs) (auto intro: ListMem.intros) qed subsubsection \Lists as Cartesian products\ text\\set_Cons A Xs\: the set of lists with head drawn from \<^term>\A\ and tail drawn from \<^term>\Xs\.\ definition set_Cons :: "'a set \ 'a list set \ 'a list set" where "set_Cons A XS = {z. \x xs. z = x # xs \ x \ A \ xs \ XS}" lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A" by (auto simp add: set_Cons_def) text\Yields the set of lists, all of the same length as the argument and with elements drawn from the corresponding element of the argument.\ primrec listset :: "'a set list \ 'a list set" where "listset [] = {[]}" | "listset (A # As) = set_Cons A (listset As)" subsection \Relations on Lists\ subsubsection \Length Lexicographic Ordering\ text\These orderings preserve well-foundedness: shorter lists precede longer lists. These ordering are not used in dictionaries.\ primrec \ \The lexicographic ordering for lists of the specified length\ lexn :: "('a \ 'a) set \ nat \ ('a list \ 'a list) set" where "lexn r 0 = {}" | "lexn r (Suc n) = (map_prod (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int {(xs, ys). length xs = Suc n \ length ys = Suc n}" definition lex :: "('a \ 'a) set \ ('a list \ 'a list) set" where "lex r = (\n. lexn r n)" \ \Holds only between lists of the same length\ definition lenlex :: "('a \ 'a) set => ('a list \ 'a list) set" where "lenlex r = inv_image (less_than <*lex*> lex r) (\xs. (length xs, xs))" \ \Compares lists by their length and then lexicographically\ lemma wf_lexn: assumes "wf r" shows "wf (lexn r n)" proof (induct n) case (Suc n) have inj: "inj (\(x, xs). x # xs)" using assms by (auto simp: inj_on_def) have wf: "wf (map_prod (\(x, xs). x # xs) (\(x, xs). x # xs) ` (r <*lex*> lexn r n))" by (simp add: Suc.hyps assms wf_lex_prod wf_map_prod_image [OF _ inj]) then show ?case by (rule wf_subset) auto qed auto lemma lexn_length: "(xs, ys) \ lexn r n \ length xs = n \ length ys = n" -by (induct n arbitrary: xs ys) auto - -lemma wf_lex [intro!]: "wf r ==> wf (lex r)" + by (induct n arbitrary: xs ys) auto + +lemma wf_lex [intro!]: + assumes "wf r" shows "wf (lex r)" unfolding lex_def - apply (rule wf_UN) - apply (simp add: wf_lexn) - apply (metis DomainE Int_emptyI RangeE lexn_length) - done +proof (rule wf_UN) + show "wf (lexn r i)" for i + by (simp add: assms wf_lexn) + show "\i j. lexn r i \ lexn r j \ Domain (lexn r i) \ Range (lexn r j) = {}" + by (metis DomainE Int_emptyI RangeE lexn_length) +qed + lemma lexn_conv: "lexn r n = {(xs,ys). length xs = n \ length ys = n \ (\xys x y xs' ys'. xs= xys @ x#xs' \ ys= xys @ y # ys' \ (x, y) \ r)}" -apply (induct n, simp) -apply (simp add: image_Collect lex_prod_def, safe, blast) - apply (rule_tac x = "ab # xys" in exI, simp) -apply (case_tac xys, simp_all, blast) -done +proof (induction n) + case (Suc n) + then show ?case + apply (simp add: image_Collect lex_prod_def, safe, blast) + apply (rule_tac x = "ab # xys" in exI, simp) + apply (case_tac xys; force) + done +qed auto text\By Mathias Fleury:\ proposition lexn_transI: assumes "trans r" shows "trans (lexn r n)" unfolding trans_def proof (intro allI impI) fix as bs cs assume asbs: "(as, bs) \ lexn r n" and bscs: "(bs, cs) \ lexn r n" obtain abs a b as' bs' where n: "length as = n" and "length bs = n" and as: "as = abs @ a # as'" and bs: "bs = abs @ b # bs'" and abr: "(a, b) \ r" using asbs unfolding lexn_conv by blast obtain bcs b' c' cs' bs' where n': "length cs = n" and "length bs = n" and bs': "bs = bcs @ b' # bs'" and cs: "cs = bcs @ c' # cs'" and b'c'r: "(b', c') \ r" using bscs unfolding lexn_conv by blast consider (le) "length bcs < length abs" | (eq) "length bcs = length abs" | (ge) "length bcs > length abs" by linarith thus "(as, cs) \ lexn r n" proof cases let ?k = "length bcs" case le hence "as ! ?k = bs ! ?k" unfolding as bs by (simp add: nth_append) hence "(as ! ?k, cs ! ?k) \ r" using b'c'r unfolding bs' cs by auto moreover have "length bcs < length as" using le unfolding as by simp from id_take_nth_drop[OF this] have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" . moreover have "length bcs < length cs" unfolding cs by simp from id_take_nth_drop[OF this] have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" . moreover have "take ?k as = take ?k cs" using le arg_cong[OF bs, of "take (length bcs)"] unfolding cs as bs' by auto ultimately show ?thesis using n n' unfolding lexn_conv by auto next let ?k = "length abs" case ge hence "bs ! ?k = cs ! ?k" unfolding bs' cs by (simp add: nth_append) hence "(as ! ?k, cs ! ?k) \ r" using abr unfolding as bs by auto moreover have "length abs < length as" using ge unfolding as by simp from id_take_nth_drop[OF this] have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" . moreover have "length abs < length cs" using n n' unfolding as by simp from id_take_nth_drop[OF this] have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" . moreover have "take ?k as = take ?k cs" using ge arg_cong[OF bs', of "take (length abs)"] unfolding cs as bs by auto ultimately show ?thesis using n n' unfolding lexn_conv by auto next let ?k = "length abs" case eq hence *: "abs = bcs" "b = b'" using bs bs' by auto hence "(a, c') \ r" using abr b'c'r assms unfolding trans_def by blast with * show ?thesis using n n' unfolding lexn_conv as bs cs by auto qed qed corollary lex_transI: assumes "trans r" shows "trans (lex r)" using lexn_transI [OF assms] by (clarsimp simp add: lex_def trans_def) (metis lexn_length) lemma lex_conv: "lex r = {(xs,ys). length xs = length ys \ (\xys x y xs' ys'. xs = xys @ x # xs' \ ys = xys @ y # ys' \ (x, y) \ r)}" by (force simp add: lex_def lexn_conv) -lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)" +lemma wf_lenlex [intro!]: "wf r \ wf (lenlex r)" by (unfold lenlex_def) blast lemma lenlex_conv: "lenlex r = {(xs,ys). length xs < length ys \ length xs = length ys \ (xs, ys) \ lex r}" by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def) lemma total_lenlex: assumes "total r" shows "total (lenlex r)" proof - have "(xs,ys) \ lexn r (length xs) \ (ys,xs) \ lexn r (length xs)" if "xs \ ys" and len: "length xs = length ys" for xs ys proof - obtain pre x xs' y ys' where "x\y" and xs: "xs = pre @ [x] @ xs'" and ys: "ys = pre @ [y] @ys'" by (meson len \xs \ ys\ same_length_different) then consider "(x,y) \ r" | "(y,x) \ r" by (meson UNIV_I assms total_on_def) then show ?thesis by cases (use len in \(force simp add: lexn_conv xs ys)+\) qed then show ?thesis by (fastforce simp: lenlex_def total_on_def lex_def) qed lemma lenlex_transI [intro]: "trans r \ trans (lenlex r)" unfolding lenlex_def by (meson lex_transI trans_inv_image trans_less_than trans_lex_prod) lemma Nil_notin_lex [iff]: "([], ys) \ lex r" -by (simp add: lex_conv) + by (simp add: lex_conv) lemma Nil2_notin_lex [iff]: "(xs, []) \ lex r" -by (simp add:lex_conv) + by (simp add:lex_conv) lemma Cons_in_lex [simp]: - "((x # xs, y # ys) \ lex r) = - ((x, y) \ r \ length xs = length ys \ x = y \ (xs, ys) \ lex r)" - apply (simp add: lex_conv) - apply (rule iffI) - prefer 2 apply (blast intro: Cons_eq_appendI, clarify) - by (metis hd_append append_Nil list.sel(1) list.sel(3) tl_append2) + "(x # xs, y # ys) \ lex r \ (x, y) \ r \ length xs = length ys \ x = y \ (xs, ys) \ lex r" + (is "?lhs = ?rhs") +proof + assume ?lhs then show ?rhs + by (simp add: lex_conv) (metis hd_append list.sel(1) list.sel(3) tl_append2) +next + assume ?rhs then show ?lhs + by (simp add: lex_conv) (blast intro: Cons_eq_appendI) +qed lemma Nil_lenlex_iff1 [simp]: "([], ns) \ lenlex r \ ns \ []" and Nil_lenlex_iff2 [simp]: "(ns,[]) \ lenlex r" by (auto simp: lenlex_def) lemma Cons_lenlex_iff: "((m # ms, n # ns) \ lenlex r) \ length ms < length ns \ length ms = length ns \ (m,n) \ r \ (m = n \ (ms,ns) \ lenlex r)" by (auto simp: lenlex_def) lemma lenlex_length: "(ms, ns) \ lenlex r \ length ms \ length ns" by (auto simp: lenlex_def) lemma lex_append_rightI: "(xs, ys) \ lex r \ length vs = length us \ (xs @ us, ys @ vs) \ lex r" -by (fastforce simp: lex_def lexn_conv) + by (fastforce simp: lex_def lexn_conv) lemma lex_append_leftI: "(ys, zs) \ lex r \ (xs @ ys, xs @ zs) \ lex r" -by (induct xs) auto + by (induct xs) auto lemma lex_append_leftD: "\x. (x,x) \ r \ (xs @ ys, xs @ zs) \ lex r \ (ys, zs) \ lex r" -by (induct xs) auto + by (induct xs) auto lemma lex_append_left_iff: "\x. (x,x) \ r \ (xs @ ys, xs @ zs) \ lex r \ (ys, zs) \ lex r" -by(metis lex_append_leftD lex_append_leftI) + by(metis lex_append_leftD lex_append_leftI) lemma lex_take_index: assumes "(xs, ys) \ lex r" - obtains i where "i < length xs" and "i < length ys" and "take i xs = -take i ys" + obtains i where "i < length xs" and "i < length ys" and "take i xs = take i ys" and "(xs ! i, ys ! i) \ r" proof - obtain n us x xs' y ys' where "(xs, ys) \ lexn r n" and "length xs = n" and "length ys = n" and "xs = us @ x # xs'" and "ys = us @ y # ys'" and "(x, y) \ r" using assms by (fastforce simp: lex_def lexn_conv) then show ?thesis by (intro that [of "length us"]) auto qed subsubsection \Lexicographic Ordering\ text \Classical lexicographic ordering on lists, ie. "a" < "ab" < "b". This ordering does \emph{not} preserve well-foundedness. Author: N. Voelker, March 2005.\ definition lexord :: "('a \ 'a) set \ ('a list \ 'a list) set" where "lexord r = {(x,y). \ a v. y = x @ a # v \ (\ u a b v w. (a,b) \ r \ x = u @ (a # v) \ y = u @ (b # w))}" lemma lexord_Nil_left[simp]: "([],y) \ lexord r = (\ a x. y = a # x)" by (unfold lexord_def, induct_tac y, auto) lemma lexord_Nil_right[simp]: "(x,[]) \ lexord r" by (unfold lexord_def, induct_tac x, auto) lemma lexord_cons_cons[simp]: - "((a # x, b # y) \ lexord r) = ((a,b)\ r \ (a = b \ (x,y)\ lexord r))" - unfolding lexord_def - apply (safe, simp_all) - apply (metis hd_append list.sel(1)) + "(a # x, b # y) \ lexord r \ (a,b)\ r \ (a = b \ (x,y)\ lexord r)" (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (simp add: lexord_def) apply (metis hd_append list.sel(1) list.sel(3) tl_append2) - apply blast - by (meson Cons_eq_appendI) + done +qed (auto simp add: lexord_def; (blast | meson Cons_eq_appendI)) lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons lemma lexord_append_rightI: "\ b z. y = b # z \ (x, x @ y) \ lexord r" -by (induct_tac x, auto) + by (induct_tac x, auto) lemma lexord_append_left_rightI: - "(a,b) \ r \ (u @ a # x, u @ b # y) \ lexord r" -by (induct_tac u, auto) + "(a,b) \ r \ (u @ a # x, u @ b # y) \ lexord r" + by (induct_tac u, auto) lemma lexord_append_leftI: " (u,v) \ lexord r \ (x @ u, x @ v) \ lexord r" -by (induct x, auto) + by (induct x, auto) lemma lexord_append_leftD: - "\ (x @ u, x @ v) \ lexord r; (\a. (a,a) \ r) \ \ (u,v) \ lexord r" -by (erule rev_mp, induct_tac x, auto) + "\(x @ u, x @ v) \ lexord r; (\a. (a,a) \ r) \ \ (u,v) \ lexord r" + by (erule rev_mp, induct_tac x, auto) lemma lexord_take_index_conv: "((x,y) \ lexord r) = ((length x < length y \ take (length x) y = x) \ (\i. i < min(length x)(length y) \ take i x = take i y \ (x!i,y!i) \ r))" - apply (unfold lexord_def Let_def, clarsimp) - apply (rule_tac f = "(% a b. a \ b)" in arg_cong2) - apply (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le) - apply auto - apply (rule_tac x ="length u" in exI, simp) - by (metis id_take_nth_drop) +proof - + have "(\a v. y = x @ a # v) = (length x < length y \ take (length x) y = x)" + by (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le) + moreover + have "(\u a b. (a, b) \ r \ (\v. x = u @ a # v) \ (\w. y = u @ b # w)) = + (\i take i x = take i y \ (x ! i, y ! i) \ r)" + apply safe + using less_iff_Suc_add apply auto[1] + by (metis id_take_nth_drop) + ultimately show ?thesis + by (auto simp: lexord_def Let_def) +qed \ \lexord is extension of partial ordering List.lex\ lemma lexord_lex: "(x,y) \ lex r = ((x,y) \ lexord r \ length x = length y)" proof (induction x arbitrary: y) case (Cons a x y) then show ?case by (cases y) (force+) qed auto lemma lexord_irreflexive: "\x. (x,x) \ r \ (xs,xs) \ lexord r" by (induct xs) auto text\By Ren\'e Thiemann:\ lemma lexord_partial_trans: "(\x y z. x \ set xs \ (x,y) \ r \ (y,z) \ r \ (x,z) \ r) \ (xs,ys) \ lexord r \ (ys,zs) \ lexord r \ (xs,zs) \ lexord r" proof (induct xs arbitrary: ys zs) case Nil from Nil(3) show ?case unfolding lexord_def by (cases zs, auto) next case (Cons x xs yys zzs) from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def by (cases yys, auto) note Cons = Cons[unfolded yys] from Cons(3) have one: "(x,y) \ r \ x = y \ (xs,ys) \ lexord r" by auto from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def by (cases zzs, auto) note Cons = Cons[unfolded zzs] from Cons(4) have two: "(y,z) \ r \ y = z \ (ys,zs) \ lexord r" by auto { assume "(xs,ys) \ lexord r" and "(ys,zs) \ lexord r" from Cons(1)[OF _ this] Cons(2) have "(xs,zs) \ lexord r" by auto } note ind1 = this { assume "(x,y) \ r" and "(y,z) \ r" from Cons(2)[OF _ this] have "(x,z) \ r" by auto } note ind2 = this from one two ind1 ind2 have "(x,z) \ r \ x = z \ (xs,zs) \ lexord r" by blast thus ?case unfolding zzs by auto qed lemma lexord_trans: "\ (x, y) \ lexord r; (y, z) \ lexord r; trans r \ \ (x, z) \ lexord r" by(auto simp: trans_def intro:lexord_partial_trans) lemma lexord_transI: "trans r \ trans (lexord r)" by (rule transI, drule lexord_trans, blast) lemma lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" proof (induction x arbitrary: y) case Nil then show ?case by (metis lexord_Nil_left list.exhaust) next case (Cons a x y) then show ?case by (cases y) (force+) qed lemma lexord_irrefl: "irrefl R \ irrefl (lexord R)" by (simp add: irrefl_def lexord_irreflexive) lemma lexord_asym: assumes "asym R" shows "asym (lexord R)" proof from assms obtain "irrefl R" by (blast elim: asym.cases) then show "irrefl (lexord R)" by (rule lexord_irrefl) next fix xs ys assume "(xs, ys) \ lexord R" then show "(ys, xs) \ lexord R" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then obtain z zs where ys: "ys = z # zs" by (cases ys) auto with assms Cons show ?case by (auto elim: asym.cases) qed qed lemma lexord_asymmetric: assumes "asym R" assumes hyp: "(a, b) \ lexord R" shows "(b, a) \ lexord R" proof - from \asym R\ have "asym (lexord R)" by (rule lexord_asym) then show ?thesis by (rule asym.cases) (auto simp add: hyp) qed text \ Predicate version of lexicographic order integrated with Isabelle's order type classes. Author: Andreas Lochbihler \ context ord begin context notes [[inductive_internals]] begin inductive lexordp :: "'a list \ 'a list \ bool" where Nil: "lexordp [] (y # ys)" | Cons: "x < y \ lexordp (x # xs) (y # ys)" | Cons_eq: "\ \ x < y; \ y < x; lexordp xs ys \ \ lexordp (x # xs) (y # ys)" end lemma lexordp_simps [simp]: "lexordp [] ys = (ys \ [])" "lexordp xs [] = False" "lexordp (x # xs) (y # ys) \ x < y \ \ y < x \ lexordp xs ys" by(subst lexordp.simps, fastforce simp add: neq_Nil_conv)+ inductive lexordp_eq :: "'a list \ 'a list \ bool" where Nil: "lexordp_eq [] ys" | Cons: "x < y \ lexordp_eq (x # xs) (y # ys)" | Cons_eq: "\ \ x < y; \ y < x; lexordp_eq xs ys \ \ lexordp_eq (x # xs) (y # ys)" lemma lexordp_eq_simps [simp]: "lexordp_eq [] ys = True" "lexordp_eq xs [] \ xs = []" "lexordp_eq (x # xs) [] = False" "lexordp_eq (x # xs) (y # ys) \ x < y \ \ y < x \ lexordp_eq xs ys" by(subst lexordp_eq.simps, fastforce)+ lemma lexordp_append_rightI: "ys \ Nil \ lexordp xs (xs @ ys)" by(induct xs)(auto simp add: neq_Nil_conv) lemma lexordp_append_left_rightI: "x < y \ lexordp (us @ x # xs) (us @ y # ys)" by(induct us) auto lemma lexordp_eq_refl: "lexordp_eq xs xs" by(induct xs) simp_all lemma lexordp_append_leftI: "lexordp us vs \ lexordp (xs @ us) (xs @ vs)" by(induct xs) auto lemma lexordp_append_leftD: "\ lexordp (xs @ us) (xs @ vs); \a. \ a < a \ \ lexordp us vs" by(induct xs) auto lemma lexordp_irreflexive: assumes irrefl: "\x. \ x < x" shows "\ lexordp xs xs" proof assume "lexordp xs xs" thus False by(induct xs ys\xs)(simp_all add: irrefl) qed lemma lexordp_into_lexordp_eq: assumes "lexordp xs ys" shows "lexordp_eq xs ys" using assms by induct simp_all end declare ord.lexordp_simps [simp, code] declare ord.lexordp_eq_simps [code, simp] lemma lexord_code [code, code_unfold]: "lexordp = ord.lexordp less" unfolding lexordp_def ord.lexordp_def .. context order begin lemma lexordp_antisym: assumes "lexordp xs ys" "lexordp ys xs" shows False using assms by induct auto lemma lexordp_irreflexive': "\ lexordp xs xs" by(rule lexordp_irreflexive) simp end context linorder begin lemma lexordp_cases [consumes 1, case_names Nil Cons Cons_eq, cases pred: lexordp]: assumes "lexordp xs ys" obtains (Nil) y ys' where "xs = []" "ys = y # ys'" | (Cons) x xs' y ys' where "xs = x # xs'" "ys = y # ys'" "x < y" | (Cons_eq) x xs' ys' where "xs = x # xs'" "ys = x # ys'" "lexordp xs' ys'" using assms by cases (fastforce simp add: not_less_iff_gr_or_eq)+ lemma lexordp_induct [consumes 1, case_names Nil Cons Cons_eq, induct pred: lexordp]: assumes major: "lexordp xs ys" and Nil: "\y ys. P [] (y # ys)" and Cons: "\x xs y ys. x < y \ P (x # xs) (y # ys)" and Cons_eq: "\x xs ys. \ lexordp xs ys; P xs ys \ \ P (x # xs) (x # ys)" shows "P xs ys" using major by induct (simp_all add: Nil Cons not_less_iff_gr_or_eq Cons_eq) lemma lexordp_iff: "lexordp xs ys \ (\x vs. ys = xs @ x # vs) \ (\us a b vs ws. a < b \ xs = us @ a # vs \ ys = us @ b # ws)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs proof induct case Cons_eq thus ?case by simp (metis append.simps(2)) qed(fastforce intro: disjI2 del: disjCI intro: exI[where x="[]"])+ next assume ?rhs thus ?lhs by(auto intro: lexordp_append_leftI[where us="[]", simplified] lexordp_append_leftI) qed lemma lexordp_conv_lexord: "lexordp xs ys \ (xs, ys) \ lexord {(x, y). x < y}" by(simp add: lexordp_iff lexord_def) lemma lexordp_eq_antisym: assumes "lexordp_eq xs ys" "lexordp_eq ys xs" shows "xs = ys" using assms by induct simp_all lemma lexordp_eq_trans: assumes "lexordp_eq xs ys" and "lexordp_eq ys zs" shows "lexordp_eq xs zs" using assms by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_trans: assumes "lexordp xs ys" "lexordp ys zs" shows "lexordp xs zs" using assms by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_linear: "lexordp xs ys \ xs = ys \ lexordp ys xs" by(induct xs arbitrary: ys; case_tac ys; fastforce) lemma lexordp_conv_lexordp_eq: "lexordp xs ys \ lexordp_eq xs ys \ \ lexordp_eq ys xs" (is "?lhs \ ?rhs") proof assume ?lhs hence "\ lexordp_eq ys xs" by induct simp_all with \?lhs\ show ?rhs by (simp add: lexordp_into_lexordp_eq) next assume ?rhs hence "lexordp_eq xs ys" "\ lexordp_eq ys xs" by simp_all thus ?lhs by induct simp_all qed lemma lexordp_eq_conv_lexord: "lexordp_eq xs ys \ xs = ys \ lexordp xs ys" by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym) lemma lexordp_eq_linear: "lexordp_eq xs ys \ lexordp_eq ys xs" by (induct xs arbitrary: ys) (case_tac ys; auto)+ lemma lexordp_linorder: "class.linorder lexordp_eq lexordp" by unfold_locales (auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl lexordp_eq_antisym intro: lexordp_eq_trans del: disjCI intro: lexordp_eq_linear) end lemma sorted_insort_is_snoc: "sorted xs \ \x \ set xs. a \ x \ insort a xs = xs @ [a]" by (induct xs) (auto dest!: insort_is_Cons) subsubsection \Lexicographic combination of measure functions\ text \These are useful for termination proofs\ definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)" lemma wf_measures[simp]: "wf (measures fs)" unfolding measures_def by blast lemma in_measures[simp]: "(x, y) \ measures [] = False" "(x, y) \ measures (f # fs) = (f x < f y \ (f x = f y \ (x, y) \ measures fs))" unfolding measures_def by auto -lemma measures_less: "f x < f y ==> (x, y) \ measures (f#fs)" +lemma measures_less: "f x < f y \ (x, y) \ measures (f#fs)" by simp -lemma measures_lesseq: "f x \ f y ==> (x, y) \ measures fs ==> (x, y) \ measures (f#fs)" +lemma measures_lesseq: "f x \ f y \ (x, y) \ measures fs \ (x, y) \ measures (f#fs)" by auto subsubsection \Lifting Relations to Lists: one element\ definition listrel1 :: "('a \ 'a) set \ ('a list \ 'a list) set" where "listrel1 r = {(xs,ys). \us z z' vs. xs = us @ z # vs \ (z,z') \ r \ ys = us @ z' # vs}" lemma listrel1I: "\ (x, y) \ r; xs = us @ x # vs; ys = us @ y # vs \ \ (xs, ys) \ listrel1 r" unfolding listrel1_def by auto lemma listrel1E: "\ (xs, ys) \ listrel1 r; !!x y us vs. \ (x, y) \ r; xs = us @ x # vs; ys = us @ y # vs \ \ P \ \ P" unfolding listrel1_def by auto lemma not_Nil_listrel1 [iff]: "([], xs) \ listrel1 r" unfolding listrel1_def by blast lemma not_listrel1_Nil [iff]: "(xs, []) \ listrel1 r" unfolding listrel1_def by blast lemma Cons_listrel1_Cons [iff]: "(x # xs, y # ys) \ listrel1 r \ (x,y) \ r \ xs = ys \ x = y \ (xs, ys) \ listrel1 r" by (simp add: listrel1_def Cons_eq_append_conv) (blast) lemma listrel1I1: "(x,y) \ r \ (x # xs, y # xs) \ listrel1 r" by fast lemma listrel1I2: "(xs, ys) \ listrel1 r \ (x # xs, x # ys) \ listrel1 r" by fast lemma append_listrel1I: "(xs, ys) \ listrel1 r \ us = vs \ xs = ys \ (us, vs) \ listrel1 r \ (xs @ us, ys @ vs) \ listrel1 r" unfolding listrel1_def by auto (blast intro: append_eq_appendI)+ lemma Cons_listrel1E1[elim!]: assumes "(x # xs, ys) \ listrel1 r" and "\y. ys = y # xs \ (x, y) \ r \ R" and "\zs. ys = x # zs \ (xs, zs) \ listrel1 r \ R" shows R using assms by (cases ys) blast+ lemma Cons_listrel1E2[elim!]: assumes "(xs, y # ys) \ listrel1 r" and "\x. xs = x # ys \ (x, y) \ r \ R" and "\zs. xs = y # zs \ (zs, ys) \ listrel1 r \ R" shows R using assms by (cases xs) blast+ lemma snoc_listrel1_snoc_iff: "(xs @ [x], ys @ [y]) \ listrel1 r \ (xs, ys) \ listrel1 r \ x = y \ xs = ys \ (x,y) \ r" (is "?L \ ?R") proof assume ?L thus ?R by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append) next assume ?R then show ?L unfolding listrel1_def by force qed lemma listrel1_eq_len: "(xs,ys) \ listrel1 r \ length xs = length ys" unfolding listrel1_def by auto lemma listrel1_mono: "r \ s \ listrel1 r \ listrel1 s" unfolding listrel1_def by blast lemma listrel1_converse: "listrel1 (r\) = (listrel1 r)\" unfolding listrel1_def by blast lemma in_listrel1_converse: "(x,y) \ listrel1 (r\) \ (x,y) \ (listrel1 r)\" unfolding listrel1_def by blast lemma listrel1_iff_update: "(xs,ys) \ (listrel1 r) \ (\y n. (xs ! n, y) \ r \ n < length xs \ ys = xs[n:=y])" (is "?L \ ?R") proof assume "?L" then obtain x y u v where "xs = u @ x # v" "ys = u @ y # v" "(x,y) \ r" unfolding listrel1_def by auto then have "ys = xs[length u := y]" and "length u < length xs" and "(xs ! length u, y) \ r" by auto then show "?R" by auto next assume "?R" then obtain x y n where "(xs!n, y) \ r" "n < size xs" "ys = xs[n:=y]" "x = xs!n" by auto then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \ r" by (auto intro: upd_conv_take_nth_drop id_take_nth_drop) then show "?L" by (auto simp: listrel1_def) qed text\Accessible part and wellfoundedness:\ lemma Cons_acc_listrel1I [intro!]: "x \ Wellfounded.acc r \ xs \ Wellfounded.acc (listrel1 r) \ (x # xs) \ Wellfounded.acc (listrel1 r)" apply (induct arbitrary: xs set: Wellfounded.acc) apply (erule thin_rl) apply (erule acc_induct) apply (rule accI) apply (blast) done lemma lists_accD: "xs \ lists (Wellfounded.acc r) \ xs \ Wellfounded.acc (listrel1 r)" proof (induct set: lists) case Nil then show ?case by (meson acc.intros not_listrel1_Nil) next case (Cons a l) then show ?case by blast qed lemma lists_accI: "xs \ Wellfounded.acc (listrel1 r) \ xs \ lists (Wellfounded.acc r)" apply (induct set: Wellfounded.acc) apply clarify apply (rule accI) apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def) done lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r" by (auto simp: wf_acc_iff intro: lists_accD lists_accI[THEN Cons_in_lists_iff[THEN iffD1, THEN conjunct1]]) subsubsection \Lifting Relations to Lists: all elements\ inductive_set listrel :: "('a \ 'b) set \ ('a list \ 'b list) set" for r :: "('a \ 'b) set" where Nil: "([],[]) \ listrel r" - | Cons: "[| (x,y) \ r; (xs,ys) \ listrel r |] ==> (x#xs, y#ys) \ listrel r" + | Cons: "\(x,y) \ r; (xs,ys) \ listrel r\ \ (x#xs, y#ys) \ listrel r" inductive_cases listrel_Nil1 [elim!]: "([],xs) \ listrel r" inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \ listrel r" inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \ listrel r" inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \ listrel r" lemma listrel_eq_len: "(xs, ys) \ listrel r \ length xs = length ys" by(induct rule: listrel.induct) auto lemma listrel_iff_zip [code_unfold]: "(xs,ys) \ listrel r \ length xs = length ys \ (\(x,y) \ set(zip xs ys). (x,y) \ r)" (is "?L \ ?R") proof assume ?L thus ?R by induct (auto intro: listrel_eq_len) next assume ?R thus ?L apply (clarify) by (induct rule: list_induct2) (auto intro: listrel.intros) qed lemma listrel_iff_nth: "(xs,ys) \ listrel r \ length xs = length ys \ (\n < length xs. (xs!n, ys!n) \ r)" (is "?L \ ?R") by (auto simp add: all_set_conv_all_nth listrel_iff_zip) lemma listrel_mono: "r \ s \ listrel r \ listrel s" by (meson listrel_iff_nth subrelI subset_eq) lemma listrel_subset: "r \ A \ A \ listrel r \ lists A \ lists A" apply clarify apply (erule listrel.induct, auto) done lemma listrel_refl_on: "refl_on A r \ refl_on (lists A) (listrel r)" apply (simp add: refl_on_def listrel_subset Ball_def) apply (rule allI) apply (induct_tac x) apply (auto intro: listrel.intros) done lemma listrel_sym: "sym r \ sym (listrel r)" by (simp add: listrel_iff_nth sym_def) lemma listrel_trans: "trans r \ trans (listrel r)" apply (simp add: trans_def) apply (intro allI) apply (rule impI) apply (erule listrel.induct) apply (blast intro: listrel.intros)+ done theorem equiv_listrel: "equiv A r \ equiv (lists A) (listrel r)" by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans) lemma listrel_rtrancl_refl[iff]: "(xs,xs) \ listrel(r\<^sup>*)" using listrel_refl_on[of UNIV, OF refl_rtrancl] by(auto simp: refl_on_def) lemma listrel_rtrancl_trans: "\(xs,ys) \ listrel(r\<^sup>*); (ys,zs) \ listrel(r\<^sup>*)\ \ (xs,zs) \ listrel(r\<^sup>*)" by (metis listrel_trans trans_def trans_rtrancl) lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}" by (blast intro: listrel.intros) lemma listrel_Cons: "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})" by (auto simp add: set_Cons_def intro: listrel.intros) text \Relating \<^term>\listrel1\, \<^term>\listrel\ and closures:\ lemma listrel1_rtrancl_subset_rtrancl_listrel1: "listrel1 (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof (rule subrelI) fix xs ys assume 1: "(xs,ys) \ listrel1 (r\<^sup>*)" { fix x y us vs have "(x,y) \ r\<^sup>* \ (us @ x # vs, us @ y # vs) \ (listrel1 r)\<^sup>*" proof(induct rule: rtrancl.induct) case rtrancl_refl show ?case by simp next case rtrancl_into_rtrancl thus ?case by (metis listrel1I rtrancl.rtrancl_into_rtrancl) qed } thus "(xs,ys) \ (listrel1 r)\<^sup>*" using 1 by(blast elim: listrel1E) qed lemma rtrancl_listrel1_eq_len: "(x,y) \ (listrel1 r)\<^sup>* \ length x = length y" by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len) lemma rtrancl_listrel1_ConsI1: "(xs,ys) \ (listrel1 r)\<^sup>* \ (x#xs,x#ys) \ (listrel1 r)\<^sup>*" apply(induct rule: rtrancl.induct) apply simp by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl) lemma rtrancl_listrel1_ConsI2: "(x,y) \ r\<^sup>* \ (xs, ys) \ (listrel1 r)\<^sup>* \ (x # xs, y # ys) \ (listrel1 r)\<^sup>*" by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1 subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1]) lemma listrel1_subset_listrel: "r \ r' \ refl r' \ listrel1 r \ listrel(r')" by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def) lemma listrel_reflcl_if_listrel1: "(xs,ys) \ listrel1 r \ (xs,ys) \ listrel(r\<^sup>*)" by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip) lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r\<^sup>*) = (listrel1 r)\<^sup>*" proof { fix x y assume "(x,y) \ listrel (r\<^sup>*)" then have "(x,y) \ (listrel1 r)\<^sup>*" by induct (auto intro: rtrancl_listrel1_ConsI2) } then show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" by (rule subrelI) next show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof(rule subrelI) fix xs ys assume "(xs,ys) \ (listrel1 r)\<^sup>*" then show "(xs,ys) \ listrel (r\<^sup>*)" proof induct case base show ?case by(auto simp add: listrel_iff_zip set_zip) next case (step ys zs) thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans) qed qed qed lemma rtrancl_listrel1_if_listrel: "(xs,ys) \ listrel r \ (xs,ys) \ (listrel1 r)\<^sup>*" by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI) lemma listrel_subset_rtrancl_listrel1: "listrel r \ (listrel1 r)\<^sup>*" by(fast intro:rtrancl_listrel1_if_listrel) subsection \Size function\ lemma [measure_function]: "is_measure f \ is_measure (size_list f)" by (rule is_measure_trivial) lemma [measure_function]: "is_measure f \ is_measure (size_option f)" by (rule is_measure_trivial) lemma size_list_estimation[termination_simp]: "x \ set xs \ y < f x \ y < size_list f xs" by (induct xs) auto lemma size_list_estimation'[termination_simp]: "x \ set xs \ y \ f x \ y \ size_list f xs" by (induct xs) auto lemma size_list_map[simp]: "size_list f (map g xs) = size_list (f \ g) xs" by (induct xs) auto lemma size_list_append[simp]: "size_list f (xs @ ys) = size_list f xs + size_list f ys" by (induct xs, auto) lemma size_list_pointwise[termination_simp]: "(\x. x \ set xs \ f x \ g x) \ size_list f xs \ size_list g xs" by (induct xs) force+ subsection \Monad operation\ definition bind :: "'a list \ ('a \ 'b list) \ 'b list" where "bind xs f = concat (map f xs)" hide_const (open) bind lemma bind_simps [simp]: "List.bind [] f = []" "List.bind (x # xs) f = f x @ List.bind xs f" by (simp_all add: bind_def) lemma list_bind_cong [fundef_cong]: assumes "xs = ys" "(\x. x \ set xs \ f x = g x)" shows "List.bind xs f = List.bind ys g" proof - from assms(2) have "List.bind xs f = List.bind xs g" by (induction xs) simp_all with assms(1) show ?thesis by simp qed lemma set_list_bind: "set (List.bind xs f) = (\x\set xs. set (f x))" by (induction xs) simp_all subsection \Code generation\ text\Optional tail recursive version of \<^const>\map\. Can avoid stack overflow in some target languages.\ fun map_tailrec_rev :: "('a \ 'b) \ 'a list \ 'b list \ 'b list" where "map_tailrec_rev f [] bs = bs" | "map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)" lemma map_tailrec_rev: "map_tailrec_rev f as bs = rev(map f as) @ bs" by(induction as arbitrary: bs) simp_all definition map_tailrec :: "('a \ 'b) \ 'a list \ 'b list" where "map_tailrec f as = rev (map_tailrec_rev f as [])" text\Code equation:\ lemma map_eq_map_tailrec: "map = map_tailrec" by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev) subsubsection \Counterparts for set-related operations\ definition member :: "'a list \ 'a \ bool" where [code_abbrev]: "member xs x \ x \ set xs" text \ Use \member\ only for generating executable code. Otherwise use \<^prop>\x \ set xs\ instead --- it is much easier to reason about. \ lemma member_rec [code]: "member (x # xs) y \ x = y \ member xs y" "member [] y \ False" by (auto simp add: member_def) lemma in_set_member (* FIXME delete candidate *): "x \ set xs \ member xs x" by (simp add: member_def) lemmas list_all_iff [code_abbrev] = fun_cong[OF list.pred_set] definition list_ex :: "('a \ bool) \ 'a list \ bool" where list_ex_iff [code_abbrev]: "list_ex P xs \ Bex (set xs) P" definition list_ex1 :: "('a \ bool) \ 'a list \ bool" where list_ex1_iff [code_abbrev]: "list_ex1 P xs \ (\! x. x \ set xs \ P x)" text \ Usually you should prefer \\x\set xs\, \\x\set xs\ and \\!x. x\set xs \ _\ over \<^const>\list_all\, \<^const>\list_ex\ and \<^const>\list_ex1\ in specifications. \ lemma list_all_simps [code]: "list_all P (x # xs) \ P x \ list_all P xs" "list_all P [] \ True" by (simp_all add: list_all_iff) lemma list_ex_simps [simp, code]: "list_ex P (x # xs) \ P x \ list_ex P xs" "list_ex P [] \ False" by (simp_all add: list_ex_iff) lemma list_ex1_simps [simp, code]: "list_ex1 P [] = False" "list_ex1 P (x # xs) = (if P x then list_all (\y. \ P y \ x = y) xs else list_ex1 P xs)" by (auto simp add: list_ex1_iff list_all_iff) lemma Ball_set_list_all: (* FIXME delete candidate *) "Ball (set xs) P \ list_all P xs" by (simp add: list_all_iff) lemma Bex_set_list_ex: (* FIXME delete candidate *) "Bex (set xs) P \ list_ex P xs" by (simp add: list_ex_iff) lemma list_all_append [simp]: "list_all P (xs @ ys) \ list_all P xs \ list_all P ys" by (auto simp add: list_all_iff) lemma list_ex_append [simp]: "list_ex P (xs @ ys) \ list_ex P xs \ list_ex P ys" by (auto simp add: list_ex_iff) lemma list_all_rev [simp]: "list_all P (rev xs) \ list_all P xs" by (simp add: list_all_iff) lemma list_ex_rev [simp]: "list_ex P (rev xs) \ list_ex P xs" by (simp add: list_ex_iff) lemma list_all_length: "list_all P xs \ (\n < length xs. P (xs ! n))" by (auto simp add: list_all_iff set_conv_nth) lemma list_ex_length: "list_ex P xs \ (\n < length xs. P (xs ! n))" by (auto simp add: list_ex_iff set_conv_nth) lemmas list_all_cong [fundef_cong] = list.pred_cong lemma list_ex_cong [fundef_cong]: "xs = ys \ (\x. x \ set ys \ f x = g x) \ list_ex f xs = list_ex g ys" by (simp add: list_ex_iff) definition can_select :: "('a \ bool) \ 'a set \ bool" where [code_abbrev]: "can_select P A = (\!x\A. P x)" lemma can_select_set_list_ex1 [code]: "can_select P (set A) = list_ex1 P A" by (simp add: list_ex1_iff can_select_def) text \Executable checks for relations on sets\ definition listrel1p :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "listrel1p r xs ys = ((xs, ys) \ listrel1 {(x, y). r x y})" lemma [code_unfold]: "(xs, ys) \ listrel1 r = listrel1p (\x y. (x, y) \ r) xs ys" unfolding listrel1p_def by auto lemma [code]: "listrel1p r [] xs = False" "listrel1p r xs [] = False" "listrel1p r (x # xs) (y # ys) \ r x y \ xs = ys \ x = y \ listrel1p r xs ys" by (simp add: listrel1p_def)+ definition lexordp :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where "lexordp r xs ys = ((xs, ys) \ lexord {(x, y). r x y})" lemma [code_unfold]: "(xs, ys) \ lexord r = lexordp (\x y. (x, y) \ r) xs ys" unfolding lexordp_def by auto lemma [code]: "lexordp r xs [] = False" "lexordp r [] (y#ys) = True" "lexordp r (x # xs) (y # ys) = (r x y \ (x = y \ lexordp r xs ys))" unfolding lexordp_def by auto text \Bounded quantification and summation over nats.\ lemma atMost_upto [code_unfold]: "{..n} = set [0..m (\m \ {0..m (\m \ {0..m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto lemma ex_nat_less [code_unfold]: "(\m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto text\Bounded \LEAST\ operator:\ definition "Bleast S P = (LEAST x. x \ S \ P x)" definition "abort_Bleast S P = (LEAST x. x \ S \ P x)" declare [[code abort: abort_Bleast]] lemma Bleast_code [code]: "Bleast (set xs) P = (case filter P (sort xs) of x#xs \ x | [] \ abort_Bleast (set xs) P)" proof (cases "filter P (sort xs)") case Nil thus ?thesis by (simp add: Bleast_def abort_Bleast_def) next case (Cons x ys) have "(LEAST x. x \ set xs \ P x) = x" proof (rule Least_equality) show "x \ set xs \ P x" by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort) next fix y assume "y \ set xs \ P y" hence "y \ set (filter P xs)" by auto thus "x \ y" by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted.simps(2) sorted_sort) qed thus ?thesis using Cons by (simp add: Bleast_def) qed declare Bleast_def[symmetric, code_unfold] text \Summation over ints.\ lemma greaterThanLessThan_upto [code_unfold]: "{i<..Optimizing by rewriting\ definition null :: "'a list \ bool" where [code_abbrev]: "null xs \ xs = []" text \ Efficient emptyness check is implemented by \<^const>\null\. \ lemma null_rec [code]: "null (x # xs) \ False" "null [] \ True" by (simp_all add: null_def) lemma eq_Nil_null: (* FIXME delete candidate *) "xs = [] \ null xs" by (simp add: null_def) lemma equal_Nil_null [code_unfold]: "HOL.equal xs [] \ null xs" "HOL.equal [] = null" by (auto simp add: equal null_def) definition maps :: "('a \ 'b list) \ 'a list \ 'b list" where [code_abbrev]: "maps f xs = concat (map f xs)" definition map_filter :: "('a \ 'b option) \ 'a list \ 'b list" where [code_post]: "map_filter f xs = map (the \ f) (filter (\x. f x \ None) xs)" text \ Operations \<^const>\maps\ and \<^const>\map_filter\ avoid intermediate lists on execution -- do not use for proving. \ lemma maps_simps [code]: "maps f (x # xs) = f x @ maps f xs" "maps f [] = []" by (simp_all add: maps_def) lemma map_filter_simps [code]: "map_filter f (x # xs) = (case f x of None \ map_filter f xs | Some y \ y # map_filter f xs)" "map_filter f [] = []" by (simp_all add: map_filter_def split: option.split) lemma concat_map_maps: (* FIXME delete candidate *) "concat (map f xs) = maps f xs" by (simp add: maps_def) lemma map_filter_map_filter [code_unfold]: "map f (filter P xs) = map_filter (\x. if P x then Some (f x) else None) xs" by (simp add: map_filter_def) text \Optimized code for \\i\{a..b::int}\ and \\n:{a.. and similiarly for \\\.\ definition all_interval_nat :: "(nat \ bool) \ nat \ nat \ bool" where "all_interval_nat P i j \ (\n \ {i.. i \ j \ P i \ all_interval_nat P (Suc i) j" proof - have *: "\n. P i \ \n\{Suc i.. i \ n \ n < j \ P n" proof - fix n assume "P i" "\n\{Suc i.. n" "n < j" then show "P n" by (cases "n = i") simp_all qed show ?thesis by (auto simp add: all_interval_nat_def intro: *) qed lemma list_all_iff_all_interval_nat [code_unfold]: "list_all P [i.. all_interval_nat P i j" by (simp add: list_all_iff all_interval_nat_def) lemma list_ex_iff_not_all_inverval_nat [code_unfold]: "list_ex P [i.. \ (all_interval_nat (Not \ P) i j)" by (simp add: list_ex_iff all_interval_nat_def) definition all_interval_int :: "(int \ bool) \ int \ int \ bool" where "all_interval_int P i j \ (\k \ {i..j}. P k)" lemma [code]: "all_interval_int P i j \ i > j \ P i \ all_interval_int P (i + 1) j" proof - have *: "\k. P i \ \k\{i+1..j}. P k \ i \ k \ k \ j \ P k" proof - fix k assume "P i" "\k\{i+1..j}. P k" "i \ k" "k \ j" then show "P k" by (cases "k = i") simp_all qed show ?thesis by (auto simp add: all_interval_int_def intro: *) qed lemma list_all_iff_all_interval_int [code_unfold]: "list_all P [i..j] \ all_interval_int P i j" by (simp add: list_all_iff all_interval_int_def) lemma list_ex_iff_not_all_inverval_int [code_unfold]: "list_ex P [i..j] \ \ (all_interval_int (Not \ P) i j)" by (simp add: list_ex_iff all_interval_int_def) text \optimized code (tail-recursive) for \<^term>\length\\ definition gen_length :: "nat \ 'a list \ nat" where "gen_length n xs = n + length xs" lemma gen_length_code [code]: "gen_length n [] = n" "gen_length n (x # xs) = gen_length (Suc n) xs" by(simp_all add: gen_length_def) declare list.size(3-4)[code del] lemma length_code [code]: "length = gen_length 0" by(simp add: gen_length_def fun_eq_iff) hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length subsubsection \Pretty lists\ ML \ (* Code generation for list literals. *) signature LIST_CODE = sig val add_literal_list: string -> theory -> theory end; structure List_Code : LIST_CODE = struct open Basic_Code_Thingol; fun implode_list t = let fun dest_cons (IConst { sym = Code_Symbol.Constant \<^const_name>\Cons\, ... } `$ t1 `$ t2) = SOME (t1, t2) | dest_cons _ = NONE; val (ts, t') = Code_Thingol.unfoldr dest_cons t; in case t' of IConst { sym = Code_Symbol.Constant \<^const_name>\Nil\, ... } => SOME ts | _ => NONE end; fun print_list (target_fxy, target_cons) pr fxy t1 t2 = Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy ( pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1, Code_Printer.str target_cons, pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2 ); fun add_literal_list target = let fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] = case Option.map (cons t1) (implode_list t2) of SOME ts => Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts) | NONE => print_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2; in Code_Target.set_printings (Code_Symbol.Constant (\<^const_name>\Cons\, [(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))])) end end; \ code_printing type_constructor list \ (SML) "_ list" and (OCaml) "_ list" and (Haskell) "![(_)]" and (Scala) "List[(_)]" | constant Nil \ (SML) "[]" and (OCaml) "[]" and (Haskell) "[]" and (Scala) "!Nil" | class_instance list :: equal \ (Haskell) - | constant "HOL.equal :: 'a list \ 'a list \ bool" \ (Haskell) infix 4 "==" setup \fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"]\ code_reserved SML list code_reserved OCaml list subsubsection \Use convenient predefined operations\ code_printing constant "(@)" \ (SML) infixr 7 "@" and (OCaml) infixr 6 "@" and (Haskell) infixr 5 "++" and (Scala) infixl 7 "++" | constant map \ (Haskell) "map" | constant filter \ (Haskell) "filter" | constant concat \ (Haskell) "concat" | constant List.maps \ (Haskell) "concatMap" | constant rev \ (Haskell) "reverse" | constant zip \ (Haskell) "zip" | constant List.null \ (Haskell) "null" | constant takeWhile \ (Haskell) "takeWhile" | constant dropWhile \ (Haskell) "dropWhile" | constant list_all \ (Haskell) "all" | constant list_ex \ (Haskell) "any" subsubsection \Implementation of sets by lists\ lemma is_empty_set [code]: "Set.is_empty (set xs) \ List.null xs" by (simp add: Set.is_empty_def null_def) lemma empty_set [code]: "{} = set []" by simp lemma UNIV_coset [code]: "UNIV = List.coset []" by simp lemma compl_set [code]: "- set xs = List.coset xs" by simp lemma compl_coset [code]: "- List.coset xs = set xs" by simp lemma [code]: "x \ set xs \ List.member xs x" "x \ List.coset xs \ \ List.member xs x" by (simp_all add: member_def) lemma insert_code [code]: "insert x (set xs) = set (List.insert x xs)" "insert x (List.coset xs) = List.coset (removeAll x xs)" by simp_all lemma remove_code [code]: "Set.remove x (set xs) = set (removeAll x xs)" "Set.remove x (List.coset xs) = List.coset (List.insert x xs)" by (simp_all add: remove_def Compl_insert) lemma filter_set [code]: "Set.filter P (set xs) = set (filter P xs)" by auto lemma image_set [code]: "image f (set xs) = set (map f xs)" by simp lemma subset_code [code]: "set xs \ B \ (\x\set xs. x \ B)" "A \ List.coset ys \ (\y\set ys. y \ A)" "List.coset [] \ set [] \ False" by auto text \A frequent case -- avoid intermediate sets\ lemma [code_unfold]: "set xs \ set ys \ list_all (\x. x \ set ys) xs" by (auto simp: list_all_iff) lemma Ball_set [code]: "Ball (set xs) P \ list_all P xs" by (simp add: list_all_iff) lemma Bex_set [code]: "Bex (set xs) P \ list_ex P xs" by (simp add: list_ex_iff) lemma card_set [code]: "card (set xs) = length (remdups xs)" proof - have "card (set (remdups xs)) = length (remdups xs)" by (rule distinct_card) simp then show ?thesis by simp qed lemma the_elem_set [code]: "the_elem (set [x]) = x" by simp lemma Pow_set [code]: "Pow (set []) = {{}}" "Pow (set (x # xs)) = (let A = Pow (set xs) in A \ insert x ` A)" by (simp_all add: Pow_insert Let_def) definition map_project :: "('a \ 'b option) \ 'a set \ 'b set" where "map_project f A = {b. \ a \ A. f a = Some b}" lemma [code]: "map_project f (set xs) = set (List.map_filter f xs)" by (auto simp add: map_project_def map_filter_def image_def) hide_const (open) map_project text \Operations on relations\ lemma product_code [code]: "Product_Type.product (set xs) (set ys) = set [(x, y). x \ xs, y \ ys]" by (auto simp add: Product_Type.product_def) lemma Id_on_set [code]: "Id_on (set xs) = set [(x, x). x \ xs]" by (auto simp add: Id_on_def) lemma [code]: "R `` S = List.map_project (\(x, y). if x \ S then Some y else None) R" unfolding map_project_def by (auto split: prod.split if_split_asm) lemma trancl_set_ntrancl [code]: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)" by (simp add: finite_trancl_ntranl) lemma set_relcomp [code]: "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" by auto (auto simp add: Bex_def image_def) lemma wf_set [code]: "wf (set xs) = acyclic (set xs)" by (simp add: wf_iff_acyclic_if_finite) subsection \Setup for Lifting/Transfer\ subsubsection \Transfer rules for the Transfer package\ context includes lifting_syntax begin lemma tl_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) tl tl" unfolding tl_def[abs_def] by transfer_prover lemma butlast_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) butlast butlast" by (rule rel_funI, erule list_all2_induct, auto) lemma map_rec: "map f xs = rec_list Nil (%x _ y. Cons (f x) y) xs" by (induct xs) auto lemma append_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> list_all2 A) append append" unfolding List.append_def by transfer_prover lemma rev_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) rev rev" unfolding List.rev_def by transfer_prover lemma filter_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) filter filter" unfolding List.filter_def by transfer_prover lemma fold_transfer [transfer_rule]: "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) fold fold" unfolding List.fold_def by transfer_prover lemma foldr_transfer [transfer_rule]: "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) foldr foldr" unfolding List.foldr_def by transfer_prover lemma foldl_transfer [transfer_rule]: "((B ===> A ===> B) ===> B ===> list_all2 A ===> B) foldl foldl" unfolding List.foldl_def by transfer_prover lemma concat_transfer [transfer_rule]: "(list_all2 (list_all2 A) ===> list_all2 A) concat concat" unfolding List.concat_def by transfer_prover lemma drop_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) drop drop" unfolding List.drop_def by transfer_prover lemma take_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) take take" unfolding List.take_def by transfer_prover lemma list_update_transfer [transfer_rule]: "(list_all2 A ===> (=) ===> A ===> list_all2 A) list_update list_update" unfolding list_update_def by transfer_prover lemma takeWhile_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) takeWhile takeWhile" unfolding takeWhile_def by transfer_prover lemma dropWhile_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> list_all2 A) dropWhile dropWhile" unfolding dropWhile_def by transfer_prover lemma zip_transfer [transfer_rule]: "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) zip zip" unfolding zip_def by transfer_prover lemma product_transfer [transfer_rule]: "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) List.product List.product" unfolding List.product_def by transfer_prover lemma product_lists_transfer [transfer_rule]: "(list_all2 (list_all2 A) ===> list_all2 (list_all2 A)) product_lists product_lists" unfolding product_lists_def by transfer_prover lemma insert_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) List.insert List.insert" unfolding List.insert_def [abs_def] by transfer_prover lemma find_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> rel_option A) List.find List.find" unfolding List.find_def by transfer_prover lemma those_transfer [transfer_rule]: "(list_all2 (rel_option P) ===> rel_option (list_all2 P)) those those" unfolding List.those_def by transfer_prover lemma remove1_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) remove1 remove1" unfolding remove1_def by transfer_prover lemma removeAll_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(A ===> list_all2 A ===> list_all2 A) removeAll removeAll" unfolding removeAll_def by transfer_prover lemma distinct_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> (=)) distinct distinct" unfolding distinct_def by transfer_prover lemma remdups_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> list_all2 A) remdups remdups" unfolding remdups_def by transfer_prover lemma remdups_adj_transfer [transfer_rule]: assumes [transfer_rule]: "bi_unique A" shows "(list_all2 A ===> list_all2 A) remdups_adj remdups_adj" proof (rule rel_funI, erule list_all2_induct) qed (auto simp: remdups_adj_Cons assms[unfolded bi_unique_def] split: list.splits) lemma replicate_transfer [transfer_rule]: "((=) ===> A ===> list_all2 A) replicate replicate" unfolding replicate_def by transfer_prover lemma length_transfer [transfer_rule]: "(list_all2 A ===> (=)) length length" unfolding size_list_overloaded_def size_list_def by transfer_prover lemma rotate1_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A) rotate1 rotate1" unfolding rotate1_def by transfer_prover lemma rotate_transfer [transfer_rule]: "((=) ===> list_all2 A ===> list_all2 A) rotate rotate" unfolding rotate_def [abs_def] by transfer_prover lemma nths_transfer [transfer_rule]: "(list_all2 A ===> rel_set (=) ===> list_all2 A) nths nths" unfolding nths_def [abs_def] by transfer_prover lemma subseqs_transfer [transfer_rule]: "(list_all2 A ===> list_all2 (list_all2 A)) subseqs subseqs" unfolding subseqs_def [abs_def] by transfer_prover lemma partition_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> rel_prod (list_all2 A) (list_all2 A)) partition partition" unfolding partition_def by transfer_prover lemma lists_transfer [transfer_rule]: "(rel_set A ===> rel_set (list_all2 A)) lists lists" apply (rule rel_funI, rule rel_setI) apply (erule lists.induct, simp) apply (simp only: rel_set_def list_all2_Cons1, metis lists.Cons) apply (erule lists.induct, simp) apply (simp only: rel_set_def list_all2_Cons2, metis lists.Cons) done lemma set_Cons_transfer [transfer_rule]: "(rel_set A ===> rel_set (list_all2 A) ===> rel_set (list_all2 A)) set_Cons set_Cons" unfolding rel_fun_def rel_set_def set_Cons_def by (fastforce simp add: list_all2_Cons1 list_all2_Cons2) lemma listset_transfer [transfer_rule]: "(list_all2 (rel_set A) ===> rel_set (list_all2 A)) listset listset" unfolding listset_def by transfer_prover lemma null_transfer [transfer_rule]: "(list_all2 A ===> (=)) List.null List.null" unfolding rel_fun_def List.null_def by auto lemma list_all_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> (=)) list_all list_all" unfolding list_all_iff [abs_def] by transfer_prover lemma list_ex_transfer [transfer_rule]: "((A ===> (=)) ===> list_all2 A ===> (=)) list_ex list_ex" unfolding list_ex_iff [abs_def] by transfer_prover lemma splice_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> list_all2 A) splice splice" apply (rule rel_funI, erule list_all2_induct, simp add: rel_fun_def, simp) apply (rule rel_funI) apply (erule_tac xs=x in list_all2_induct, simp, simp add: rel_fun_def) done lemma shuffles_transfer [transfer_rule]: "(list_all2 A ===> list_all2 A ===> rel_set (list_all2 A)) shuffles shuffles" proof (intro rel_funI, goal_cases) case (1 xs xs' ys ys') thus ?case proof (induction xs ys arbitrary: xs' ys' rule: shuffles.induct) case (3 x xs y ys xs' ys') from "3.prems" obtain x' xs'' where xs': "xs' = x' # xs''" by (cases xs') auto from "3.prems" obtain y' ys'' where ys': "ys' = y' # ys''" by (cases ys') auto have [transfer_rule]: "A x x'" "A y y'" "list_all2 A xs xs''" "list_all2 A ys ys''" using "3.prems" by (simp_all add: xs' ys') have [transfer_rule]: "rel_set (list_all2 A) (shuffles xs (y # ys)) (shuffles xs'' ys')" and [transfer_rule]: "rel_set (list_all2 A) (shuffles (x # xs) ys) (shuffles xs' ys'')" using "3.prems" by (auto intro!: "3.IH" simp: xs' ys') have "rel_set (list_all2 A) ((#) x ` shuffles xs (y # ys) \ (#) y ` shuffles (x # xs) ys) ((#) x' ` shuffles xs'' ys' \ (#) y' ` shuffles xs' ys'')" by transfer_prover thus ?case by (simp add: xs' ys') qed (auto simp: rel_set_def) qed lemma rtrancl_parametric [transfer_rule]: assumes [transfer_rule]: "bi_unique A" "bi_total A" shows "(rel_set (rel_prod A A) ===> rel_set (rel_prod A A)) rtrancl rtrancl" unfolding rtrancl_def by transfer_prover lemma monotone_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone" unfolding monotone_def[abs_def] by transfer_prover lemma fun_ord_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total C" shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord" unfolding fun_ord_def[abs_def] by transfer_prover lemma fun_lub_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" "bi_unique A" shows "((rel_set A ===> B) ===> rel_set (C ===> A) ===> C ===> B) fun_lub fun_lub" unfolding fun_lub_def[abs_def] by transfer_prover end end diff --git a/src/HOL/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2533 +1,2533 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: - assumes "P 0" - and "\n. P n \ P (Suc n)" + assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" - using assms - apply (unfold Zero_nat_def Suc_def) - apply (rule Rep_Nat_inverse [THEN subst]) \ \types force good instantiation\ - apply (erule Nat_Rep_Nat [THEN Nat.induct]) - apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst]) - done +proof - + have "P (Abs_Nat (Rep_Nat n))" + using assms unfolding Zero_nat_def Suc_def + by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) + then show ?thesis + by (simp add: Rep_Nat_inverse) +qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" apply atomize_elim apply (rename_tac n, induct_tac n rule: nat_induct0, auto) apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) apply (simp only: Suc_not_Zero) done \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc - apply (erule nat_induct0) - apply assumption - apply (rule nat.inject) - apply (rule nat.distinct(1)) - done + by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" - apply (rule trans) - apply (rule_tac [2] mult_eq_1_iff) - apply fastforce - done - -lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" - for m n :: nat - unfolding One_nat_def by (rule mult_eq_1_iff) - -lemma nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" - for m n :: nat - unfolding One_nat_def by (rule one_eq_mult_iff) + by (simp add: eq_commute flip: mult_eq_1_iff) + +lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" + and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat + by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P - apply (rule major [THEN lessE]) - apply (rule eq) - apply blast - apply (rule less) - apply blast - done +proof (rule major [THEN lessE]) + show "Suc n = Suc m \ P" + using eq by blast + show "\j. \m < j; Suc n = Suc j\ \ P" + by (blast intro: less) +qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P - apply (rule major [THEN lessE]) - apply (erule lessI [THEN minor]) - apply (erule Suc_lessD [THEN minor]) - apply assumption - done +proof (rule major [THEN lessE]) + show "k = Suc (Suc i) \ P" + using lessI minor by iprover + show "\j. \Suc i < j; k = Suc j\ \ P" + using Suc_lessD minor by iprover +qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ -lemma add_less_mono: "i < j \ k < l \ i + k < j + l" - for i j k l :: nat - apply (rule add_less_mono1 [THEN less_trans], assumption+) - apply (induct j) - apply simp_all - done +lemma add_less_mono: + fixes i j k l :: nat + assumes "i < j" "k < l" shows "i + k < j + l" +proof - + have "i + k < j + k" + by (simp add: add_less_mono1 assms) + also have "... < j + l" + using \i < j\ by (induction j) (auto simp: assms) + finally show ?thesis . +qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat by (rule max_absorb2) simp lemma max_0R [simp]: "max n 0 = n" for n :: nat by (rule max_absorb1) simp lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) -lemma Least_Suc: "P n \ \ P 0 \ (LEAST n. P n) = Suc (LEAST m. P (Suc m))" - apply (cases n) - apply auto - apply (frule LeastI) - apply (drule_tac P = "\x. P (Suc x)" in LeastI) - apply (subgoal_tac " (LEAST x. P x) \ Suc (LEAST x. P (Suc x))") - apply (erule_tac [2] Least_le) - apply (cases "LEAST x. P x") - apply auto - apply (drule_tac P = "\x. P (Suc x)" in Least_le) - apply (blast intro: order_antisym) - done +lemma Least_Suc: + assumes "P n" "\ P 0" + shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" +proof (cases n) + case (Suc m) + show ?thesis + proof (rule antisym) + show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" + using assms Suc by (force intro: LeastI Least_le) + have \
: "P (LEAST x. P x)" + by (blast intro: LeastI assms) + show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" + proof (cases "(LEAST n. P n)") + case 0 + then show ?thesis + using \
by (simp add: assms) + next + case Suc + with \
show ?thesis + by (auto simp: Least_le) + qed + qed +qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp -lemma ex_least_nat_le: "\ P 0 \ P n \ \k\n. (\i P i) \ P k" - for P :: "nat \ bool" - apply (cases n) - apply blast - apply (rule_tac x="LEAST k. P k" in exI) - apply (blast intro: Least_le dest: not_less_Least intro: LeastI_ex) - done - -lemma ex_least_nat_less: "\ P 0 \ P n \ \ki\k. \ P i) \ P (k + 1)" - for P :: "nat \ bool" - apply (cases n) - apply blast - apply (frule (1) ex_least_nat_le) - apply (erule exE) - apply (case_tac k) - apply simp - apply (rename_tac k1) - apply (rule_tac x=k1 in exI) - apply (auto simp add: less_eq_Suc_le) - done +lemma ex_least_nat_le: + fixes P :: "nat \ bool" + assumes "P n" "\ P 0" + shows "\k\n. (\i P i) \ P k" +proof (cases n) + case (Suc m) + with assms show ?thesis + by (blast intro: Least_le LeastI_ex dest: not_less_Least) +qed (use assms in auto) + +lemma ex_least_nat_less: + fixes P :: "nat \ bool" + assumes "P n" "\ P 0" + shows "\ki\k. \ P i) \ P (Suc k)" +proof (cases n) + case (Suc m) + then obtain k where k: "k \ n" "\i P i" "P k" + using ex_least_nat_le [OF assms] by blast + show ?thesis + by (cases k) (use assms k less_eq_Suc_le in auto) +qed (use assms in auto) + lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" - apply (rule infinite_descent) - using assms - apply (case_tac "n > 0") - apply auto - done +proof (rule infinite_descent) + show "\n. \ P n \ \m P m" + using assms by (case_tac "n > 0") auto +qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent, auto) show "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat - apply (rule notI) - apply (drule add_lessD1) - apply (erule less_irrefl [THEN notE]) - done + by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat - by (simp add: add.commute) + by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat - apply (simp add: add.commute) - apply (erule add_leD1) - done + by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" - apply (induct m n rule: diff_induct) - apply (erule_tac [3] less_SucE) - apply (simp_all add: less_Suc_eq) - done + by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat - apply (safe intro!: mult_less_mono1) - apply (cases k) - apply auto - apply (simp add: linorder_not_le [symmetric]) - apply (blast intro: mult_le_mono1) - done +proof (intro iffI conjI) + assume m: "m * k < n * k" + then show "0 < k" + by (cases k) auto + show "m < n" + proof (cases k) + case 0 + then show ?thesis + using m by auto + next + case (Suc k') + then show ?thesis + using m + by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) + qed +next + assume "0 < k \ m < n" + then show "m * k < n * k" + by (blast intro: mult_less_mono1) +qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ -lemma mult_eq_self_implies_10: "m = m * n \ n = 1 \ m = 0" - for m n :: nat - apply (drule sym) - apply (rule disjCI) - apply (rule linorder_cases) - defer - apply assumption - apply (drule mult_less_mono2) - apply auto - done +lemma mult_eq_self_implies_10: + fixes m n :: nat + assumes "m = m * n" shows "n = 1 \ m = 0" +proof (rule disjCI) + assume "m \ 0" + show "n = 1" + proof (cases n "1::nat" rule: linorder_cases) + case greater + show ?thesis + using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto + qed (use assms \m \ 0\ in auto) +qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) hide_const (open) funpow lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case by (simp add: add.commute) qed lemma of_nat_of_bool [simp]: "of_nat (of_bool P) = of_bool P" by auto end declare of_nat_code [code] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) lemma sgn_of_nat [simp]: "sgn (of_nat n) = of_bool (n > 0)" by simp end lemma of_nat_id [simp]: "of_nat n = n" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" - apply (simp add: Nats_def) - apply (rule range_eqI) - apply (rule of_nat_0 [symmetric]) - done + using of_nat_0 [symmetric] unfolding Nats_def + by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" - apply (simp add: Nats_def) - apply (rule range_eqI) - apply (rule of_nat_1 [symmetric]) - done + using of_nat_1 [symmetric] unfolding Nats_def + by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" - apply (auto simp add: Nats_def) - apply (rule range_eqI) - apply (rule of_nat_add [symmetric]) - done + unfolding Nats_def using of_nat_add [symmetric] + by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" - apply (auto simp add: Nats_def) - apply (rule range_eqI) - apply (rule of_nat_mult [symmetric]) - done + unfolding Nats_def using of_nat_mult [symmetric] + by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = \fn phi => try o Nat_Arith.cancel_eq_conv\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = \fn phi => try o Nat_Arith.cancel_less_conv\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = \fn phi => try o Nat_Arith.cancel_le_conv\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = \fn phi => try o Nat_Arith.cancel_diff_conv\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed -lemma GreatestI_nat: - "\ P(k::nat); \y. P y \ y \ b \ \ P (Greatest P)" -apply(drule (1) ex_has_greatest_nat) -using GreatestI2_order by auto - -lemma Greatest_le_nat: - "\ P(k::nat); \y. P y \ y \ b \ \ k \ (Greatest P)" -apply(frule (1) ex_has_greatest_nat) -using GreatestI2_order[where P=P and Q=\\x. k \ x\] by auto +lemma + fixes k::nat + assumes "P k" and minor: "\y. P y \ y \ b" + shows GreatestI_nat: "P (Greatest P)" + and Greatest_le_nat: "k \ Greatest P" +proof - + obtain x where "P x" "\y. P y \ y \ x" + using assms ex_has_greatest_nat by blast + with \P k\ show "P (Greatest P)" "k \ Greatest P" + using GreatestI2_order by blast+ +qed lemma GreatestI_ex_nat: - "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" -apply (erule exE) -apply (erule (1) GreatestI_nat) -done + "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" + by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) -lemma dvd_diffD: "k dvd m - n \ k dvd n \ n \ m \ k dvd m" - for k m n :: nat - apply (erule linorder_not_less [THEN iffD2, THEN add_diff_inverse, THEN subst]) - apply (blast intro: dvd_add) - done +lemma dvd_diffD: + fixes k m n :: nat + assumes "k dvd m - n" "k dvd n" "n \ m" + shows "k dvd m" +proof - + have "k dvd n + (m - n)" + using assms by (blast intro: dvd_add) + with assms show ?thesis + by simp +qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed -lemma dvd_mult_cancel1: "0 < m \ m * n dvd m \ n = 1" - for m n :: nat - apply auto - apply (subgoal_tac "m * n dvd m * 1") - apply (drule dvd_mult_cancel) - apply auto - done +lemma dvd_mult_cancel1: + fixes m n :: nat + assumes "0 < m" + shows "m * n dvd m \ n = 1" +proof + assume "m * n dvd m" + then have "m * n dvd m * 1" + by simp + then have "n dvd 1" + by (iprover intro: assms dvd_mult_cancel) + then show "n = 1" + by auto +qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) - -lemma nat_add_left_cancel: "k + m = k + n \ m = n" - for k m n :: nat - by (fact add_left_cancel) - -lemma nat_add_right_cancel: "m + k = n + k \ m = n" - for k m n :: nat - by (fact add_right_cancel) - lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') -lemma le_add_diff: "k \ n \ m \ n + m - k" - for k m n :: nat - by (fact le_add_diff) (* FIXME delete *) - -lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" - for i j k :: nat - by (fact le_diff_conv2) (* FIXME delete *) - lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Transcendental.thy b/src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy +++ b/src/HOL/Transcendental.thy @@ -1,7354 +1,7355 @@ (* Title: HOL/Transcendental.thy Author: Jacques D. Fleuriot, University of Cambridge, University of Edinburgh Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Power Series, Transcendental Functions etc.\ theory Transcendental imports Series Deriv NthRoot begin text \A theorem about the factcorial function on the reals.\ lemma square_fact_le_2_fact: "fact n * fact n \ (fact (2 * n) :: real)" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(fact (Suc n)) * (fact (Suc n)) = of_nat (Suc n) * of_nat (Suc n) * (fact n * fact n :: real)" by (simp add: field_simps) also have "\ \ of_nat (Suc n) * of_nat (Suc n) * fact (2 * n)" by (rule mult_left_mono [OF Suc]) simp also have "\ \ of_nat (Suc (Suc (2 * n))) * of_nat (Suc (2 * n)) * fact (2 * n)" by (rule mult_right_mono)+ (auto simp: field_simps) also have "\ = fact (2 * Suc n)" by (simp add: field_simps) finally show ?case . qed lemma fact_in_Reals: "fact n \ \" by (induction n) auto lemma of_real_fact [simp]: "of_real (fact n) = fact n" by (metis of_nat_fact of_real_of_nat_eq) lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_prod) lemma norm_fact [simp]: "norm (fact n :: 'a::real_normed_algebra_1) = fact n" proof - have "(fact n :: 'a) = of_real (fact n)" by simp also have "norm \ = fact n" by (subst norm_of_real) simp finally show ?thesis . qed lemma root_test_convergence: fixes f :: "nat \ 'a::banach" assumes f: "(\n. root n (norm (f n))) \ x" \ \could be weakened to lim sup\ and "x < 1" shows "summable f" proof - have "0 \ x" by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1]) from \x < 1\ obtain z where z: "x < z" "z < 1" by (metis dense) from f \x < z\ have "eventually (\n. root n (norm (f n)) < z) sequentially" by (rule order_tendstoD) then have "eventually (\n. norm (f n) \ z^n) sequentially" using eventually_ge_at_top proof eventually_elim fix n assume less: "root n (norm (f n)) < z" and n: "1 \ n" from power_strict_mono[OF less, of n] n show "norm (f n) \ z ^ n" by simp qed then show "summable f" unfolding eventually_sequentially using z \0 \ x\ by (auto intro!: summable_comparison_test[OF _ summable_geometric]) qed subsection \More facts about binomial coefficients\ text \ These facts could have been proven before, but having real numbers makes the proofs a lot easier. \ lemma central_binomial_odd: "odd n \ n choose (Suc (n div 2)) = n choose (n div 2)" proof - assume "odd n" hence "Suc (n div 2) \ n" by presburger hence "n choose (Suc (n div 2)) = n choose (n - Suc (n div 2))" by (rule binomial_symmetric) also from \odd n\ have "n - Suc (n div 2) = n div 2" by presburger finally show ?thesis . qed lemma binomial_less_binomial_Suc: assumes k: "k < n div 2" shows "n choose k < n choose (Suc k)" proof - from k have k': "k \ n" "Suc k \ n" by simp_all from k' have "real (n choose k) = fact n / (fact k * fact (n - k))" by (simp add: binomial_fact) also from k' have "n - k = Suc (n - Suc k)" by simp also from k' have "fact \ = (real n - real k) * fact (n - Suc k)" by (subst fact_Suc) (simp_all add: of_nat_diff) also from k have "fact k = fact (Suc k) / (real k + 1)" by (simp add: field_simps) also have "fact n / (fact (Suc k) / (real k + 1) * ((real n - real k) * fact (n - Suc k))) = (n choose (Suc k)) * ((real k + 1) / (real n - real k))" using k by (simp add: field_split_simps binomial_fact) also from assms have "(real k + 1) / (real n - real k) < 1" by simp finally show ?thesis using k by (simp add: mult_less_cancel_left) qed lemma binomial_strict_mono: assumes "k < k'" "2*k' \ n" shows "n choose k < n choose k'" proof - from assms have "k \ k' - 1" by simp thus ?thesis proof (induction rule: inc_induct) case base with assms binomial_less_binomial_Suc[of "k' - 1" n] show ?case by simp next case (step k) from step.prems step.hyps assms have "n choose k < n choose (Suc k)" by (intro binomial_less_binomial_Suc) simp_all also have "\ < n choose k'" by (rule step.IH) finally show ?case . qed qed lemma binomial_mono: assumes "k \ k'" "2*k' \ n" shows "n choose k \ n choose k'" using assms binomial_strict_mono[of k k' n] by (cases "k = k'") simp_all lemma binomial_strict_antimono: assumes "k < k'" "2 * k \ n" "k' \ n" shows "n choose k > n choose k'" proof - from assms have "n choose (n - k) > n choose (n - k')" by (intro binomial_strict_mono) (simp_all add: algebra_simps) with assms show ?thesis by (simp add: binomial_symmetric [symmetric]) qed lemma binomial_antimono: assumes "k \ k'" "k \ n div 2" "k' \ n" shows "n choose k \ n choose k'" proof (cases "k = k'") case False note not_eq = False show ?thesis proof (cases "k = n div 2 \ odd n") case False with assms(2) have "2*k \ n" by presburger with not_eq assms binomial_strict_antimono[of k k' n] show ?thesis by simp next case True have "n choose k' \ n choose (Suc (n div 2))" proof (cases "k' = Suc (n div 2)") case False with assms True not_eq have "Suc (n div 2) < k'" by simp with assms binomial_strict_antimono[of "Suc (n div 2)" k' n] True show ?thesis by auto qed simp_all also from True have "\ = n choose k" by (simp add: central_binomial_odd) finally show ?thesis . qed qed simp_all lemma binomial_maximum: "n choose k \ n choose (n div 2)" proof - have "k \ n div 2 \ 2*k \ n" by linarith consider "2*k \ n" | "2*k \ n" "k \ n" | "k > n" by linarith thus ?thesis proof cases case 1 thus ?thesis by (intro binomial_mono) linarith+ next case 2 thus ?thesis by (intro binomial_antimono) simp_all qed (simp_all add: binomial_eq_0) qed lemma binomial_maximum': "(2*n) choose k \ (2*n) choose n" using binomial_maximum[of "2*n"] by simp lemma central_binomial_lower_bound: assumes "n > 0" shows "4^n / (2*real n) \ real ((2*n) choose n)" proof - from binomial[of 1 1 "2*n"] have "4 ^ n = (\k\2*n. (2*n) choose k)" by (simp add: power_mult power2_eq_square One_nat_def [symmetric] del: One_nat_def) also have "{..2*n} = {0<..<2*n} \ {0,2*n}" by auto also have "(\k\\. (2*n) choose k) = (\k\{0<..<2*n}. (2*n) choose k) + (\k\{0,2*n}. (2*n) choose k)" by (subst sum.union_disjoint) auto also have "(\k\{0,2*n}. (2*n) choose k) \ (\k\1. (n choose k)\<^sup>2)" by (cases n) simp_all also from assms have "\ \ (\k\n. (n choose k)\<^sup>2)" by (intro sum_mono2) auto also have "\ = (2*n) choose n" by (rule choose_square_sum) also have "(\k\{0<..<2*n}. (2*n) choose k) \ (\k\{0<..<2*n}. (2*n) choose n)" by (intro sum_mono binomial_maximum') also have "\ = card {0<..<2*n} * ((2*n) choose n)" by simp also have "card {0<..<2*n} \ 2*n - 1" by (cases n) simp_all also have "(2 * n - 1) * (2 * n choose n) + (2 * n choose n) = ((2*n) choose n) * (2*n)" using assms by (simp add: algebra_simps) finally have "4 ^ n \ (2 * n choose n) * (2 * n)" by simp_all hence "real (4 ^ n) \ real ((2 * n choose n) * (2 * n))" by (subst of_nat_le_iff) with assms show ?thesis by (simp add: field_simps) qed subsection \Properties of Power Series\ lemma powser_zero [simp]: "(\n. f n * 0 ^ n) = f 0" for f :: "nat \ 'a::real_normed_algebra_1" proof - have "(\n<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" by (subst suminf_finite[where N="{0}"]) (auto simp: power_0_left) then show ?thesis by simp qed lemma powser_sums_zero: "(\n. a n * 0^n) sums a 0" for a :: "nat \ 'a::real_normed_div_algebra" using sums_finite [of "{0}" "\n. a n * 0 ^ n"] by simp lemma powser_sums_zero_iff [simp]: "(\n. a n * 0^n) sums x \ a 0 = x" for a :: "nat \ 'a::real_normed_div_algebra" using powser_sums_zero sums_unique2 by blast text \ Power series has a circle or radius of convergence: if it sums for \x\, then it sums absolutely for \z\ with \<^term>\\z\ < \x\\.\ lemma powser_insidea: fixes x z :: "'a::real_normed_div_algebra" assumes 1: "summable (\n. f n * x^n)" and 2: "norm z < norm x" shows "summable (\n. norm (f n * z ^ n))" proof - from 2 have x_neq_0: "x \ 0" by clarsimp from 1 have "(\n. f n * x^n) \ 0" by (rule summable_LIMSEQ_zero) then have "convergent (\n. f n * x^n)" by (rule convergentI) then have "Cauchy (\n. f n * x^n)" by (rule convergent_Cauchy) then have "Bseq (\n. f n * x^n)" by (rule Cauchy_Bseq) then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x^n) \ K" by (auto simp: Bseq_def) have "\N. \n\N. norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" proof (intro exI allI impI) fix n :: nat assume "0 \ n" have "norm (norm (f n * z ^ n)) * norm (x^n) = norm (f n * x^n) * norm (z ^ n)" by (simp add: norm_mult abs_mult) also have "\ \ K * norm (z ^ n)" by (simp only: mult_right_mono 4 norm_ge_zero) also have "\ = K * norm (z ^ n) * (inverse (norm (x^n)) * norm (x^n))" by (simp add: x_neq_0) also have "\ = K * norm (z ^ n) * inverse (norm (x^n)) * norm (x^n)" by (simp only: mult.assoc) finally show "norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" by (simp add: mult_le_cancel_right x_neq_0) qed moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" proof - from 2 have "norm (norm (z * inverse x)) < 1" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse divide_inverse [where 'a=real, symmetric]) then have "summable (\n. norm (z * inverse x) ^ n)" by (rule summable_geometric) then have "summable (\n. K * norm (z * inverse x) ^ n)" by (rule summable_mult) then show "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib power_inverse norm_power mult.assoc) qed ultimately show "summable (\n. norm (f n * z ^ n))" by (rule summable_comparison_test) qed lemma powser_inside: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" shows "summable (\n. f n * (x^n)) \ norm z < norm x \ summable (\n. f n * (z ^ n))" by (rule powser_insidea [THEN summable_norm_cancel]) lemma powser_times_n_limit_0: fixes x :: "'a::{real_normed_div_algebra,banach}" assumes "norm x < 1" shows "(\n. of_nat n * x ^ n) \ 0" proof - have "norm x / (1 - norm x) \ 0" using assms by (auto simp: field_split_simps) moreover obtain N where N: "norm x / (1 - norm x) < of_int N" using ex_le_of_int by (meson ex_less_of_int) ultimately have N0: "N>0" by auto then have *: "real_of_int (N + 1) * norm x / real_of_int N < 1" using N assms by (auto simp: field_simps) have **: "real_of_int N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) \ real_of_nat n * (norm x * ((1 + N) * norm (x ^ n)))" if "N \ int n" for n :: nat proof - from that have "real_of_int N * real_of_nat (Suc n) \ real_of_nat n * real_of_int (1 + N)" by (simp add: algebra_simps) then have "(real_of_int N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) \ (real_of_nat n * (1 + N)) * (norm x * norm (x ^ n))" using N0 mult_mono by fastforce then show ?thesis by (simp add: algebra_simps) qed show ?thesis using * by (rule summable_LIMSEQ_zero [OF summable_ratio_test, where N1="nat N"]) (simp add: N0 norm_mult field_simps ** del: of_nat_Suc of_int_add) qed corollary lim_n_over_pown: fixes x :: "'a::{real_normed_field,banach}" shows "1 < norm x \ ((\n. of_nat n / x^n) \ 0) sequentially" using powser_times_n_limit_0 [of "inverse x"] by (simp add: norm_divide field_split_simps) lemma sum_split_even_odd: fixes f :: "nat \ real" shows "(\i<2 * n. if even i then f i else g i) = (\iii<2 * Suc n. if even i then f i else g i) = (\ii = (\ii real" assumes "g sums x" shows "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" unfolding sums_def proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from \g sums x\[unfolded sums_def, THEN LIMSEQ_D, OF this] obtain no where no_eq: "\n. n \ no \ (norm (sum g {.. 2 * no" for m proof - from that have "m div 2 \ no" by auto have sum_eq: "?SUM (2 * (m div 2)) = sum g {..< m div 2}" using sum_split_even_odd by auto then have "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using \m div 2 \ no\ by auto moreover have "?SUM (2 * (m div 2)) = ?SUM m" proof (cases "even m") case True then show ?thesis by (auto simp: even_two_times_div_two) next case False then have eq: "Suc (2 * (m div 2)) = m" by simp then have "even (2 * (m div 2))" using \odd m\ by auto have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq .. also have "\ = ?SUM (2 * (m div 2))" using \even (2 * (m div 2))\ by auto finally show ?thesis by auto qed ultimately show ?thesis by auto qed then show "\no. \ m \ no. norm (?SUM m - x) < r" by blast qed lemma sums_if: fixes g :: "nat \ real" assumes "g sums x" and "f sums y" shows "(\ n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)" proof - let ?s = "\ n. if even n then 0 else f ((n - 1) div 2)" have if_sum: "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)" for B T E by (cases B) auto have g_sums: "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF \g sums x\] . have if_eq: "\B T E. (if \ B then T else E) = (if B then E else T)" by auto have "?s sums y" using sums_if'[OF \f sums y\] . from this[unfolded sums_def, THEN LIMSEQ_Suc] have "(\n. if even n then f (n div 2) else 0) sums y" by (simp add: lessThan_Suc_eq_insert_0 sum.atLeast1_atMost_eq image_Suc_lessThan if_eq sums_def cong del: if_weak_cong) from sums_add[OF g_sums this] show ?thesis by (simp only: if_sum) qed subsection \Alternating series test / Leibniz formula\ (* FIXME: generalise these results from the reals via type classes? *) lemma sums_alternating_upper_lower: fixes a :: "nat \ real" assumes mono: "\n. a (Suc n) \ a n" and a_pos: "\n. 0 \ a n" and "a \ 0" shows "\l. ((\n. (\i<2*n. (- 1)^i*a i) \ l) \ (\ n. \i<2*n. (- 1)^i*a i) \ l) \ ((\n. l \ (\i<2*n + 1. (- 1)^i*a i)) \ (\ n. \i<2*n + 1. (- 1)^i*a i) \ l)" (is "\l. ((\n. ?f n \ l) \ _) \ ((\n. l \ ?g n) \ _)") proof (rule nested_sequence_unique) have fg_diff: "\n. ?f n - ?g n = - a (2 * n)" by auto show "\n. ?f n \ ?f (Suc n)" proof show "?f n \ ?f (Suc n)" for n using mono[of "2*n"] by auto qed show "\n. ?g (Suc n) \ ?g n" proof show "?g (Suc n) \ ?g n" for n using mono[of "Suc (2*n)"] by auto qed show "\n. ?f n \ ?g n" proof show "?f n \ ?g n" for n using fg_diff a_pos by auto qed show "(\n. ?f n - ?g n) \ 0" unfolding fg_diff proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \a \ 0\[THEN LIMSEQ_D] obtain N where "\ n. n \ N \ norm (a n - 0) < r" by auto then have "\n \ N. norm (- a (2 * n) - 0) < r" by auto then show "\N. \n \ N. norm (- a (2 * n) - 0) < r" by auto qed qed lemma summable_Leibniz': fixes a :: "nat \ real" assumes a_zero: "a \ 0" and a_pos: "\n. 0 \ a n" and a_monotone: "\n. a (Suc n) \ a n" shows summable: "summable (\ n. (-1)^n * a n)" and "\n. (\i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "(\n. \i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "\n. (\i. (-1)^i*a i) \ (\i<2*n+1. (-1)^i*a i)" and "(\n. \i<2*n+1. (-1)^i*a i) \ (\i. (-1)^i*a i)" proof - let ?S = "\n. (-1)^n * a n" let ?P = "\n. \i n. ?f n \ l" and "?f \ l" and above_l: "\ n. l \ ?g n" and "?g \ l" using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast let ?Sa = "\m. \n l" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \?f \ l\[THEN LIMSEQ_D] obtain f_no where f: "\n. n \ f_no \ norm (?f n - l) < r" by auto from \0 < r\ \?g \ l\[THEN LIMSEQ_D] obtain g_no where g: "\n. n \ g_no \ norm (?g n - l) < r" by auto have "norm (?Sa n - l) < r" if "n \ (max (2 * f_no) (2 * g_no))" for n proof - from that have "n \ 2 * f_no" and "n \ 2 * g_no" by auto show ?thesis proof (cases "even n") case True then have n_eq: "2 * (n div 2) = n" by (simp add: even_two_times_div_two) with \n \ 2 * f_no\ have "n div 2 \ f_no" by auto from f[OF this] show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost . next case False then have "even (n - 1)" by simp then have n_eq: "2 * ((n - 1) div 2) = n - 1" by (simp add: even_two_times_div_two) then have range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto from n_eq \n \ 2 * g_no\ have "(n - 1) div 2 \ g_no" by auto from g[OF this] show ?thesis by (simp only: n_eq range_eq) qed qed then show "\no. \n \ no. norm (?Sa n - l) < r" by blast qed then have sums_l: "(\i. (-1)^i * a i) sums l" by (simp only: sums_def) then show "summable ?S" by (auto simp: summable_def) have "l = suminf ?S" by (rule sums_unique[OF sums_l]) fix n show "suminf ?S \ ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto show "?f n \ suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto show "?g \ suminf ?S" using \?g \ l\ \l = suminf ?S\ by auto show "?f \ suminf ?S" using \?f \ l\ \l = suminf ?S\ by auto qed theorem summable_Leibniz: fixes a :: "nat \ real" assumes a_zero: "a \ 0" and "monoseq a" shows "summable (\ n. (-1)^n * a n)" (is "?summable") and "0 < a 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n. (- 1)^i * a i .. \i<2*n+1. (- 1)^i * a i})" (is "?pos") and "a 0 < 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n+1. (- 1)^i * a i .. \i<2*n. (- 1)^i * a i})" (is "?neg") and "(\n. \i<2*n. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?f") and "(\n. \i<2*n+1. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?g") proof - have "?summable \ ?pos \ ?neg \ ?f \ ?g" proof (cases "(\n. 0 \ a n) \ (\m. \n\m. a n \ a m)") case True then have ord: "\n m. m \ n \ a n \ a m" and ge0: "\n. 0 \ a n" by auto have mono: "a (Suc n) \ a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF \a \ 0\ ge0] from leibniz[OF mono] show ?thesis using \0 \ a 0\ by auto next let ?a = "\n. - a n" case False with monoseq_le[OF \monoseq a\ \a \ 0\] have "(\ n. a n \ 0) \ (\m. \n\m. a m \ a n)" by auto then have ord: "\n m. m \ n \ ?a n \ ?a m" and ge0: "\ n. 0 \ ?a n" by auto have monotone: "?a (Suc n) \ ?a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF _ ge0, of "\x. x", OF tendsto_minus[OF \a \ 0\, unfolded minus_zero] monotone] have "summable (\ n. (-1)^n * ?a n)" using leibniz(1) by auto then obtain l where "(\ n. (-1)^n * ?a n) sums l" unfolding summable_def by auto from this[THEN sums_minus] have "(\ n. (-1)^n * a n) sums -l" by auto then have ?summable by (auto simp: summable_def) moreover have "\- a - - b\ = \a - b\" for a b :: real unfolding minus_diff_minus by auto from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus] have move_minus: "(\n. - ((- 1) ^ n * a n)) = - (\n. (- 1) ^ n * a n)" by auto have ?pos using \0 \ ?a 0\ by auto moreover have ?neg using leibniz(2,4) unfolding mult_minus_right sum_negf move_minus neg_le_iff_le by auto moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right sum_negf move_minus, THEN tendsto_minus_cancel] by auto ultimately show ?thesis by auto qed then show ?summable and ?pos and ?neg and ?f and ?g by safe qed subsection \Term-by-Term Differentiability of Power Series\ definition diffs :: "(nat \ 'a::ring_1) \ nat \ 'a" where "diffs c = (\n. of_nat (Suc n) * c (Suc n))" text \Lemma about distributing negation over it.\ lemma diffs_minus: "diffs (\n. - c n) = (\n. - diffs c n)" by (simp add: diffs_def) lemma diffs_equiv: fixes x :: "'a::{real_normed_vector,ring_1}" shows "summable (\n. diffs c n * x^n) \ (\n. of_nat n * c n * x^(n - Suc 0)) sums (\n. diffs c n * x^n)" unfolding diffs_def by (simp add: summable_sums sums_Suc_imp) lemma lemma_termdiff1: fixes z :: "'a :: {monoid_mult,comm_ring}" shows "(\ppipp 0" shows "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = h * (\p< n - Suc 0. \q< n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") proof (cases n) - case (Suc n) + case (Suc m) have 0: "\x k. (\njiijiijppip::nat. p < n \ f p \ K" - and K: "0 \ K" + assumes f: "\p::nat. p < n \ f p \ K" and K: "0 \ K" shows "sum f {.. of_nat n * K" - apply (rule order_trans [OF sum_mono [OF f]]) - apply (auto simp: mult_right_mono K) - done +proof - + have "sum f {.. (\i of_nat n * K" + by (auto simp: mult_right_mono K) + finally show ?thesis . +qed lemma lemma_termdiff3: fixes h z :: "'a::real_normed_field" assumes 1: "h \ 0" and 2: "norm z \ K" and 3: "norm (z + h) \ K" shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" proof - have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = norm (\pq \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" proof (rule mult_right_mono [OF _ norm_ge_zero]) from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) - have le_Kn: "\i j n. i + j = n \ norm ((z + h) ^ i * z ^ j) \ K ^ n" - apply (erule subst) - apply (simp only: norm_mult norm_power power_add) - apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) - done + have le_Kn: "norm ((z + h) ^ i * z ^ j) \ K ^ n" if "i + j = n" for i j n + proof - + have "norm (z + h) ^ i * norm z ^ j \ K ^ i * K ^ j" + by (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) + also have "... = K^n" + by (metis power_add that) + finally show ?thesis + by (simp add: norm_mult norm_power) + qed + then have "\p q. + \p < n; q < n - Suc 0\ \ norm ((z + h) ^ q * z ^ (n - 2 - q)) \ K ^ (n - 2)" + by simp + then show "norm (\pq of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" - apply (intro - order_trans [OF norm_sum] - real_sum_nat_ivl_bounded2 - mult_nonneg_nonneg - of_nat_0_le_iff - zero_le_power K) - apply (rule le_Kn, simp) - done + by (intro order_trans [OF norm_sum] + real_sum_nat_ivl_bounded2 mult_nonneg_nonneg of_nat_0_le_iff zero_le_power K) qed also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" by (simp only: mult.assoc) finally show ?thesis . qed lemma lemma_termdiff4: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" and k :: real assumes k: "0 < k" and le: "\h. h \ 0 \ norm h < k \ norm (f h) \ K * norm h" shows "f \0\ 0" proof (rule tendsto_norm_zero_cancel) show "(\h. norm (f h)) \0\ 0" proof (rule real_tendsto_sandwich) show "eventually (\h. 0 \ norm (f h)) (at 0)" by simp show "eventually (\h. norm (f h) \ K * norm h) (at 0)" using k by (auto simp: eventually_at dist_norm le) show "(\h. 0) \(0::'a)\ (0::real)" by (rule tendsto_const) have "(\h. K * norm h) \(0::'a)\ K * norm (0::'a)" by (intro tendsto_intros) then show "(\h. K * norm h) \(0::'a)\ 0" by simp qed qed lemma lemma_termdiff5: fixes g :: "'a::real_normed_vector \ nat \ 'b::banach" and k :: real assumes k: "0 < k" and f: "summable f" and le: "\h n. h \ 0 \ norm h < k \ norm (g h n) \ f n * norm h" shows "(\h. suminf (g h)) \0\ 0" proof (rule lemma_termdiff4 [OF k]) fix h :: 'a assume "h \ 0" and "norm h < k" then have 1: "\n. norm (g h n) \ f n * norm h" by (simp add: le) then have "\N. \n\N. norm (norm (g h n)) \ f n * norm h" by simp moreover from f have 2: "summable (\n. f n * norm h)" by (rule summable_mult2) ultimately have 3: "summable (\n. norm (g h n))" by (rule summable_comparison_test) then have "norm (suminf (g h)) \ (\n. norm (g h n))" by (rule summable_norm) also from 1 3 2 have "(\n. norm (g h n)) \ (\n. f n * norm h)" by (rule suminf_le) also from f have "(\n. f n * norm h) = suminf f * norm h" by (rule suminf_mult2 [symmetric]) finally show "norm (suminf (g h)) \ suminf f * norm h" . qed (* FIXME: Long proofs *) lemma termdiffs_aux: fixes x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" and 2: "norm x < norm K" shows "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" proof - from dense [OF 2] obtain r where r1: "norm x < r" and r2: "r < norm K" by fast from norm_ge_zero r1 have r: "0 < r" by (rule order_le_less_trans) then have r_neq_0: "r \ 0" by simp show ?thesis proof (rule lemma_termdiff5) show "0 < r - norm x" using r1 by simp from r r2 have "norm (of_real r::'a) < norm K" by simp with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" by (rule powser_insidea) then have "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" using r by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) then have "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) = - (\n. diffs (\m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" - apply (rule ext) - apply (case_tac n) - apply (simp_all add: diffs_def r_neq_0) - done + (\n. diffs (\m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" + by (simp add: diffs_def r_neq_0 fun_eq_iff split: nat_diff_split) finally have "summable (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0)) = (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" - apply (rule ext) - apply (case_tac n, simp) - apply (rename_tac nat) - apply (case_tac nat, simp) - apply (simp add: r_neq_0) - done + by (rule ext) (simp add: r_neq_0 split: nat_diff_split) finally show "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . next - fix h :: 'a - fix n :: nat + fix h :: 'a and n assume h: "h \ 0" assume "norm h < r - norm x" then have "norm x + norm h < r" by simp - with norm_triangle_ineq have xh: "norm (x + h) < r" + with norm_triangle_ineq + have xh: "norm (x + h) < r" by (rule order_le_less_trans) - show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \ + have "norm (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)) + \ real n * (real (n - Suc 0) * (r ^ (n - 2) * norm h))" + by (metis (mono_tags, lifting) h mult.assoc lemma_termdiff3 less_eq_real_def r1 xh) + then show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" - apply (simp only: norm_mult mult.assoc) - apply (rule mult_left_mono [OF _ norm_ge_zero]) - apply (simp add: mult.assoc [symmetric]) - apply (metis h lemma_termdiff3 less_eq_real_def r1 xh) - done + by (simp only: norm_mult mult.assoc mult_left_mono [OF _ norm_ge_zero]) qed qed lemma termdiffs: fixes K x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. c n * K ^ n)" and 2: "summable (\n. (diffs c) n * K ^ n)" and 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" and 4: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. (diffs c) n * x^n)" unfolding DERIV_def proof (rule LIM_zero_cancel) show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x^n)) / h - suminf (\n. diffs c n * x^n)) \0\ 0" proof (rule LIM_equal2) show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq) next fix h :: 'a assume "norm (h - 0) < norm K - norm x" then have "norm x + norm h < norm K" by simp then have 5: "norm (x + h) < norm K" by (rule norm_triangle_ineq [THEN order_le_less_trans]) have "summable (\n. c n * x^n)" and "summable (\n. c n * (x + h) ^ n)" and "summable (\n. diffs c n * x^n)" using 1 2 4 5 by (auto elim: powser_inside) then have "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. (c n * (x + h) ^ n - c n * x^n) / h - of_nat n * c n * x ^ (n - Suc 0))" by (intro sums_unique sums_diff sums_divide diffs_equiv summable_sums) then show "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0)))" by (simp add: algebra_simps) next show "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" by (rule termdiffs_aux [OF 3 4]) qed qed subsection \The Derivative of a Power Series Has the Same Radius of Convergence\ lemma termdiff_converges: fixes x :: "'a::{real_normed_field,banach}" assumes K: "norm x < K" and sm: "\x. norm x < K \ summable(\n. c n * x ^ n)" shows "summable (\n. diffs c n * x ^ n)" proof (cases "x = 0") case True then show ?thesis using powser_sums_zero sums_summable by auto next case False then have "K > 0" using K less_trans zero_less_norm_iff by blast then obtain r :: real where r: "norm x < norm r" "norm r < K" "r > 0" using K False by (auto simp: field_simps abs_less_iff add_pos_pos intro: that [of "(norm x + K) / 2"]) have to0: "(\n. of_nat n * (x / of_real r) ^ n) \ 0" using r by (simp add: norm_divide powser_times_n_limit_0 [of "x / of_real r"]) obtain N where N: "\n. n\N \ real_of_nat n * norm x ^ n < r ^ n" using r LIMSEQ_D [OF to0, of 1] by (auto simp: norm_divide norm_mult norm_power field_simps) have "summable (\n. (of_nat n * c n) * x ^ n)" proof (rule summable_comparison_test') show "summable (\n. norm (c n * of_real r ^ n))" apply (rule powser_insidea [OF sm [of "of_real ((r+K)/2)"]]) using N r norm_of_real [of "r + K", where 'a = 'a] by auto show "\n. N \ n \ norm (of_nat n * c n * x ^ n) \ norm (c n * of_real r ^ n)" using N r by (fastforce simp add: norm_mult norm_power less_eq_real_def) qed then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ Suc n)" using summable_iff_shift [of "\n. of_nat n * c n * x ^ n" 1] by simp then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ n)" using False summable_mult2 [of "\n. (of_nat (Suc n) * c(Suc n) * x ^ n) * x" "inverse x"] by (simp add: mult.assoc) (auto simp: ac_simps) then show ?thesis by (simp add: diffs_def) qed lemma termdiff_converges_all: fixes x :: "'a::{real_normed_field,banach}" assumes "\x. summable (\n. c n * x^n)" shows "summable (\n. diffs c n * x^n)" by (rule termdiff_converges [where K = "1 + norm x"]) (use assms in auto) lemma termdiffs_strong: fixes K x :: "'a::{real_normed_field,banach}" assumes sm: "summable (\n. c n * K ^ n)" and K: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. diffs c n * x^n)" proof - - have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K" - using K - apply (auto simp: norm_divide field_simps) - apply (rule le_less_trans [of _ "of_real (norm K) + of_real (norm x)"]) - apply (auto simp: mult_2_right norm_triangle_mono) - done + have "norm K + norm x < norm K + norm K" + using K by force + then have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K" + by (auto simp: norm_triangle_lt norm_divide field_simps) then have [simp]: "norm ((of_real (norm K) + of_real (norm x)) :: 'a) < norm K * 2" by simp have "summable (\n. c n * (of_real (norm x + norm K) / 2) ^ n)" by (metis K2 summable_norm_cancel [OF powser_insidea [OF sm]] add.commute of_real_add) moreover have "\x. norm x < norm K \ summable (\n. diffs c n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) moreover have "\x. norm x < norm K \ summable (\n. diffs(diffs c) n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) ultimately show ?thesis - apply (rule termdiffs [where K = "of_real (norm x + norm K) / 2"]) - using K - apply (auto simp: field_simps) - apply (simp flip: of_real_add) - done + by (rule termdiffs [where K = "of_real (norm x + norm K) / 2"]) + (use K in \auto simp: field_simps simp flip: of_real_add\) qed lemma termdiffs_strong_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "((\x. \n. c n * x^n) has_field_derivative (\n. diffs c n * x^n)) (at x)" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force simp del: of_real_add) lemma termdiffs_strong': fixes z :: "'a :: {real_normed_field,banach}" assumes "\z. norm z < K \ summable (\n. c n * z ^ n)" assumes "norm z < K" shows "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" proof (rule termdiffs_strong) define L :: real where "L = (norm z + K) / 2" have "0 \ norm z" by simp also note \norm z < K\ finally have K: "K \ 0" by simp from assms K have L: "L \ 0" "norm z < L" "L < K" by (simp_all add: L_def) from L show "norm z < norm (of_real L :: 'a)" by simp from L show "summable (\n. c n * of_real L ^ n)" by (intro assms(1)) simp_all qed lemma termdiffs_sums_strong: fixes z :: "'a :: {banach,real_normed_field}" assumes sums: "\z. norm z < K \ (\n. c n * z ^ n) sums f z" assumes deriv: "(f has_field_derivative f') (at z)" assumes norm: "norm z < K" shows "(\n. diffs c n * z ^ n) sums f'" proof - have summable: "summable (\n. diffs c n * z^n)" by (intro termdiff_converges[OF norm] sums_summable[OF sums]) from norm have "eventually (\z. z \ norm -` {..z. (\n. c n * z^n) = f z) (nhds z)" by eventually_elim (insert sums, simp add: sums_iff) have "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" by (intro termdiffs_strong'[OF _ norm] sums_summable[OF sums]) hence "(f has_field_derivative (\n. diffs c n * z^n)) (at z)" by (subst (asm) DERIV_cong_ev[OF refl eq refl]) from this and deriv have "(\n. diffs c n * z^n) = f'" by (rule DERIV_unique) with summable show ?thesis by (simp add: sums_iff) qed lemma isCont_powser: fixes K x :: "'a::{real_normed_field,banach}" assumes "summable (\n. c n * K ^ n)" assumes "norm x < norm K" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms] by (blast intro!: DERIV_isCont) lemmas isCont_powser' = isCont_o2[OF _ isCont_powser] lemma isCont_powser_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force intro!: DERIV_isCont simp del: of_real_add) lemma powser_limit_0: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have "norm (of_real s / 2 :: 'a) < s" using s by (auto simp: norm_divide) then have "summable (\n. a n * (of_real s / 2) ^ n)" by (rule sums_summable [OF sm]) then have "((\x. \n. a n * x ^ n) has_field_derivative (\n. diffs a n * 0 ^ n)) (at 0)" by (rule termdiffs_strong) (use s in \auto simp: norm_divide\) then have "isCont (\x. \n. a n * x ^ n) 0" by (blast intro: DERIV_continuous) then have "((\x. \n. a n * x ^ n) \ a 0) (at 0)" by (simp add: continuous_within) - then show ?thesis - apply (rule Lim_transform) + moreover have "(\x. f x - (\n. a n * x ^ n)) \0\ 0" apply (clarsimp simp: LIM_eq) apply (rule_tac x=s in exI) using s sm sums_unique by fastforce + ultimately show ?thesis + by (rule Lim_transform) qed lemma powser_limit_0_strong: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. x \ 0 \ norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have *: "((\x. if x = 0 then a 0 else f x) \ a 0) (at 0)" by (rule powser_limit_0 [OF s]) (auto simp: powser_sums_zero sm) show ?thesis - apply (subst LIM_equal [where g = "(\x. if x = 0 then a 0 else f x)"]) - apply (simp_all add: *) - done + by (simp add: * LIM_equal [where g = "(\x. if x = 0 then a 0 else f x)"]) qed subsection \Derivability of power series\ lemma DERIV_series': fixes f :: "real \ nat \ real" assumes DERIV_f: "\ n. DERIV (\ x. f x n) x0 :> (f' x0 n)" and allf_summable: "\ x. x \ {a <..< b} \ summable (f x)" and x0_in_I: "x0 \ {a <..< b}" and "summable (f' x0)" and "summable L" and L_def: "\n x y. x \ {a <..< b} \ y \ {a <..< b} \ \f x n - f y n\ \ L n * \x - y\" shows "DERIV (\ x. suminf (f x)) x0 :> (suminf (f' x0))" unfolding DERIV_def proof (rule LIM_I) fix r :: real assume "0 < r" then have "0 < r/3" by auto obtain N_L where N_L: "\ n. N_L \ n \ \ \ i. L (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable L\] by auto obtain N_f' where N_f': "\ n. N_f' \ n \ \ \ i. f' x0 (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable (f' x0)\] by auto let ?N = "Suc (max N_L N_f')" have "\ \ i. f' x0 (i + ?N) \ < r/3" (is "?f'_part < r/3") and L_estimate: "\ \ i. L (i + ?N) \ < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto let ?diff = "\i x. (f (x0 + x) i - f x0 i) / x" let ?r = "r / (3 * real ?N)" from \0 < r\ have "0 < ?r" by simp let ?s = "\n. SOME s. 0 < s \ (\ x. x \ 0 \ \ x \ < s \ \ ?diff n x - f' x0 n \ < ?r)" define S' where "S' = Min (?s ` {..< ?N })" have "0 < S'" unfolding S'_def proof (rule iffD2[OF Min_gr_iff]) show "\x \ (?s ` {..< ?N }). 0 < x" proof fix x assume "x \ ?s ` {.. {..0 < ?r\, unfolded real_norm_def] obtain s where s_bound: "0 < s \ (\x. x \ 0 \ \x\ < s \ \?diff n x - f' x0 n\ < ?r)" by auto have "0 < ?s n" by (rule someI2[where a=s]) (auto simp: s_bound simp del: of_nat_Suc) then show "0 < x" by (simp only: \x = ?s n\) qed qed auto define S where "S = min (min (x0 - a) (b - x0)) S'" then have "0 < S" and S_a: "S \ x0 - a" and S_b: "S \ b - x0" and "S \ S'" using x0_in_I and \0 < S'\ by auto have "\(suminf (f (x0 + x)) - suminf (f x0)) / x - suminf (f' x0)\ < r" if "x \ 0" and "\x\ < S" for x proof - from that have x_in_I: "x0 + x \ {a <..< b}" using S_a S_b by auto note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] note div_smbl = summable_divide[OF diff_smbl] note all_smbl = summable_diff[OF div_smbl \summable (f' x0)\] note ign = summable_ignore_initial_segment[where k="?N"] note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]] note div_shft_smbl = summable_divide[OF diff_shft_smbl] note all_shft_smbl = summable_diff[OF div_smbl ign[OF \summable (f' x0)\]] have 1: "\(\?diff (n + ?N) x\)\ \ L (n + ?N)" for n proof - have "\?diff (n + ?N) x\ \ L (n + ?N) * \(x0 + x) - x0\ / \x\" using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] by (simp only: abs_divide) with \x \ 0\ show ?thesis by auto qed note 2 = summable_rabs_comparison_test[OF _ ign[OF \summable L\]] from 1 have "\ \ i. ?diff (i + ?N) x \ \ (\ i. L (i + ?N))" by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF \summable L\]]]) then have "\\i. ?diff (i + ?N) x\ \ r / 3" (is "?L_part \ r/3") using L_estimate by auto have "\\n \ (\n?diff n x - f' x0 n\)" .. also have "\ < (\n {..< ?N}" have "\x\ < S" using \\x\ < S\ . also have "S \ S'" using \S \ S'\ . also have "S' \ ?s n" unfolding S'_def proof (rule Min_le_iff[THEN iffD2]) have "?s n \ (?s ` {.. ?s n \ ?s n" using \n \ {..< ?N}\ by auto then show "\ a \ (?s ` {.. ?s n" by blast qed auto finally have "\x\ < ?s n" . from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \0 < ?r\, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2] have "\x. x \ 0 \ \x\ < ?s n \ \?diff n x - f' x0 n\ < ?r" . with \x \ 0\ and \\x\ < ?s n\ show "\?diff n x - f' x0 n\ < ?r" by blast qed auto also have "\ = of_nat (card {.. = real ?N * ?r" by simp also have "\ = r/3" by (auto simp del: of_nat_Suc) finally have "\\n < r / 3" (is "?diff_part < r / 3") . from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] have "\(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\ = \\n. ?diff n x - f' x0 n\" unfolding suminf_diff[OF div_smbl \summable (f' x0)\, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto also have "\ \ ?diff_part + \(\n. ?diff (n + ?N) x) - (\ n. f' x0 (n + ?N))\" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF \summable (f' x0)\]] apply (simp only: add.commute) using abs_triangle_ineq by blast also have "\ \ ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto also have "\ < r /3 + r/3 + r/3" using \?diff_part < r/3\ \?L_part \ r/3\ and \?f'_part < r/3\ by (rule add_strict_mono [OF add_less_le_mono]) finally show ?thesis by auto qed then show "\s > 0. \ x. x \ 0 \ norm (x - 0) < s \ norm (((\n. f (x0 + x) n) - (\n. f x0 n)) / x - (\n. f' x0 n)) < r" using \0 < S\ by auto qed lemma DERIV_power_series': fixes f :: "nat \ real" assumes converges: "\x. x \ {-R <..< R} \ summable (\n. f n * real (Suc n) * x^n)" and x0_in_I: "x0 \ {-R <..< R}" and "0 < R" shows "DERIV (\x. (\n. f n * x^(Suc n))) x0 :> (\n. f n * real (Suc n) * x0^n)" (is "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)") proof - have for_subinterval: "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)" if "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'" for R' proof - from that have "x0 \ {-R' <..< R'}" and "R' \ {-R <..< R}" and "x0 \ {-R <..< R}" by auto show ?thesis proof (rule DERIV_series') show "summable (\ n. \f n * real (Suc n) * R'^n\)" proof - have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) then have in_Rball: "(R' + R) / 2 \ {-R <..< R}" using \R' < R\ by auto have "norm R' < norm ((R' + R) / 2)" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto qed next fix n x y assume "x \ {-R' <..< R'}" and "y \ {-R' <..< R'}" show "\?f x n - ?f y n\ \ \f n * real (Suc n) * R'^n\ * \x-y\" proof - have "\f n * x ^ (Suc n) - f n * y ^ (Suc n)\ = (\f n\ * \x-y\) * \\p" unfolding right_diff_distrib[symmetric] diff_power_eq_sum abs_mult by auto also have "\ \ (\f n\ * \x-y\) * (\real (Suc n)\ * \R' ^ n\)" proof (rule mult_left_mono) have "\\p \ (\px ^ p * y ^ (n - p)\)" by (rule sum_abs) also have "\ \ (\p {.. n" by auto have "\x^n\ \ R'^n" if "x \ {-R'<..x\ \ R'" by auto then show ?thesis unfolding power_abs by (rule power_mono) auto qed from mult_mono[OF this[OF \x \ {-R'<.., of p] this[OF \y \ {-R'<.., of "n-p"]] and \0 < R'\ have "\x^p * y^(n - p)\ \ R'^p * R'^(n - p)" unfolding abs_mult by auto then show "\x^p * y^(n - p)\ \ R'^n" unfolding power_add[symmetric] using \p \ n\ by auto qed also have "\ = real (Suc n) * R' ^ n" unfolding sum_constant card_atLeastLessThan by auto finally show "\\p \ \real (Suc n)\ * \R' ^ n\" unfolding abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] by linarith show "0 \ \f n\ * \x - y\" unfolding abs_mult[symmetric] by auto qed also have "\ = \f n * real (Suc n) * R' ^ n\ * \x - y\" unfolding abs_mult mult.assoc[symmetric] by algebra finally show ?thesis . qed next show "DERIV (\x. ?f x n) x0 :> ?f' x0 n" for n by (auto intro!: derivative_eq_intros simp del: power_Suc) next fix x assume "x \ {-R' <..< R'}" then have "R' \ {-R <..< R}" and "norm x < norm R'" using assms \R' < R\ by auto have "summable (\n. f n * x^n)" proof (rule summable_comparison_test, intro exI allI impI) fix n have le: "\f n\ * 1 \ \f n\ * real (Suc n)" by (rule mult_left_mono) auto show "norm (f n * x^n) \ norm (f n * real (Suc n) * x^n)" unfolding real_norm_def abs_mult using le mult_right_mono by fastforce qed (rule powser_insidea[OF converges[OF \R' \ {-R <..< R}\] \norm x < norm R'\]) from this[THEN summable_mult2[where c=x], simplified mult.assoc, simplified mult.commute] show "summable (?f x)" by auto next show "summable (?f' x0)" using converges[OF \x0 \ {-R <..< R}\] . show "x0 \ {-R' <..< R'}" using \x0 \ {-R' <..< R'}\ . qed qed let ?R = "(R + \x0\) / 2" have "\x0\ < ?R" using assms by (auto simp: field_simps) then have "- ?R < x0" proof (cases "x0 < 0") case True then have "- x0 < ?R" using \\x0\ < ?R\ by auto then show ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto next case False have "- ?R < 0" using assms by auto also have "\ \ x0" using False by auto finally show ?thesis . qed then have "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by (auto simp: field_simps) from for_subinterval[OF this] show ?thesis . qed lemma geometric_deriv_sums: fixes z :: "'a :: {real_normed_field,banach}" assumes "norm z < 1" shows "(\n. of_nat (Suc n) * z ^ n) sums (1 / (1 - z)^2)" proof - have "(\n. diffs (\n. 1) n * z^n) sums (1 / (1 - z)^2)" proof (rule termdiffs_sums_strong) fix z :: 'a assume "norm z < 1" thus "(\n. 1 * z^n) sums (1 / (1 - z))" by (simp add: geometric_sums) qed (insert assms, auto intro!: derivative_eq_intros simp: power2_eq_square) thus ?thesis unfolding diffs_def by simp qed lemma isCont_pochhammer [continuous_intros]: "isCont (\z. pochhammer z n) z" for z :: "'a::real_normed_field" by (induct n) (auto simp: pochhammer_rec') lemma continuous_on_pochhammer [continuous_intros]: "continuous_on A (\z. pochhammer z n)" for A :: "'a::real_normed_field set" by (intro continuous_at_imp_continuous_on ballI isCont_pochhammer) lemmas continuous_on_pochhammer' [continuous_intros] = continuous_on_compose2[OF continuous_on_pochhammer _ subset_UNIV] subsection \Exponential Function\ definition exp :: "'a \ 'a::{real_normed_algebra_1,banach}" where "exp = (\x. \n. x^n /\<^sub>R fact n)" lemma summable_exp_generic: fixes x :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \n. x^n /\<^sub>R fact n" shows "summable S" proof - have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R (Suc n)" unfolding S_def by (simp del: mult_Suc) obtain r :: real where r0: "0 < r" and r1: "r < 1" using dense [OF zero_less_one] by fast obtain N :: nat where N: "norm x < real N * r" using ex_less_of_nat_mult r0 by auto from r1 show ?thesis proof (rule summable_ratio_test [rule_format]) fix n :: nat assume n: "N \ n" have "norm x \ real N * r" using N by (rule order_less_imp_le) also have "real N * r \ real (Suc n) * r" using r0 n by (simp add: mult_right_mono) finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" using norm_ge_zero by (rule mult_right_mono) then have "norm (x * S n) \ real (Suc n) * r * norm (S n)" by (rule order_trans [OF norm_mult_ineq]) then have "norm (x * S n) / real (Suc n) \ r * norm (S n)" by (simp add: pos_divide_le_eq ac_simps) then show "norm (S (Suc n)) \ r * norm (S n)" by (simp add: S_Suc inverse_eq_divide) qed qed lemma summable_norm_exp: "summable (\n. norm (x^n /\<^sub>R fact n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_norm_comparison_test [OF exI, rule_format]) show "summable (\n. norm x^n /\<^sub>R fact n)" by (rule summable_exp_generic) show "norm (x^n /\<^sub>R fact n) \ norm x^n /\<^sub>R fact n" for n by (simp add: norm_power_ineq) qed lemma summable_exp: "summable (\n. inverse (fact n) * x^n)" for x :: "'a::{real_normed_field,banach}" using summable_exp_generic [where x=x] by (simp add: scaleR_conv_of_real nonzero_of_real_inverse) lemma exp_converges: "(\n. x^n /\<^sub>R fact n) sums exp x" unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) lemma exp_fdiffs: "diffs (\n. inverse (fact n)) = (\n. inverse (fact n :: 'a::{real_normed_field,banach}))" by (simp add: diffs_def mult_ac nonzero_inverse_mult_distrib nonzero_of_real_inverse del: mult_Suc of_nat_Suc) lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" by (simp add: diffs_def) lemma DERIV_exp [simp]: "DERIV exp x :> exp x" unfolding exp_def scaleR_conv_of_real proof (rule DERIV_cong) have sinv: "summable (\n. of_real (inverse (fact n)) * x ^ n)" for x::'a by (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real]) note xx = exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real] show "((\x. \n. of_real (inverse (fact n)) * x ^ n) has_field_derivative (\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n)) (at x)" by (rule termdiffs [where K="of_real (1 + norm x)"]) (simp_all only: diffs_of_real exp_fdiffs sinv norm_of_real) show "(\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n) = (\n. of_real (inverse (fact n)) * x ^ n)" by (simp add: diffs_of_real exp_fdiffs) qed declare DERIV_exp[THEN DERIV_chain2, derivative_intros] and DERIV_exp[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_exp[derivative_intros] = DERIV_exp[THEN DERIV_compose_FDERIV] lemma norm_exp: "norm (exp x) \ exp (norm x)" proof - from summable_norm[OF summable_norm_exp, of x] have "norm (exp x) \ (\n. inverse (fact n) * norm (x^n))" by (simp add: exp_def) also have "\ \ exp (norm x)" using summable_exp_generic[of "norm x"] summable_norm_exp[of x] by (auto simp: exp_def intro!: suminf_le norm_power_ineq) finally show ?thesis . qed lemma isCont_exp: "isCont exp x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_exp [THEN DERIV_isCont]) lemma isCont_exp' [simp]: "isCont f a \ isCont (\x. exp (f x)) a" for f :: "_ \'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_exp]) lemma tendsto_exp [tendsto_intros]: "(f \ a) F \ ((\x. exp (f x)) \ exp a) F" for f:: "_ \'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_exp]) lemma continuous_exp [continuous_intros]: "continuous F f \ continuous F (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_exp) lemma continuous_on_exp [continuous_intros]: "continuous_on s f \ continuous_on s (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_on_def by (auto intro: tendsto_exp) subsubsection \Properties of the Exponential Function\ lemma exp_zero [simp]: "exp 0 = 1" unfolding exp_def by (simp add: scaleR_conv_of_real) lemma exp_series_add_commuting: fixes x y :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \x n. x^n /\<^sub>R fact n" assumes comm: "x * y = y * x" shows "S (x + y) n = (\i\n. S x i * S y (n - i))" proof (induct n) case 0 show ?case unfolding S_def by simp next case (Suc n) have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" unfolding S_def by (simp del: mult_Suc) then have times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" by simp have S_comm: "\n. S x n * y = y * S x n" by (simp add: power_commuting_commutes comm S_def) have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * S (x + y) n" by (simp only: times_S) also have "\ = (x + y) * (\i\n. S x i * S y (n - i))" by (simp only: Suc) also have "\ = x * (\i\n. S x i * S y (n - i)) + y * (\i\n. S x i * S y (n - i))" by (rule distrib_right) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * y * S y (n - i))" by (simp add: sum_distrib_left ac_simps S_comm) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * (y * S y (n - i)))" by (simp add: ac_simps) also have "\ = (\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) + (\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp add: times_S Suc_diff_le) also have "(\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) = (\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i)))" by (subst sum.atMost_Suc_shift) simp also have "(\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by simp also have "(\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i))) + (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp only: sum.distrib [symmetric] scaleR_left_distrib [symmetric] of_nat_add [symmetric]) simp also have "\ = real (Suc n) *\<^sub>R (\i\Suc n. S x i * S y (Suc n - i))" by (simp only: scaleR_right.sum) finally show "S (x + y) (Suc n) = (\i\Suc n. S x i * S y (Suc n - i))" by (simp del: sum.cl_ivl_Suc) qed lemma exp_add_commuting: "x * y = y * x \ exp (x + y) = exp x * exp y" by (simp only: exp_def Cauchy_product summable_norm_exp exp_series_add_commuting) lemma exp_times_arg_commute: "exp A * A = A * exp A" by (simp add: exp_def suminf_mult[symmetric] summable_exp_generic power_commutes suminf_mult2) lemma exp_add: "exp (x + y) = exp x * exp y" for x y :: "'a::{real_normed_field,banach}" by (rule exp_add_commuting) (simp add: ac_simps) lemma exp_double: "exp(2 * z) = exp z ^ 2" by (simp add: exp_add_commuting mult_2 power2_eq_square) lemmas mult_exp_exp = exp_add [symmetric] lemma exp_of_real: "exp (of_real x) = of_real (exp x)" unfolding exp_def apply (subst suminf_of_real [OF summable_exp_generic]) apply (simp add: scaleR_conv_of_real) done lemmas of_real_exp = exp_of_real[symmetric] corollary exp_in_Reals [simp]: "z \ \ \ exp z \ \" by (metis Reals_cases Reals_of_real exp_of_real) lemma exp_not_eq_zero [simp]: "exp x \ 0" proof have "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) also assume "exp x = 0" finally show False by simp qed lemma exp_minus_inverse: "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) lemma exp_minus: "exp (- x) = inverse (exp x)" for x :: "'a::{real_normed_field,banach}" by (intro inverse_unique [symmetric] exp_minus_inverse) lemma exp_diff: "exp (x - y) = exp x / exp y" for x :: "'a::{real_normed_field,banach}" using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse) lemma exp_of_nat_mult: "exp (of_nat n * x) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (induct n) (auto simp: distrib_left exp_add mult.commute) corollary exp_of_nat2_mult: "exp (x * of_nat n) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (metis exp_of_nat_mult mult_of_nat_commute) lemma exp_sum: "finite I \ exp (sum f I) = prod (\x. exp (f x)) I" by (induct I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) lemma exp_divide_power_eq: fixes x :: "'a::{real_normed_field,banach}" assumes "n > 0" shows "exp (x / of_nat n) ^ n = exp x" using assms proof (induction n arbitrary: x) case (Suc n) show ?case proof (cases "n = 0") case True then show ?thesis by simp next case False have [simp]: "1 + (of_nat n * of_nat n + of_nat n * 2) \ (0::'a)" using of_nat_eq_iff [of "1 + n * n + n * 2" "0"] by simp from False have [simp]: "x * of_nat n / (1 + of_nat n) / of_nat n = x / (1 + of_nat n)" by simp have [simp]: "x / (1 + of_nat n) + x * of_nat n / (1 + of_nat n) = x" using of_nat_neq_0 by (auto simp add: field_split_simps) show ?thesis using Suc.IH [of "x * of_nat n / (1 + of_nat n)"] False by (simp add: exp_add [symmetric]) qed qed simp subsubsection \Properties of the Exponential Function on Reals\ text \Comparisons of \<^term>\exp x\ with zero.\ text \Proof: because every exponential can be seen as a square.\ lemma exp_ge_zero [simp]: "0 \ exp x" for x :: real proof - have "0 \ exp (x/2) * exp (x/2)" by simp then show ?thesis by (simp add: exp_add [symmetric]) qed lemma exp_gt_zero [simp]: "0 < exp x" for x :: real by (simp add: order_less_le) lemma not_exp_less_zero [simp]: "\ exp x < 0" for x :: real by (simp add: not_less) lemma not_exp_le_zero [simp]: "\ exp x \ 0" for x :: real by (simp add: not_le) lemma abs_exp_cancel [simp]: "\exp x\ = exp x" for x :: real by simp text \Strict monotonicity of exponential.\ lemma exp_ge_add_one_self_aux: fixes x :: real assumes "0 \ x" shows "1 + x \ exp x" using order_le_imp_less_or_eq [OF assms] proof assume "0 < x" have "1 + x \ (\n<2. inverse (fact n) * x^n)" by (auto simp: numeral_2_eq_2) also have "\ \ (\n. inverse (fact n) * x^n)" - apply (rule sum_le_suminf [OF summable_exp]) - using \0 < x\ - apply (auto simp add: zero_le_mult_iff) - done + proof (rule sum_le_suminf [OF summable_exp]) + show "\m\- {..<2}. 0 \ inverse (fact m) * x ^ m" + using \0 < x\ by (auto simp add: zero_le_mult_iff) + qed auto finally show "1 + x \ exp x" by (simp add: exp_def) qed auto lemma exp_gt_one: "0 < x \ 1 < exp x" for x :: real proof - assume x: "0 < x" then have "1 < 1 + x" by simp also from x have "1 + x \ exp x" by (simp add: exp_ge_add_one_self_aux) finally show ?thesis . qed lemma exp_less_mono: fixes x y :: real assumes "x < y" shows "exp x < exp y" proof - from \x < y\ have "0 < y - x" by simp then have "1 < exp (y - x)" by (rule exp_gt_one) then have "1 < exp y / exp x" by (simp only: exp_diff) then show "exp x < exp y" by simp qed lemma exp_less_cancel: "exp x < exp y \ x < y" for x y :: real unfolding linorder_not_le [symmetric] by (auto simp: order_le_less exp_less_mono) lemma exp_less_cancel_iff [iff]: "exp x < exp y \ x < y" for x y :: real by (auto intro: exp_less_mono exp_less_cancel) lemma exp_le_cancel_iff [iff]: "exp x \ exp y \ x \ y" for x y :: real by (auto simp: linorder_not_less [symmetric]) lemma exp_inj_iff [iff]: "exp x = exp y \ x = y" for x y :: real by (simp add: order_eq_iff) text \Comparisons of \<^term>\exp x\ with one.\ lemma one_less_exp_iff [simp]: "1 < exp x \ 0 < x" for x :: real using exp_less_cancel_iff [where x = 0 and y = x] by simp lemma exp_less_one_iff [simp]: "exp x < 1 \ x < 0" for x :: real using exp_less_cancel_iff [where x = x and y = 0] by simp lemma one_le_exp_iff [simp]: "1 \ exp x \ 0 \ x" for x :: real using exp_le_cancel_iff [where x = 0 and y = x] by simp lemma exp_le_one_iff [simp]: "exp x \ 1 \ x \ 0" for x :: real using exp_le_cancel_iff [where x = x and y = 0] by simp lemma exp_eq_one_iff [simp]: "exp x = 1 \ x = 0" for x :: real using exp_inj_iff [where x = x and y = 0] by simp lemma lemma_exp_total: "1 \ y \ \x. 0 \ x \ x \ y - 1 \ exp x = y" for y :: real proof (rule IVT) assume "1 \ y" then have "0 \ y - 1" by simp then have "1 + (y - 1) \ exp (y - 1)" by (rule exp_ge_add_one_self_aux) then show "y \ exp (y - 1)" by simp qed (simp_all add: le_diff_eq) lemma exp_total: "0 < y \ \x. exp x = y" for y :: real proof (rule linorder_le_cases [of 1 y]) assume "1 \ y" then show "\x. exp x = y" by (fast dest: lemma_exp_total) next assume "0 < y" and "y \ 1" then have "1 \ inverse y" by (simp add: one_le_inverse_iff) then obtain x where "exp x = inverse y" by (fast dest: lemma_exp_total) then have "exp (- x) = y" by (simp add: exp_minus) then show "\x. exp x = y" .. qed subsection \Natural Logarithm\ class ln = real_normed_algebra_1 + banach + fixes ln :: "'a \ 'a" assumes ln_one [simp]: "ln 1 = 0" definition powr :: "'a \ 'a \ 'a::ln" (infixr "powr" 80) \ \exponentation via ln and exp\ where "x powr a \ if x = 0 then 0 else exp (a * ln x)" lemma powr_0 [simp]: "0 powr z = 0" by (simp add: powr_def) instantiation real :: ln begin definition ln_real :: "real \ real" where "ln_real x = (THE u. exp u = x)" instance by intro_classes (simp add: ln_real_def) end lemma powr_eq_0_iff [simp]: "w powr z = 0 \ w = 0" by (simp add: powr_def) lemma ln_exp [simp]: "ln (exp x) = x" for x :: real by (simp add: ln_real_def) lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" for x :: real by (auto dest: exp_total) lemma exp_ln_iff [simp]: "exp (ln x) = x \ 0 < x" for x :: real by (metis exp_gt_zero exp_ln) lemma ln_unique: "exp y = x \ ln x = y" for x :: real by (erule subst) (rule ln_exp) lemma ln_mult: "0 < x \ 0 < y \ ln (x * y) = ln x + ln y" for x :: real by (rule ln_unique) (simp add: exp_add) lemma ln_prod: "finite I \ (\i. i \ I \ f i > 0) \ ln (prod f I) = sum (\x. ln(f x)) I" for f :: "'a \ real" by (induct I rule: finite_induct) (auto simp: ln_mult prod_pos) lemma ln_inverse: "0 < x \ ln (inverse x) = - ln x" for x :: real by (rule ln_unique) (simp add: exp_minus) lemma ln_div: "0 < x \ 0 < y \ ln (x / y) = ln x - ln y" for x :: real by (rule ln_unique) (simp add: exp_diff) lemma ln_realpow: "0 < x \ ln (x^n) = real n * ln x" by (rule ln_unique) (simp add: exp_of_nat_mult) lemma ln_less_cancel_iff [simp]: "0 < x \ 0 < y \ ln x < ln y \ x < y" for x :: real by (subst exp_less_cancel_iff [symmetric]) simp lemma ln_le_cancel_iff [simp]: "0 < x \ 0 < y \ ln x \ ln y \ x \ y" for x :: real by (simp add: linorder_not_less [symmetric]) lemma ln_inj_iff [simp]: "0 < x \ 0 < y \ ln x = ln y \ x = y" for x :: real by (simp add: order_eq_iff) lemma ln_add_one_self_le_self: "0 \ x \ ln (1 + x) \ x" for x :: real by (rule exp_le_cancel_iff [THEN iffD1]) (simp add: exp_ge_add_one_self_aux) lemma ln_less_self [simp]: "0 < x \ ln x < x" for x :: real by (rule order_less_le_trans [where y = "ln (1 + x)"]) (simp_all add: ln_add_one_self_le_self) lemma ln_ge_iff: "\x::real. 0 < x \ y \ ln x \ exp y \ x" using exp_le_cancel_iff exp_total by force lemma ln_ge_zero [simp]: "1 \ x \ 0 \ ln x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_imp_ge_one: "0 \ ln x \ 0 < x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_iff [simp]: "0 < x \ 0 \ ln x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_less_zero_iff [simp]: "0 < x \ ln x < 0 \ x < 1" for x :: real using ln_less_cancel_iff [of x 1] by simp lemma ln_le_zero_iff [simp]: "0 < x \ ln x \ 0 \ x \ 1" for x :: real by (metis less_numeral_extra(1) ln_le_cancel_iff ln_one) lemma ln_gt_zero: "1 < x \ 0 < ln x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_imp_gt_one: "0 < ln x \ 0 < x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_iff [simp]: "0 < x \ 0 < ln x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_eq_zero_iff [simp]: "0 < x \ ln x = 0 \ x = 1" for x :: real using ln_inj_iff [of x 1] by simp lemma ln_less_zero: "0 < x \ x < 1 \ ln x < 0" for x :: real by simp lemma ln_neg_is_const: "x \ 0 \ ln x = (THE x. False)" for x :: real by (auto simp: ln_real_def intro!: arg_cong[where f = The]) lemma powr_eq_one_iff [simp]: "a powr x = 1 \ x = 0" if "a > 1" for a x :: real using that by (auto simp: powr_def split: if_splits) lemma isCont_ln: fixes x :: real assumes "x \ 0" shows "isCont ln x" proof (cases "0 < x") case True then have "isCont ln (exp (ln x))" by (intro isCont_inverse_function[where d = "\x\" and f = exp]) auto with True show ?thesis by simp next case False with \x \ 0\ show "isCont ln x" unfolding isCont_def by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\_. ln 0"]) (auto simp: ln_neg_is_const not_less eventually_at dist_real_def intro!: exI[of _ "\x\"]) qed lemma tendsto_ln [tendsto_intros]: "(f \ a) F \ a \ 0 \ ((\x. ln (f x)) \ ln a) F" for a :: real by (rule isCont_tendsto_compose [OF isCont_ln]) lemma continuous_ln: "continuous F f \ f (Lim F (\x. x)) \ 0 \ continuous F (\x. ln (f x :: real))" unfolding continuous_def by (rule tendsto_ln) lemma isCont_ln' [continuous_intros]: "continuous (at x) f \ f x \ 0 \ continuous (at x) (\x. ln (f x :: real))" unfolding continuous_at by (rule tendsto_ln) lemma continuous_within_ln [continuous_intros]: "continuous (at x within s) f \ f x \ 0 \ continuous (at x within s) (\x. ln (f x :: real))" unfolding continuous_within by (rule tendsto_ln) lemma continuous_on_ln [continuous_intros]: "continuous_on s f \ (\x\s. f x \ 0) \ continuous_on s (\x. ln (f x :: real))" unfolding continuous_on_def by (auto intro: tendsto_ln) lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" for x :: real by (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) (auto intro: DERIV_cong [OF DERIV_exp exp_ln] isCont_ln) lemma DERIV_ln_divide: "0 < x \ DERIV ln x :> 1 / x" for x :: real by (rule DERIV_ln[THEN DERIV_cong]) (simp_all add: divide_inverse) declare DERIV_ln_divide[THEN DERIV_chain2, derivative_intros] and DERIV_ln_divide[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_ln[derivative_intros] = DERIV_ln[THEN DERIV_compose_FDERIV] lemma ln_series: assumes "0 < x" and "x < 2" shows "ln x = (\ n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))") proof - let ?f' = "\x n. (-1)^n * (x - 1)^n" have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))" proof (rule DERIV_isconst3 [where x = x]) fix x :: real assume "x \ {0 <..< 2}" then have "0 < x" and "x < 2" by auto have "norm (1 - x) < 1" using \0 < x\ and \x < 2\ by auto have "1 / x = 1 / (1 - (1 - x))" by auto also have "\ = (\ n. (1 - x)^n)" using geometric_sums[OF \norm (1 - x) < 1\] by (rule sums_unique) also have "\ = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="(^)"], auto) finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF \0 < x\] unfolding divide_inverse by auto moreover have repos: "\ h x :: real. h - 1 + x = h + x - 1" by auto have "DERIV (\x. suminf (?f x)) (x - 1) :> (\n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)" proof (rule DERIV_power_series') show "x - 1 \ {- 1<..<1}" and "(0 :: real) < 1" using \0 < x\ \x < 2\ by auto next fix x :: real assume "x \ {- 1<..<1}" then have "norm (-x) < 1" by auto show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" unfolding One_nat_def by (auto simp: power_mult_distrib[symmetric] summable_geometric[OF \norm (-x) < 1\]) qed then have "DERIV (\x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto then have "DERIV (\x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_def repos . ultimately have "DERIV (\x. ln x - suminf (?f (x - 1))) x :> suminf (?f' x) - suminf (?f' x)" by (rule DERIV_diff) then show "DERIV (\x. ln x - suminf (?f (x - 1))) x :> 0" by auto qed (auto simp: assms) then show ?thesis by auto qed lemma exp_first_terms: fixes x :: "'a::{real_normed_algebra_1,banach}" shows "exp x = (\nR (x ^ n)) + (\n. inverse(fact (n + k)) *\<^sub>R (x ^ (n + k)))" proof - have "exp x = suminf (\n. inverse(fact n) *\<^sub>R (x^n))" by (simp add: exp_def) also from summable_exp_generic have "\ = (\ n. inverse(fact(n+k)) *\<^sub>R (x ^ (n + k))) + (\ n::natR (x^n))" (is "_ = _ + ?a") by (rule suminf_split_initial_segment) finally show ?thesis by simp qed lemma exp_first_term: "exp x = 1 + (\n. inverse (fact (Suc n)) *\<^sub>R (x ^ Suc n))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 1] by simp lemma exp_first_two_terms: "exp x = 1 + x + (\n. inverse (fact (n + 2)) *\<^sub>R (x ^ (n + 2)))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 2] by (simp add: eval_nat_numeral) lemma exp_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "exp x \ 1 + x + x\<^sup>2" proof - have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ x\<^sup>2" proof - have "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1 / 2)))" by (intro sums_mult geometric_sums) simp then have sumsx: "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums x\<^sup>2" by simp have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ suminf (\n. (x\<^sup>2/2) * ((1/2)^n))" proof (intro suminf_le allI) show "inverse (fact (n + 2)) * x ^ (n + 2) \ (x\<^sup>2/2) * ((1/2)^n)" for n :: nat proof - have "(2::nat) * 2 ^ n \ fact (n + 2)" by (induct n) simp_all then have "real ((2::nat) * 2 ^ n) \ real_of_nat (fact (n + 2))" by (simp only: of_nat_le_iff) then have "((2::real) * 2 ^ n) \ fact (n + 2)" unfolding of_nat_fact by simp then have "inverse (fact (n + 2)) \ inverse ((2::real) * 2 ^ n)" by (rule le_imp_inverse_le) simp then have "inverse (fact (n + 2)) \ 1/(2::real) * (1/2)^n" by (simp add: power_inverse [symmetric]) then have "inverse (fact (n + 2)) * (x^n * x\<^sup>2) \ 1/2 * (1/2)^n * (1 * x\<^sup>2)" by (rule mult_mono) (rule mult_mono, simp_all add: power_le_one a b) then show ?thesis unfolding power_add by (simp add: ac_simps del: fact_Suc) qed show "summable (\n. inverse (fact (n + 2)) * x ^ (n + 2))" by (rule summable_exp [THEN summable_ignore_initial_segment]) show "summable (\n. x\<^sup>2 / 2 * (1 / 2) ^ n)" by (rule sums_summable [OF sumsx]) qed also have "\ = x\<^sup>2" by (rule sums_unique [THEN sym]) (rule sumsx) finally show ?thesis . qed then show ?thesis unfolding exp_first_two_terms by auto qed corollary exp_half_le2: "exp(1/2) \ (2::real)" using exp_bound [of "1/2"] by (simp add: field_simps) corollary exp_le: "exp 1 \ (3::real)" using exp_bound [of 1] by (simp add: field_simps) lemma exp_bound_half: "norm z \ 1/2 \ norm (exp z) \ 2" by (blast intro: order_trans intro!: exp_half_le2 norm_exp) lemma exp_bound_lemma: assumes "norm z \ 1/2" shows "norm (exp z) \ 1 + 2 * norm z" proof - have *: "(norm z)\<^sup>2 \ norm z * 1" unfolding power2_eq_square by (rule mult_left_mono) (use assms in auto) have "norm (exp z) \ exp (norm z)" by (rule norm_exp) also have "\ \ 1 + (norm z) + (norm z)\<^sup>2" using assms exp_bound by auto also have "\ \ 1 + 2 * norm z" using * by auto finally show ?thesis . qed lemma real_exp_bound_lemma: "0 \ x \ x \ 1/2 \ exp x \ 1 + 2 * x" for x :: real using exp_bound_lemma [of x] by simp lemma ln_one_minus_pos_upper_bound: fixes x :: real assumes a: "0 \ x" and b: "x < 1" shows "ln (1 - x) \ - x" proof - have "(1 - x) * (1 + x + x\<^sup>2) = 1 - x^3" by (simp add: algebra_simps power2_eq_square power3_eq_cube) also have "\ \ 1" by (auto simp: a) finally have "(1 - x) * (1 + x + x\<^sup>2) \ 1" . moreover have c: "0 < 1 + x + x\<^sup>2" by (simp add: add_pos_nonneg a) ultimately have "1 - x \ 1 / (1 + x + x\<^sup>2)" by (elim mult_imp_le_div_pos) also have "\ \ 1 / exp x" by (metis a abs_one b exp_bound exp_gt_zero frac_le less_eq_real_def real_sqrt_abs real_sqrt_pow2_iff real_sqrt_power) also have "\ = exp (- x)" by (auto simp: exp_minus divide_inverse) finally have "1 - x \ exp (- x)" . also have "1 - x = exp (ln (1 - x))" by (metis b diff_0 exp_ln_iff less_iff_diff_less_0 minus_diff_eq) finally have "exp (ln (1 - x)) \ exp (- x)" . then show ?thesis by (auto simp only: exp_le_cancel_iff) qed lemma exp_ge_add_one_self [simp]: "1 + x \ exp x" for x :: real proof (cases "0 \ x \ x \ -1") case True then show ?thesis - apply (rule disjE) - apply (simp add: exp_ge_add_one_self_aux) - using exp_ge_zero order_trans real_add_le_0_iff by blast + by (meson exp_ge_add_one_self_aux exp_ge_zero order.trans real_add_le_0_iff) next case False then have ln1: "ln (1 + x) \ x" using ln_one_minus_pos_upper_bound [of "-x"] by simp have "1 + x = exp (ln (1 + x))" using False by auto also have "\ \ exp x" by (simp add: ln1) finally show ?thesis . qed lemma ln_one_plus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "x - x\<^sup>2 \ ln (1 + x)" proof - have "exp (x - x\<^sup>2) = exp x / exp (x\<^sup>2)" by (rule exp_diff) also have "\ \ (1 + x + x\<^sup>2) / exp (x \<^sup>2)" by (metis a b divide_right_mono exp_bound exp_ge_zero) also have "\ \ (1 + x + x\<^sup>2) / (1 + x\<^sup>2)" by (simp add: a divide_left_mono add_pos_nonneg) also from a have "\ \ 1 + x" by (simp add: field_simps add_strict_increasing zero_le_mult_iff) finally have "exp (x - x\<^sup>2) \ 1 + x" . also have "\ = exp (ln (1 + x))" proof - from a have "0 < 1 + x" by auto then show ?thesis by (auto simp only: exp_ln_iff [THEN sym]) qed finally have "exp (x - x\<^sup>2) \ exp (ln (1 + x))" . then show ?thesis by (metis exp_le_cancel_iff) qed lemma ln_one_minus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1 / 2" shows "- x - 2 * x\<^sup>2 \ ln (1 - x)" proof - from b have c: "x < 1" by auto then have "ln (1 - x) = - ln (1 + x / (1 - x))" by (auto simp: ln_inverse [symmetric] field_simps intro: arg_cong [where f=ln]) also have "- (x / (1 - x)) \ \" proof - have "ln (1 + x / (1 - x)) \ x / (1 - x)" using a c by (intro ln_add_one_self_le_self) auto then show ?thesis by auto qed also have "- (x / (1 - x)) = - x / (1 - x)" by auto finally have d: "- x / (1 - x) \ ln (1 - x)" . have "0 < 1 - x" using a b by simp then have e: "- x - 2 * x\<^sup>2 \ - x / (1 - x)" using mult_right_le_one_le[of "x * x" "2 * x"] a b by (simp add: field_simps power2_eq_square) from e d show "- x - 2 * x\<^sup>2 \ ln (1 - x)" by (rule order_trans) qed lemma ln_add_one_self_le_self2: fixes x :: real shows "-1 < x \ ln (1 + x) \ x" by (metis diff_gt_0_iff_gt diff_minus_eq_add exp_ge_add_one_self exp_le_cancel_iff exp_ln minus_less_iff) lemma abs_ln_one_plus_x_minus_x_bound_nonneg: fixes x :: real assumes x: "0 \ x" and x1: "x \ 1" shows "\ln (1 + x) - x\ \ x\<^sup>2" proof - from x have "ln (1 + x) \ x" by (rule ln_add_one_self_le_self) then have "ln (1 + x) - x \ 0" by simp then have "\ln(1 + x) - x\ = - (ln(1 + x) - x)" by (rule abs_of_nonpos) also have "\ = x - ln (1 + x)" by simp also have "\ \ x\<^sup>2" proof - from x x1 have "x - x\<^sup>2 \ ln (1 + x)" by (intro ln_one_plus_pos_lower_bound) then show ?thesis by simp qed finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound_nonpos: fixes x :: real assumes a: "-(1 / 2) \ x" and b: "x \ 0" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof - have *: "- (-x) - 2 * (-x)\<^sup>2 \ ln (1 - (- x))" by (metis a b diff_zero ln_one_minus_pos_lower_bound minus_diff_eq neg_le_iff_le) have "\ln (1 + x) - x\ = x - ln (1 - (- x))" using a ln_add_one_self_le_self2 [of x] by (simp add: abs_if) also have "\ \ 2 * x\<^sup>2" using * by (simp add: algebra_simps) finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound: fixes x :: real assumes "\x\ \ 1 / 2" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof (cases "0 \ x") case True then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonneg assms by fastforce next case False then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonpos assms by auto qed lemma ln_x_over_x_mono: fixes x :: real assumes x: "exp 1 \ x" "x \ y" shows "ln y / y \ ln x / x" proof - note x moreover have "0 < exp (1::real)" by simp ultimately have a: "0 < x" and b: "0 < y" by (fast intro: less_le_trans order_trans)+ have "x * ln y - x * ln x = x * (ln y - ln x)" by (simp add: algebra_simps) also have "\ = x * ln (y / x)" by (simp only: ln_div a b) also have "y / x = (x + (y - x)) / x" by simp also have "\ = 1 + (y - x) / x" using x a by (simp add: field_simps) also have "x * ln (1 + (y - x) / x) \ x * ((y - x) / x)" using x a by (intro mult_left_mono ln_add_one_self_le_self) simp_all also have "\ = y - x" using a by simp also have "\ = (y - x) * ln (exp 1)" by simp also have "\ \ (y - x) * ln x" using a x exp_total of_nat_1 x(1) by (fastforce intro: mult_left_mono) also have "\ = y * ln x - x * ln x" by (rule left_diff_distrib) finally have "x * ln y \ y * ln x" by arith then have "ln y \ (y * ln x) / x" using a by (simp add: field_simps) also have "\ = y * (ln x / x)" by simp finally show ?thesis using b by (simp add: field_simps) qed lemma ln_le_minus_one: "0 < x \ ln x \ x - 1" for x :: real using exp_ge_add_one_self[of "ln x"] by simp corollary ln_diff_le: "0 < x \ 0 < y \ ln x - ln y \ (x - y) / y" for x :: real by (simp add: ln_div [symmetric] diff_divide_distrib ln_le_minus_one) lemma ln_eq_minus_one: fixes x :: real assumes "0 < x" "ln x = x - 1" shows "x = 1" proof - let ?l = "\y. ln y - y + 1" have D: "\x::real. 0 < x \ DERIV ?l x :> (1 / x - 1)" by (auto intro!: derivative_eq_intros) show ?thesis proof (cases rule: linorder_cases) assume "x < 1" from dense[OF \x < 1\] obtain a where "x < a" "a < 1" by blast from \x < a\ have "?l x < ?l a" proof (rule DERIV_pos_imp_increasing) fix y assume "x \ y" "y \ a" with \0 < x\ \a < 1\ have "0 < 1 / y - 1" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ 0 < z" by blast qed also have "\ \ 0" using ln_le_minus_one \0 < x\ \x < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "1 < x" from dense[OF this] obtain a where "1 < a" "a < x" by blast from \a < x\ have "?l x < ?l a" proof (rule DERIV_neg_imp_decreasing) fix y assume "a \ y" "y \ x" with \1 < a\ have "1 / y - 1 < 0" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ z < 0" by blast qed also have "\ \ 0" using ln_le_minus_one \1 < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "x = 1" then show ?thesis by simp qed qed lemma ln_x_over_x_tendsto_0: "((\x::real. ln x / x) \ 0) at_top" proof (rule lhospital_at_top_at_top[where f' = inverse and g' = "\_. 1"]) from eventually_gt_at_top[of "0::real"] show "\\<^sub>F x in at_top. (ln has_real_derivative inverse x) (at x)" by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) qed (use tendsto_inverse_0 in \auto simp: filterlim_ident dest!: tendsto_mono[OF at_top_le_at_infinity]\) lemma exp_ge_one_plus_x_over_n_power_n: assumes "x \ - real n" "n > 0" shows "(1 + x / of_nat n) ^ n \ exp x" proof (cases "x = - of_nat n") case False from assms False have "(1 + x / of_nat n) ^ n = exp (of_nat n * ln (1 + x / of_nat n))" by (subst exp_of_nat_mult, subst exp_ln) (simp_all add: field_simps) also from assms False have "ln (1 + x / real n) \ x / real n" by (intro ln_add_one_self_le_self2) (simp_all add: field_simps) with assms have "exp (of_nat n * ln (1 + x / of_nat n)) \ exp x" by (simp add: field_simps) finally show ?thesis . next case True then show ?thesis by (simp add: zero_power) qed lemma exp_ge_one_minus_x_over_n_power_n: assumes "x \ real n" "n > 0" shows "(1 - x / of_nat n) ^ n \ exp (-x)" using exp_ge_one_plus_x_over_n_power_n[of n "-x"] assms by simp lemma exp_at_bot: "(exp \ (0::real)) at_bot" unfolding tendsto_Zfun_iff proof (rule ZfunI, simp add: eventually_at_bot_dense) fix r :: real assume "0 < r" have "exp x < r" if "x < ln r" for x by (metis \0 < r\ exp_less_mono exp_ln that) then show "\k. \n at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g=ln]) (auto intro: eventually_gt_at_top) lemma lim_exp_minus_1: "((\z::'a. (exp(z) - 1) / z) \ 1) (at 0)" for x :: "'a::{real_normed_field,banach}" proof - have "((\z::'a. exp(z) - 1) has_field_derivative 1) (at 0)" by (intro derivative_eq_intros | simp)+ then show ?thesis by (simp add: Deriv.has_field_derivative_iff) qed lemma ln_at_0: "LIM x at_right 0. ln (x::real) :> at_bot" by (rule filterlim_at_bot_at_right[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto simp: eventually_at_filter) lemma ln_at_top: "LIM x at_top. ln (x::real) :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto intro: eventually_gt_at_top) lemma filtermap_ln_at_top: "filtermap (ln::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of exp] exp_at_top ln_at_top) auto lemma filtermap_exp_at_top: "filtermap (exp::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of ln] exp_at_top ln_at_top) (auto simp: eventually_at_top_dense) lemma filtermap_ln_at_right: "filtermap ln (at_right (0::real)) = at_bot" by (auto intro!: filtermap_fun_inverse[where g="\x. exp x"] ln_at_0 simp: filterlim_at exp_at_bot) lemma tendsto_power_div_exp_0: "((\x. x ^ k / exp x) \ (0::real)) at_top" proof (induct k) case 0 show "((\x. x ^ 0 / exp x) \ (0::real)) at_top" by (simp add: inverse_eq_divide[symmetric]) (metis filterlim_compose[OF tendsto_inverse_0] exp_at_top filterlim_mono at_top_le_at_infinity order_refl) next case (Suc k) show ?case proof (rule lhospital_at_top_at_top) show "eventually (\x. DERIV (\x. x ^ Suc k) x :> (real (Suc k) * x^k)) at_top" by eventually_elim (intro derivative_eq_intros, auto) show "eventually (\x. DERIV exp x :> exp x) at_top" by eventually_elim auto show "eventually (\x. exp x \ 0) at_top" by auto from tendsto_mult[OF tendsto_const Suc, of "real (Suc k)"] show "((\x. real (Suc k) * x ^ k / exp x) \ 0) at_top" by simp qed (rule exp_at_top) qed subsubsection\ A couple of simple bounds\ lemma exp_plus_inverse_exp: fixes x::real shows "2 \ exp x + inverse (exp x)" proof - have "2 \ exp x + exp (-x)" using exp_ge_add_one_self [of x] exp_ge_add_one_self [of "-x"] by linarith then show ?thesis by (simp add: exp_minus) qed lemma real_le_x_sinh: fixes x::real assumes "0 \ x" shows "x \ (exp x - inverse(exp x)) / 2" proof - have *: "exp a - inverse(exp a) - 2*a \ exp b - inverse(exp b) - 2*b" if "a \ b" for a b::real using exp_plus_inverse_exp by (fastforce intro: derivative_eq_intros DERIV_nonneg_imp_nondecreasing [OF that]) show ?thesis using*[OF assms] by simp qed lemma real_le_abs_sinh: fixes x::real shows "abs x \ abs((exp x - inverse(exp x)) / 2)" proof (cases "0 \ x") case True show ?thesis using real_le_x_sinh [OF True] True by (simp add: abs_if) next case False have "-x \ (exp(-x) - inverse(exp(-x))) / 2" by (meson False linear neg_le_0_iff_le real_le_x_sinh) also have "\ \ \(exp x - inverse (exp x)) / 2\" by (metis (no_types, hide_lams) abs_divide abs_le_iff abs_minus_cancel add.inverse_inverse exp_minus minus_diff_eq order_refl) finally show ?thesis using False by linarith qed subsection\The general logarithm\ definition log :: "real \ real \ real" \ \logarithm of \<^term>\x\ to base \<^term>\a\\ where "log a x = ln x / ln a" lemma tendsto_log [tendsto_intros]: "(f \ a) F \ (g \ b) F \ 0 < a \ a \ 1 \ 0 < b \ ((\x. log (f x) (g x)) \ log a b) F" unfolding log_def by (intro tendsto_intros) auto lemma continuous_log: assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" and "f (Lim F (\x. x)) \ 1" and "0 < g (Lim F (\x. x))" shows "continuous F (\x. log (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_log) lemma continuous_at_within_log[continuous_intros]: assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "0 < f a" and "f a \ 1" and "0 < g a" shows "continuous (at a within s) (\x. log (f x) (g x))" using assms unfolding continuous_within by (rule tendsto_log) lemma isCont_log[continuous_intros, simp]: assumes "isCont f a" "isCont g a" "0 < f a" "f a \ 1" "0 < g a" shows "isCont (\x. log (f x) (g x)) a" using assms unfolding continuous_at by (rule tendsto_log) lemma continuous_on_log[continuous_intros]: assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" "\x\s. f x \ 1" "\x\s. 0 < g x" shows "continuous_on s (\x. log (f x) (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_log) lemma powr_one_eq_one [simp]: "1 powr a = 1" by (simp add: powr_def) lemma powr_zero_eq_one [simp]: "x powr 0 = (if x = 0 then 0 else 1)" by (simp add: powr_def) lemma powr_one_gt_zero_iff [simp]: "x powr 1 = x \ 0 \ x" for x :: real by (auto simp: powr_def) declare powr_one_gt_zero_iff [THEN iffD2, simp] lemma powr_diff: fixes w:: "'a::{ln,real_normed_field}" shows "w powr (z1 - z2) = w powr z1 / w powr z2" by (simp add: powr_def algebra_simps exp_diff) lemma powr_mult: "0 \ x \ 0 \ y \ (x * y) powr a = (x powr a) * (y powr a)" for a x y :: real by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left) lemma powr_ge_pzero [simp]: "0 \ x powr y" for x y :: real by (simp add: powr_def) lemma powr_non_neg[simp]: "\a powr x < 0" for a x::real using powr_ge_pzero[of a x] by arith +lemma inverse_powr: "\y::real. 0 \ y \ inverse y powr a = inverse (y powr a)" + by (simp add: exp_minus ln_inverse powr_def) + lemma powr_divide: "\0 \ x; 0 \ y\ \ (x / y) powr a = (x powr a) / (y powr a)" for a b x :: real - apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult) - apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse) - done + by (simp add: divide_inverse powr_mult inverse_powr) lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" for a b x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_add [symmetric] distrib_right) lemma powr_mult_base: "0 \ x \x * x powr y = x powr (1 + y)" for x :: real by (auto simp: powr_add) lemma powr_powr: "(x powr a) powr b = x powr (a * b)" for a b x :: real by (simp add: powr_def) lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" for a b x :: real by (simp add: powr_powr mult.commute) lemma powr_minus: "x powr (- a) = inverse (x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_minus [symmetric]) lemma powr_minus_divide: "x powr (- a) = 1/(x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: divide_inverse powr_minus) lemma divide_powr_uminus: "a / b powr c = a * b powr (- c)" for a b c :: real by (simp add: powr_minus_divide) lemma powr_less_mono: "a < b \ 1 < x \ x powr a < x powr b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel: "x powr a < x powr b \ 1 < x \ a < b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel_iff [simp]: "1 < x \ x powr a < x powr b \ a < b" for a b x :: real by (blast intro: powr_less_cancel powr_less_mono) lemma powr_le_cancel_iff [simp]: "1 < x \ x powr a \ x powr b \ a \ b" for a b x :: real by (simp add: linorder_not_less [symmetric]) lemma powr_realpow: "0 < x \ x powr (real n) = x^n" by (induction n) (simp_all add: ac_simps powr_add) lemma log_ln: "ln x = log (exp(1)) x" by (simp add: log_def) lemma DERIV_log: assumes "x > 0" shows "DERIV (\y. log b y) x :> 1 / (ln b * x)" proof - define lb where "lb = 1 / ln b" moreover have "DERIV (\y. lb * ln y) x :> lb / x" using \x > 0\ by (auto intro!: derivative_eq_intros) ultimately show ?thesis by (simp add: log_def) qed lemmas DERIV_log[THEN DERIV_chain2, derivative_intros] and DERIV_log[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma powr_log_cancel [simp]: "0 < a \ a \ 1 \ 0 < x \ a powr (log a x) = x" by (simp add: powr_def log_def) lemma log_powr_cancel [simp]: "0 < a \ a \ 1 \ log a (a powr y) = y" by (simp add: log_def powr_def) lemma log_mult: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x * y) = log a x + log a y" by (simp add: log_def ln_mult divide_inverse distrib_right) lemma log_eq_div_ln_mult_log: "0 < a \ a \ 1 \ 0 < b \ b \ 1 \ 0 < x \ log a x = (ln b/ln a) * log b x" by (simp add: log_def divide_inverse) text\Base 10 logarithms\ lemma log_base_10_eq1: "0 < x \ log 10 x = (ln (exp 1) / ln 10) * ln x" by (simp add: log_def) lemma log_base_10_eq2: "0 < x \ log 10 x = (log 10 (exp 1)) * ln x" by (simp add: log_def) lemma log_one [simp]: "log a 1 = 0" by (simp add: log_def) lemma log_eq_one [simp]: "0 < a \ a \ 1 \ log a a = 1" by (simp add: log_def) lemma log_inverse: "0 < a \ a \ 1 \ 0 < x \ log a (inverse x) = - log a x" using ln_inverse log_def by auto lemma log_divide: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x/y) = log a x - log a y" by (simp add: log_mult divide_inverse log_inverse) lemma powr_gt_zero [simp]: "0 < x powr a \ x \ 0" for a x :: real by (simp add: powr_def) lemma powr_nonneg_iff[simp]: "a powr x \ 0 \ a = 0" for a x::real by (meson not_less powr_gt_zero) lemma log_add_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x + y = log b (x * b powr y)" and add_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y + log b x = log b (b powr y * x)" and log_minus_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x - y = log b (x * b powr -y)" and minus_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y - log b x = log b (b powr y / x)" by (simp_all add: log_mult log_divide) lemma log_less_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x < log a y \ x < y" using powr_less_cancel_iff [of a] powr_log_cancel [of a x] powr_log_cancel [of a y] by (metis less_eq_real_def less_trans not_le zero_less_one) lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}" proof (rule inj_onI, simp) fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y" show "x = y" proof (cases rule: linorder_cases) assume "x = y" then show ?thesis by simp next assume "x < y" then have "log b x < log b y" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp next assume "y < x" then have "log b y < log b x" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp qed qed lemma log_le_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x \ log a y \ x \ y" by (simp add: linorder_not_less [symmetric]) lemma zero_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 < log a x \ 1 < x" using log_less_cancel_iff[of a 1 x] by simp lemma zero_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 \ log a x \ 1 \ x" using log_le_cancel_iff[of a 1 x] by simp lemma log_less_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 0 \ x < 1" using log_less_cancel_iff[of a x 1] by simp lemma log_le_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 0 \ x \ 1" using log_le_cancel_iff[of a x 1] by simp lemma one_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 < log a x \ a < x" using log_less_cancel_iff[of a a x] by simp lemma one_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 \ log a x \ a \ x" using log_le_cancel_iff[of a a x] by simp lemma log_less_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 1 \ x < a" using log_less_cancel_iff[of a x a] by simp lemma log_le_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 1 \ x \ a" using log_le_cancel_iff[of a x a] by simp lemma le_log_iff: fixes b x y :: real assumes "1 < b" "x > 0" shows "y \ log b x \ b powr y \ x" using assms by (metis less_irrefl less_trans powr_le_cancel_iff powr_log_cancel zero_less_one) lemma less_log_iff: assumes "1 < b" "x > 0" shows "y < log b x \ b powr y < x" by (metis assms dual_order.strict_trans less_irrefl powr_less_cancel_iff powr_log_cancel zero_less_one) lemma assumes "1 < b" "x > 0" shows log_less_iff: "log b x < y \ x < b powr y" and log_le_iff: "log b x \ y \ x \ b powr y" using le_log_iff[OF assms, of y] less_log_iff[OF assms, of y] by auto lemmas powr_le_iff = le_log_iff[symmetric] and powr_less_iff = less_log_iff[symmetric] and less_powr_iff = log_less_iff[symmetric] and le_powr_iff = log_le_iff[symmetric] lemma le_log_of_power: assumes "b ^ n \ m" "1 < b" shows "n \ log b m" proof - from assms have "0 < m" by (metis less_trans zero_less_power less_le_trans zero_less_one) thus ?thesis using assms by (simp add: le_log_iff powr_realpow) qed lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" for m n :: nat using le_log_of_power[of 2] by simp lemma log_of_power_le: "\ m \ b ^ n; b > 1; m > 0 \ \ log b (real m) \ n" by (simp add: log_le_iff powr_realpow) lemma log2_of_power_le: "\ m \ 2 ^ n; m > 0 \ \ log 2 m \ n" for m n :: nat using log_of_power_le[of _ 2] by simp lemma log_of_power_less: "\ m < b ^ n; b > 1; m > 0 \ \ log b (real m) < n" by (simp add: log_less_iff powr_realpow) lemma log2_of_power_less: "\ m < 2 ^ n; m > 0 \ \ log 2 m < n" for m n :: nat using log_of_power_less[of _ 2] by simp lemma less_log_of_power: assumes "b ^ n < m" "1 < b" shows "n < log b m" proof - have "0 < m" by (metis assms less_trans zero_less_power zero_less_one) thus ?thesis using assms by (simp add: less_log_iff powr_realpow) qed lemma less_log2_of_power: "2 ^ n < m \ n < log 2 m" for m n :: nat using less_log_of_power[of 2] by simp lemma gr_one_powr[simp]: fixes x y :: real shows "\ x > 1; y > 0 \ \ 1 < x powr y" by(simp add: less_powr_iff) lemma log_pow_cancel [simp]: "a > 0 \ a \ 1 \ log a (a ^ b) = b" by (simp add: ln_realpow log_def) lemma floor_log_eq_powr_iff: "x > 0 \ b > 1 \ \log b x\ = k \ b powr k \ x \ x < b powr (k + 1)" by (auto simp: floor_eq_iff powr_le_iff less_powr_iff) lemma floor_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ floor (log b (real k)) = n \ b^n \ k \ k < b^(n+1)" by (auto simp: floor_log_eq_powr_iff powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma floor_log_nat_eq_if: fixes b n k :: nat assumes "b^n \ k" "k < b^(n+1)" "b \ 2" shows "floor (log b (real k)) = n" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: floor_log_nat_eq_powr_iff) qed lemma ceiling_log_eq_powr_iff: "\ x > 0; b > 1 \ \ \log b x\ = int k + 1 \ b powr k < x \ x \ b powr (k + 1)" by (auto simp: ceiling_eq_iff powr_less_iff le_powr_iff) lemma ceiling_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ ceiling (log b (real k)) = int n + 1 \ (b^n < k \ k \ b^(n+1))" using ceiling_log_eq_powr_iff by (auto simp: powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma ceiling_log_nat_eq_if: fixes b n k :: nat assumes "b^n < k" "k \ b^(n+1)" "b \ 2" shows "ceiling (log b (real k)) = int n + 1" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: ceiling_log_nat_eq_powr_iff) qed lemma floor_log2_div2: fixes n :: nat assumes "n \ 2" shows "floor(log 2 n) = floor(log 2 (n div 2)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "n div 2" assume "n\2" hence "1 \ ?m" using assms by arith then obtain i where i: "2 ^ i \ ?m" "?m < 2 ^ (i + 1)" using ex_power_ivl1[of 2 ?m] by auto have "2^(i+1) \ 2*?m" using i(1) by simp also have "2*?m \ n" by arith finally have *: "2^(i+1) \ \" . have "n < 2^(i+1+1)" using i(2) by simp from floor_log_nat_eq_if[OF * this] floor_log_nat_eq_if[OF i] show ?thesis by simp qed lemma ceiling_log2_div2: assumes "n \ 2" shows "ceiling(log 2 (real n)) = ceiling(log 2 ((n-1) div 2 + 1)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "(n-1) div 2 + 1" assume "n\2" hence "2 \ ?m" using assms by arith then obtain i where i: "2 ^ i < ?m" "?m \ 2 ^ (i + 1)" using ex_power_ivl2[of 2 ?m] by auto have "n \ 2*?m" by arith also have "2*?m \ 2 ^ ((i+1)+1)" using i(2) by simp finally have *: "n \ \" . have "2^(i+1) < n" using i(1) by (auto simp: less_Suc_eq_0_disj) from ceiling_log_nat_eq_if[OF this *] ceiling_log_nat_eq_if[OF i] show ?thesis by simp qed lemma powr_real_of_int: "x > 0 \ x powr real_of_int n = (if n \ 0 then x ^ nat n else inverse (x ^ nat (- n)))" using powr_realpow[of x "nat n"] powr_realpow[of x "nat (-n)"] by (auto simp: field_simps powr_minus) lemma powr_numeral [simp]: "0 \ x \ x powr (numeral n :: real) = x ^ (numeral n)" by (metis less_le power_zero_numeral powr_0 of_nat_numeral powr_realpow) lemma powr_int: assumes "x > 0" shows "x powr i = (if i \ 0 then x ^ nat i else 1 / x ^ nat (-i))" proof (cases "i < 0") case True have r: "x powr i = 1 / x powr (- i)" by (simp add: powr_minus field_simps) show ?thesis using \i < 0\ \x > 0\ by (simp add: r field_simps powr_realpow[symmetric]) next case False then show ?thesis by (simp add: assms powr_realpow[symmetric]) qed definition powr_real :: "real \ real \ real" where [code_abbrev, simp]: "powr_real = Transcendental.powr" lemma compute_powr_real [code]: "powr_real b i = (if b \ 0 then Code.abort (STR ''powr_real with nonpositive base'') (\_. powr_real b i) else if \i\ = i then (if 0 \ i then b ^ nat \i\ else 1 / b ^ nat \- i\) else Code.abort (STR ''powr_real with non-integer exponent'') (\_. powr_real b i))" for b i :: real by (auto simp: powr_int) lemma powr_one: "0 \ x \ x powr 1 = x" for x :: real using powr_realpow [of x 1] by simp lemma powr_neg_one: "0 < x \ x powr - 1 = 1 / x" for x :: real using powr_int [of x "- 1"] by simp lemma powr_neg_numeral: "0 < x \ x powr - numeral n = 1 / x ^ numeral n" for x :: real using powr_int [of x "- numeral n"] by simp lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr) lemma ln_powr: "x \ 0 \ ln (x powr y) = y * ln x" for x :: real by (simp add: powr_def) lemma ln_root: "n > 0 \ b > 0 \ ln (root n b) = ln b / n" by (simp add: root_powr_inverse ln_powr) lemma ln_sqrt: "0 < x \ ln (sqrt x) = ln x / 2" by (simp add: ln_powr ln_powr[symmetric] mult.commute) lemma log_root: "n > 0 \ a > 0 \ log b (root n a) = log b a / n" by (simp add: log_def ln_root) lemma log_powr: "x \ 0 \ log b (x powr y) = y * log b x" by (simp add: log_def ln_powr) (* [simp] is not worth it, interferes with some proofs *) lemma log_nat_power: "0 < x \ log b (x^n) = real n * log b x" by (simp add: log_powr powr_realpow [symmetric]) lemma log_of_power_eq: assumes "m = b ^ n" "b > 1" shows "n = log b (real m)" proof - have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) also have "\ = log b m" using assms by simp finally show ?thesis . qed lemma log2_of_power_eq: "m = 2 ^ n \ n = log 2 m" for m n :: nat using log_of_power_eq[of _ 2] by simp lemma log_base_change: "0 < a \ a \ 1 \ log b x = log a x / log a b" by (simp add: log_def) lemma log_base_pow: "0 < a \ log (a ^ n) x = log a x / n" by (simp add: log_def ln_realpow) lemma log_base_powr: "a \ 0 \ log (a powr b) x = log a x / b" by (simp add: log_def ln_powr) lemma log_base_root: "n > 0 \ b > 0 \ log (root n b) x = n * (log b x)" by (simp add: log_def ln_root) lemma ln_bound: "0 < x \ ln x \ x" for x :: real using ln_le_minus_one by force lemma powr_mono: fixes x :: real assumes "a \ b" and "1 \ x" shows "x powr a \ x powr b" using assms less_eq_real_def by auto lemma ge_one_powr_ge_zero: "1 \ x \ 0 \ a \ 1 \ x powr a" for x :: real using powr_mono by fastforce lemma powr_less_mono2: "0 < a \ 0 \ x \ x < y \ x powr a < y powr a" for x :: real by (simp add: powr_def) lemma powr_less_mono2_neg: "a < 0 \ 0 < x \ x < y \ y powr a < x powr a" for x :: real by (simp add: powr_def) lemma powr_mono2: "x powr a \ y powr a" if "0 \ a" "0 \ x" "x \ y" for x :: real using less_eq_real_def powr_less_mono2 that by auto lemma powr_le1: "0 \ a \ 0 \ x \ x \ 1 \ x powr a \ 1" for x :: real using powr_mono2 by fastforce lemma powr_mono2': fixes a x y :: real assumes "a \ 0" "x > 0" "x \ y" shows "x powr a \ y powr a" proof - from assms have "x powr - a \ y powr - a" by (intro powr_mono2) simp_all with assms show ?thesis by (auto simp: powr_minus field_simps) qed lemma powr_mono_both: fixes x :: real assumes "0 \ a" "a \ b" "1 \ x" "x \ y" shows "x powr a \ y powr b" by (meson assms order.trans powr_mono powr_mono2 zero_le_one) lemma powr_inj: "0 < a \ a \ 1 \ a powr x = a powr y \ x = y" for x :: real unfolding powr_def exp_inj_iff by simp lemma powr_half_sqrt: "0 \ x \ x powr (1/2) = sqrt x" by (simp add: powr_def root_powr_inverse sqrt_def) lemma square_powr_half [simp]: fixes x::real shows "x\<^sup>2 powr (1/2) = \x\" by (simp add: powr_half_sqrt) lemma ln_powr_bound: "1 \ x \ 0 < a \ ln x \ (x powr a) / a" for x :: real by (metis exp_gt_zero linear ln_eq_zero_iff ln_exp ln_less_self ln_powr mult.commute mult_imp_le_div_pos not_less powr_gt_zero) lemma ln_powr_bound2: fixes x :: real assumes "1 < x" and "0 < a" shows "(ln x) powr a \ (a powr a) * x" proof - from assms have "ln x \ (x powr (1 / a)) / (1 / a)" by (metis less_eq_real_def ln_powr_bound zero_less_divide_1_iff) also have "\ = a * (x powr (1 / a))" by simp finally have "(ln x) powr a \ (a * (x powr (1 / a))) powr a" by (metis assms less_imp_le ln_gt_zero powr_mono2) also have "\ = (a powr a) * ((x powr (1 / a)) powr a)" using assms powr_mult by auto also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" by (rule powr_powr) also have "\ = x" using assms by auto finally show ?thesis . qed lemma tendsto_powr: fixes a b :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0" shows "((\x. f x powr g x) \ a powr b) F" unfolding powr_def proof (rule filterlim_If) from f show "((\x. 0) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x = 0}))" by simp (auto simp: filterlim_iff eventually_inf_principal elim: eventually_mono dest: t1_space_nhds) from f g a show "((\x. exp (g x * ln (f x))) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x \ 0}))" by (auto intro!: tendsto_intros intro: tendsto_mono inf_le1) qed lemma tendsto_powr'[tendsto_intros]: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0 \ (b > 0 \ eventually (\x. f x \ 0) F)" shows "((\x. f x powr g x) \ a powr b) F" proof - from a consider "a \ 0" | "a = 0" "b > 0" "eventually (\x. f x \ 0) F" by auto then show ?thesis proof cases case 1 with f g show ?thesis by (rule tendsto_powr) next case 2 have "((\x. if f x = 0 then 0 else exp (g x * ln (f x))) \ 0) F" proof (intro filterlim_If) have "filterlim f (principal {0<..}) (inf F (principal {z. f z \ 0}))" using \eventually (\x. f x \ 0) F\ by (auto simp: filterlim_iff eventually_inf_principal eventually_principal elim: eventually_mono) moreover have "filterlim f (nhds a) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ f]) simp_all ultimately have f: "filterlim f (at_right 0) (inf F (principal {x. f x \ 0}))" by (simp add: at_within_def filterlim_inf \a = 0\) have g: "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all show "((\x. exp (g x * ln (f x))) \ 0) (inf F (principal {x. f x \ 0}))" by (rule filterlim_compose[OF exp_at_bot] filterlim_tendsto_pos_mult_at_bot filterlim_compose[OF ln_at_0] f g \b > 0\)+ qed simp_all with \a = 0\ show ?thesis by (simp add: powr_def) qed qed lemma continuous_powr: assumes "continuous F f" and "continuous F g" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) powr (g x :: real))" using assms unfolding continuous_def by (rule tendsto_powr) lemma continuous_at_within_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "f a \ 0" shows "continuous (at a within s) (\x. (f x) powr (g x))" using assms unfolding continuous_within by (rule tendsto_powr) lemma isCont_powr[continuous_intros, simp]: fixes f g :: "_ \ real" assumes "isCont f a" "isCont g a" "f a \ 0" shows "isCont (\x. (f x) powr g x) a" using assms unfolding continuous_at by (rule tendsto_powr) lemma continuous_on_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0" shows "continuous_on s (\x. (f x) powr (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_powr) lemma tendsto_powr2: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and "\\<^sub>F x in F. 0 \ f x" and b: "0 < b" shows "((\x. f x powr g x) \ a powr b) F" using tendsto_powr'[of f a F g b] assms by auto lemma has_derivative_powr[derivative_intros]: assumes g[derivative_intros]: "(g has_derivative g') (at x within X)" and f[derivative_intros]:"(f has_derivative f') (at x within X)" assumes pos: "0 < g x" and "x \ X" shows "((\x. g x powr f x::real) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" proof - have "\\<^sub>F x in at x within X. g x > 0" by (rule order_tendstoD[OF _ pos]) (rule has_derivative_continuous[OF g, unfolded continuous_within]) then obtain d where "d > 0" and pos': "\x'. x' \ X \ dist x' x < d \ 0 < g x'" using pos unfolding eventually_at by force have "((\x. exp (f x * ln (g x))) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" using pos by (auto intro!: derivative_eq_intros simp: field_split_simps powr_def) then show ?thesis by (rule has_derivative_transform_within[OF _ \d > 0\ \x \ X\]) (auto simp: powr_def dest: pos') qed lemma DERIV_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" and f: "DERIV f x :> r" shows "DERIV (\x. g x powr f x) x :> (g x powr f x) * (r * ln (g x) + m * f x / g x)" using assms by (auto intro!: derivative_eq_intros ext simp: has_field_derivative_def algebra_simps) lemma DERIV_fun_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" shows "DERIV (\x. (g x) powr r) x :> r * (g x) powr (r - of_nat 1) * m" using DERIV_powr[OF g pos DERIV_const, of r] pos by (simp add: powr_diff field_simps) lemma has_real_derivative_powr: assumes "z > 0" shows "((\z. z powr r) has_real_derivative r * z powr (r - 1)) (at z)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "eventually (\z. z \ 0) (nhds z)" by (intro t1_space_nhds) auto then show "eventually (\z. z powr r = exp (r * ln z)) (nhds z)" unfolding powr_def by eventually_elim simp from assms show "((\z. exp (r * ln z)) has_real_derivative r * z powr (r - 1)) (at z)" by (auto intro!: derivative_eq_intros simp: powr_def field_simps exp_diff) qed declare has_real_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma tendsto_zero_powrI: assumes "(f \ (0::real)) F" "(g \ b) F" "\\<^sub>F x in F. 0 \ f x" "0 < b" shows "((\x. f x powr g x) \ 0) F" using tendsto_powr2[OF assms] by simp lemma continuous_on_powr': fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0 \ (f x = 0 \ g x > 0)" shows "continuous_on s (\x. (f x) powr (g x))" unfolding continuous_on_def proof fix x assume x: "x \ s" from assms x show "((\x. f x powr g x) \ f x powr g x) (at x within s)" proof (cases "f x = 0") case True from assms(3) have "eventually (\x. f x \ 0) (at x within s)" by (auto simp: at_within_def eventually_inf_principal) with True x assms show ?thesis by (auto intro!: tendsto_zero_powrI[of f _ g "g x"] simp: continuous_on_def) next case False with assms x show ?thesis by (auto intro!: tendsto_powr' simp: continuous_on_def) qed qed lemma tendsto_neg_powr: assumes "s < 0" and f: "LIM x F. f x :> at_top" shows "((\x. f x powr s) \ (0::real)) F" proof - have "((\x. exp (s * ln (f x))) \ (0::real)) F" (is "?X") by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_top] filterlim_tendsto_neg_mult_at_bot assms) also have "?X \ ((\x. f x powr s) \ (0::real)) F" using f filterlim_at_top_dense[of f F] by (intro filterlim_cong[OF refl refl]) (auto simp: neq_iff powr_def elim: eventually_mono) finally show ?thesis . qed lemma tendsto_exp_limit_at_right: "((\y. (1 + x * y) powr (1 / y)) \ exp x) (at_right 0)" for x :: real proof (cases "x = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + x * y)::real) has_real_derivative 1 * x) (at 0)" by (auto intro!: derivative_eq_intros) then have "((\y. ln (1 + x * y) / y) \ x) (at 0)" by (auto simp: has_field_derivative_def field_has_derivative_at) then have *: "((\y. exp (ln (1 + x * y) / y)) \ exp x) (at 0)" by (rule tendsto_intros) then show ?thesis proof (rule filterlim_mono_eventually) show "eventually (\xa. exp (ln (1 + x * xa) / xa) = (1 + x * xa) powr (1 / xa)) (at_right 0)" unfolding eventually_at_right[OF zero_less_one] using False by (intro exI[of _ "1 / \x\"]) (auto simp: field_simps powr_def abs_if add_nonneg_eq_0_iff) qed (simp_all add: at_eq_sup_left_right) qed lemma tendsto_exp_limit_at_top: "((\y. (1 + x / y) powr y) \ exp x) at_top" for x :: real by (simp add: filterlim_at_top_to_right inverse_eq_divide tendsto_exp_limit_at_right) lemma tendsto_exp_limit_sequentially: "(\n. (1 + x / n) ^ n) \ exp x" for x :: real proof (rule filterlim_mono_eventually) from reals_Archimedean2 [of "\x\"] obtain n :: nat where *: "real n > \x\" .. then have "eventually (\n :: nat. 0 < 1 + x / real n) at_top" by (intro eventually_sequentiallyI [of n]) (auto simp: field_split_simps) then show "eventually (\n. (1 + x / n) powr n = (1 + x / n) ^ n) at_top" by (rule eventually_mono) (erule powr_realpow) show "(\n. (1 + x / real n) powr real n) \ exp x" by (rule filterlim_compose [OF tendsto_exp_limit_at_top filterlim_real_sequentially]) qed auto subsection \Sine and Cosine\ definition sin_coeff :: "nat \ real" where "sin_coeff = (\n. if even n then 0 else (- 1) ^ ((n - Suc 0) div 2) / (fact n))" definition cos_coeff :: "nat \ real" where "cos_coeff = (\n. if even n then ((- 1) ^ (n div 2)) / (fact n) else 0)" definition sin :: "'a \ 'a::{real_normed_algebra_1,banach}" where "sin = (\x. \n. sin_coeff n *\<^sub>R x^n)" definition cos :: "'a \ 'a::{real_normed_algebra_1,banach}" where "cos = (\x. \n. cos_coeff n *\<^sub>R x^n)" lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" unfolding sin_coeff_def by simp lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" unfolding cos_coeff_def by simp lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) (auto elim: oddE) lemma summable_norm_sin: "summable (\n. norm (sin_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" - unfolding sin_coeff_def - apply (rule summable_comparison_test [OF _ summable_norm_exp [where x=x]]) - apply (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) - done +proof (rule summable_comparison_test [OF _ summable_norm_exp]) + show "\N. \n\N. norm (norm (sin_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" + unfolding sin_coeff_def + by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +qed lemma summable_norm_cos: "summable (\n. norm (cos_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" - unfolding cos_coeff_def - apply (rule summable_comparison_test [OF _ summable_norm_exp [where x=x]]) - apply (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) - done +proof (rule summable_comparison_test [OF _ summable_norm_exp]) + show "\N. \n\N. norm (norm (cos_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" + unfolding cos_coeff_def + by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +qed + lemma sin_converges: "(\n. sin_coeff n *\<^sub>R x^n) sums sin x" unfolding sin_def by (metis (full_types) summable_norm_cancel summable_norm_sin summable_sums) lemma cos_converges: "(\n. cos_coeff n *\<^sub>R x^n) sums cos x" unfolding cos_def by (metis (full_types) summable_norm_cancel summable_norm_cos summable_sums) lemma sin_of_real: "sin (of_real x) = of_real (sin x)" for x :: real proof - have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) = (\n. sin_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (sin_coeff n *\<^sub>R x^n) = sin_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (sin (of_real x))" by (rule sin_converges) finally have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) sums (sin (of_real x))" . then show ?thesis - using sums_unique2 sums_of_real [OF sin_converges] - by blast + using sums_unique2 sums_of_real [OF sin_converges] by blast qed corollary sin_in_Reals [simp]: "z \ \ \ sin z \ \" by (metis Reals_cases Reals_of_real sin_of_real) lemma cos_of_real: "cos (of_real x) = of_real (cos x)" for x :: real proof - have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) = (\n. cos_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (cos_coeff n *\<^sub>R x^n) = cos_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (cos (of_real x))" by (rule cos_converges) finally have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) sums (cos (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF cos_converges] by blast qed corollary cos_in_Reals [simp]: "z \ \ \ cos z \ \" by (metis Reals_cases Reals_of_real cos_of_real) lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" by (simp add: diffs_def sin_coeff_Suc del: of_nat_Suc) lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" by (simp add: diffs_def cos_coeff_Suc del: of_nat_Suc) lemma sin_int_times_real: "sin (of_int m * of_real x) = of_real (sin (of_int m * x))" by (metis sin_of_real of_real_mult of_real_of_int_eq) lemma cos_int_times_real: "cos (of_int m * of_real x) = of_real (cos (of_int m * x))" by (metis cos_of_real of_real_mult of_real_of_int_eq) text \Now at last we can get the derivatives of exp, sin and cos.\ lemma DERIV_sin [simp]: "DERIV sin x :> cos x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_sin[THEN DERIV_chain2, derivative_intros] and DERIV_sin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_sin[derivative_intros] = DERIV_sin[THEN DERIV_compose_FDERIV] lemma DERIV_cos [simp]: "DERIV cos x :> - sin x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_minus suminf_minus diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_cos[THEN DERIV_chain2, derivative_intros] and DERIV_cos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_cos[derivative_intros] = DERIV_cos[THEN DERIV_compose_FDERIV] lemma isCont_sin: "isCont sin x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_sin [THEN DERIV_isCont]) lemma continuous_on_sin_real: "continuous_on {a..b} sin" for a::real using continuous_at_imp_continuous_on isCont_sin by blast lemma isCont_cos: "isCont cos x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cos [THEN DERIV_isCont]) lemma continuous_on_cos_real: "continuous_on {a..b} cos" for a::real using continuous_at_imp_continuous_on isCont_cos by blast + +context + fixes f :: "'a::t2_space \ 'b::{real_normed_field,banach}" +begin + lemma isCont_sin' [simp]: "isCont f a \ isCont (\x. sin (f x)) a" - for f :: "_ \ 'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_sin]) -(* FIXME a context for f would be better *) - lemma isCont_cos' [simp]: "isCont f a \ isCont (\x. cos (f x)) a" - for f :: "_ \ 'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_cos]) lemma tendsto_sin [tendsto_intros]: "(f \ a) F \ ((\x. sin (f x)) \ sin a) F" - for f :: "_ \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_sin]) lemma tendsto_cos [tendsto_intros]: "(f \ a) F \ ((\x. cos (f x)) \ cos a) F" - for f :: "_ \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_cos]) lemma continuous_sin [continuous_intros]: "continuous F f \ continuous F (\x. sin (f x))" - for f :: "_ \ 'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_sin) lemma continuous_on_sin [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sin (f x))" - for f :: "_ \ 'a::{real_normed_field,banach}" unfolding continuous_on_def by (auto intro: tendsto_sin) +lemma continuous_cos [continuous_intros]: "continuous F f \ continuous F (\x. cos (f x))" + unfolding continuous_def by (rule tendsto_cos) + +lemma continuous_on_cos [continuous_intros]: "continuous_on s f \ continuous_on s (\x. cos (f x))" + unfolding continuous_on_def by (auto intro: tendsto_cos) + +end + lemma continuous_within_sin: "continuous (at z within s) sin" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_sin) -lemma continuous_cos [continuous_intros]: "continuous F f \ continuous F (\x. cos (f x))" - for f :: "_ \ 'a::{real_normed_field,banach}" - unfolding continuous_def by (rule tendsto_cos) - -lemma continuous_on_cos [continuous_intros]: "continuous_on s f \ continuous_on s (\x. cos (f x))" - for f :: "_ \ 'a::{real_normed_field,banach}" - unfolding continuous_on_def by (auto intro: tendsto_cos) - lemma continuous_within_cos: "continuous (at z within s) cos" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_cos) subsection \Properties of Sine and Cosine\ lemma sin_zero [simp]: "sin 0 = 0" by (simp add: sin_def sin_coeff_def scaleR_conv_of_real) lemma cos_zero [simp]: "cos 0 = 1" by (simp add: cos_def cos_coeff_def scaleR_conv_of_real) lemma DERIV_fun_sin: "DERIV g x :> m \ DERIV (\x. sin (g x)) x :> cos (g x) * m" - by (auto intro!: derivative_intros) + by (fact derivative_intros) lemma DERIV_fun_cos: "DERIV g x :> m \ DERIV (\x. cos(g x)) x :> - sin (g x) * m" - by (auto intro!: derivative_eq_intros) + by (fact derivative_intros) subsection \Deriving the Addition Formulas\ text \The product of two cosine series.\ lemma cos_x_cos_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (cos x * cos y)" proof - have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p - n)) = (if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p - n) else 0)" if "n \ p" for n p :: nat proof - from that have *: "even n \ even p \ (-1) ^ (n div 2) * (-1) ^ ((p - n) div 2) = (-1 :: real) ^ (p div 2)" by (metis div_add power_add le_add_diff_inverse odd_add) with that show ?thesis by (auto simp: algebra_simps cos_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (cos_coeff n *\<^sub>R x^n) * (cos_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (cos x * cos y)" using summable_norm_cos by (auto simp: cos_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed text \The product of two sine series.\ lemma sin_x_sin_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (sin x * sin y)" proof - have "(sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) = (if even p \ odd n then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat proof - have "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))" if np: "odd n" "even p" proof - + have "p > 0" + using \n \ p\ neq0_conv that(1) by blast + then have \
: "(- 1::real) ^ (p div 2 - Suc 0) = - ((- 1) ^ (p div 2))" + using \even p\ by (auto simp add: dvd_def power_eq_if) from \n \ p\ np have *: "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \ p" by arith+ have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0" by simp - with \n \ p\ np * show ?thesis - apply (simp add: power_add [symmetric] div_add [symmetric] del: div_add) - apply (metis (no_types) One_nat_def Suc_1 le_div_geq minus_minus - mult.left_neutral mult_minus_left power.simps(2) zero_less_Suc) - done + with \n \ p\ np \
* show ?thesis + by (simp add: flip: div_add power_add) qed then show ?thesis using \n\p\ by (auto simp: algebra_simps sin_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (sin_coeff n *\<^sub>R x^n) * (sin_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (sin x * sin y)" using summable_norm_sin by (auto simp: sin_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed lemma sums_cos_x_plus_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums cos (x + y)" proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = cos_coeff p *\<^sub>R ((x + y) ^ p)" for p :: nat proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then \n\p. ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" by simp also have "\ = (if even p then of_real ((-1) ^ (p div 2) / (fact p)) * (\n\p. (p choose n) *\<^sub>R (x^n) * y^(p-n)) else 0)" by (auto simp: sum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide) also have "\ = cos_coeff p *\<^sub>R ((x + y) ^ p)" by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost) finally show ?thesis . qed then have "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. cos_coeff p *\<^sub>R ((x+y)^p))" by simp also have "\ sums cos (x + y)" by (rule cos_converges) finally show ?thesis . qed theorem cos_add: fixes x :: "'a::{real_normed_field,banach}" shows "cos (x + y) = cos x * cos y - sin x * sin y" proof - have "(if even p \ even n then ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) - (if even p \ odd n then - ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat by simp then have "(\p. \n\p. (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)) sums (cos x * cos y - sin x * sin y)" using sums_diff [OF cos_x_cos_y [of x y] sin_x_sin_y [of x y]] by (simp add: sum_subtractf [symmetric]) then show ?thesis by (blast intro: sums_cos_x_plus_y sums_unique2) qed lemma sin_minus_converges: "(\n. - (sin_coeff n *\<^sub>R (-x)^n)) sums sin x" proof - have [simp]: "\n. - (sin_coeff n *\<^sub>R (-x)^n) = (sin_coeff n *\<^sub>R x^n)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: sin_def summable_norm_sin [THEN summable_norm_cancel, THEN summable_sums]) qed lemma sin_minus [simp]: "sin (- x) = - sin x" for x :: "'a::{real_normed_algebra_1,banach}" using sin_minus_converges [of x] by (auto simp: sin_def summable_norm_sin [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma cos_minus_converges: "(\n. (cos_coeff n *\<^sub>R (-x)^n)) sums cos x" proof - have [simp]: "\n. (cos_coeff n *\<^sub>R (-x)^n) = (cos_coeff n *\<^sub>R x^n)" by (auto simp: Transcendental.cos_coeff_def elim!: evenE) show ?thesis by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel, THEN summable_sums]) qed lemma cos_minus [simp]: "cos (-x) = cos x" for x :: "'a::{real_normed_algebra_1,banach}" using cos_minus_converges [of x] by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma sin_cos_squared_add [simp]: "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "-x"] by (simp add: power2_eq_square algebra_simps) lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" by (subst add.commute, rule sin_cos_squared_add) lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" for x :: "'a::{real_normed_field,banach}" using sin_cos_squared_add2 [unfolded power2_eq_square] . lemma sin_squared_eq: "(sin x)\<^sup>2 = 1 - (cos x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add) lemma cos_squared_eq: "(cos x)\<^sup>2 = 1 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add2) lemma abs_sin_le_one [simp]: "\sin x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: sin_squared_eq) lemma sin_ge_minus_one [simp]: "- 1 \ sin x" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma sin_le_one [simp]: "sin x \ 1" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma abs_cos_le_one [simp]: "\cos x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: cos_squared_eq) lemma cos_ge_minus_one [simp]: "- 1 \ cos x" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_le_one [simp]: "cos x \ 1" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "- y"] by simp lemma cos_double: "cos(2*x) = (cos x)\<^sup>2 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" using cos_add [where x=x and y=x] by (simp add: power2_eq_square) lemma sin_cos_le1: "\sin x * sin y + cos x * cos y\ \ 1" for x :: real using cos_diff [of x y] by (metis abs_cos_le_one add.commute) lemma DERIV_fun_pow: "DERIV g x :> m \ DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" by (auto intro!: derivative_eq_intros simp:) lemma DERIV_fun_exp: "DERIV g x :> m \ DERIV (\x. exp (g x)) x :> exp (g x) * m" by (auto intro!: derivative_intros) subsection \The Constant Pi\ definition pi :: real where "pi = 2 * (THE x. 0 \ x \ x \ 2 \ cos x = 0)" text \Show that there's a least positive \<^term>\x\ with \<^term>\cos x = 0\; hence define pi.\ lemma sin_paired: "(\n. (- 1) ^ n / (fact (2 * n + 1)) * x ^ (2 * n + 1)) sums sin x" for x :: real proof - have "(\n. \k = n*2..n. 0 < ?f n" proof fix n :: nat let ?k2 = "real (Suc (Suc (4 * n)))" let ?k3 = "real (Suc (Suc (Suc (4 * n))))" have "x * x < ?k2 * ?k3" using assms by (intro mult_strict_mono', simp_all) then have "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)" by (intro mult_strict_right_mono zero_less_power \0 < x\) then show "0 < ?f n" by (simp add: ac_simps divide_less_eq) qed have sums: "?f sums sin x" by (rule sin_paired [THEN sums_group]) simp show "0 < sin x" unfolding sums_unique [OF sums] using sums_summable [OF sums] pos by (rule suminf_pos) qed lemma cos_double_less_one: "0 < x \ x < 2 \ cos (2 * x) < 1" for x :: real using sin_gt_zero_02 [where x = x] by (auto simp: cos_squared_eq cos_double) lemma cos_paired: "(\n. (- 1) ^ n / (fact (2 * n)) * x ^ (2 * n)) sums cos x" for x :: real proof - have "(\n. \k = n * 2.. real" assumes f: "summable f" and fplus: "\d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc (Suc 0) * d) + 1))" shows "sum f {..n. \n = n * Suc (Suc 0)..n. f (n + k))" proof (rule sums_group) show "(\n. f (n + k)) sums (\n. f (n + k))" by (simp add: f summable_iff_shift summable_sums) qed auto with fplus have "0 < (\n. f (n + k))" apply (simp add: add.commute) apply (metis (no_types, lifting) suminf_pos summable_def sums_unique) done then show ?thesis by (simp add: f suminf_minus_initial_segment) qed lemma cos_two_less_zero [simp]: "cos 2 < (0::real)" proof - note fact_Suc [simp del] from sums_minus [OF cos_paired] have *: "(\n. - ((- 1) ^ n * 2 ^ (2 * n) / fact (2 * n))) sums - cos (2::real)" by simp then have sm: "summable (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_summable) have "0 < (\nnn. - ((- 1) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" proof - { fix d let ?six4d = "Suc (Suc (Suc (Suc (Suc (Suc (4 * d))))))" have "(4::real) * (fact (?six4d)) < (Suc (Suc (?six4d)) * fact (Suc (?six4d)))" unfolding of_nat_mult by (rule mult_strict_mono) (simp_all add: fact_less_mono) then have "(4::real) * (fact (?six4d)) < (fact (Suc (Suc (?six4d))))" by (simp only: fact_Suc [of "Suc (?six4d)"] of_nat_mult of_nat_fact) then have "(4::real) * inverse (fact (Suc (Suc (?six4d)))) < inverse (fact (?six4d))" by (simp add: inverse_eq_divide less_divide_eq) } then show ?thesis by (force intro!: sum_pos_lt_pair [OF sm] simp add: divide_inverse algebra_simps) qed ultimately have "0 < (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule order_less_trans) moreover from * have "- cos 2 = (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_unique) ultimately have "(0::real) < - cos 2" by simp then show ?thesis by simp qed lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] lemma cos_is_zero: "\!x::real. 0 \ x \ x \ 2 \ cos x = 0" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ 2 \ cos x = 0" by (rule IVT2) simp_all next fix a b :: real assume ab: "0 \ a \ a \ 2 \ cos a = 0" "0 \ b \ b \ 2 \ cos b = 0" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero_02) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero_02) qed auto qed lemma pi_half: "pi/2 = (THE x. 0 \ x \ x \ 2 \ cos x = 0)" by (simp add: pi_def) lemma cos_pi_half [simp]: "cos (pi/2) = 0" by (simp add: pi_half cos_is_zero [THEN theI']) lemma cos_of_real_pi_half [simp]: "cos ((of_real pi/2) :: 'a) = 0" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" by (metis cos_pi_half cos_of_real eq_numeral_simps(4) nonzero_of_real_divide of_real_0 of_real_numeral) lemma pi_half_gt_zero [simp]: "0 < pi/2" proof - have "0 \ pi/2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_zero less_eq_real_def one_neq_zero) qed lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] lemma pi_half_less_two [simp]: "pi/2 < 2" proof - have "pi/2 \ 2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_two_neq_zero le_less) qed lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] lemma pi_gt_zero [simp]: "0 < pi" using pi_half_gt_zero by simp lemma pi_ge_zero [simp]: "0 \ pi" by (rule pi_gt_zero [THEN order_less_imp_le]) lemma pi_neq_zero [simp]: "pi \ 0" by (rule pi_gt_zero [THEN less_imp_neq, symmetric]) lemma pi_not_less_zero [simp]: "\ pi < 0" by (simp add: linorder_not_less) lemma minus_pi_half_less_zero: "-(pi/2) < 0" by simp lemma m2pi_less_pi: "- (2*pi) < pi" by simp lemma sin_pi_half [simp]: "sin(pi/2) = 1" using sin_cos_squared_add2 [where x = "pi/2"] using sin_gt_zero_02 [OF pi_half_gt_zero pi_half_less_two] by (simp add: power2_eq_1_iff) lemma sin_of_real_pi_half [simp]: "sin ((of_real pi/2) :: 'a) = 1" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" using sin_pi_half by (metis sin_pi_half eq_numeral_simps(4) nonzero_of_real_divide of_real_1 of_real_numeral sin_of_real) lemma sin_cos_eq: "sin x = cos (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_diff) lemma minus_sin_cos_eq: "- sin x = cos (x + of_real pi/2)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_add nonzero_of_real_divide) lemma cos_sin_eq: "cos x = sin (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" using sin_cos_eq [of "of_real pi/2 - x"] by simp lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of "of_real pi/2 - x" "-y"] by (simp add: cos_sin_eq) (simp add: sin_cos_eq) lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" for x :: "'a::{real_normed_field,banach}" using sin_add [of x "- y"] by simp lemma sin_double: "sin(2 * x) = 2 * sin x * cos x" for x :: "'a::{real_normed_field,banach}" using sin_add [where x=x and y=x] by simp lemma cos_of_real_pi [simp]: "cos (of_real pi) = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by (simp add: cos_of_real) lemma sin_of_real_pi [simp]: "sin (of_real pi) = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by (simp add: sin_of_real) lemma cos_pi [simp]: "cos pi = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_pi [simp]: "sin pi = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" by (simp add: sin_add) lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" by (simp add: sin_add) lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" by (simp add: cos_add) lemma cos_periodic_pi2 [simp]: "cos (pi + x) = - cos x" by (simp add: cos_add) lemma sin_periodic [simp]: "sin (x + 2 * pi) = sin x" by (simp add: sin_add sin_double cos_double) lemma cos_periodic [simp]: "cos (x + 2 * pi) = cos x" by (simp add: cos_add sin_double cos_double) lemma cos_npi [simp]: "cos (real n * pi) = (- 1) ^ n" by (induct n) (auto simp: distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = (- 1) ^ n" by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real n * pi) = 0" for n :: nat by (induct n) (auto simp: distrib_right) lemma sin_npi2 [simp]: "sin (pi * real n) = 0" for n :: nat by (simp add: mult.commute [of pi]) lemma cos_two_pi [simp]: "cos (2 * pi) = 1" by (simp add: cos_double) lemma sin_two_pi [simp]: "sin (2 * pi) = 0" by (simp add: sin_double) +context + fixes w :: "'a::{real_normed_field,banach}" + +begin + lemma sin_times_sin: "sin w * sin z = (cos (w - z) - cos (w + z)) / 2" - for w :: "'a::{real_normed_field,banach}" by (simp add: cos_diff cos_add) lemma sin_times_cos: "sin w * cos z = (sin (w + z) + sin (w - z)) / 2" - for w :: "'a::{real_normed_field,banach}" by (simp add: sin_diff sin_add) lemma cos_times_sin: "cos w * sin z = (sin (w + z) - sin (w - z)) / 2" - for w :: "'a::{real_normed_field,banach}" by (simp add: sin_diff sin_add) lemma cos_times_cos: "cos w * cos z = (cos (w - z) + cos (w + z)) / 2" - for w :: "'a::{real_normed_field,banach}" by (simp add: cos_diff cos_add) +lemma cos_double_cos: "cos (2 * w) = 2 * cos w ^ 2 - 1" + by (simp add: cos_double sin_squared_eq) + +lemma cos_double_sin: "cos (2 * w) = 1 - 2 * sin w ^ 2" + by (simp add: cos_double sin_squared_eq) + +end + lemma sin_plus_sin: "sin w + sin z = 2 * sin ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma sin_diff_sin: "sin w - sin z = 2 * sin ((w - z) / 2) * cos ((w + z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma cos_plus_cos: "cos w + cos z = 2 * cos ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc cos_times_cos) apply (simp add: field_simps) done lemma cos_diff_cos: "cos w - cos z = 2 * sin ((w + z) / 2) * sin ((z - w) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc sin_times_sin) apply (simp add: field_simps) done -lemma cos_double_cos: "cos (2 * z) = 2 * cos z ^ 2 - 1" - for z :: "'a::{real_normed_field,banach}" - by (simp add: cos_double sin_squared_eq) - -lemma cos_double_sin: "cos (2 * z) = 1 - 2 * sin z ^ 2" - for z :: "'a::{real_normed_field,banach}" - by (simp add: cos_double sin_squared_eq) - lemma sin_pi_minus [simp]: "sin (pi - x) = sin x" by (metis sin_minus sin_periodic_pi minus_minus uminus_add_conv_diff) lemma cos_pi_minus [simp]: "cos (pi - x) = - (cos x)" by (metis cos_minus cos_periodic_pi uminus_add_conv_diff) lemma sin_minus_pi [simp]: "sin (x - pi) = - (sin x)" by (simp add: sin_diff) lemma cos_minus_pi [simp]: "cos (x - pi) = - (cos x)" by (simp add: cos_diff) lemma sin_2pi_minus [simp]: "sin (2 * pi - x) = - (sin x)" by (metis sin_periodic_pi2 add_diff_eq mult_2 sin_pi_minus) lemma cos_2pi_minus [simp]: "cos (2 * pi - x) = cos x" by (metis (no_types, hide_lams) cos_add cos_minus cos_two_pi sin_minus sin_two_pi diff_0_right minus_diff_eq mult_1 mult_zero_left uminus_add_conv_diff) lemma sin_gt_zero2: "0 < x \ x < pi/2 \ 0 < sin x" by (metis sin_gt_zero_02 order_less_trans pi_half_less_two) lemma sin_less_zero: assumes "- pi/2 < x" and "x < 0" shows "sin x < 0" proof - have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2) then show ?thesis by simp qed lemma pi_less_4: "pi < 4" using pi_half_less_two by auto lemma cos_gt_zero: "0 < x \ x < pi/2 \ 0 < cos x" by (simp add: cos_sin_eq sin_gt_zero2) lemma cos_gt_zero_pi: "-(pi/2) < x \ x < pi/2 \ 0 < cos x" using cos_gt_zero [of x] cos_gt_zero [of "-x"] by (cases rule: linorder_cases [of x 0]) auto lemma cos_ge_zero: "-(pi/2) \ x \ x \ pi/2 \ 0 \ cos x" by (auto simp: order_le_less cos_gt_zero_pi) (metis cos_pi_half eq_divide_eq eq_numeral_simps(4)) lemma sin_gt_zero: "0 < x \ x < pi \ 0 < sin x" by (simp add: sin_cos_eq cos_gt_zero_pi) lemma sin_lt_zero: "pi < x \ x < 2 * pi \ sin x < 0" using sin_gt_zero [of "x - pi"] by (simp add: sin_diff) lemma pi_ge_two: "2 \ pi" proof (rule ccontr) assume "\ ?thesis" then have "pi < 2" by auto have "\y > pi. y < 2 \ y < 2 * pi" proof (cases "2 < 2 * pi") case True with dense[OF \pi < 2\] show ?thesis by auto next case False have "pi < 2 * pi" by auto from dense[OF this] and False show ?thesis by auto qed then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast then have "0 < sin y" using sin_gt_zero_02 by auto moreover have "sin y < 0" using sin_gt_zero[of "y - pi"] \pi < y\ and \y < 2 * pi\ sin_periodic_pi[of "y - pi"] by auto ultimately show False by auto qed lemma sin_ge_zero: "0 \ x \ x \ pi \ 0 \ sin x" by (auto simp: order_le_less sin_gt_zero) lemma sin_le_zero: "pi \ x \ x < 2 * pi \ sin x \ 0" using sin_ge_zero [of "x - pi"] by (simp add: sin_diff) lemma sin_pi_divide_n_ge_0 [simp]: assumes "n \ 0" shows "0 \ sin (pi / real n)" by (rule sin_ge_zero) (use assms in \simp_all add: field_split_simps\) lemma sin_pi_divide_n_gt_0: assumes "2 \ n" shows "0 < sin (pi / real n)" by (rule sin_gt_zero) (use assms in \simp_all add: field_split_simps\) text\Proof resembles that of \cos_is_zero\ but with \<^term>\pi\ for the upper bound\ lemma cos_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. 0 \ x \ x \ pi \ cos x = y" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ pi \ cos x = y" by (rule IVT2) (simp_all add: y) next fix a b :: real assume ab: "0 \ a \ a \ pi \ cos a = y" "0 \ b \ b \ pi \ cos b = y" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero) qed auto qed lemma sin_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. - (pi/2) \ x \ x \ pi/2 \ sin x = y" proof - from cos_total [OF y] obtain x where x: "0 \ x" "x \ pi" "cos x = y" and uniq: "\x'. 0 \ x' \ x' \ pi \ cos x' = y \ x' = x " by blast show ?thesis unfolding sin_cos_eq proof (rule ex1I [where a="pi/2 - x"]) show "- (pi/2) \ z \ z \ pi/2 \ cos (of_real pi/2 - z) = y \ z = pi/2 - x" for z using uniq [of "pi/2 -z"] by auto qed (use x in auto) qed lemma cos_zero_lemma: assumes "0 \ x" "cos x = 0" - shows "\n. odd n \ x = of_nat n * (pi/2) \ n > 0" + shows "\n. odd n \ x = of_nat n * (pi/2)" proof - have xle: "x < (1 + real_of_int \x/pi\) * pi" using floor_correct [of "x/pi"] by (simp add: add.commute divide_less_eq) obtain n where "real n * pi \ x" "x < real (Suc n) * pi" proof show "real (nat \x / pi\) * pi \ x" using assms floor_divide_lower [of pi x] by auto show "x < real (Suc (nat \x / pi\)) * pi" using assms floor_divide_upper [of pi x] by (simp add: xle) qed then have x: "0 \ x - n * pi" "(x - n * pi) \ pi" "cos (x - n * pi) = 0" by (auto simp: algebra_simps cos_diff assms) then have "\!x. 0 \ x \ x \ pi \ cos x = 0" by (auto simp: intro!: cos_total) then obtain \ where \: "0 \ \" "\ \ pi" "cos \ = 0" and uniq: "\\. 0 \ \ \ \ \ pi \ cos \ = 0 \ \ = \" by blast then have "x - real n * pi = \" using x by blast moreover have "pi/2 = \" using pi_half_ge_zero uniq by fastforce ultimately show ?thesis by (rule_tac x = "Suc (2 * n)" in exI) (simp add: algebra_simps) qed -lemma sin_zero_lemma: "0 \ x \ sin x = 0 \ \n::nat. even n \ x = real n * (pi/2)" - using cos_zero_lemma [of "x + pi/2"] - apply (clarsimp simp add: cos_add) - apply (rule_tac x = "n - 1" in exI) - apply (simp add: algebra_simps of_nat_diff) - done +lemma sin_zero_lemma: + assumes "0 \ x" "sin x = 0" + shows "\n::nat. even n \ x = real n * (pi/2)" +proof - + obtain n where "odd n" and n: "x + pi/2 = of_nat n * (pi/2)" "n > 0" + using cos_zero_lemma [of "x + pi/2"] assms by (auto simp add: cos_add) + then have "x = real (n - 1) * (pi / 2)" + by (simp add: algebra_simps of_nat_diff) + then show ?thesis + by (simp add: \odd n\) +qed lemma cos_zero_iff: "cos x = 0 \ ((\n. odd n \ x = real n * (pi/2)) \ (\n. odd n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof - have *: "cos (real n * pi/2) = 0" if "odd n" for n :: nat proof - from that obtain m where "n = 2 * m + 1" .. then show ?thesis by (simp add: field_simps) (simp add: cos_add add_divide_distrib) qed show ?thesis proof show ?rhs if ?lhs using that cos_zero_lemma [of x] cos_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto dest: * simp del: eq_divide_eq_numeral1) qed qed lemma sin_zero_iff: "sin x = 0 \ ((\n. even n \ x = real n * (pi/2)) \ (\n. even n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof show ?rhs if ?lhs using that sin_zero_lemma [of x] sin_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto elim: evenE) qed lemma sin_zero_pi_iff: fixes x::real assumes "\x\ < pi" shows "sin x = 0 \ x = 0" proof show "x = 0" if "sin x = 0" using that assms by (auto simp: sin_zero_iff) qed auto -lemma cos_zero_iff_int: "cos x = 0 \ (\n. odd n \ x = of_int n * (pi/2))" +lemma cos_zero_iff_int: "cos x = 0 \ (\i. odd i \ x = of_int i * (pi/2))" proof - have 1: "\n. odd n \ \i. odd i \ real n = real_of_int i" by (metis even_of_nat of_int_of_nat_eq) have 2: "\n. odd n \ \i. odd i \ - (real n * pi) = real_of_int i * pi" by (metis even_minus even_of_nat mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) have 3: "\odd i; \n. even n \ real_of_int i \ - (real n)\ \ \n. odd n \ real_of_int i = real n" for i by (cases i rule: int_cases2) auto show ?thesis by (force simp: cos_zero_iff intro!: 1 2 3) qed -lemma sin_zero_iff_int: "sin x = 0 \ (\n. even n \ x = of_int n * (pi/2))" +lemma sin_zero_iff_int: "sin x = 0 \ (\i. even i \ x = of_int i * (pi/2))" (is "?lhs = ?rhs") proof safe - assume "sin x = 0" + assume ?lhs + then consider (plus) n where "even n" "x = real n * (pi/2)" | (minus) n where "even n" "x = - (real n * (pi/2))" + using sin_zero_iff by auto then show "\n. even n \ x = of_int n * (pi/2)" - apply (simp add: sin_zero_iff, safe) - apply (metis even_of_nat of_int_of_nat_eq) - apply (rule_tac x="- (int n)" in exI) - apply simp - done + proof cases + case plus + then show ?rhs + by (metis even_of_nat of_int_of_nat_eq) + next + case minus + then show ?thesis + by (rule_tac x="- (int n)" in exI) simp + qed next fix i :: int assume "even i" then show "sin (of_int i * (pi/2)) = 0" by (cases i rule: int_cases2, simp_all add: sin_zero_iff) qed -lemma sin_zero_iff_int2: "sin x = 0 \ (\n::int. x = of_int n * pi)" - apply (simp only: sin_zero_iff_int) - apply (safe elim!: evenE) - apply (simp_all add: field_simps) - using dvd_triv_left apply fastforce - done +lemma sin_zero_iff_int2: "sin x = 0 \ (\i::int. x = of_int i * pi)" +proof - + have "sin x = 0 \ (\i. even i \ x = real_of_int i * (pi / 2))" + by (auto simp: sin_zero_iff_int) + also have "... = (\j. x = real_of_int (2*j) * (pi / 2))" + using dvd_triv_left by blast + also have "... = (\i::int. x = of_int i * pi)" + by auto + finally show ?thesis . +qed lemma sin_npi_int [simp]: "sin (pi * of_int n) = 0" by (simp add: sin_zero_iff_int2) lemma cos_monotone_0_pi: assumes "0 \ y" and "y < x" and "x \ pi" shows "cos x < cos y" proof - have "- (x - y) < 0" using assms by auto from MVT2[OF \y < x\ DERIV_cos] obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto then have "0 < z" and "z < pi" using assms by auto then have "0 < sin z" using sin_gt_zero by auto then have "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using \- (x - y) < 0\ by (rule mult_pos_neg2) then show ?thesis by auto qed lemma cos_monotone_0_pi_le: assumes "0 \ y" and "y \ x" and "x \ pi" shows "cos x \ cos y" proof (cases "y < x") case True show ?thesis using cos_monotone_0_pi[OF \0 \ y\ True \x \ pi\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma cos_monotone_minus_pi_0: assumes "- pi \ y" and "y < x" and "x \ 0" shows "cos y < cos x" proof - have "0 \ - x" and "- x < - y" and "- y \ pi" using assms by auto from cos_monotone_0_pi[OF this] show ?thesis unfolding cos_minus . qed lemma cos_monotone_minus_pi_0': assumes "- pi \ y" and "y \ x" and "x \ 0" shows "cos y \ cos x" proof (cases "y < x") case True show ?thesis using cos_monotone_minus_pi_0[OF \-pi \ y\ True \x \ 0\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma sin_monotone_2pi: assumes "- (pi/2) \ y" and "y < x" and "x \ pi/2" shows "sin y < sin x" unfolding sin_cos_eq using assms by (auto intro: cos_monotone_0_pi) lemma sin_monotone_2pi_le: assumes "- (pi/2) \ y" and "y \ x" and "x \ pi/2" shows "sin y \ sin x" by (metis assms le_less sin_monotone_2pi) lemma sin_x_le_x: fixes x :: real - assumes x: "x \ 0" + assumes "x \ 0" shows "sin x \ x" proof - let ?f = "\x. x - sin x" - from x have "?f x \ ?f 0" - apply (rule DERIV_nonneg_imp_nondecreasing) - apply (intro allI impI exI[of _ "1 - cos x" for x]) - apply (auto intro!: derivative_eq_intros simp: field_simps) - done + have "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 - cos u) (at u)" + by (auto intro!: derivative_eq_intros simp: field_simps) + then have "?f x \ ?f 0" + by (metis cos_le_one diff_ge_0_iff_ge DERIV_nonneg_imp_nondecreasing [OF assms]) then show "sin x \ x" by simp qed lemma sin_x_ge_neg_x: fixes x :: real assumes x: "x \ 0" shows "sin x \ - x" proof - let ?f = "\x. x + sin x" - from x have "?f x \ ?f 0" - apply (rule DERIV_nonneg_imp_nondecreasing) - apply (intro allI impI exI[of _ "1 + cos x" for x]) - apply (auto intro!: derivative_eq_intros simp: field_simps real_0_le_add_iff) - done + have \
: "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 + cos u) (at u)" + by (auto intro!: derivative_eq_intros simp: field_simps) + have "?f x \ ?f 0" + by (rule DERIV_nonneg_imp_nondecreasing [OF assms]) (use \
real_0_le_add_iff in force) then show "sin x \ -x" by simp qed lemma abs_sin_x_le_abs_x: "\sin x\ \ \x\" for x :: real using sin_x_ge_neg_x [of x] sin_x_le_x [of x] sin_x_ge_neg_x [of "-x"] sin_x_le_x [of "-x"] by (auto simp: abs_real_def) subsection \More Corollaries about Sine and Cosine\ lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi/2) = (-1) ^ n" proof - have "sin ((real n + 1/2) * pi) = cos (real n * pi)" by (auto simp: algebra_simps sin_add) then show ?thesis by (simp add: distrib_right add_divide_distrib add.commute mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real n * pi) = 1" for n :: nat by (cases "even n") (simp_all add: cos_double mult.assoc) lemma cos_3over2_pi [simp]: "cos (3/2*pi) = 0" proof - have "cos (3/2*pi) = cos (pi + pi/2)" by simp also have "... = 0" by (subst cos_add, simp) finally show ?thesis . qed lemma sin_2npi [simp]: "sin (2 * real n * pi) = 0" for n :: nat by (auto simp: mult.assoc sin_double) lemma sin_3over2_pi [simp]: "sin (3/2*pi) = - 1" proof - have "sin (3/2*pi) = sin (pi + pi/2)" by simp also have "... = -1" by (subst sin_add, simp) finally show ?thesis . qed lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" by (simp only: cos_add sin_add of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) lemma DERIV_cos_add [simp]: "DERIV (\x. cos (x + k)) xa :> - sin (xa + k)" by (auto intro!: derivative_eq_intros) lemma sin_zero_norm_cos_one: fixes x :: "'a::{real_normed_field,banach}" assumes "sin x = 0" shows "norm (cos x) = 1" using sin_cos_squared_add [of x, unfolded assms] by (simp add: square_norm_one) lemma sin_zero_abs_cos_one: "sin x = 0 \ \cos x\ = (1::real)" using sin_zero_norm_cos_one by fastforce lemma cos_one_sin_zero: fixes x :: "'a::{real_normed_field,banach}" assumes "cos x = 1" shows "sin x = 0" using sin_cos_squared_add [of x, unfolded assms] by simp lemma sin_times_pi_eq_0: "sin (x * pi) = 0 \ x \ \" by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int) lemma cos_one_2pi: "cos x = 1 \ (\n::nat. x = n * 2 * pi) \ (\n::nat. x = - (n * 2 * pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin x = 0" by (simp add: cos_one_sin_zero) then show ?rhs proof (simp only: sin_zero_iff, elim exE disjE conjE) fix n :: nat assume n: "even n" "x = real n * (pi/2)" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) next fix n :: nat assume n: "even n" "x = - (real n * (pi/2))" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) qed next assume ?rhs then show "cos x = 1" by (metis cos_2npi cos_minus mult.assoc mult.left_commute) qed lemma cos_one_2pi_int: "cos x = 1 \ (\n::int. x = n * 2 * pi)" (is "?lhs = ?rhs") proof assume "cos x = 1" then show ?rhs by (metis cos_one_2pi mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) next assume ?rhs then show "cos x = 1" by (clarsimp simp add: cos_one_2pi) (metis mult_minus_right of_int_of_nat) qed lemma cos_npi_int [simp]: fixes n::int shows "cos (pi * of_int n) = (if even n then 1 else -1)" by (auto simp: algebra_simps cos_one_2pi_int elim!: oddE evenE) lemma sin_cos_sqrt: "0 \ sin x \ sin x = sqrt (1 - (cos(x) ^ 2))" using sin_squared_eq real_sqrt_unique by fastforce lemma sin_eq_0_pi: "- pi < x \ x < pi \ sin x = 0 \ x = 0" by (metis sin_gt_zero sin_minus minus_less_iff neg_0_less_iff_less not_less_iff_gr_or_eq) lemma cos_treble_cos: "cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x" for x :: "'a::{real_normed_field,banach}" proof - have *: "(sin x * (sin x * 3)) = 3 - (cos x * (cos x * 3))" by (simp add: mult.assoc [symmetric] sin_squared_eq [unfolded power2_eq_square]) have "cos(3 * x) = cos(2*x + x)" by simp also have "\ = 4 * cos x ^ 3 - 3 * cos x" - apply (simp only: cos_add cos_double sin_double) - apply (simp add: * field_simps power2_eq_square power3_eq_cube) - done + unfolding cos_add cos_double sin_double + by (simp add: * field_simps power2_eq_square power3_eq_cube) finally show ?thesis . qed lemma cos_45: "cos (pi / 4) = sqrt 2 / 2" proof - let ?c = "cos (pi / 4)" let ?s = "sin (pi / 4)" have nonneg: "0 \ ?c" by (simp add: cos_ge_zero) have "0 = cos (pi / 4 + pi / 4)" by simp also have "cos (pi / 4 + pi / 4) = ?c\<^sup>2 - ?s\<^sup>2" by (simp only: cos_add power2_eq_square) also have "\ = 2 * ?c\<^sup>2 - 1" by (simp add: sin_squared_eq) finally have "?c\<^sup>2 = (sqrt 2 / 2)\<^sup>2" by (simp add: power_divide) then show ?thesis using nonneg by (rule power2_eq_imp_eq) simp qed lemma cos_30: "cos (pi / 6) = sqrt 3/2" proof - let ?c = "cos (pi / 6)" let ?s = "sin (pi / 6)" have pos_c: "0 < ?c" by (rule cos_gt_zero) simp_all have "0 = cos (pi / 6 + pi / 6 + pi / 6)" by simp also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" by (simp only: cos_add sin_add) also have "\ = ?c * (?c\<^sup>2 - 3 * ?s\<^sup>2)" by (simp add: algebra_simps power2_eq_square) finally have "?c\<^sup>2 = (sqrt 3/2)\<^sup>2" using pos_c by (simp add: sin_squared_eq power_divide) then show ?thesis using pos_c [THEN order_less_imp_le] by (rule power2_eq_imp_eq) simp qed lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" by (simp add: sin_cos_eq cos_45) lemma sin_60: "sin (pi / 3) = sqrt 3/2" by (simp add: sin_cos_eq cos_30) lemma cos_60: "cos (pi / 3) = 1 / 2" proof - have "0 \ cos (pi / 3)" by (rule cos_ge_zero) (use pi_half_ge_zero in \linarith+\) then show ?thesis by (simp add: cos_squared_eq sin_60 power_divide power2_eq_imp_eq) qed lemma sin_30: "sin (pi / 6) = 1 / 2" by (simp add: sin_cos_eq cos_60) lemma cos_integer_2pi: "n \ \ \ cos(2 * pi * n) = 1" by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute) lemma sin_integer_2pi: "n \ \ \ sin(2 * pi * n) = 0" by (metis sin_two_pi Ints_mult mult.assoc mult.commute sin_times_pi_eq_0) lemma cos_int_2pin [simp]: "cos ((2 * pi) * of_int n) = 1" by (simp add: cos_one_2pi_int) lemma sin_int_2pin [simp]: "sin ((2 * pi) * of_int n) = 0" by (metis Ints_of_int sin_integer_2pi) lemma sincos_principal_value: "\y. (- pi < y \ y \ pi) \ (sin y = sin x \ cos y = cos x)" - apply (rule exI [where x="pi - (2 * pi) * frac ((pi - x) / (2 * pi))"]) - apply (auto simp: field_simps frac_lt_1) - apply (simp_all add: frac_def field_simps) - apply (simp_all add: add_divide_distrib diff_divide_distrib) - apply (simp_all add: sin_add cos_add mult.assoc [symmetric]) - done +proof - + define y where "y \ pi - (2 * pi) * frac ((pi - x) / (2 * pi))" + have "-pi < y"" y \ pi" + by (auto simp: field_simps frac_lt_1 y_def) + moreover + have "sin y = sin x" "cos y = cos x" + unfolding y_def + apply (simp_all add: frac_def divide_simps sin_add cos_add) + by (metis sin_int_2pin cos_int_2pin diff_zero add.right_neutral mult.commute mult.left_neutral mult_zero_left)+ + ultimately + show ?thesis by metis +qed subsection \Tangent\ definition tan :: "'a \ 'a::{real_normed_field,banach}" where "tan = (\x. sin x / cos x)" lemma tan_of_real: "of_real (tan x) = (tan (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: tan_def sin_of_real cos_of_real) lemma tan_in_Reals [simp]: "z \ \ \ tan z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: tan_def) lemma tan_zero [simp]: "tan 0 = 0" by (simp add: tan_def) lemma tan_pi [simp]: "tan pi = 0" by (simp add: tan_def) lemma tan_npi [simp]: "tan (real n * pi) = 0" for n :: nat by (simp add: tan_def) lemma tan_minus [simp]: "tan (- x) = - tan x" by (simp add: tan_def) lemma tan_periodic [simp]: "tan (x + 2 * pi) = tan x" by (simp add: tan_def) lemma lemma_tan_add1: "cos x \ 0 \ cos y \ 0 \ 1 - tan x * tan y = cos (x + y)/(cos x * cos y)" by (simp add: tan_def cos_add field_simps) lemma add_tan_eq: "cos x \ 0 \ cos y \ 0 \ tan x + tan y = sin(x + y)/(cos x * cos y)" for x :: "'a::{real_normed_field,banach}" by (simp add: tan_def sin_add field_simps) lemma tan_add: "cos x \ 0 \ cos y \ 0 \ cos (x + y) \ 0 \ tan (x + y) = (tan x + tan y)/(1 - tan x * tan y)" for x :: "'a::{real_normed_field,banach}" by (simp add: add_tan_eq lemma_tan_add1 field_simps) (simp add: tan_def) lemma tan_double: "cos x \ 0 \ cos (2 * x) \ 0 \ tan (2 * x) = (2 * tan x) / (1 - (tan x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x x] by (simp add: power2_eq_square) lemma tan_gt_zero: "0 < x \ x < pi/2 \ 0 < tan x" by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma tan_less_zero: assumes "- pi/2 < x" and "x < 0" shows "tan x < 0" proof - have "0 < tan (- x)" using assms by (simp only: tan_gt_zero) then show ?thesis by simp qed lemma tan_half: "tan x = sin (2 * x) / (cos (2 * x) + 1)" for x :: "'a::{real_normed_field,banach,field}" unfolding tan_def sin_double cos_double sin_squared_eq by (simp add: power2_eq_square) lemma tan_30: "tan (pi / 6) = 1 / sqrt 3" unfolding tan_def by (simp add: sin_30 cos_30) lemma tan_45: "tan (pi / 4) = 1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_60: "tan (pi / 3) = sqrt 3" unfolding tan_def by (simp add: sin_60 cos_60) lemma DERIV_tan [simp]: "cos x \ 0 \ DERIV tan x :> inverse ((cos x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding tan_def by (auto intro!: derivative_eq_intros, simp add: divide_inverse power2_eq_square) declare DERIV_tan[THEN DERIV_chain2, derivative_intros] and DERIV_tan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_tan[derivative_intros] = DERIV_tan[THEN DERIV_compose_FDERIV] lemma isCont_tan: "cos x \ 0 \ isCont tan x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_tan [THEN DERIV_isCont]) lemma isCont_tan' [simp,continuous_intros]: fixes a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" shows "isCont f a \ cos (f a) \ 0 \ isCont (\x. tan (f x)) a" by (rule isCont_o2 [OF _ isCont_tan]) lemma tendsto_tan [tendsto_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "(f \ a) F \ cos a \ 0 \ ((\x. tan (f x)) \ tan a) F" by (rule isCont_tendsto_compose [OF isCont_tan]) lemma continuous_tan: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous F f \ cos (f (Lim F (\x. x))) \ 0 \ continuous F (\x. tan (f x))" unfolding continuous_def by (rule tendsto_tan) lemma continuous_on_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. cos (f x) \ 0) \ continuous_on s (\x. tan (f x))" unfolding continuous_on_def by (auto intro: tendsto_tan) lemma continuous_within_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ cos (f x) \ 0 \ continuous (at x within s) (\x. tan (f x))" unfolding continuous_within by (rule tendsto_tan) lemma LIM_cos_div_sin: "(\x. cos(x)/sin(x)) \pi/2\ 0" by (rule tendsto_cong_limit, (rule tendsto_intros)+, simp_all) lemma lemma_tan_total: assumes "0 < y" shows "\x. 0 < x \ x < pi/2 \ y < tan x" proof - obtain s where "0 < s" and s: "\x. \x \ pi/2; norm (x - pi/2) < s\ \ norm (cos x / sin x - 0) < inverse y" using LIM_D [OF LIM_cos_div_sin, of "inverse y"] that assms by force obtain e where e: "0 < e" "e < s" "e < pi/2" using \0 < s\ field_lbound_gt_zero pi_half_gt_zero by blast show ?thesis proof (intro exI conjI) have "0 < sin e" "0 < cos e" using e by (auto intro: cos_gt_zero sin_gt_zero2 simp: mult.commute) then show "y < tan (pi/2 - e)" using s [of "pi/2 - e"] e assms by (simp add: tan_def sin_diff cos_diff) (simp add: field_simps split: if_split_asm) qed (use e in auto) qed lemma tan_total_pos: assumes "0 \ y" shows "\x. 0 \ x \ x < pi/2 \ tan x = y" proof (cases "y = 0") case True then show ?thesis using pi_half_gt_zero tan_zero by blast next case False with assms have "y > 0" by linarith obtain x where x: "0 < x" "x < pi/2" "y < tan x" using lemma_tan_total \0 < y\ by blast have "\u\0. u \ x \ tan u = y" proof (intro IVT allI impI) show "isCont tan u" if "0 \ u \ u \ x" for u proof - have "cos u \ 0" using antisym_conv2 cos_gt_zero that x(2) by fastforce with assms show ?thesis by (auto intro!: DERIV_tan [THEN DERIV_isCont]) qed qed (use assms x in auto) then show ?thesis using x(2) by auto qed lemma lemma_tan_total1: "\x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof (cases "0::real" y rule: le_cases) case le then show ?thesis by (meson less_le_trans minus_pi_half_less_zero tan_total_pos) next case ge with tan_total_pos [of "-y"] obtain x where "0 \ x" "x < pi / 2" "tan x = - y" by force then show ?thesis by (rule_tac x="-x" in exI) auto qed proposition tan_total: "\! x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof - have "u = v" if u: "- (pi / 2) < u" "u < pi / 2" and v: "- (pi / 2) < v" "v < pi / 2" and eq: "tan u = tan v" for u v proof (cases u v rule: linorder_cases) case less have "\x. u \ x \ x \ v \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(1) v(2)) then have "continuous_on {u..v} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. u < x \ x < v \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(1) v(2)) ultimately obtain z where "u < z" "z < v" "DERIV tan z :> 0" by (metis less Rolle eq) moreover have "cos z \ 0" by (metis (no_types) \u < z\ \z < v\ cos_gt_zero_pi less_le_trans linorder_not_less not_less_iff_gr_or_eq u(1) v(2)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce next case greater have "\x. v \ x \ x \ u \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(2) v(1)) then have "continuous_on {v..u} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. v < x \ x < u \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(2) v(1)) ultimately obtain z where "v < z" "z < u" "DERIV tan z :> 0" by (metis greater Rolle eq) moreover have "cos z \ 0" by (metis \v < z\ \z < u\ cos_gt_zero_pi less_eq_real_def less_le_trans order_less_irrefl u(2) v(1)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce qed auto then have "\!x. - (pi / 2) < x \ x < pi / 2 \ tan x = y" if x: "- (pi / 2) < x" "x < pi / 2" "tan x = y" for x using that by auto then show ?thesis using lemma_tan_total1 [where y = y] by auto qed lemma tan_monotone: assumes "- (pi/2) < y" and "y < x" and "x < pi/2" shows "tan y < tan x" proof - have "DERIV tan x' :> inverse ((cos x')\<^sup>2)" if "y \ x'" "x' \ x" for x' proof - have "-(pi/2) < x'" and "x' < pi/2" using that assms by auto with cos_gt_zero_pi have "cos x' \ 0" by force then show "DERIV tan x' :> inverse ((cos x')\<^sup>2)" by (rule DERIV_tan) qed from MVT2[OF \y < x\ this] obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto then have "- (pi/2) < z" and "z < pi/2" using assms by auto then have "0 < cos z" using cos_gt_zero_pi by auto then have inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto have "0 < x - y" using \y < x\ by auto with inv_pos have "0 < tan x - tan y" unfolding tan_diff by auto then show ?thesis by auto qed lemma tan_monotone': assumes "- (pi/2) < y" and "y < pi/2" and "- (pi/2) < x" and "x < pi/2" shows "y < x \ tan y < tan x" proof assume "y < x" then show "tan y < tan x" using tan_monotone and \- (pi/2) < y\ and \x < pi/2\ by auto next assume "tan y < tan x" show "y < x" proof (rule ccontr) assume "\ ?thesis" then have "x \ y" by auto then have "tan x \ tan y" proof (cases "x = y") case True then show ?thesis by auto next case False then have "x < y" using \x \ y\ by auto from tan_monotone[OF \- (pi/2) < x\ this \y < pi/2\] show ?thesis by auto qed then show False using \tan y < tan x\ by auto qed qed lemma tan_inverse: "1 / (tan y) = tan (pi/2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" by (simp add: tan_def) lemma tan_periodic_nat[simp]: "tan (x + real n * pi) = tan x" for n :: nat proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 of_nat_add distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma tan_periodic_int[simp]: "tan (x + of_int i * pi) = tan x" proof (cases "0 \ i") case True then have i_nat: "of_int i = of_int (nat i)" by auto show ?thesis unfolding i_nat by (metis of_int_of_nat_eq tan_periodic_nat) next case False then have i_nat: "of_int i = - of_int (nat (- i))" by auto have "tan x = tan (x + of_int i * pi - of_int i * pi)" by auto also have "\ = tan (x + of_int i * pi)" unfolding i_nat mult_minus_left diff_minus_eq_add by (metis of_int_of_nat_eq tan_periodic_nat) finally show ?thesis by auto qed lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" using tan_periodic_int[of _ "numeral n" ] by simp lemma tan_minus_45: "tan (-(pi/4)) = -1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_diff: "cos x \ 0 \ cos y \ 0 \ cos (x - y) \ 0 \ tan (x - y) = (tan x - tan y)/(1 + tan x * tan y)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x "-y"] by simp lemma tan_pos_pi2_le: "0 \ x \ x < pi/2 \ 0 \ tan x" using less_eq_real_def tan_gt_zero by auto lemma cos_tan: "\x\ < pi/2 \ cos x = 1 / sqrt (1 + tan x ^ 2)" using cos_gt_zero_pi [of x] by (simp add: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma sin_tan: "\x\ < pi/2 \ sin x = tan x / sqrt (1 + tan x ^ 2)" using cos_gt_zero [of "x"] cos_gt_zero [of "-x"] by (force simp: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) lemma tan_mono_le: "-(pi/2) < x \ x \ y \ y < pi/2 \ tan x \ tan y" using less_eq_real_def tan_monotone by auto lemma tan_mono_lt_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x < tan y \ x < y" using tan_monotone' by blast lemma tan_mono_le_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x \ tan y \ x \ y" by (meson tan_mono_le not_le tan_monotone) lemma tan_bound_pi2: "\x\ < pi/4 \ \tan x\ < 1" using tan_45 tan_monotone [of x "pi/4"] tan_monotone [of "-x" "pi/4"] by (auto simp: abs_if split: if_split_asm) lemma tan_cot: "tan(pi/2 - x) = inverse(tan x)" by (simp add: tan_def sin_diff cos_diff) subsection \Cotangent\ definition cot :: "'a \ 'a::{real_normed_field,banach}" where "cot = (\x. cos x / sin x)" lemma cot_of_real: "of_real (cot x) = (cot (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: cot_def sin_of_real cos_of_real) lemma cot_in_Reals [simp]: "z \ \ \ cot z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: cot_def) lemma cot_zero [simp]: "cot 0 = 0" by (simp add: cot_def) lemma cot_pi [simp]: "cot pi = 0" by (simp add: cot_def) lemma cot_npi [simp]: "cot (real n * pi) = 0" for n :: nat by (simp add: cot_def) lemma cot_minus [simp]: "cot (- x) = - cot x" by (simp add: cot_def) lemma cot_periodic [simp]: "cot (x + 2 * pi) = cot x" by (simp add: cot_def) lemma cot_altdef: "cot x = inverse (tan x)" by (simp add: cot_def tan_def) lemma tan_altdef: "tan x = inverse (cot x)" by (simp add: cot_def tan_def) lemma tan_cot': "tan (pi/2 - x) = cot x" by (simp add: tan_cot cot_altdef) lemma cot_gt_zero: "0 < x \ x < pi/2 \ 0 < cot x" by (simp add: cot_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma cot_less_zero: assumes lb: "- pi/2 < x" and "x < 0" shows "cot x < 0" proof - have "0 < cot (- x)" using assms by (simp only: cot_gt_zero) then show ?thesis by simp qed lemma DERIV_cot [simp]: "sin x \ 0 \ DERIV cot x :> -inverse ((sin x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding cot_def using cos_squared_eq[of x] by (auto intro!: derivative_eq_intros) (simp add: divide_inverse power2_eq_square) lemma isCont_cot: "sin x \ 0 \ isCont cot x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cot [THEN DERIV_isCont]) lemma isCont_cot' [simp,continuous_intros]: "isCont f a \ sin (f a) \ 0 \ isCont (\x. cot (f x)) a" for a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" by (rule isCont_o2 [OF _ isCont_cot]) lemma tendsto_cot [tendsto_intros]: "(f \ a) F \ sin a \ 0 \ ((\x. cot (f x)) \ cot a) F" for f :: "'a \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_cot]) lemma continuous_cot: "continuous F f \ sin (f (Lim F (\x. x))) \ 0 \ continuous F (\x. cot (f x))" for f :: "'a \ 'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_cot) lemma continuous_on_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. sin (f x) \ 0) \ continuous_on s (\x. cot (f x))" unfolding continuous_on_def by (auto intro: tendsto_cot) lemma continuous_within_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ sin (f x) \ 0 \ continuous (at x within s) (\x. cot (f x))" unfolding continuous_within by (rule tendsto_cot) subsection \Inverse Trigonometric Functions\ definition arcsin :: "real \ real" where "arcsin y = (THE x. -(pi/2) \ x \ x \ pi/2 \ sin x = y)" definition arccos :: "real \ real" where "arccos y = (THE x. 0 \ x \ x \ pi \ cos x = y)" definition arctan :: "real \ real" where "arctan y = (THE x. -(pi/2) < x \ x < pi/2 \ tan x = y)" lemma arcsin: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2 \ sin (arcsin y) = y" unfolding arcsin_def by (rule theI' [OF sin_total]) lemma arcsin_pi: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi \ sin (arcsin y) = y" by (drule (1) arcsin) (force intro: order_trans) lemma sin_arcsin [simp]: "- 1 \ y \ y \ 1 \ sin (arcsin y) = y" by (blast dest: arcsin) lemma arcsin_bounded: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lbound: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y" by (blast dest: arcsin) lemma arcsin_ubound: "- 1 \ y \ y \ 1 \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lt_bounded: assumes "- 1 < y" "y < 1" shows "- (pi/2) < arcsin y \ arcsin y < pi/2" proof - have "arcsin y \ pi/2" by (metis arcsin assms not_less not_less_iff_gr_or_eq sin_pi_half) moreover have "arcsin y \ - pi/2" by (metis arcsin assms minus_divide_left not_less not_less_iff_gr_or_eq sin_minus sin_pi_half) ultimately show ?thesis using arcsin_bounded [of y] assms by auto qed lemma arcsin_sin: "- (pi/2) \ x \ x \ pi/2 \ arcsin (sin x) = x" unfolding arcsin_def using the1_equality [OF sin_total] by simp lemma arcsin_0 [simp]: "arcsin 0 = 0" using arcsin_sin [of 0] by simp lemma arcsin_1 [simp]: "arcsin 1 = pi/2" using arcsin_sin [of "pi/2"] by simp lemma arcsin_minus_1 [simp]: "arcsin (- 1) = - (pi/2)" using arcsin_sin [of "- pi/2"] by simp lemma arcsin_minus: "- 1 \ x \ x \ 1 \ arcsin (- x) = - arcsin x" by (metis (no_types, hide_lams) arcsin arcsin_sin minus_minus neg_le_iff_le sin_minus) lemma arcsin_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arcsin x = arcsin y \ x = y" by (metis abs_le_iff arcsin minus_le_iff) lemma cos_arcsin_nonzero: "- 1 < x \ x < 1 \ cos (arcsin x) \ 0" using arcsin_lt_bounded cos_gt_zero_pi by force lemma arccos: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi \ cos (arccos y) = y" unfolding arccos_def by (rule theI' [OF cos_total]) lemma cos_arccos [simp]: "- 1 \ y \ y \ 1 \ cos (arccos y) = y" by (blast dest: arccos) lemma arccos_bounded: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lbound: "- 1 \ y \ y \ 1 \ 0 \ arccos y" by (blast dest: arccos) lemma arccos_ubound: "- 1 \ y \ y \ 1 \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lt_bounded: assumes "- 1 < y" "y < 1" shows "0 < arccos y \ arccos y < pi" proof - have "arccos y \ 0" by (metis (no_types) arccos assms(1) assms(2) cos_zero less_eq_real_def less_irrefl) moreover have "arccos y \ -pi" by (metis arccos assms(1) assms(2) cos_minus cos_pi not_less not_less_iff_gr_or_eq) ultimately show ?thesis using arccos_bounded [of y] assms by (metis arccos cos_pi not_less not_less_iff_gr_or_eq) qed lemma arccos_cos: "0 \ x \ x \ pi \ arccos (cos x) = x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma arccos_cos2: "x \ 0 \ - pi \ x \ arccos (cos x) = -x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma cos_arcsin: assumes "- 1 \ x" "x \ 1" shows "cos (arcsin x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(cos (arcsin x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms cos_squared_eq) show "0 \ cos (arcsin x)" using arcsin assms cos_ge_zero by blast show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma sin_arccos: assumes "- 1 \ x" "x \ 1" shows "sin (arccos x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(sin (arccos x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms sin_squared_eq) show "0 \ sin (arccos x)" by (simp add: arccos_bounded assms sin_ge_zero) show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma arccos_0 [simp]: "arccos 0 = pi/2" by (metis arccos_cos cos_gt_zero cos_pi cos_pi_half pi_gt_zero pi_half_ge_zero not_le not_zero_less_neg_numeral numeral_One) lemma arccos_1 [simp]: "arccos 1 = 0" using arccos_cos by force lemma arccos_minus_1 [simp]: "arccos (- 1) = pi" by (metis arccos_cos cos_pi order_refl pi_ge_zero) lemma arccos_minus: "-1 \ x \ x \ 1 \ arccos (- x) = pi - arccos x" by (metis arccos_cos arccos_cos2 cos_minus_pi cos_total diff_le_0_iff_le le_add_same_cancel1 minus_diff_eq uminus_add_conv_diff) corollary arccos_minus_abs: assumes "\x\ \ 1" shows "arccos (- x) = pi - arccos x" using assms by (simp add: arccos_minus) lemma sin_arccos_nonzero: "- 1 < x \ x < 1 \ sin (arccos x) \ 0" using arccos_lt_bounded sin_gt_zero by force lemma arctan: "- (pi/2) < arctan y \ arctan y < pi/2 \ tan (arctan y) = y" unfolding arctan_def by (rule theI' [OF tan_total]) lemma tan_arctan: "tan (arctan y) = y" by (simp add: arctan) lemma arctan_bounded: "- (pi/2) < arctan y \ arctan y < pi/2" by (auto simp only: arctan) lemma arctan_lbound: "- (pi/2) < arctan y" by (simp add: arctan) lemma arctan_ubound: "arctan y < pi/2" by (auto simp only: arctan) lemma arctan_unique: assumes "-(pi/2) < x" and "x < pi/2" and "tan x = y" shows "arctan y = x" using assms arctan [of y] tan_total [of y] by (fast elim: ex1E) lemma arctan_tan: "-(pi/2) < x \ x < pi/2 \ arctan (tan x) = x" by (rule arctan_unique) simp_all lemma arctan_zero_zero [simp]: "arctan 0 = 0" by (rule arctan_unique) simp_all lemma arctan_minus: "arctan (- x) = - arctan x" using arctan [of "x"] by (auto simp: arctan_unique) lemma cos_arctan_not_zero [simp]: "cos (arctan x) \ 0" by (intro less_imp_neq [symmetric] cos_gt_zero_pi arctan_lbound arctan_ubound) lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<^sup>2)" proof (rule power2_eq_imp_eq) have "0 < 1 + x\<^sup>2" by (simp add: add_pos_nonneg) show "0 \ 1 / sqrt (1 + x\<^sup>2)" by simp show "0 \ cos (arctan x)" by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound) have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1" unfolding tan_def by (simp add: distrib_left power_divide) then show "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2" using \0 < 1 + x\<^sup>2\ by (simp add: arctan power_divide eq_divide_eq) qed lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)" using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]] using tan_arctan [of x] unfolding tan_def cos_arctan by (simp add: eq_divide_eq) lemma tan_sec: "cos x \ 0 \ 1 + (tan x)\<^sup>2 = (inverse (cos x))\<^sup>2" for x :: "'a::{real_normed_field,banach,field}" by (simp add: add_divide_eq_iff inverse_eq_divide power2_eq_square tan_def) lemma arctan_less_iff: "arctan x < arctan y \ x < y" by (metis tan_monotone' arctan_lbound arctan_ubound tan_arctan) lemma arctan_le_iff: "arctan x \ arctan y \ x \ y" by (simp only: not_less [symmetric] arctan_less_iff) lemma arctan_eq_iff: "arctan x = arctan y \ x = y" by (simp only: eq_iff [where 'a=real] arctan_le_iff) lemma zero_less_arctan_iff [simp]: "0 < arctan x \ 0 < x" using arctan_less_iff [of 0 x] by simp lemma arctan_less_zero_iff [simp]: "arctan x < 0 \ x < 0" using arctan_less_iff [of x 0] by simp lemma zero_le_arctan_iff [simp]: "0 \ arctan x \ 0 \ x" using arctan_le_iff [of 0 x] by simp lemma arctan_le_zero_iff [simp]: "arctan x \ 0 \ x \ 0" using arctan_le_iff [of x 0] by simp lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \ x = 0" using arctan_eq_iff [of x 0] by simp lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin" proof - have "continuous_on (sin ` {- pi/2 .. pi/2}) arcsin" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arcsin_sin) also have "sin ` {- pi/2 .. pi/2} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ sin ` {- pi/2..pi/2}" using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arcsin [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arcsin (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arcsin']] by (auto simp: comp_def subset_eq) lemma isCont_arcsin: "-1 < x \ x < 1 \ isCont arcsin x" using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos" proof - have "continuous_on (cos ` {0 .. pi}) arccos" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arccos_cos) also have "cos ` {0 .. pi} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ cos ` {0..pi}" using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arccos [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arccos (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arccos']] by (auto simp: comp_def subset_eq) lemma isCont_arccos: "-1 < x \ x < 1 \ isCont arccos x" using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma isCont_arctan: "isCont arctan x" proof - obtain u where u: "- (pi / 2) < u" "u < arctan x" by (meson arctan arctan_less_iff linordered_field_no_lb) obtain v where v: "arctan x < v" "v < pi / 2" by (meson arctan_less_iff arctan_ubound linordered_field_no_ub) have "isCont arctan (tan (arctan x))" proof (rule isCont_inverse_function2 [of u "arctan x" v]) show "\z. \u \ z; z \ v\ \ arctan (tan z) = z" using arctan_unique u(1) v(2) by auto then show "\z. \u \ z; z \ v\ \ isCont tan z" by (metis arctan cos_gt_zero_pi isCont_tan less_irrefl) qed (use u v in auto) then show ?thesis by (simp add: arctan) qed lemma tendsto_arctan [tendsto_intros]: "(f \ x) F \ ((\x. arctan (f x)) \ arctan x) F" by (rule isCont_tendsto_compose [OF isCont_arctan]) lemma continuous_arctan [continuous_intros]: "continuous F f \ continuous F (\x. arctan (f x))" unfolding continuous_def by (rule tendsto_arctan) lemma continuous_on_arctan [continuous_intros]: "continuous_on s f \ continuous_on s (\x. arctan (f x))" unfolding continuous_on_def by (auto intro: tendsto_arctan) lemma DERIV_arcsin: assumes "- 1 < x" "x < 1" shows "DERIV arcsin x :> inverse (sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(sin has_real_derivative sqrt (1 - x\<^sup>2)) (at (arcsin x))" by (rule derivative_eq_intros | use assms cos_arcsin in force)+ show "sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arcsin in auto) lemma DERIV_arccos: assumes "- 1 < x" "x < 1" shows "DERIV arccos x :> inverse (- sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(cos has_real_derivative - sqrt (1 - x\<^sup>2)) (at (arccos x))" by (rule derivative_eq_intros | use assms sin_arccos in force)+ show "- sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arccos in auto) lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)" -proof (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"]) - show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))" - apply (rule derivative_eq_intros | simp)+ +proof (rule DERIV_inverse_function) + have "inverse ((cos (arctan x))\<^sup>2) = 1 + x\<^sup>2" by (metis arctan cos_arctan_not_zero power_inverse tan_sec) + then show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))" + by (auto intro!: derivative_eq_intros) show "\y. \x - 1 < y; y < x + 1\ \ tan (arctan y) = y" using tan_arctan by blast show "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) qed (use isCont_arctan in auto) declare DERIV_arcsin[THEN DERIV_chain2, derivative_intros] DERIV_arcsin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arccos[THEN DERIV_chain2, derivative_intros] DERIV_arccos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arctan[THEN DERIV_chain2, derivative_intros] DERIV_arctan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_arctan[derivative_intros] = DERIV_arctan[THEN DERIV_compose_FDERIV] and has_derivative_arccos[derivative_intros] = DERIV_arccos[THEN DERIV_compose_FDERIV] and has_derivative_arcsin[derivative_intros] = DERIV_arcsin[THEN DERIV_compose_FDERIV] lemma filterlim_tan_at_right: "filterlim tan at_bot (at_right (- (pi/2)))" by (rule filterlim_at_bot_at_right[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma filterlim_tan_at_left: "filterlim tan at_top (at_left (pi/2))" by (rule filterlim_at_top_at_left[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma tendsto_arctan_at_top: "(arctan \ (pi/2)) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" define y where "y = pi/2 - min (pi/2) e" then have y: "0 \ y" "y < pi/2" "pi/2 \ e + y" using \0 < e\ by auto show "eventually (\x. dist (arctan x) (pi/2) < e) at_top" proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI) fix x assume "tan y < x" then have "arctan (tan y) < arctan x" by (simp add: arctan_less_iff) with y have "y < arctan x" by (subst (asm) arctan_tan) simp_all with arctan_ubound[of x, arith] y \0 < e\ show "dist (arctan x) (pi/2) < e" by (simp add: dist_real_def) qed qed lemma tendsto_arctan_at_bot: "(arctan \ - (pi/2)) at_bot" unfolding filterlim_at_bot_mirror arctan_minus by (intro tendsto_minus tendsto_arctan_at_top) subsection \Prove Totality of the Trigonometric Functions\ lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" by (simp add: abs_le_iff) lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\<^sup>2)" by (simp add: sin_arccos abs_le_iff) lemma sin_mono_less_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x < sin y \ x < y" by (metis not_less_iff_gr_or_eq sin_monotone_2pi) lemma sin_mono_le_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x \ sin y \ x \ y" by (meson leD le_less_linear sin_monotone_2pi sin_monotone_2pi_le) lemma sin_inj_pi: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x = sin y \ x = y" by (metis arcsin_sin) lemma arcsin_le_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma le_arcsin_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma cos_mono_less_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x < cos y \ y < x" by (meson cos_monotone_0_pi cos_monotone_0_pi_le leD le_less_linear) lemma cos_mono_le_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x \ cos y \ y \ x" by (metis arccos_cos cos_monotone_0_pi_le eq_iff linear) lemma cos_inj_pi: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x = cos y \ x = y" by (metis arccos_cos) lemma arccos_le_pi2: "\0 \ y; y \ 1\ \ arccos y \ pi/2" by (metis (mono_tags) arccos_0 arccos cos_le_one cos_monotone_0_pi_le cos_pi cos_pi_half pi_half_ge_zero antisym_conv less_eq_neg_nonpos linear minus_minus order.trans order_refl) lemma sincos_total_pi_half: assumes "0 \ x" "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi/2 \ x = cos t \ y = sin t" proof - have x1: "x \ 1" using assms by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) with assms have *: "0 \ arccos x" "cos (arccos x) = x" by (auto simp: arccos) from assms have "y = sqrt (1 - x\<^sup>2)" by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs) with x1 * assms arccos_le_pi2 [of x] show ?thesis by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos) qed lemma sincos_total_pi: assumes "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 x]) case le from sincos_total_pi_half [OF le] show ?thesis by (metis pi_ge_two pi_half_le_two add.commute add_le_cancel_left add_mono assms) next case ge then have "0 \ -x" by simp then obtain t where t: "t\0" "t \ pi/2" "-x = cos t" "y = sin t" using sincos_total_pi_half assms by auto (metis \0 \ - x\ power2_minus) show ?thesis by (rule exI [where x = "pi -t"]) (use t in auto) qed lemma sincos_total_2pi_le: assumes "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ 2 * pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 y]) case le from sincos_total_pi [OF le] show ?thesis by (metis assms le_add_same_cancel1 mult.commute mult_2_right order.trans) next case ge then have "0 \ -y" by simp then obtain t where t: "t\0" "t \ pi" "x = cos t" "-y = sin t" using sincos_total_pi assms by auto (metis \0 \ - y\ power2_minus) show ?thesis by (rule exI [where x = "2 * pi - t"]) (use t in auto) qed lemma sincos_total_2pi: assumes "x\<^sup>2 + y\<^sup>2 = 1" obtains t where "0 \ t" "t < 2*pi" "x = cos t" "y = sin t" proof - from sincos_total_2pi_le [OF assms] obtain t where t: "0 \ t" "t \ 2*pi" "x = cos t" "y = sin t" by blast show ?thesis by (cases "t = 2 * pi") (use t that in \force+\) qed lemma arcsin_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x < arcsin y \ x < y" by (rule trans [OF sin_mono_less_eq [symmetric]]) (use arcsin_ubound arcsin_lbound in auto) lemma arcsin_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x \ arcsin y \ x \ y" using arcsin_less_mono not_le by blast lemma arcsin_less_arcsin: "- 1 \ x \ x < y \ y \ 1 \ arcsin x < arcsin y" using arcsin_less_mono by auto lemma arcsin_le_arcsin: "- 1 \ x \ x \ y \ y \ 1 \ arcsin x \ arcsin y" using arcsin_le_mono by auto lemma arccos_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x < arccos y \ y < x" by (rule trans [OF cos_mono_less_eq [symmetric]]) (use arccos_ubound arccos_lbound in auto) lemma arccos_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x \ arccos y \ y \ x" using arccos_less_mono [of y x] by (simp add: not_le [symmetric]) lemma arccos_less_arccos: "- 1 \ x \ x < y \ y \ 1 \ arccos y < arccos x" using arccos_less_mono by auto lemma arccos_le_arccos: "- 1 \ x \ x \ y \ y \ 1 \ arccos y \ arccos x" using arccos_le_mono by auto lemma arccos_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arccos x = arccos y \ x = y" using cos_arccos_abs by fastforce lemma arccos_cos_eq_abs: assumes "\\\ \ pi" shows "arccos (cos \) = \\\" unfolding arccos_def proof (intro the_equality conjI; clarify?) show "cos \\\ = cos \" by (simp add: abs_real_def) show "x = \\\" if "cos x = cos \" "0 \ x" "x \ pi" for x by (simp add: \cos \\\ = cos \\ assms cos_inj_pi that) qed (use assms in auto) lemma arccos_cos_eq_abs_2pi: obtains k where "arccos (cos \) = \\ - of_int k * (2 * pi)\" proof - define k where "k \ \(\ + pi) / (2 * pi)\" have lepi: "\\ - of_int k * (2 * pi)\ \ pi" using floor_divide_lower [of "2*pi" "\ + pi"] floor_divide_upper [of "2*pi" "\ + pi"] by (auto simp: k_def abs_if algebra_simps) have "arccos (cos \) = arccos (cos (\ - of_int k * (2 * pi)))" using cos_int_2pin sin_int_2pin by (simp add: cos_diff mult.commute) also have "\ = \\ - of_int k * (2 * pi)\" using arccos_cos_eq_abs lepi by blast finally show ?thesis using that by metis qed lemma cos_limit_1: assumes "(\j. cos (\ j)) \ 1" shows "\k. (\j. \ j - of_int (k j) * (2 * pi)) \ 0" proof - have "\\<^sub>F j in sequentially. cos (\ j) \ {- 1..1}" by auto then have "(\j. arccos (cos (\ j))) \ arccos 1" using continuous_on_tendsto_compose [OF continuous_on_arccos' assms] by auto moreover have "\j. \k. arccos (cos (\ j)) = \\ j - of_int k * (2 * pi)\" using arccos_cos_eq_abs_2pi by metis then have "\k. \j. arccos (cos (\ j)) = \\ j - of_int (k j) * (2 * pi)\" by metis ultimately have "\k. (\j. \\ j - of_int (k j) * (2 * pi)\) \ 0" by auto then show ?thesis by (simp add: tendsto_rabs_zero_iff) qed lemma cos_diff_limit_1: assumes "(\j. cos (\ j - \)) \ 1" obtains k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - obtain k where "(\j. (\ j - \) - of_int (k j) * (2 * pi)) \ 0" using cos_limit_1 [OF assms] by auto then have "(\j. \ + ((\ j - \) - of_int (k j) * (2 * pi))) \ \ + 0" by (rule tendsto_add [OF tendsto_const]) with that show ?thesis by auto qed subsection \Machin's formula\ lemma arctan_one: "arctan 1 = pi / 4" by (rule arctan_unique) (simp_all add: tan_45 m2pi_less_pi) lemma tan_total_pi4: assumes "\x\ < 1" shows "\z. - (pi / 4) < z \ z < pi / 4 \ tan z = x" proof show "- (pi / 4) < arctan x \ arctan x < pi / 4 \ tan (arctan x) = x" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_less_iff using assms by (auto simp: arctan) qed lemma arctan_add: assumes "\x\ \ 1" "\y\ < 1" shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) have "- (pi / 4) \ arctan x" "- (pi / 4) < arctan y" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 1: "- (pi/2) < arctan x + arctan y" by simp have "arctan x \ pi / 4" "arctan y < pi / 4" unfolding arctan_one [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 2: "arctan x + arctan y < pi/2" by simp show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 1 2] by (simp add: arctan tan_add) qed lemma arctan_double: "\x\ < 1 \ 2 * arctan x = arctan ((2 * x) / (1 - x\<^sup>2))" by (metis arctan_add linear mult_2 not_less power2_eq_square) theorem machin: "pi / 4 = 4 * arctan (1 / 5) - arctan (1 / 239)" proof - have "\1 / 5\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto moreover have "\5 / 12\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto moreover have "\1\ \ (1::real)" and "\1 / 239\ < (1::real)" by auto from arctan_add[OF this] have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" by auto ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" by auto then show ?thesis unfolding arctan_one by algebra qed lemma machin_Euler: "5 * arctan (1 / 7) + 2 * arctan (3 / 79) = pi / 4" proof - have 17: "\1 / 7\ < (1 :: real)" by auto with arctan_double have "2 * arctan (1 / 7) = arctan (7 / 24)" by simp (simp add: field_simps) moreover have "\7 / 24\ < (1 :: real)" by auto with arctan_double have "2 * arctan (7 / 24) = arctan (336 / 527)" by simp (simp add: field_simps) moreover have "\336 / 527\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF 17] this] have "arctan(1/7) + arctan (336 / 527) = arctan (2879 / 3353)" by auto ultimately have I: "5 * arctan (1 / 7) = arctan (2879 / 3353)" by auto have 379: "\3 / 79\ < (1 :: real)" by auto with arctan_double have II: "2 * arctan (3 / 79) = arctan (237 / 3116)" by simp (simp add: field_simps) have *: "\2879 / 3353\ < (1 :: real)" by auto have "\237 / 3116\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF *] this] have "arctan (2879/3353) + arctan (237/3116) = pi/4" by (simp add: arctan_one) with I II show ?thesis by auto qed (*But could also prove MACHIN_GAUSS: 12 * arctan(1/18) + 8 * arctan(1/57) - 5 * arctan(1/239) = pi/4*) subsection \Introducing the inverse tangent power series\ lemma monoseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "monoseq (\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1))" (is "monoseq ?a") proof (cases "x = 0") case True then show ?thesis by (auto simp: monoseq_def) next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "monoseq ?a" proof - have mono: "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" if "0 \ x" and "x \ 1" for n and x :: real proof (rule mult_mono) show "1 / real (Suc (Suc n * 2)) \ 1 / real (Suc (n * 2))" by (rule frac_le) simp_all show "0 \ 1 / real (Suc (n * 2))" by auto show "x ^ Suc (Suc n * 2) \ x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: \0 \ x\ \x \ 1\) show "0 \ x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: \0 \ x\) qed show ?thesis proof (cases "0 \ x") case True from mono[OF this \x \ 1\, THEN allI] show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI2) next case False then have "0 \ - x" and "- x \ 1" using \-1 \ x\ by auto from mono[OF this] have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" for n using \0 \ -x\ by auto then show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI]) qed qed qed lemma zeroseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "(\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1)) \ 0" (is "?a \ 0") proof (cases "x = 0") case True then show ?thesis by simp next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "?a \ 0" proof (cases "\x\ < 1") case True then have "norm x < 1" by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF \norm x < 1\, THEN LIMSEQ_Suc]] have "(\n. 1 / real (n + 1) * x ^ (n + 1)) \ 0" unfolding inverse_eq_divide Suc_eq_plus1 by simp then show ?thesis using pos2 by (rule LIMSEQ_linear) next case False then have "x = -1 \ x = 1" using \\x\ \ 1\ by auto then have n_eq: "\ n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]] show ?thesis unfolding n_eq Suc_eq_plus1 by auto qed qed lemma summable_arctan_series: fixes n :: nat assumes "\x\ \ 1" shows "summable (\ k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)") by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms]) lemma DERIV_arctan_series: assumes "\x\ < 1" shows "DERIV (\x'. \k. (-1)^k * (1 / real (k * 2 + 1) * x' ^ (k * 2 + 1))) x :> (\k. (-1)^k * x^(k * 2))" (is "DERIV ?arctan _ :> ?Int") proof - let ?f = "\n. if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0" have n_even: "even n \ 2 * (n div 2) = n" for n :: nat by presburger then have if_eq: "?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" for n x' by auto have summable_Integral: "summable (\ n. (- 1) ^ n * x^(2 * n))" if "\x\ < 1" for x :: real proof - from that have "x\<^sup>2 < 1" by (simp add: abs_square_less_1) have "summable (\ n. (- 1) ^ n * (x\<^sup>2) ^n)" by (rule summable_Leibniz(1)) (auto intro!: LIMSEQ_realpow_zero monoseq_realpow \x\<^sup>2 < 1\ order_less_imp_le[OF \x\<^sup>2 < 1\]) then show ?thesis by (simp only: power_mult) qed have sums_even: "(sums) f = (sums) (\ n. if even n then f (n div 2) else 0)" for f :: "nat \ real" proof - have "f sums x = (\ n. if even n then f (n div 2) else 0) sums x" for x :: real proof assume "f sums x" from sums_if[OF sums_zero this] show "(\n. if even n then f (n div 2) else 0) sums x" by auto next assume "(\ n. if even n then f (n div 2) else 0) sums x" from LIMSEQ_linear[OF this[simplified sums_def] pos2, simplified sum_split_even_odd[simplified mult.commute]] show "f sums x" unfolding sums_def by auto qed then show ?thesis .. qed have Int_eq: "(\n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult.commute[of _ 2] suminf_def sums_even[of "\ n. (- 1) ^ n * x ^ (2 * n)", symmetric] by auto have arctan_eq: "(\n. ?f n * x^(Suc n)) = ?arctan x" for x proof - have if_eq': "\n. (if even n then (- 1) ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = (if even n then (- 1) ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)" using n_even by auto have idx_eq: "\n. n * 2 + 1 = Suc (2 * n)" by auto then show ?thesis unfolding if_eq' idx_eq suminf_def sums_even[of "\ n. (- 1) ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric] by auto qed have "DERIV (\ x. \ n. ?f n * x^(Suc n)) x :> (\n. ?f n * real (Suc n) * x^n)" proof (rule DERIV_power_series') show "x \ {- 1 <..< 1}" using \\ x \ < 1\ by auto show "summable (\ n. ?f n * real (Suc n) * x'^n)" if x'_bounds: "x' \ {- 1 <..< 1}" for x' :: real proof - from that have "\x'\ < 1" by auto then show ?thesis using that sums_summable sums_if [OF sums_0 [of "\x. 0"] summable_sums [OF summable_Integral]] by (auto simp add: if_distrib [of "\x. x * y" for y] cong: if_cong) qed qed auto then show ?thesis by (simp only: Int_eq arctan_eq) qed lemma arctan_series: assumes "\x\ \ 1" shows "arctan x = (\k. (-1)^k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" (is "_ = suminf (\ n. ?c x n)") proof - let ?c' = "\x n. (-1)^n * x^(n*2)" have DERIV_arctan_suminf: "DERIV (\ x. suminf (?c x)) x :> (suminf (?c' x))" if "0 < r" and "r < 1" and "\x\ < r" for r x :: real proof (rule DERIV_arctan_series) from that show "\x\ < 1" using \r < 1\ and \\x\ < r\ by auto qed { fix x :: real assume "\x\ \ 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] } note arctan_series_borders = this have when_less_one: "arctan x = (\k. ?c x k)" if "\x\ < 1" for x :: real proof - obtain r where "\x\ < r" and "r < 1" using dense[OF \\x\ < 1\] by blast then have "0 < r" and "- r < x" and "x < r" by auto have suminf_eq_arctan_bounded: "suminf (?c x) - arctan x = suminf (?c a) - arctan a" if "-r < a" and "b < r" and "a < b" and "a \ x" and "x \ b" for x a b proof - from that have "\x\ < r" by auto show "suminf (?c x) - arctan x = suminf (?c a) - arctan a" proof (rule DERIV_isconst2[of "a" "b"]) show "a < b" and "a \ x" and "x \ b" using \a < b\ \a \ x\ \x \ b\ by auto have "\x. - r < x \ x < r \ DERIV (\ x. suminf (?c x) - arctan x) x :> 0" proof (rule allI, rule impI) fix x assume "-r < x \ x < r" then have "\x\ < r" by auto with \r < 1\ have "\x\ < 1" by auto have "\- (x\<^sup>2)\ < 1" using abs_square_less_1 \\x\ < 1\ by auto then have "(\n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) then have "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))" unfolding power_mult_distrib[symmetric] power_mult mult.commute[of _ 2] by auto then have suminf_c'_eq_geom: "inverse (1 + x\<^sup>2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto have "DERIV (\ x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))" unfolding suminf_c'_eq_geom by (rule DERIV_arctan_suminf[OF \0 < r\ \r < 1\ \\x\ < r\]) from DERIV_diff [OF this DERIV_arctan] show "DERIV (\x. suminf (?c x) - arctan x) x :> 0" by auto qed then have DERIV_in_rball: "\y. a \ y \ y \ b \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \-r < a\ \b < r\ by auto then show "\y. \a < y; y < b\ \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \\x\ < r\ by auto show "continuous_on {a..b} (\x. suminf (?c x) - arctan x)" using DERIV_in_rball DERIV_atLeastAtMost_imp_continuous_on by blast qed qed have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0" unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto have "suminf (?c x) - arctan x = 0" proof (cases "x = 0") case True then show ?thesis using suminf_arctan_zero by auto next case False then have "0 < \x\" and "- \x\ < \x\" by auto have "suminf (?c (- \x\)) - arctan (- \x\) = suminf (?c 0) - arctan 0" by (rule suminf_eq_arctan_bounded[where x1=0 and a1="-\x\" and b1="\x\", symmetric]) (simp_all only: \\x\ < r\ \-\x\ < \x\\ neg_less_iff_less) moreover have "suminf (?c x) - arctan x = suminf (?c (- \x\)) - arctan (- \x\)" by (rule suminf_eq_arctan_bounded[where x1=x and a1="- \x\" and b1="\x\"]) (simp_all only: \\x\ < r\ \- \x\ < \x\\ neg_less_iff_less) ultimately show ?thesis using suminf_arctan_zero by auto qed then show ?thesis by auto qed show "arctan x = suminf (\n. ?c x n)" proof (cases "\x\ < 1") case True then show ?thesis by (rule when_less_one) next case False then have "\x\ = 1" using \\x\ \ 1\ by auto let ?a = "\x n. \1 / real (n * 2 + 1) * x^(n * 2 + 1)\" let ?diff = "\x n. \arctan x - (\i" have "?diff 1 n \ ?a 1 n" for n :: nat proof - have "0 < (1 :: real)" by auto moreover have "?diff x n \ ?a x n" if "0 < x" and "x < 1" for x :: real proof - from that have "\x\ \ 1" and "\x\ < 1" by auto from \0 < x\ have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto note bounds = mp[OF arctan_series_borders(2)[OF \\x\ \ 1\] this, unfolded when_less_one[OF \\x\ < 1\, symmetric], THEN spec] have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos) (simp_all only: zero_less_power[OF \0 < x\], auto) then have a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos) show ?thesis proof (cases "even n") case True then have sgn_pos: "(-1)^n = (1::real)" by auto from \even n\ obtain m where "n = 2 * m" .. then have "2 * m = n" .. from bounds[of m, unfolded this atLeastAtMost_iff] have "\arctan x - (\i \ (\ii = ?c x n" by auto also have "\ = ?a x n" unfolding sgn_pos a_pos by auto finally show ?thesis . next case False then have sgn_neg: "(-1)^n = (-1::real)" by auto from \odd n\ obtain m where "n = 2 * m + 1" .. then have m_def: "2 * m + 1 = n" .. then have m_plus: "2 * (m + 1) = n + 1" by auto from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2] have "\arctan x - (\i \ (\ii = - ?c x n" by auto also have "\ = ?a x n" unfolding sgn_neg a_pos by auto finally show ?thesis . qed qed hence "\x \ { 0 <..< 1 }. 0 \ ?a x n - ?diff x n" by auto moreover have "isCont (\ x. ?a x n - ?diff x n) x" for x unfolding diff_conv_add_uminus divide_inverse by (auto intro!: isCont_add isCont_rabs continuous_ident isCont_minus isCont_arctan continuous_at_within_inverse isCont_mult isCont_power continuous_const isCont_sum simp del: add_uminus_conv_diff) ultimately have "0 \ ?a 1 n - ?diff 1 n" by (rule LIM_less_bound) then show ?thesis by auto qed have "?a 1 \ 0" unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: of_nat_Suc) have "?diff 1 \ 0" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" obtain N :: nat where N_I: "N \ n \ ?a 1 n < r" for n using LIMSEQ_D[OF \?a 1 \ 0\ \0 < r\] by auto have "norm (?diff 1 n - 0) < r" if "N \ n" for n using \?diff 1 n \ ?a 1 n\ N_I[OF that] by auto then show "\N. \ n \ N. norm (?diff 1 n - 0) < r" by blast qed from this [unfolded tendsto_rabs_zero_iff, THEN tendsto_add [OF _ tendsto_const], of "- arctan 1", THEN tendsto_minus] have "(?c 1) sums (arctan 1)" unfolding sums_def by auto then have "arctan 1 = (\i. ?c 1 i)" by (rule sums_unique) show ?thesis proof (cases "x = 1") case True then show ?thesis by (simp add: \arctan 1 = (\ i. ?c 1 i)\) next case False then have "x = -1" using \\x\ = 1\ by auto have "- (pi/2) < 0" using pi_gt_zero by auto have "- (2 * pi) < 0" using pi_gt_zero by auto have c_minus_minus: "?c (- 1) i = - ?c 1 i" for i by auto have "arctan (- 1) = arctan (tan (-(pi / 4)))" unfolding tan_45 tan_minus .. also have "\ = - (pi / 4)" by (rule arctan_tan) (auto simp: order_less_trans[OF \- (pi/2) < 0\ pi_gt_zero]) also have "\ = - (arctan (tan (pi / 4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric]) (auto simp: order_less_trans[OF \- (2 * pi) < 0\ pi_gt_zero]) also have "\ = - (arctan 1)" unfolding tan_45 .. also have "\ = - (\ i. ?c 1 i)" using \arctan 1 = (\ i. ?c 1 i)\ by auto also have "\ = (\ i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF \(?c 1) sums (arctan 1)\]] unfolding c_minus_minus by auto finally show ?thesis using \x = -1\ by auto qed qed qed lemma arctan_half: "arctan x = 2 * arctan (x / (1 + sqrt(1 + x\<^sup>2)))" for x :: real proof - obtain y where low: "- (pi/2) < y" and high: "y < pi/2" and y_eq: "tan y = x" using tan_total by blast then have low2: "- (pi/2) < y / 2" and high2: "y / 2 < pi/2" by auto have "0 < cos y" by (rule cos_gt_zero_pi[OF low high]) then have "cos y \ 0" and cos_sqrt: "sqrt ((cos y)\<^sup>2) = cos y" by auto have "1 + (tan y)\<^sup>2 = 1 + (sin y)\<^sup>2 / (cos y)\<^sup>2" unfolding tan_def power_divide .. also have "\ = (cos y)\<^sup>2 / (cos y)\<^sup>2 + (sin y)\<^sup>2 / (cos y)\<^sup>2" using \cos y \ 0\ by auto also have "\ = 1 / (cos y)\<^sup>2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 .. finally have "1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2" . have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def using \cos y \ 0\ by (simp add: field_simps) also have "\ = tan y / (1 + 1 / cos y)" using \cos y \ 0\ unfolding add_divide_distrib by auto also have "\ = tan y / (1 + 1 / sqrt ((cos y)\<^sup>2))" unfolding cos_sqrt .. also have "\ = tan y / (1 + sqrt (1 / (cos y)\<^sup>2))" unfolding real_sqrt_divide by auto finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)\<^sup>2))" unfolding \1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2\ . have "arctan x = y" using arctan_tan low high y_eq by auto also have "\ = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto also have "\ = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half by auto finally show ?thesis unfolding eq \tan y = x\ . qed lemma arctan_monotone: "x < y \ arctan x < arctan y" by (simp only: arctan_less_iff) lemma arctan_monotone': "x \ y \ arctan x \ arctan y" by (simp only: arctan_le_iff) lemma arctan_inverse: assumes "x \ 0" shows "arctan (1 / x) = sgn x * pi/2 - arctan x" proof (rule arctan_unique) + have \
: "x > 0 \ arctan x < pi" + using arctan_bounded [of x] by linarith show "- (pi/2) < sgn x * pi/2 - arctan x" - using arctan_bounded [of x] assms - unfolding sgn_real_def - apply (auto simp: arctan algebra_simps) - apply (drule zero_less_arctan_iff [THEN iffD2], arith) - done + using assms by (auto simp: sgn_real_def arctan algebra_simps \
) show "sgn x * pi/2 - arctan x < pi/2" using arctan_bounded [of "- x"] assms - unfolding sgn_real_def arctan_minus - by (auto simp: algebra_simps) + by (auto simp: algebra_simps sgn_real_def arctan_minus) show "tan (sgn x * pi/2 - arctan x) = 1 / x" - unfolding tan_inverse [of "arctan x", unfolded tan_arctan] - unfolding sgn_real_def + unfolding tan_inverse [of "arctan x", unfolded tan_arctan] sgn_real_def by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff) qed theorem pi_series: "pi / 4 = (\k. (-1)^k * 1 / real (k * 2 + 1))" (is "_ = ?SUM") proof - have "pi / 4 = arctan 1" using arctan_one by auto also have "\ = ?SUM" using arctan_series[of 1] by auto finally show ?thesis by auto qed subsection \Existence of Polar Coordinates\ lemma cos_x_y_le_one: "\x / sqrt (x\<^sup>2 + y\<^sup>2)\ \ 1" by (rule power2_le_imp_le [OF _ zero_le_one]) (simp add: power_divide divide_le_eq not_sum_power2_lt_zero) -lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one] - -lemmas sin_arccos_lemma1 = sin_arccos_abs [OF cos_x_y_le_one] - lemma polar_Ex: "\r::real. \a. x = r * cos a \ y = r * sin a" proof - - have polar_ex1: "0 < y \ \r a. x = r * cos a \ y = r * sin a" for y - apply (rule exI [where x = "sqrt (x\<^sup>2 + y\<^sup>2)"]) - apply (rule exI [where x = "arccos (x / sqrt (x\<^sup>2 + y\<^sup>2))"]) - apply (simp add: cos_arccos_lemma1 sin_arccos_lemma1 power_divide - real_sqrt_mult [symmetric] right_diff_distrib) - done + have polar_ex1: "\r a. x = r * cos a \ y = r * sin a" if "0 < y" for y + proof - + have "x = sqrt (x\<^sup>2 + y\<^sup>2) * cos (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" + by (simp add: cos_arccos_abs [OF cos_x_y_le_one]) + moreover have "y = sqrt (x\<^sup>2 + y\<^sup>2) * sin (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" + using that + by (simp add: sin_arccos_abs [OF cos_x_y_le_one] power_divide right_diff_distrib flip: real_sqrt_mult) + ultimately show ?thesis + by blast + qed show ?thesis proof (cases "0::real" y rule: linorder_cases) case less then show ?thesis by (rule polar_ex1) next case equal then show ?thesis by (force simp: intro!: cos_zero sin_zero) next case greater with polar_ex1 [where y="-y"] show ?thesis by auto (metis cos_minus minus_minus minus_mult_right sin_minus) qed qed subsection \Basics about polynomial functions: products, extremal behaviour and root counts\ lemma pairs_le_eq_Sigma: "{(i, j). i + j \ m} = Sigma (atMost m) (\r. atMost (m - r))" for m :: nat by auto lemma sum_up_index_split: "(\k\m + n. f k) = (\k\m. f k) + (\k = Suc m..m + n. f k)" by (metis atLeast0AtMost Suc_eq_plus1 le0 sum.ub_add_nat) lemma Sigma_interval_disjoint: "(SIGMA i:A. {..v i}) \ (SIGMA i:A.{v i<..w}) = {}" for w :: "'a::order" by auto lemma product_atMost_eq_Un: "A \ {..m} = (SIGMA i:A.{..m - i}) \ (SIGMA i:A.{m - i<..m})" for m :: nat by auto lemma polynomial_product: (*with thanks to Chaitanya Mangla*) fixes x :: "'a::idom" assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = - (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" -proof - + (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" +proof - + have "\i j. \m + n - i < j; a i \ 0\ \ b j = 0" + by (meson le_add_diff leI le_less_trans m n) + then have \
: "(\(i,j)\(SIGMA i:{..m+n}. {m+n - i<..m+n}). a i * x ^ i * (b j * x ^ j)) = 0" + by (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral) have "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\i\m. \j\n. (a i * x ^ i) * (b j * x ^ j))" by (rule sum_product) also have "\ = (\i\m + n. \j\n + m. a i * x ^ i * (b j * x ^ j))" using assms by (auto simp: sum_up_index_split) also have "\ = (\r\m + n. \j\m + n - r. a r * x ^ r * (b j * x ^ j))" - apply (simp add: add_ac sum.Sigma product_atMost_eq_Un) - apply (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral) - apply (metis add_diff_assoc2 add.commute add_lessD1 leD m n nat_le_linear neqE) - done + by (simp add: add_ac sum.Sigma product_atMost_eq_Un sum_Un Sigma_interval_disjoint \
) also have "\ = (\(i,j)\{(i,j). i+j \ m+n}. (a i * x ^ i) * (b j * x ^ j))" by (auto simp: pairs_le_eq_Sigma sum.Sigma) + also have "... = (\k\m + n. \i\k. a i * x ^ i * (b (k - i) * x ^ (k - i)))" + by (rule sum.triangle_reindex_eq) also have "\ = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" - apply (subst sum.triangle_reindex_eq) - apply (auto simp: algebra_simps sum_distrib_left intro!: sum.cong) - apply (metis le_add_diff_inverse power_add) - done + by (auto simp: algebra_simps sum_distrib_left simp flip: power_add intro!: sum.cong) finally show ?thesis . qed lemma polynomial_product_nat: fixes x :: nat assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = - (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" + (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" using polynomial_product [of m a n b x] assms by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric] of_nat_eq_iff Int.int_sum [symmetric]) lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" proof - have h: "bij_betw (\(i,j). (j,i)) ((SIGMA i : atMost n. lessThan i)) (SIGMA j : lessThan n. {Suc j..n})" by (auto simp: bij_betw_def inj_on_def) have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: right_diff_distrib sum_subtractf) also have "\ = (\i\n. a i * (x - y) * (\j = (\i\n. \j = (\(i,j) \ (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (\(j,i) \ (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))" by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) also have "\ = (\ji=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" by (simp add: sum_distrib_left mult_ac) finally show ?thesis . qed lemma polyfun_diff_alt: (*COMPLEX_SUB_POLYFUN_ALT in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * ((\jki=Suc j..n. a i * y^(i - j - 1)) = (\ki. i - (j + 1)) {Suc j..n} (lessThan (n-j))" - apply (auto simp: bij_betw_def inj_on_def) - apply (rule_tac x="x + Suc j" in image_eqI, auto) - done + have "\k. k < n - j \ k \ (\i. i - Suc j) ` {Suc j..n}" + by (rule_tac x="k + Suc j" in image_eqI, auto) + then have h: "bij_betw (\i. i - (j + 1)) {Suc j..n} (lessThan (n-j))" + by (auto simp: bij_betw_def inj_on_def) then show ?thesis by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) qed then show ?thesis by (simp add: polyfun_diff [OF assms] sum_distrib_right) qed lemma polyfun_linear_factor: (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*) fixes a :: "'a::idom" shows "\b. \z. (\i\n. c(i) * z^i) = (z - a) * (\ii\n. c(i) * a^i)" proof (cases "n = 0") case True then show ?thesis by simp next case False have "(\b. \z. (\i\n. c i * z^i) = (z - a) * (\ii\n. c i * a^i)) \ (\b. \z. (\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\i \ (\b. \z. (z - a) * (\ji = Suc j..n. c i * a^(i - Suc j)) * z^j) = (z - a) * (\i = True" by auto finally show ?thesis by simp qed lemma polyfun_linear_factor_root: (*COMPLEX_POLYFUN_LINEAR_FACTOR_ROOT in HOL Light*) fixes a :: "'a::idom" assumes "(\i\n. c(i) * a^i) = 0" obtains b where "\z. (\i\n. c i * z^i) = (z - a) * (\iw. \i\n. c i * w^i) a" for c :: "nat \ 'a::real_normed_div_algebra" by simp lemma zero_polynom_imp_zero_coeffs: fixes c :: "nat \ 'a::{ab_semigroup_mult,real_normed_div_algebra}" assumes "\w. (\i\n. c i * w^i) = 0" "k \ n" shows "c k = 0" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc n c k) have [simp]: "c 0 = 0" using Suc.prems(1) [of 0] by simp have "(\i\Suc n. c i * w^i) = w * (\i\n. c (Suc i) * w^i)" for w proof - have "(\i\Suc n. c i * w^i) = (\i\n. c (Suc i) * w ^ Suc i)" unfolding Set_Interval.sum.atMost_Suc_shift by simp also have "\ = w * (\i\n. c (Suc i) * w^i)" by (simp add: sum_distrib_left ac_simps) finally show ?thesis . qed then have w: "\w. w \ 0 \ (\i\n. c (Suc i) * w^i) = 0" using Suc by auto then have "(\h. \i\n. c (Suc i) * h^i) \0\ 0" by (simp cong: LIM_cong) \ \the case \w = 0\ by continuity\ then have "(\i\n. c (Suc i) * 0^i) = 0" using isCont_polynom [of 0 "\i. c (Suc i)" n] LIM_unique by (force simp: Limits.isCont_iff) then have "\w. (\i\n. c (Suc i) * w^i) = 0" using w by metis then have "\i. i \ n \ c (Suc i) = 0" using Suc.IH [of "\i. c (Suc i)"] by blast then show ?case using \k \ Suc n\ by (cases k) auto qed lemma polyfun_rootbound: (*COMPLEX_POLYFUN_ROOTBOUND in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows "finite {z. (\i\n. c(i) * z^i) = 0} \ card {z. (\i\n. c(i) * z^i) = 0} \ n" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc m c k) let ?succase = ?case show ?case proof (cases "{z. (\i\Suc m. c(i) * z^i) = 0} = {}") case True then show ?succase by simp next case False then obtain z0 where z0: "(\i\Suc m. c(i) * z0^i) = 0" by blast then obtain b where b: "\w. (\i\Suc m. c i * w^i) = (w - z0) * (\i\m. b i * w^i)" using polyfun_linear_factor_root [OF z0, unfolded lessThan_Suc_atMost] by blast then have eq: "{z. (\i\Suc m. c i * z^i) = 0} = insert z0 {z. (\i\m. b i * z^i) = 0}" by auto have "\ (\k\m. b k = 0)" proof assume [simp]: "\k\m. b k = 0" then have "\w. (\i\m. b i * w^i) = 0" by simp then have "\w. (\i\Suc m. c i * w^i) = 0" using b by simp then have "\k. k \ Suc m \ c k = 0" using zero_polynom_imp_zero_coeffs by blast then show False using Suc.prems by blast qed then obtain k' where bk': "b k' \ 0" "k' \ m" by blast show ?succase using Suc.IH [of b k'] bk' by (simp add: eq card_insert_if del: sum.atMost_Suc) qed qed lemma fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows polyfun_roots_finite: "finite {z. (\i\n. c(i) * z^i) = 0}" and polyfun_roots_card: "card {z. (\i\n. c(i) * z^i) = 0} \ n" using polyfun_rootbound assms by auto lemma polyfun_finite_roots: (*COMPLEX_POLYFUN_FINITE_ROOTS in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "finite {x. (\i\n. c i * x^i) = 0} \ (\i\n. c i \ 0)" (is "?lhs = ?rhs") proof assume ?lhs moreover have "\ finite {x. (\i\n. c i * x^i) = 0}" if "\i\n. c i = 0" proof - from that have "\x. (\i\n. c i * x^i) = 0" by simp then show ?thesis using ex_new_if_finite [OF infinite_UNIV_char_0 [where 'a='a]] by auto qed ultimately show ?rhs by metis next assume ?rhs with polyfun_rootbound show ?lhs by blast qed lemma polyfun_eq_0: "(\x. (\i\n. c i * x^i) = 0) \ (\i\n. c i = 0)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" (*COMPLEX_POLYFUN_EQ_0 in HOL Light*) using zero_polynom_imp_zero_coeffs by auto lemma polyfun_eq_coeffs: "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\i\n. c i = d i)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" proof - have "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\x. (\i\n. (c i - d i) * x^i) = 0)" by (simp add: left_diff_distrib Groups_Big.sum_subtractf) also have "\ \ (\i\n. c i - d i = 0)" by (rule polyfun_eq_0) finally show ?thesis by simp qed lemma polyfun_eq_const: (*COMPLEX_POLYFUN_EQ_CONST in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "(\x. (\i\n. c i * x^i) = k) \ c 0 = k \ (\i \ {1..n}. c i = 0)" (is "?lhs = ?rhs") proof - have *: "\x. (\i\n. (if i=0 then k else 0) * x^i) = k" by (induct n) auto show ?thesis proof assume ?lhs with * have "(\i\n. c i = (if i=0 then k else 0))" by (simp add: polyfun_eq_coeffs [symmetric]) then show ?rhs by simp next assume ?rhs then show ?lhs by (induct n) auto qed qed lemma root_polyfun: fixes z :: "'a::idom" assumes "1 \ n" shows "z^n = a \ (\i\n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0" using assms by (cases n) (simp_all add: sum.atLeast_Suc_atMost atLeast0AtMost [symmetric]) lemma assumes "SORT_CONSTRAINT('a::{idom,real_normed_div_algebra})" and "1 \ n" shows finite_roots_unity: "finite {z::'a. z^n = 1}" and card_roots_unity: "card {z::'a. z^n = 1} \ n" using polyfun_rootbound [of "\i. if i = 0 then -1 else if i=n then 1 else 0" n n] assms(2) by (auto simp: root_polyfun [OF assms(2)]) subsection \Hyperbolic functions\ definition sinh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "sinh x = (exp x - exp (-x)) /\<^sub>R 2" definition cosh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "cosh x = (exp x + exp (-x)) /\<^sub>R 2" definition tanh :: "'a :: {banach, real_normed_field} \ 'a" where "tanh x = sinh x / cosh x" definition arsinh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arsinh x = ln (x + (x^2 + 1) powr of_real (1/2))" definition arcosh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arcosh x = ln (x + (x^2 - 1) powr of_real (1/2))" definition artanh :: "'a :: {banach, real_normed_field, ln} \ 'a" where "artanh x = ln ((1 + x) / (1 - x)) / 2" lemma arsinh_0 [simp]: "arsinh 0 = 0" by (simp add: arsinh_def) lemma arcosh_1 [simp]: "arcosh 1 = 0" by (simp add: arcosh_def) lemma artanh_0 [simp]: "artanh 0 = 0" by (simp add: artanh_def) lemma tanh_altdef: "tanh x = (exp x - exp (-x)) / (exp x + exp (-x))" proof - have "tanh x = (2 *\<^sub>R sinh x) / (2 *\<^sub>R cosh x)" by (simp add: tanh_def scaleR_conv_of_real) also have "2 *\<^sub>R sinh x = exp x - exp (-x)" by (simp add: sinh_def) also have "2 *\<^sub>R cosh x = exp x + exp (-x)" by (simp add: cosh_def) finally show ?thesis . qed lemma tanh_real_altdef: "tanh (x::real) = (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))" proof - have [simp]: "exp (2 * x) = exp x * exp x" "exp (x * 2) = exp x * exp x" by (subst exp_add [symmetric]; simp)+ have "tanh x = (2 * exp (-x) * sinh x) / (2 * exp (-x) * cosh x)" by (simp add: tanh_def) also have "2 * exp (-x) * sinh x = 1 - exp (-2*x)" by (simp add: exp_minus field_simps sinh_def) also have "2 * exp (-x) * cosh x = 1 + exp (-2*x)" by (simp add: exp_minus field_simps cosh_def) finally show ?thesis . qed lemma sinh_converges: "(\n. if even n then 0 else x ^ n /\<^sub>R fact n) sums sinh x" proof - have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums sinh x" unfolding sinh_def by (intro sums_scaleR_right sums_diff exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then 0 else x ^ n /\<^sub>R fact n)" by auto finally show ?thesis . qed lemma cosh_converges: "(\n. if even n then x ^ n /\<^sub>R fact n else 0) sums cosh x" proof - have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums cosh x" unfolding cosh_def by (intro sums_scaleR_right sums_add exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then x ^ n /\<^sub>R fact n else 0)" by auto finally show ?thesis . qed lemma sinh_0 [simp]: "sinh 0 = 0" by (simp add: sinh_def) lemma cosh_0 [simp]: "cosh 0 = 1" proof - have "cosh 0 = (1/2) *\<^sub>R (1 + 1)" by (simp add: cosh_def) also have "\ = 1" by (rule scaleR_half_double) finally show ?thesis . qed lemma tanh_0 [simp]: "tanh 0 = 0" by (simp add: tanh_def) lemma sinh_minus [simp]: "sinh (- x) = -sinh x" by (simp add: sinh_def algebra_simps) lemma cosh_minus [simp]: "cosh (- x) = cosh x" by (simp add: cosh_def algebra_simps) lemma tanh_minus [simp]: "tanh (-x) = -tanh x" by (simp add: tanh_def) lemma sinh_ln_real: "x > 0 \ sinh (ln x :: real) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus) lemma cosh_ln_real: "x > 0 \ cosh (ln x :: real) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus) lemma tanh_ln_real: "tanh (ln x :: real) = (x ^ 2 - 1) / (x ^ 2 + 1)" if "x > 0" proof - from that have "(x * 2 - inverse x * 2) * (x\<^sup>2 + 1) = (x\<^sup>2 - 1) * (2 * x + 2 * inverse x)" by (simp add: field_simps power2_eq_square) moreover have "x\<^sup>2 + 1 > 0" using that by (simp add: ac_simps add_pos_nonneg) moreover have "2 * x + 2 * inverse x > 0" using that by (simp add: add_pos_pos) ultimately have "(x * 2 - inverse x * 2) / (2 * x + 2 * inverse x) = (x\<^sup>2 - 1) / (x\<^sup>2 + 1)" by (simp add: frac_eq_eq) with that show ?thesis by (simp add: tanh_def sinh_ln_real cosh_ln_real) qed lemma has_field_derivative_scaleR_right [derivative_intros]: "(f has_field_derivative D) F \ ((\x. c *\<^sub>R f x) has_field_derivative (c *\<^sub>R D)) F" unfolding has_field_derivative_def using has_derivative_scaleR_right[of f "\x. D * x" F c] by (simp add: mult_scaleR_left [symmetric] del: mult_scaleR_left) lemma has_field_derivative_sinh [THEN DERIV_chain2, derivative_intros]: "(sinh has_field_derivative cosh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_cosh [THEN DERIV_chain2, derivative_intros]: "(cosh has_field_derivative sinh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_tanh [THEN DERIV_chain2, derivative_intros]: "cosh x \ 0 \ (tanh has_field_derivative 1 - tanh x ^ 2) (at (x :: 'a :: {banach, real_normed_field}))" unfolding tanh_def by (auto intro!: derivative_eq_intros simp: power2_eq_square field_split_simps) lemma has_derivative_sinh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\x. Db * x)) (at x within s)" shows "((\x. sinh (g x)) has_derivative (\y. (cosh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\x. (-Db) * x)" by (simp add: fun_eq_iff) finally have "((\x. sinh (g x)) has_derivative (\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding sinh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (cosh (g x) * Db) * y)" by (simp add: fun_eq_iff cosh_def algebra_simps) finally show ?thesis . qed lemma has_derivative_cosh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\y. Db * y)) (at x within s)" shows "((\x. cosh (g x)) has_derivative (\y. (sinh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\y. (-Db) * y)" by (simp add: fun_eq_iff) finally have "((\x. cosh (g x)) has_derivative (\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding cosh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (sinh (g x) * Db) * y)" by (simp add: fun_eq_iff sinh_def algebra_simps) finally show ?thesis . qed lemma sinh_plus_cosh: "sinh x + cosh x = exp x" proof - have "sinh x + cosh x = (1 / 2) *\<^sub>R (exp x + exp x)" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp x" by (rule scaleR_half_double) finally show ?thesis . qed lemma cosh_plus_sinh: "cosh x + sinh x = exp x" by (subst add.commute) (rule sinh_plus_cosh) lemma cosh_minus_sinh: "cosh x - sinh x = exp (-x)" proof - have "cosh x - sinh x = (1 / 2) *\<^sub>R (exp (-x) + exp (-x))" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp (-x)" by (rule scaleR_half_double) finally show ?thesis . qed lemma sinh_minus_cosh: "sinh x - cosh x = -exp (-x)" using cosh_minus_sinh[of x] by (simp add: algebra_simps) context fixes x :: "'a :: {real_normed_field, banach}" begin lemma sinh_zero_iff: "sinh x = 0 \ exp x \ {1, -1}" by (auto simp: sinh_def field_simps exp_minus power2_eq_square square_eq_1_iff) lemma cosh_zero_iff: "cosh x = 0 \ exp x ^ 2 = -1" by (auto simp: cosh_def exp_minus field_simps power2_eq_square eq_neg_iff_add_eq_0) lemma cosh_square_eq: "cosh x ^ 2 = sinh x ^ 2 + 1" by (simp add: cosh_def sinh_def algebra_simps power2_eq_square exp_add [symmetric] scaleR_conv_of_real) lemma sinh_square_eq: "sinh x ^ 2 = cosh x ^ 2 - 1" by (simp add: cosh_square_eq) lemma hyperbolic_pythagoras: "cosh x ^ 2 - sinh x ^ 2 = 1" by (simp add: cosh_square_eq) lemma sinh_add: "sinh (x + y) = sinh x * cosh y + cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma sinh_diff: "sinh (x - y) = sinh x * cosh y - cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_add: "cosh (x + y) = cosh x * cosh y + sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_diff: "cosh (x - y) = cosh x * cosh y - sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma tanh_add: "tanh (x + y) = (tanh x + tanh y) / (1 + tanh x * tanh y)" if "cosh x \ 0" "cosh y \ 0" proof - have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (cosh x * cosh y + sinh x * sinh y) * ((sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x))" using that by (simp add: field_split_simps) also have "(sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x) = sinh x / cosh x + sinh y / cosh y" using that by (simp add: field_split_simps) finally have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (sinh x / cosh x + sinh y / cosh y) * (cosh x * cosh y + sinh x * sinh y)" by simp then show ?thesis using that by (auto simp add: tanh_def sinh_add cosh_add eq_divide_eq) (simp_all add: field_split_simps) qed lemma sinh_double: "sinh (2 * x) = 2 * sinh x * cosh x" using sinh_add[of x] by simp lemma cosh_double: "cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2" using cosh_add[of x] by (simp add: power2_eq_square) end lemma sinh_field_def: "sinh z = (exp z - exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: sinh_def scaleR_conv_of_real) lemma cosh_field_def: "cosh z = (exp z + exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: cosh_def scaleR_conv_of_real) subsubsection \More specific properties of the real functions\ lemma sinh_real_zero_iff [simp]: "sinh (x::real) = 0 \ x = 0" proof - have "(-1 :: real) < 0" by simp also have "0 < exp x" by simp finally have "exp x \ -1" by (intro notI) simp thus ?thesis by (subst sinh_zero_iff) simp qed lemma plus_inverse_ge_2: fixes x :: real assumes "x > 0" shows "x + inverse x \ 2" proof - have "0 \ (x - 1) ^ 2" by simp also have "\ = x^2 - 2*x + 1" by (simp add: power2_eq_square algebra_simps) finally show ?thesis using assms by (simp add: field_simps power2_eq_square) qed lemma sinh_real_nonneg_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_pos_iff [simp]: "sinh (x :: real) > 0 \ x > 0" by (simp add: sinh_def) lemma sinh_real_nonpos_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_neg_iff [simp]: "sinh (x :: real) < 0 \ x < 0" by (simp add: sinh_def) lemma cosh_real_ge_1: "cosh (x :: real) \ 1" using plus_inverse_ge_2[of "exp x"] by (simp add: cosh_def exp_minus) lemma cosh_real_pos [simp]: "cosh (x :: real) > 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonneg[simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonzero [simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma tanh_real_nonneg_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_pos_iff [simp]: "tanh (x :: real) > 0 \ x > 0" by (simp add: tanh_def field_simps) lemma tanh_real_nonpos_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_neg_iff [simp]: "tanh (x :: real) < 0 \ x < 0" by (simp add: tanh_def field_simps) lemma tanh_real_zero_iff [simp]: "tanh (x :: real) = 0 \ x = 0" by (simp add: tanh_def field_simps) lemma arsinh_real_def: "arsinh (x::real) = ln (x + sqrt (x^2 + 1))" by (simp add: arsinh_def powr_half_sqrt) lemma arcosh_real_def: "x \ 1 \ arcosh (x::real) = ln (x + sqrt (x^2 - 1))" by (simp add: arcosh_def powr_half_sqrt) lemma arsinh_real_aux: "0 < x + sqrt (x ^ 2 + 1 :: real)" proof (cases "x < 0") case True have "(-x) ^ 2 = x ^ 2" by simp also have "x ^ 2 < x ^ 2 + 1" by simp finally have "sqrt ((-x) ^ 2) < sqrt (x ^ 2 + 1)" by (rule real_sqrt_less_mono) thus ?thesis using True by simp qed (auto simp: add_nonneg_pos) lemma arsinh_minus_real [simp]: "arsinh (-x::real) = -arsinh x" proof - have "arsinh (-x) = ln (sqrt (x\<^sup>2 + 1) - x)" by (simp add: arsinh_real_def) also have "sqrt (x^2 + 1) - x = inverse (sqrt (x^2 + 1) + x)" using arsinh_real_aux[of x] by (simp add: field_split_simps algebra_simps power2_eq_square) also have "ln \ = -arsinh x" using arsinh_real_aux[of x] by (simp add: arsinh_real_def ln_inverse) finally show ?thesis . qed lemma artanh_minus_real [simp]: assumes "abs x < 1" shows "artanh (-x::real) = -artanh x" using assms by (simp add: artanh_def ln_div field_simps) lemma sinh_less_cosh_real: "sinh (x :: real) < cosh x" by (simp add: sinh_def cosh_def) lemma sinh_le_cosh_real: "sinh (x :: real) \ cosh x" by (simp add: sinh_def cosh_def) lemma tanh_real_lt_1: "tanh (x :: real) < 1" by (simp add: tanh_def sinh_less_cosh_real) lemma tanh_real_gt_neg1: "tanh (x :: real) > -1" proof - have "- cosh x < sinh x" by (simp add: sinh_def cosh_def field_split_simps) thus ?thesis by (simp add: tanh_def field_simps) qed lemma tanh_real_bounds: "tanh (x :: real) \ {-1<..<1}" using tanh_real_lt_1 tanh_real_gt_neg1 by simp context fixes x :: real begin lemma arsinh_sinh_real: "arsinh (sinh x) = x" by (simp add: arsinh_real_def powr_def sinh_square_eq sinh_plus_cosh) lemma arcosh_cosh_real: "x \ 0 \ arcosh (cosh x) = x" by (simp add: arcosh_real_def powr_def cosh_square_eq cosh_real_ge_1 cosh_plus_sinh) lemma artanh_tanh_real: "artanh (tanh x) = x" proof - have "artanh (tanh x) = ln (cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x))) / 2" by (simp add: artanh_def tanh_def field_split_simps) also have "cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x)) = (cosh x + sinh x) / (cosh x - sinh x)" by simp also have "\ = (exp x)^2" by (simp add: cosh_plus_sinh cosh_minus_sinh exp_minus field_simps power2_eq_square) also have "ln ((exp x)^2) / 2 = x" by (simp add: ln_realpow) finally show ?thesis . qed end lemma sinh_real_strict_mono: "strict_mono (sinh :: real \ real)" by (rule pos_deriv_imp_strict_mono derivative_intros)+ auto lemma cosh_real_strict_mono: assumes "0 \ x" and "x < (y::real)" shows "cosh x < cosh y" proof - from assms have "\z>x. z < y \ cosh y - cosh x = (y - x) * sinh z" by (intro MVT2) (auto dest: connectedD_interval intro!: derivative_eq_intros) then obtain z where z: "z > x" "z < y" "cosh y - cosh x = (y - x) * sinh z" by blast note \cosh y - cosh x = (y - x) * sinh z\ also from \z > x\ and assms have "(y - x) * sinh z > 0" by (intro mult_pos_pos) auto finally show "cosh x < cosh y" by simp qed lemma tanh_real_strict_mono: "strict_mono (tanh :: real \ real)" proof - have *: "tanh x ^ 2 < 1" for x :: real using tanh_real_bounds[of x] by (simp add: abs_square_less_1 abs_if) show ?thesis by (rule pos_deriv_imp_strict_mono) (insert *, auto intro!: derivative_intros) qed lemma sinh_real_abs [simp]: "sinh (abs x :: real) = abs (sinh x)" by (simp add: abs_if) lemma cosh_real_abs [simp]: "cosh (abs x :: real) = cosh x" by (simp add: abs_if) lemma tanh_real_abs [simp]: "tanh (abs x :: real) = abs (tanh x)" by (auto simp: abs_if) lemma sinh_real_eq_iff [simp]: "sinh x = sinh y \ x = (y :: real)" using sinh_real_strict_mono by (simp add: strict_mono_eq) lemma tanh_real_eq_iff [simp]: "tanh x = tanh y \ x = (y :: real)" using tanh_real_strict_mono by (simp add: strict_mono_eq) lemma cosh_real_eq_iff [simp]: "cosh x = cosh y \ abs x = abs (y :: real)" proof - have "cosh x = cosh y \ x = y" if "x \ 0" "y \ 0" for x y :: real using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] that by (cases x y rule: linorder_cases) auto from this[of "abs x" "abs y"] show ?thesis by simp qed lemma sinh_real_le_iff [simp]: "sinh x \ sinh y \ x \ (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less_eq) lemma cosh_real_nonneg_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_nonneg_le_iff[of "-x" "-y"] by simp lemma tanh_real_le_iff [simp]: "tanh x \ tanh y \ x \ (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less_eq) lemma sinh_real_less_iff [simp]: "sinh x < sinh y \ x < (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less) lemma cosh_real_nonneg_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x < (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x > (y::real)" using cosh_real_nonneg_less_iff[of "-x" "-y"] by simp lemma tanh_real_less_iff [simp]: "tanh x < tanh y \ x < (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less) subsubsection \Limits\ lemma sinh_real_at_top: "filterlim (sinh :: real \ real) at_top at_top" proof - have *: "((\x. - exp (- x)) \ (-0::real)) at_top" by (intro tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1 / 2) * (-exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1 / 2) * (-exp (-x) + exp x) :: real) = sinh" by (simp add: fun_eq_iff sinh_def) finally show ?thesis . qed lemma sinh_real_at_bot: "filterlim (sinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x. -sinh x :: real) at_bot at_top" by (simp add: filterlim_uminus_at_top [symmetric] sinh_real_at_top) also have "(\x. -sinh x :: real) = (\x. sinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma cosh_real_at_top: "filterlim (cosh :: real \ real) at_top at_top" proof - have *: "((\x. exp (- x)) \ (0::real)) at_top" by (intro filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1 / 2) * (exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1 / 2) * (exp (-x) + exp x) :: real) = cosh" by (simp add: fun_eq_iff cosh_def) finally show ?thesis . qed lemma cosh_real_at_bot: "filterlim (cosh :: real \ real) at_top at_bot" proof - have "filterlim (\x. cosh (-x) :: real) at_top at_top" by (simp add: cosh_real_at_top) thus ?thesis by (subst filterlim_at_bot_mirror) qed lemma tanh_real_at_top: "(tanh \ (1::real)) at_top" proof - have "((\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) \ (1 - 0) / (1 + 0)) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_ident) auto also have "(\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) = tanh" by (rule ext) (simp add: tanh_real_altdef) finally show ?thesis by simp qed lemma tanh_real_at_bot: "(tanh \ (-1::real)) at_bot" proof - have "((\x::real. -tanh x) \ -1) at_top" by (intro tendsto_minus tanh_real_at_top) also have "(\x. -tanh x :: real) = (\x. tanh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed subsubsection \Properties of the inverse hyperbolic functions\ lemma isCont_sinh: "isCont sinh (x :: 'a :: {real_normed_field, banach})" unfolding sinh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_cosh: "isCont cosh (x :: 'a :: {real_normed_field, banach})" unfolding cosh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_tanh: "cosh x \ 0 \ isCont tanh (x :: 'a :: {real_normed_field, banach})" unfolding tanh_def [abs_def] by (auto intro!: continuous_intros isCont_divide isCont_sinh isCont_cosh) lemma continuous_on_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_on_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_on_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" "\x. x \ A \ cosh (f x) \ 0" shows "continuous_on A (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros) auto lemma continuous_at_within_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous (at x within A) f" "cosh (f x) \ 0" shows "continuous (at x within A) (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma continuous_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" "cosh (f (Lim F (\x. x))) \ 0" shows "continuous F (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma tendsto_sinh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. sinh (f x)) \ sinh a) F" by (rule isCont_tendsto_compose [OF isCont_sinh]) lemma tendsto_cosh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. cosh (f x)) \ cosh a) F" by (rule isCont_tendsto_compose [OF isCont_cosh]) lemma tendsto_tanh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ cosh a \ 0 \ ((\x. tanh (f x)) \ tanh a) F" by (rule isCont_tendsto_compose [OF isCont_tanh]) lemma arsinh_real_has_field_derivative [derivative_intros]: fixes x :: real shows "(arsinh has_field_derivative (1 / (sqrt (x ^ 2 + 1)))) (at x within A)" proof - have pos: "1 + x ^ 2 > 0" by (intro add_pos_nonneg) auto from pos arsinh_real_aux[of x] show ?thesis unfolding arsinh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps) qed lemma arcosh_real_has_field_derivative [derivative_intros]: fixes x :: real assumes "x > 1" shows "(arcosh has_field_derivative (1 / (sqrt (x ^ 2 - 1)))) (at x within A)" proof - from assms have "x + sqrt (x\<^sup>2 - 1) > 0" by (simp add: add_pos_pos) thus ?thesis using assms unfolding arcosh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps power2_eq_1_iff) qed lemma artanh_real_has_field_derivative [derivative_intros]: "(artanh has_field_derivative (1 / (1 - x ^ 2))) (at x within A)" if "\x\ < 1" for x :: real proof - from that have "- 1 < x" "x < 1" by linarith+ hence "(artanh has_field_derivative (4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4)) (at x within A)" unfolding artanh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt) also have "(4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4) = 1 / ((1 + x) * (1 - x))" using \-1 < x\ \x < 1\ by (simp add: frac_eq_eq) also have "(1 + x) * (1 - x) = 1 - x ^ 2" by (simp add: algebra_simps power2_eq_square) finally show ?thesis . qed lemma continuous_on_arsinh [continuous_intros]: "continuous_on A (arsinh :: real \ real)" by (rule DERIV_continuous_on derivative_intros)+ lemma continuous_on_arcosh [continuous_intros]: assumes "A \ {1..}" shows "continuous_on A (arcosh :: real \ real)" proof - have pos: "x + sqrt (x ^ 2 - 1) > 0" if "x \ 1" for x using that by (intro add_pos_nonneg) auto show ?thesis unfolding arcosh_def [abs_def] by (intro continuous_on_subset [OF _ assms] continuous_on_ln continuous_on_add continuous_on_id continuous_on_powr') (auto dest: pos simp: powr_half_sqrt intro!: continuous_intros) qed lemma continuous_on_artanh [continuous_intros]: assumes "A \ {-1<..<1}" shows "continuous_on A (artanh :: real \ real)" unfolding artanh_def [abs_def] by (intro continuous_on_subset [OF _ assms]) (auto intro!: continuous_intros) lemma continuous_on_arsinh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" shows "continuous_on A (\x. arsinh (f x))" by (rule continuous_on_compose2[OF continuous_on_arsinh assms]) auto lemma continuous_on_arcosh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ 1" shows "continuous_on A (\x. arcosh (f x))" by (rule continuous_on_compose2[OF continuous_on_arcosh assms(1) order.refl]) (use assms(2) in auto) lemma continuous_on_artanh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ {-1<..<1}" shows "continuous_on A (\x. artanh (f x))" by (rule continuous_on_compose2[OF continuous_on_artanh assms(1) order.refl]) (use assms(2) in auto) lemma isCont_arsinh [continuous_intros]: "isCont arsinh (x :: real)" using continuous_on_arsinh[of UNIV] by (auto simp: continuous_on_eq_continuous_at) lemma isCont_arcosh [continuous_intros]: assumes "x > 1" shows "isCont arcosh (x :: real)" proof - have "continuous_on {1::real<..} arcosh" by (rule continuous_on_arcosh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma isCont_artanh [continuous_intros]: assumes "x > -1" "x < 1" shows "isCont artanh (x :: real)" proof - have "continuous_on {-1<..<(1::real)} artanh" by (rule continuous_on_artanh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma tendsto_arsinh [tendsto_intros]: "(f \ a) F \ ((\x. arsinh (f x)) \ arsinh a) F" for f :: "_ \ real" by (rule isCont_tendsto_compose [OF isCont_arsinh]) lemma tendsto_arcosh_strong [tendsto_intros]: fixes f :: "_ \ real" assumes "(f \ a) F" "a \ 1" "eventually (\x. f x \ 1) F" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule continuous_on_tendsto_compose[OF continuous_on_arcosh[OF order.refl]]) (use assms in auto) lemma tendsto_arcosh: fixes f :: "_ \ real" assumes "(f \ a) F" "a > 1" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule isCont_tendsto_compose [OF isCont_arcosh]) (use assms in auto) lemma tendsto_arcosh_at_left_1: "(arcosh \ 0) (at_right (1::real))" proof - have "(arcosh \ arcosh 1) (at_right (1::real))" by (rule tendsto_arcosh_strong) (auto simp: eventually_at intro!: exI[of _ 1]) thus ?thesis by simp qed lemma tendsto_artanh [tendsto_intros]: fixes f :: "'a \ real" assumes "(f \ a) F" "a > -1" "a < 1" shows "((\x. artanh (f x)) \ artanh a) F" by (rule isCont_tendsto_compose [OF isCont_artanh]) (use assms in auto) lemma continuous_arsinh [continuous_intros]: "continuous F f \ continuous F (\x. arsinh (f x :: real))" unfolding continuous_def by (rule tendsto_arsinh) (* TODO: This rule does not work for one-sided continuity at 1 *) lemma continuous_arcosh_strong [continuous_intros]: assumes "continuous F f" "eventually (\x. f x \ 1) F" shows "continuous F (\x. arcosh (f x :: real))" proof (cases "F = bot") case False show ?thesis unfolding continuous_def proof (intro tendsto_arcosh_strong) show "1 \ f (Lim F (\x. x))" using assms False unfolding continuous_def by (rule tendsto_lowerbound) qed (insert assms, auto simp: continuous_def) qed auto lemma continuous_arcosh: "continuous F f \ f (Lim F (\x. x)) > 1 \ continuous F (\x. arcosh (f x :: real))" unfolding continuous_def by (rule tendsto_arcosh) auto lemma continuous_artanh [continuous_intros]: "continuous F f \ f (Lim F (\x. x)) \ {-1<..<1} \ continuous F (\x. artanh (f x :: real))" unfolding continuous_def by (rule tendsto_artanh) auto lemma arsinh_real_at_top: "filterlim (arsinh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arsinh_real_def add_ac) lemma arsinh_real_at_bot: "filterlim (arsinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x::real. -arsinh x) at_bot at_top" by (subst filterlim_uminus_at_top [symmetric]) (rule arsinh_real_at_top) also have "(\x::real. -arsinh x) = (\x. arsinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma arcosh_real_at_top: "filterlim (arcosh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (-1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arcosh_real_def) lemma artanh_real_at_left_1: "filterlim (artanh :: real \ real) at_top (at_left 1)" proof - have *: "filterlim (\x::real. (1 + x) / (1 - x)) at_top (at_left 1)" by (rule LIM_at_top_divide) (auto intro!: tendsto_eq_intros eventually_mono[OF eventually_at_left_real[of 0]]) have "filterlim (\x::real. (1/2) * ln ((1 + x) / (1 - x))) at_top (at_left 1)" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] * filterlim_compose[OF ln_at_top]) auto also have "(\x::real. (1/2) * ln ((1 + x) / (1 - x))) = artanh" by (simp add: artanh_def [abs_def]) finally show ?thesis . qed lemma artanh_real_at_right_1: "filterlim (artanh :: real \ real) at_bot (at_right (-1))" proof - have "?thesis \ filterlim (\x::real. -artanh x) at_top (at_right (-1))" by (simp add: filterlim_uminus_at_bot) also have "\ \ filterlim (\x::real. artanh (-x)) at_top (at_right (-1))" by (intro filterlim_cong refl eventually_mono[OF eventually_at_right_real[of "-1" "1"]]) auto also have "\ \ filterlim (artanh :: real \ real) at_top (at_left 1)" by (simp add: filterlim_at_left_to_right) also have \ by (rule artanh_real_at_left_1) finally show ?thesis . qed subsection \Simprocs for root and power literals\ lemma numeral_powr_numeral_real [simp]: "numeral m powr numeral n = (numeral m ^ numeral n :: real)" by (simp add: powr_numeral) context begin private lemma sqrt_numeral_simproc_aux: assumes "m * m \ n" shows "sqrt (numeral n :: real) \ numeral m" proof - have "numeral n \ numeral m * (numeral m :: real)" by (simp add: assms [symmetric]) moreover have "sqrt \ \ numeral m" by (subst real_sqrt_abs2) simp ultimately show "sqrt (numeral n :: real) \ numeral m" by simp qed private lemma root_numeral_simproc_aux: assumes "Num.pow m n \ x" shows "root (numeral n) (numeral x :: real) \ numeral m" by (subst assms [symmetric], subst numeral_pow, subst real_root_pos2) simp_all private lemma powr_numeral_simproc_aux: assumes "Num.pow y n = x" shows "numeral x powr (m / numeral n :: real) \ numeral y powr m" by (subst assms [symmetric], subst numeral_pow, subst powr_numeral [symmetric]) (simp, subst powr_powr, simp_all) private lemma numeral_powr_inverse_eq: "numeral x powr (inverse (numeral n)) = numeral x powr (1 / numeral n :: real)" by simp ML \ signature ROOT_NUMERAL_SIMPROC = sig val sqrt : int option -> int -> int option val sqrt' : int option -> int -> int option val nth_root : int option -> int -> int -> int option val nth_root' : int option -> int -> int -> int option val sqrt_simproc : Proof.context -> cterm -> thm option val root_simproc : int * int -> Proof.context -> cterm -> thm option val powr_simproc : int * int -> Proof.context -> cterm -> thm option end structure Root_Numeral_Simproc : ROOT_NUMERAL_SIMPROC = struct fun iterate NONE p f x = let fun go x = if p x then x else go (f x) in SOME (go x) end | iterate (SOME threshold) p f x = let fun go (threshold, x) = if p x then SOME x else if threshold = 0 then NONE else go (threshold - 1, f x) in go (threshold, x) end fun nth_root _ 1 x = SOME x | nth_root _ _ 0 = SOME 0 | nth_root _ _ 1 = SOME 1 | nth_root threshold n x = let fun newton_step y = ((n - 1) * y + x div Integer.pow (n - 1) y) div n fun is_root y = Integer.pow n y <= x andalso x < Integer.pow n (y + 1) in if x < n then SOME 1 else if x < Integer.pow n 2 then SOME 1 else let val y = Real.floor (Math.pow (Real.fromInt x, Real.fromInt 1 / Real.fromInt n)) in if is_root y then SOME y else iterate threshold is_root newton_step ((x + n - 1) div n) end end fun nth_root' _ 1 x = SOME x | nth_root' _ _ 0 = SOME 0 | nth_root' _ _ 1 = SOME 1 | nth_root' threshold n x = if x < n then NONE else if x < Integer.pow n 2 then NONE else case nth_root threshold n x of NONE => NONE | SOME y => if Integer.pow n y = x then SOME y else NONE fun sqrt _ 0 = SOME 0 | sqrt _ 1 = SOME 1 | sqrt threshold n = let fun aux (a, b) = if n >= b * b then aux (b, b * b) else (a, b) val (lower_root, lower_n) = aux (1, 2) fun newton_step x = (x + n div x) div 2 fun is_sqrt r = r*r <= n andalso n < (r+1)*(r+1) val y = Real.floor (Math.sqrt (Real.fromInt n)) in if is_sqrt y then SOME y else Option.mapPartial (iterate threshold is_sqrt newton_step o (fn x => x * lower_root)) (sqrt threshold (n div lower_n)) end fun sqrt' threshold x = case sqrt threshold x of NONE => NONE | SOME y => if y * y = x then SOME y else NONE fun sqrt_simproc ctxt ct = let val n = ct |> Thm.term_of |> dest_comb |> snd |> dest_comb |> snd |> HOLogic.dest_numeral in case sqrt' (SOME 10000) n of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n]) @{thm sqrt_numeral_simproc_aux}) end handle TERM _ => NONE fun root_simproc (threshold1, threshold2) ctxt ct = let val [n, x] = ct |> Thm.term_of |> strip_comb |> snd |> map (dest_comb #> snd #> HOLogic.dest_numeral) in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n, x]) @{thm root_numeral_simproc_aux}) end handle TERM _ => NONE | Match => NONE fun powr_simproc (threshold1, threshold2) ctxt ct = let val eq_thm = Conv.try_conv (Conv.rewr_conv @{thm numeral_powr_inverse_eq}) ct val ct = Thm.dest_equals_rhs (Thm.cprop_of eq_thm) val (_, [x, t]) = strip_comb (Thm.term_of ct) val (_, [m, n]) = strip_comb t val [x, n] = map (dest_comb #> snd #> HOLogic.dest_numeral) [x, n] in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME y => let val [y, n, x] = map HOLogic.mk_numeral [y, n, x] val thm = Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) [y, n, x, m]) @{thm powr_numeral_simproc_aux} in SOME (@{thm transitive} OF [eq_thm, thm]) end end handle TERM _ => NONE | Match => NONE end \ end simproc_setup sqrt_numeral ("sqrt (numeral n)") = \K Root_Numeral_Simproc.sqrt_simproc\ simproc_setup root_numeral ("root (numeral n) (numeral x)") = \K (Root_Numeral_Simproc.root_simproc (200, Integer.pow 200 2))\ simproc_setup powr_divide_numeral ("numeral x powr (m / numeral n :: real)" | "numeral x powr (inverse (numeral n) :: real)") = \K (Root_Numeral_Simproc.powr_simproc (200, Integer.pow 200 2))\ lemma "root 100 1267650600228229401496703205376 = 2" by simp lemma "sqrt 196 = 14" by simp lemma "256 powr (7 / 4 :: real) = 16384" by simp lemma "27 powr (inverse 3) = (3::real)" by simp end