diff --git a/src/HOL/Algebra/Finite_Extensions.thy b/src/HOL/Algebra/Finite_Extensions.thy --- a/src/HOL/Algebra/Finite_Extensions.thy +++ b/src/HOL/Algebra/Finite_Extensions.thy @@ -1,726 +1,744 @@ (* Title: HOL/Algebra/Finite_Extensions.thy Author: Paulo Emílio de Vilhena *) theory Finite_Extensions imports Embedded_Algebras Polynomials Polynomial_Divisibility begin section \Finite Extensions\ subsection \Definitions\ definition (in ring) transcendental :: "'a set \ 'a \ bool" where "transcendental K x \ inj_on (\p. eval p x) (carrier (K[X]))" abbreviation (in ring) algebraic :: "'a set \ 'a \ bool" where "algebraic K x \ \ transcendental K x" definition (in ring) Irr :: "'a set \ 'a \ 'a list" where "Irr K x = (THE p. p \ carrier (K[X]) \ pirreducible K p \ eval p x = \ \ lead_coeff p = \)" inductive_set (in ring) simple_extension :: "'a set \ 'a \ 'a set" for K and x where zero [simp, intro]: "\ \ simple_extension K x" | lin: "\ k1 \ simple_extension K x; k2 \ K \ \ (k1 \ x) \ k2 \ simple_extension K x" fun (in ring) finite_extension :: "'a set \ 'a list \ 'a set" where "finite_extension K xs = foldr (\x K'. simple_extension K' x) xs K" subsection \Basic Properties\ lemma (in ring) transcendental_consistent: assumes "subring K R" shows "transcendental = ring.transcendental (R \ carrier := K \)" unfolding transcendental_def ring.transcendental_def[OF subring_is_ring[OF assms]] univ_poly_consistent[OF assms] eval_consistent[OF assms] .. lemma (in ring) algebraic_consistent: assumes "subring K R" shows "algebraic = ring.algebraic (R \ carrier := K \)" unfolding over_def transcendental_consistent[OF assms] .. lemma (in ring) eval_transcendental: assumes "(transcendental over K) x" "p \ carrier (K[X])" "eval p x = \" shows "p = []" proof - have "[] \ carrier (K[X])" and "eval [] x = \" by (auto simp add: univ_poly_def) thus ?thesis using assms unfolding over_def transcendental_def inj_on_def by auto qed lemma (in ring) transcendental_imp_trivial_ker: shows "(transcendental over K) x \ a_kernel (K[X]) R (\p. eval p x) = { [] }" using eval_transcendental unfolding a_kernel_def' by (auto simp add: univ_poly_def) lemma (in ring) non_trivial_ker_imp_algebraic: shows "a_kernel (K[X]) R (\p. eval p x) \ { [] } \ (algebraic over K) x" using transcendental_imp_trivial_ker unfolding over_def by auto lemma (in domain) trivial_ker_imp_transcendental: assumes "subring K R" and "x \ carrier R" shows "a_kernel (K[X]) R (\p. eval p x) = { [] } \ (transcendental over K) x" using ring_hom_ring.trivial_ker_imp_inj[OF eval_ring_hom[OF assms]] unfolding transcendental_def over_def by (simp add: univ_poly_zero) lemma (in domain) algebraic_imp_non_trivial_ker: assumes "subring K R" and "x \ carrier R" shows "(algebraic over K) x \ a_kernel (K[X]) R (\p. eval p x) \ { [] }" using trivial_ker_imp_transcendental[OF assms] unfolding over_def by auto lemma (in domain) algebraicE: assumes "subring K R" and "x \ carrier R" "(algebraic over K) x" obtains p where "p \ carrier (K[X])" "p \ []" "eval p x = \" proof - have "[] \ a_kernel (K[X]) R (\p. eval p x)" unfolding a_kernel_def' univ_poly_def by auto then obtain p where "p \ carrier (K[X])" "p \ []" "eval p x = \" using algebraic_imp_non_trivial_ker[OF assms] unfolding a_kernel_def' by blast thus thesis using that by auto qed lemma (in ring) algebraicI: assumes "p \ carrier (K[X])" "p \ []" and "eval p x = \" shows "(algebraic over K) x" using assms non_trivial_ker_imp_algebraic unfolding a_kernel_def' by auto lemma (in ring) transcendental_mono: assumes "K \ K'" "(transcendental over K') x" shows "(transcendental over K) x" proof - have "carrier (K[X]) \ carrier (K'[X])" using assms(1) unfolding univ_poly_def polynomial_def by auto thus ?thesis using assms unfolding over_def transcendental_def by (metis inj_on_subset) qed corollary (in ring) algebraic_mono: assumes "K \ K'" "(algebraic over K) x" shows "(algebraic over K') x" using transcendental_mono[OF assms(1)] assms(2) unfolding over_def by blast lemma (in domain) zero_is_algebraic: assumes "subring K R" shows "(algebraic over K) \" using algebraicI[OF var_closed(1)[OF assms]] unfolding var_def by auto lemma (in domain) algebraic_self: assumes "subring K R" and "k \ K" shows "(algebraic over K) k" proof (rule algebraicI[of "[ \, \ k ]"]) show "[ \, \ k ] \ carrier (K [X])" and "[ \, \ k ] \ []" using subringE(2-3,5)[OF assms(1)] assms(2) unfolding univ_poly_def polynomial_def by auto have "k \ carrier R" using subringE(1)[OF assms(1)] assms(2) by auto thus "eval [ \, \ k ] k = \" by (auto, algebra) qed lemma (in domain) ker_diff_carrier: assumes "subring K R" shows "a_kernel (K[X]) R (\p. eval p x) \ carrier (K[X])" proof - have "eval [ \ ] x \ \" and "[ \ ] \ carrier (K[X])" using subringE(3)[OF assms] unfolding univ_poly_def polynomial_def by auto thus ?thesis unfolding a_kernel_def' by blast qed subsection \Minimal Polynomial\ lemma (in domain) minimal_polynomial_is_unique: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" shows "\!p \ carrier (K[X]). pirreducible K p \ eval p x = \ \ lead_coeff p = \" (is "\!p. ?minimal_poly p") proof - interpret UP: principal_domain "K[X]" using univ_poly_is_principal[OF assms(1)] . let ?ker_gen = "\p. p \ carrier (K[X]) \ pirreducible K p \ lead_coeff p = \ \ a_kernel (K[X]) R (\p. eval p x) = PIdl\<^bsub>K[X]\<^esub> p" obtain p where p: "?ker_gen p" and unique: "\q. ?ker_gen q \ q = p" using exists_unique_pirreducible_gen[OF assms(1) eval_ring_hom[OF _ assms(2)] algebraic_imp_non_trivial_ker[OF _ assms(2-3)] ker_diff_carrier] subfieldE(1)[OF assms(1)] by auto hence "?minimal_poly p" using UP.cgenideal_self p unfolding a_kernel_def' by auto moreover have "\q. ?minimal_poly q \ q = p" proof - fix q assume q: "?minimal_poly q" then have "q \ PIdl\<^bsub>K[X]\<^esub> p" using p unfolding a_kernel_def' by auto hence "p \\<^bsub>K[X]\<^esub> q" using cgenideal_pirreducible[OF assms(1)] p q by simp hence "a_kernel (K[X]) R (\p. eval p x) = PIdl\<^bsub>K[X]\<^esub> q" using UP.associated_iff_same_ideal q p by simp thus "q = p" using unique q by simp qed ultimately show ?thesis by blast qed lemma (in domain) IrrE: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" shows "Irr K x \ carrier (K[X])" and "pirreducible K (Irr K x)" and "lead_coeff (Irr K x) = \" and "eval (Irr K x) x = \" using theI'[OF minimal_polynomial_is_unique[OF assms]] unfolding Irr_def by auto lemma (in domain) Irr_generates_ker: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" shows "a_kernel (K[X]) R (\p. eval p x) = PIdl\<^bsub>K[X]\<^esub> (Irr K x)" proof - obtain q where q: "q \ carrier (K[X])" "pirreducible K q" and ker: "a_kernel (K[X]) R (\p. eval p x) = PIdl\<^bsub>K[X]\<^esub> q" using exists_unique_pirreducible_gen[OF assms(1) eval_ring_hom[OF _ assms(2)] algebraic_imp_non_trivial_ker[OF _ assms(2-3)] ker_diff_carrier] subfieldE(1)[OF assms(1)] by auto have "Irr K x \ PIdl\<^bsub>K[X]\<^esub> q" using IrrE(1,4)[OF assms] ker unfolding a_kernel_def' by auto thus ?thesis using cgenideal_pirreducible[OF assms(1) q(1-2) IrrE(2)[OF assms]] q(1) IrrE(1)[OF assms] cring.associated_iff_same_ideal[OF univ_poly_is_cring[OF subfieldE(1)[OF assms(1)]]] unfolding ker by simp qed lemma (in domain) Irr_minimal: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" and "p \ carrier (K[X])" "eval p x = \" shows "(Irr K x) pdivides p" proof - interpret UP: principal_domain "K[X]" using univ_poly_is_principal[OF assms(1)] . have "p \ PIdl\<^bsub>K[X]\<^esub> (Irr K x)" using Irr_generates_ker[OF assms(1-3)] assms(4-5) unfolding a_kernel_def' by auto hence "(Irr K x) divides\<^bsub>K[X]\<^esub> p" using UP.to_contain_is_to_divide IrrE(1)[OF assms(1-3)] by (meson UP.cgenideal_ideal UP.cgenideal_minimal assms(4)) thus ?thesis unfolding pdivides_iff_shell[OF assms(1) IrrE(1)[OF assms(1-3)] assms(4)] . qed lemma (in domain) rupture_of_Irr: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" shows "field (Rupt K (Irr K x))" using rupture_is_field_iff_pirreducible[OF assms(1)] IrrE(1-2)[OF assms] by simp subsection \Simple Extensions\ lemma (in ring) simple_extension_consistent: assumes "subring K R" shows "ring.simple_extension (R \ carrier := K \) = simple_extension" proof - interpret K: ring "R \ carrier := K \" using subring_is_ring[OF assms] . have "\K' x. K.simple_extension K' x \ simple_extension K' x" proof fix K' x a show "a \ K.simple_extension K' x \ a \ simple_extension K' x" by (induction rule: K.simple_extension.induct) (auto simp add: simple_extension.lin) qed moreover have "\K' x. simple_extension K' x \ K.simple_extension K' x" proof fix K' x a assume a: "a \ simple_extension K' x" thus "a \ K.simple_extension K' x" using K.simple_extension.zero K.simple_extension.lin by (induction rule: simple_extension.induct) (simp)+ qed ultimately show ?thesis by blast qed lemma (in ring) mono_simple_extension: assumes "K \ K'" shows "simple_extension K x \ simple_extension K' x" proof fix a assume "a \ simple_extension K x" thus "a \ simple_extension K' x" proof (induct a rule: simple_extension.induct, simp) case lin thus ?case using simple_extension.lin assms by blast qed qed lemma (in ring) simple_extension_incl: assumes "K \ carrier R" and "x \ carrier R" shows "K \ simple_extension K x" proof fix k assume "k \ K" thus "k \ simple_extension K x" using simple_extension.lin[OF simple_extension.zero, of k K x] assms by auto qed lemma (in ring) simple_extension_mem: assumes "subring K R" and "x \ carrier R" shows "x \ simple_extension K x" proof - have "\ \ simple_extension K x" using simple_extension_incl[OF _ assms(2)] subringE(1,3)[OF assms(1)] by auto thus ?thesis using simple_extension.lin[OF _ subringE(2)[OF assms(1)], of \ x] assms(2) by auto qed lemma (in ring) simple_extension_carrier: assumes "x \ carrier R" shows "simple_extension (carrier R) x = carrier R" proof show "carrier R \ simple_extension (carrier R) x" using simple_extension_incl[OF _ assms] by auto next show "simple_extension (carrier R) x \ carrier R" proof fix a assume "a \ simple_extension (carrier R) x" thus "a \ carrier R" by (induct a rule: simple_extension.induct) (auto simp add: assms) qed qed lemma (in ring) simple_extension_in_carrier: assumes "K \ carrier R" and "x \ carrier R" shows "simple_extension K x \ carrier R" using mono_simple_extension[OF assms(1), of x] simple_extension_carrier[OF assms(2)] by auto lemma (in ring) simple_extension_subring_incl: assumes "subring K' R" and "K \ K'" "x \ K'" shows "simple_extension K x \ K'" using ring.simple_extension_in_carrier[OF subring_is_ring[OF assms(1)]] assms(2-3) unfolding simple_extension_consistent[OF assms(1)] by simp lemma (in ring) simple_extension_as_eval_img: assumes "K \ carrier R" "x \ carrier R" shows "simple_extension K x = (\p. eval p x) ` carrier (K[X])" proof show "simple_extension K x \ (\p. eval p x) ` carrier (K[X])" proof fix a assume "a \ simple_extension K x" thus "a \ (\p. eval p x) ` carrier (K[X])" proof (induction rule: simple_extension.induct) case zero have "polynomial K []" and "eval [] x = \" unfolding polynomial_def by simp+ thus ?case unfolding univ_poly_carrier by force next case (lin k1 k2) then obtain p where p: "p \ carrier (K[X])" "polynomial K p" "eval p x = k1" by (auto simp add: univ_poly_carrier) hence "set p \ carrier R" and "k2 \ carrier R" using assms(1) lin(2) unfolding polynomial_def by auto hence "eval (normalize (p @ [ k2 ])) x = k1 \ x \ k2" using eval_append_aux[of p k2 x] eval_normalize[of "p @ [ k2 ]" x] assms(2) p(3) by auto moreover have "set (p @ [k2]) \ K" using polynomial_incl[OF p(2)] \k2 \ K\ by auto then have "local.normalize (p @ [k2]) \ carrier (K [X])" using normalize_gives_polynomial univ_poly_carrier by blast ultimately show ?case unfolding univ_poly_carrier by force qed qed next show "(\p. eval p x) ` carrier (K[X]) \ simple_extension K x" proof fix a assume "a \ (\p. eval p x) ` carrier (K[X])" then obtain p where p: "set p \ K" "eval p x = a" using polynomial_incl unfolding univ_poly_def by auto thus "a \ simple_extension K x" proof (induct "length p" arbitrary: p a) case 0 thus ?case using simple_extension.zero by simp next case (Suc n) obtain p' k where p: "p = p' @ [ k ]" using Suc(2) by (metis list.size(3) nat.simps(3) rev_exhaust) hence "a = (eval p' x) \ x \ k" using eval_append_aux[of p' k x] Suc(3-4) assms unfolding p by auto moreover have "eval p' x \ simple_extension K x" using Suc(1-3) unfolding p by auto ultimately show ?case using simple_extension.lin Suc(3) unfolding p by auto qed qed qed corollary (in domain) simple_extension_is_subring: assumes "subring K R" "x \ carrier R" shows "subring (simple_extension K x) R" using ring_hom_ring.img_is_subring[OF eval_ring_hom[OF assms] ring.carrier_is_subring[OF univ_poly_is_ring[OF assms(1)]]] simple_extension_as_eval_img[OF subringE(1)[OF assms(1)] assms(2)] by simp corollary (in domain) simple_extension_minimal: assumes "subring K R" "x \ carrier R" shows "simple_extension K x = \ { K'. subring K' R \ K \ K' \ x \ K' }" using simple_extension_is_subring[OF assms] simple_extension_mem[OF assms] simple_extension_incl[OF subringE(1)[OF assms(1)] assms(2)] simple_extension_subring_incl by blast corollary (in domain) simple_extension_isomorphism: assumes "subring K R" "x \ carrier R" shows "(K[X]) Quot (a_kernel (K[X]) R (\p. eval p x)) \ R \ carrier := simple_extension K x \" using ring_hom_ring.FactRing_iso_set_aux[OF eval_ring_hom[OF assms]] simple_extension_as_eval_img[OF subringE(1)[OF assms(1)] assms(2)] unfolding is_ring_iso_def by auto corollary (in domain) simple_extension_of_algebraic: assumes "subfield K R" and "x \ carrier R" "(algebraic over K) x" shows "Rupt K (Irr K x) \ R \ carrier := simple_extension K x \" using simple_extension_isomorphism[OF subfieldE(1)[OF assms(1)] assms(2)] unfolding Irr_generates_ker[OF assms] rupture_def by simp corollary (in domain) simple_extension_of_transcendental: assumes "subring K R" and "x \ carrier R" "(transcendental over K) x" shows "K[X] \ R \ carrier := simple_extension K x \" using simple_extension_isomorphism[OF _ assms(2), of K] assms(1) ring_iso_trans[OF ring.FactRing_zeroideal(2)[OF univ_poly_is_ring]] unfolding transcendental_imp_trivial_ker[OF assms(3)] univ_poly_zero by auto proposition (in domain) simple_extension_subfield_imp_algebraic: assumes "subring K R" "x \ carrier R" shows "subfield (simple_extension K x) R \ (algebraic over K) x" proof - assume simple_ext: "subfield (simple_extension K x) R" show "(algebraic over K) x" proof (rule ccontr) assume "\ (algebraic over K) x" then have "(transcendental over K) x" unfolding over_def by simp then obtain h where h: "h \ ring_iso (R \ carrier := simple_extension K x \) (K[X])" using ring_iso_sym[OF univ_poly_is_ring simple_extension_of_transcendental] assms unfolding is_ring_iso_def by blast then interpret Hom: ring_hom_ring "R \ carrier := simple_extension K x \" "K[X]" h using subring_is_ring[OF simple_extension_is_subring[OF assms]] univ_poly_is_ring[OF assms(1)] assms h by (auto simp add: ring_hom_ring_def ring_hom_ring_axioms_def ring_iso_def) have "field (K[X])" using field.ring_iso_imp_img_field[OF subfield_iff(2)[OF simple_ext] h] unfolding Hom.hom_one Hom.hom_zero by simp moreover have "\ field (K[X])" using univ_poly_not_field[OF assms(1)] . ultimately show False by simp qed qed proposition (in domain) simple_extension_is_subfield: assumes "subfield K R" "x \ carrier R" shows "subfield (simple_extension K x) R \ (algebraic over K) x" proof assume alg: "(algebraic over K) x" then obtain h where h: "h \ ring_iso (Rupt K (Irr K x)) (R \ carrier := simple_extension K x \)" using simple_extension_of_algebraic[OF assms] unfolding is_ring_iso_def by blast have rupt_field: "field (Rupt K (Irr K x))" and "ring (R \ carrier := simple_extension K x \)" using subring_is_ring[OF simple_extension_is_subring[OF subfieldE(1)]] rupture_of_Irr[OF assms alg] assms by simp+ then interpret Hom: ring_hom_ring "Rupt K (Irr K x)" "R \ carrier := simple_extension K x \" h using h cring.axioms(1)[OF domain.axioms(1)[OF field.axioms(1)]] by (auto simp add: ring_hom_ring_def ring_hom_ring_axioms_def ring_iso_def) show "subfield (simple_extension K x) R" using field.ring_iso_imp_img_field[OF rupt_field h] subfield_iff(1)[OF _ simple_extension_in_carrier[OF subfieldE(3)[OF assms(1)] assms(2)]] by simp next assume simple_ext: "subfield (simple_extension K x) R" thus "(algebraic over K) x" using simple_extension_subfield_imp_algebraic[OF subfieldE(1)[OF assms(1)] assms(2)] by simp qed subsection \Link between dimension of K-algebras and algebraic extensions\ lemma (in domain) exp_base_independent: assumes "subfield K R" "x \ carrier R" "(algebraic over K) x" shows "independent K (exp_base x (degree (Irr K x)))" proof - have "\n. n \ degree (Irr K x) \ independent K (exp_base x n)" proof - fix n show "n \ degree (Irr K x) \ independent K (exp_base x n)" proof (induct n, simp add: exp_base_def) case (Suc n) have "x [^] n \ Span K (exp_base x n)" proof (rule ccontr) assume "\ x [^] n \ Span K (exp_base x n)" then obtain a Ks where Ks: "a \ K - { \ }" "set Ks \ K" "length Ks = n" "combine (a # Ks) (exp_base x (Suc n)) = \" using Span_mem_imp_non_trivial_combine[OF assms(1) exp_base_closed[OF assms(2), of n]] by (auto simp add: exp_base_def) hence "eval (a # Ks) x = \" using combine_eq_eval by (auto simp add: exp_base_def) moreover have "(a # Ks) \ carrier (K[X]) - { [] }" unfolding univ_poly_def polynomial_def using Ks(1-2) by auto ultimately have "degree (Irr K x) \ n" using pdivides_imp_degree_le[OF subfieldE(1)[OF assms(1)] IrrE(1)[OF assms] _ _ Irr_minimal[OF assms, of "a # Ks"]] Ks(3) by auto from \Suc n \ degree (Irr K x)\ and this show False by simp qed thus ?case using independent.li_Cons assms(2) Suc by (auto simp add: exp_base_def) qed qed thus ?thesis by simp qed lemma (in ring) Span_eq_eval_img: assumes "subfield K R" "x \ carrier R" shows "Span K (exp_base x n) = (\p. eval p x) ` { p \ carrier (K[X]). length p \ n }" (is "?Span = ?eval_img") proof show "?Span \ ?eval_img" proof fix u assume "u \ Span K (exp_base x n)" then obtain Ks where Ks: "set Ks \ K" "length Ks = n" "u = combine Ks (exp_base x n)" using Span_eq_combine_set_length_version[OF assms(1) exp_base_closed[OF assms(2)]] by (auto simp add: exp_base_def) hence "u = eval (normalize Ks) x" using combine_eq_eval eval_normalize[OF _ assms(2)] subfieldE(3)[OF assms(1)] by auto moreover have "normalize Ks \ carrier (K[X])" using normalize_gives_polynomial[OF Ks(1)] unfolding univ_poly_def by auto moreover have "length (normalize Ks) \ n" using normalize_length_le[of Ks] Ks(2) by auto ultimately show "u \ ?eval_img" by auto qed next show "?eval_img \ ?Span" proof fix u assume "u \ ?eval_img" then obtain p where p: "p \ carrier (K[X])" "length p \ n" "u = eval p x" by blast hence "combine p (exp_base x (length p)) = u" using combine_eq_eval by auto moreover have set_p: "set p \ K" using polynomial_incl[of K p] p(1) unfolding univ_poly_carrier by auto hence "set p \ carrier R" using subfieldE(3)[OF assms(1)] by auto moreover have "drop (n - length p) (exp_base x n) = exp_base x (length p)" using p(2) drop_exp_base by auto ultimately have "combine ((replicate (n - length p) \) @ p) (exp_base x n) = u" using combine_prepend_replicate[OF _ exp_base_closed[OF assms(2), of n]] by auto moreover have "set ((replicate (n - length p) \) @ p) \ K" using subringE(2)[OF subfieldE(1)[OF assms(1)]] set_p by auto ultimately show "u \ ?Span" using Span_eq_combine_set[OF assms(1) exp_base_closed[OF assms(2), of n]] by blast qed qed lemma (in domain) Span_exp_base: assumes "subfield K R" "x \ carrier R" "(algebraic over K) x" shows "Span K (exp_base x (degree (Irr K x))) = simple_extension K x" unfolding simple_extension_as_eval_img[OF subfieldE(3)[OF assms(1)] assms(2)] Span_eq_eval_img[OF assms(1-2)] proof (auto) interpret UP: principal_domain "K[X]" using univ_poly_is_principal[OF assms(1)] . note hom_simps = ring_hom_memE[OF eval_is_hom[OF subfieldE(1)[OF assms(1)] assms(2)]] fix p assume p: "p \ carrier (K[X])" have Irr: "Irr K x \ carrier (K[X])" "Irr K x \ []" using IrrE(1-2)[OF assms] unfolding ring_irreducible_def univ_poly_zero by auto then obtain q r where q: "q \ carrier (K[X])" and r: "r \ carrier (K[X])" and dvd: "p = Irr K x \\<^bsub>K [X]\<^esub> q \\<^bsub>K [X]\<^esub> r" "r = [] \ degree r < degree (Irr K x)" using subfield_long_division_theorem_shell[OF assms(1) p Irr(1)] unfolding univ_poly_zero by auto hence "eval p x = (eval (Irr K x) x) \ (eval q x) \ (eval r x)" using hom_simps(2-3) Irr(1) by simp hence "eval p x = eval r x" using hom_simps(1) q r unfolding IrrE(4)[OF assms] by simp moreover have "length r < length (Irr K x)" using dvd(2) Irr(2) by auto ultimately show "eval p x \ (\p. local.eval p x) ` { p \ carrier (K [X]). length p \ length (Irr K x) - Suc 0 }" using r by auto qed corollary (in domain) dimension_simple_extension: assumes "subfield K R" "x \ carrier R" "(algebraic over K) x" shows "dimension (degree (Irr K x)) K (simple_extension K x)" using dimension_independent[OF exp_base_independent[OF assms]] Span_exp_base[OF assms] by (simp add: exp_base_def) lemma (in ring) finite_dimension_imp_algebraic: assumes "subfield K R" "subring F R" and "finite_dimension K F" shows "x \ F \ (algebraic over K) x" proof - let ?Us = "\n. map (\i. x [^] i) (rev [0..< Suc n])" assume x: "x \ F" then have in_carrier: "x \ carrier R" using subringE[OF assms(2)] by auto obtain n where n: "dimension n K F" using assms(3) by auto have set_Us: "set (?Us n) \ F" using x subringE(3,6)[OF assms(2)] by (induct n) (auto) hence "set (?Us n) \ carrier R" using subringE(1)[OF assms(2)] by auto moreover have "dependent K (?Us n)" using independent_length_le_dimension[OF assms(1) n _ set_Us] by auto ultimately obtain Ks where Ks: "length Ks = Suc n" "combine Ks (?Us n) = \" "set Ks \ K" "set Ks \ { \ }" using dependent_imp_non_trivial_combine[OF assms(1), of "?Us n"] by auto have "set Ks \ carrier R" using subring_props(1)[OF assms(1)] Ks(3) by auto hence "eval (normalize Ks) x = \" using combine_eq_eval[of Ks] eval_normalize[OF _ in_carrier] Ks(1-2) by (simp add: exp_base_def) moreover have "normalize Ks = [] \ set Ks \ { \ }" by (induct Ks) (auto, meson list.discI, metis all_not_in_conv list.discI list.sel(3) singletonD subset_singletonD) hence "normalize Ks \ []" using Ks(1,4) by (metis list.size(3) nat.distinct(1) set_empty subset_singleton_iff) moreover have "normalize Ks \ carrier (K[X])" using normalize_gives_polynomial[OF Ks(3)] unfolding univ_poly_def by auto ultimately show ?thesis using algebraicI by auto qed corollary (in domain) simple_extension_dim: assumes "subfield K R" "x \ carrier R" "(algebraic over K) x" shows "(dim over K) (simple_extension K x) = degree (Irr K x)" using dimI[OF assms(1) dimension_simple_extension[OF assms]] . corollary (in domain) finite_dimension_simple_extension: assumes "subfield K R" "x \ carrier R" shows "finite_dimension K (simple_extension K x) \ (algebraic over K) x" using finite_dimensionI[OF dimension_simple_extension[OF assms]] finite_dimension_imp_algebraic[OF _ simple_extension_is_subring[OF subfieldE(1)]] simple_extension_mem[OF subfieldE(1)] assms by auto subsection \Finite Extensions\ lemma (in ring) finite_extension_consistent: assumes "subring K R" shows "ring.finite_extension (R \ carrier := K \) = finite_extension" proof - have "\K' xs. ring.finite_extension (R \ carrier := K \) K' xs = finite_extension K' xs" proof - fix K' xs show "ring.finite_extension (R \ carrier := K \) K' xs = finite_extension K' xs" using ring.finite_extension.simps[OF subring_is_ring[OF assms]] simple_extension_consistent[OF assms] by (induct xs) (auto) qed thus ?thesis by blast qed lemma (in ring) mono_finite_extension: assumes "K \ K'" shows "finite_extension K xs \ finite_extension K' xs" using mono_simple_extension assms by (induct xs) (auto) lemma (in ring) finite_extension_carrier: assumes "set xs \ carrier R" shows "finite_extension (carrier R) xs = carrier R" using assms simple_extension_carrier by (induct xs) (auto) lemma (in ring) finite_extension_in_carrier: assumes "K \ carrier R" and "set xs \ carrier R" shows "finite_extension K xs \ carrier R" using assms simple_extension_in_carrier by (induct xs) (auto) lemma (in ring) finite_extension_subring_incl: assumes "subring K' R" and "K \ K'" "set xs \ K'" shows "finite_extension K xs \ K'" using ring.finite_extension_in_carrier[OF subring_is_ring[OF assms(1)]] assms(2-3) unfolding finite_extension_consistent[OF assms(1)] by simp lemma (in ring) finite_extension_incl_aux: assumes "K \ carrier R" and "x \ carrier R" "set xs \ carrier R" shows "finite_extension K xs \ finite_extension K (x # xs)" using simple_extension_incl[OF finite_extension_in_carrier[OF assms(1,3)] assms(2)] by simp lemma (in ring) finite_extension_incl: assumes "K \ carrier R" and "set xs \ carrier R" shows "K \ finite_extension K xs" using finite_extension_incl_aux[OF assms(1)] assms(2) by (induct xs) (auto) lemma (in ring) finite_extension_as_eval_img: assumes "K \ carrier R" and "x \ carrier R" "set xs \ carrier R" shows "finite_extension K (x # xs) = (\p. eval p x) ` carrier ((finite_extension K xs) [X])" using simple_extension_as_eval_img[OF finite_extension_in_carrier[OF assms(1,3)] assms(2)] by simp lemma (in domain) finite_extension_is_subring: assumes "subring K R" "set xs \ carrier R" shows "subring (finite_extension K xs) R" using assms simple_extension_is_subring by (induct xs) (auto) corollary (in domain) finite_extension_mem: - assumes "subring K R" "set xs \ carrier R" shows "set xs \ finite_extension K xs" -proof - - { fix x xs assume "x \ carrier R" "set xs \ carrier R" - hence "x \ finite_extension K (x # xs)" - using simple_extension_mem[OF finite_extension_is_subring[OF assms(1), of xs]] by simp } - note aux_lemma = this - show ?thesis - using aux_lemma finite_extension_incl_aux[OF subringE(1)[OF assms(1)]] assms(2) - by (induct xs) (simp, smt insert_subset list.simps(15) subset_trans) + assumes subring: "subring K R" + shows "set xs \ carrier R \ set xs \ finite_extension K xs" +proof (induct xs) + case Nil + then show ?case by simp +next + case (Cons a xs) + from Cons(2) have a: "a \ carrier R" and xs: "set xs \ carrier R" by auto + show ?case + proof + fix x assume "x \ set (a # xs)" + then consider "x = a" | "x \ set xs" by auto + then show "x \ finite_extension K (a # xs)" + proof cases + case 1 + with a have "x \ carrier R" by simp + with xs have "x \ finite_extension K (x # xs)" + using simple_extension_mem[OF finite_extension_is_subring[OF subring]] by simp + with 1 show ?thesis by simp + next + case 2 + with Cons have *: "x \ finite_extension K xs" by auto + from a xs have "finite_extension K xs \ finite_extension K (a # xs)" + by (rule finite_extension_incl_aux[OF subringE(1)[OF subring]]) + with * show ?thesis by auto + qed + qed qed corollary (in domain) finite_extension_minimal: assumes "subring K R" "set xs \ carrier R" shows "finite_extension K xs = \ { K'. subring K' R \ K \ K' \ set xs \ K' }" using finite_extension_is_subring[OF assms] finite_extension_mem[OF assms] finite_extension_incl[OF subringE(1)[OF assms(1)] assms(2)] finite_extension_subring_incl by blast corollary (in domain) finite_extension_same_set: assumes "subring K R" "set xs \ carrier R" "set xs = set ys" shows "finite_extension K xs = finite_extension K ys" using finite_extension_minimal[OF assms(1)] assms(2-3) by auto text \The reciprocal is also true, but it is more subtle.\ proposition (in domain) finite_extension_is_subfield: assumes "subfield K R" "set xs \ carrier R" shows "(\x. x \ set xs \ (algebraic over K) x) \ subfield (finite_extension K xs) R" using simple_extension_is_subfield algebraic_mono assms by (induct xs) (auto, metis finite_extension.simps finite_extension_incl subring_props(1)) proposition (in domain) finite_extension_finite_dimension: assumes "subfield K R" "set xs \ carrier R" shows "(\x. x \ set xs \ (algebraic over K) x) \ finite_dimension K (finite_extension K xs)" and "finite_dimension K (finite_extension K xs) \ (\x. x \ set xs \ (algebraic over K) x)" proof - show "finite_dimension K (finite_extension K xs) \ (\x. x \ set xs \ (algebraic over K) x)" using finite_dimension_imp_algebraic[OF assms(1) finite_extension_is_subring[OF subfieldE(1)[OF assms(1)] assms(2)]] finite_extension_mem[OF subfieldE(1)[OF assms(1)] assms(2)] by auto next show "(\x. x \ set xs \ (algebraic over K) x) \ finite_dimension K (finite_extension K xs)" using assms(2) proof (induct xs, simp add: finite_dimensionI[OF dimension_one[OF assms(1)]]) case (Cons x xs) hence "finite_dimension K (finite_extension K xs)" by auto moreover have "(algebraic over (finite_extension K xs)) x" using algebraic_mono[OF finite_extension_incl[OF subfieldE(3)[OF assms(1)]]] Cons(2-3) by auto moreover have "subfield (finite_extension K xs) R" using finite_extension_is_subfield[OF assms(1)] Cons(2-3) by auto ultimately show ?case using telescopic_base_dim(1)[OF assms(1) _ _ finite_dimensionI[OF dimension_simple_extension, of _ x]] Cons(3) by auto qed qed corollary (in domain) finite_extesion_mem_imp_algebraic: assumes "subfield K R" "set xs \ carrier R" and "\x. x \ set xs \ (algebraic over K) x" shows "y \ finite_extension K xs \ (algebraic over K) y" using finite_dimension_imp_algebraic[OF assms(1) finite_extension_is_subring[OF subfieldE(1)[OF assms(1)] assms(2)]] finite_extension_finite_dimension(1)[OF assms(1-2)] assms(3) by auto corollary (in domain) simple_extesion_mem_imp_algebraic: assumes "subfield K R" "x \ carrier R" "(algebraic over K) x" shows "y \ simple_extension K x \ (algebraic over K) y" using finite_extesion_mem_imp_algebraic[OF assms(1), of "[ x ]"] assms(2-3) by auto subsection \Arithmetic of algebraic numbers\ text \We show that the set of algebraic numbers of a field over a subfield K is a subfield itself.\ lemma (in field) subfield_of_algebraics: assumes "subfield K R" shows "subfield { x \ carrier R. (algebraic over K) x } R" proof - let ?set_of_algebraics = "{ x \ carrier R. (algebraic over K) x }" show ?thesis proof (rule subfieldI'[OF subringI]) show "?set_of_algebraics \ carrier R" and "\ \ ?set_of_algebraics" using algebraic_self[OF _ subringE(3)] subfieldE(1)[OF assms(1)] by auto next fix x y assume x: "x \ ?set_of_algebraics" and y: "y \ ?set_of_algebraics" have "\ x \ simple_extension K x" using subringE(5)[OF simple_extension_is_subring[OF subfieldE(1)]] simple_extension_mem[OF subfieldE(1)] assms(1) x by auto thus "\ x \ ?set_of_algebraics" using simple_extesion_mem_imp_algebraic[OF assms] x by auto have "x \ y \ finite_extension K [ x, y ]" and "x \ y \ finite_extension K [ x, y ]" using subringE(6-7)[OF finite_extension_is_subring[OF subfieldE(1)[OF assms(1)]], of "[ x, y ]"] finite_extension_mem[OF subfieldE(1)[OF assms(1)], of "[ x, y ]"] x y by auto thus "x \ y \ ?set_of_algebraics" and "x \ y \ ?set_of_algebraics" using finite_extesion_mem_imp_algebraic[OF assms, of "[ x, y ]"] x y by auto next fix z assume z: "z \ ?set_of_algebraics - { \ }" have "inv z \ simple_extension K z" using subfield_m_inv(1)[of "simple_extension K z"] simple_extension_is_subfield[OF assms, of z] simple_extension_mem[OF subfieldE(1)] assms(1) z by auto thus "inv z \ ?set_of_algebraics" using simple_extesion_mem_imp_algebraic[OF assms] field_Units z by auto qed qed end \ No newline at end of file diff --git a/src/HOL/Data_Structures/Interval_Tree.thy b/src/HOL/Data_Structures/Interval_Tree.thy --- a/src/HOL/Data_Structures/Interval_Tree.thy +++ b/src/HOL/Data_Structures/Interval_Tree.thy @@ -1,301 +1,301 @@ (* Author: Bohua Zhan, with modifications by Tobias Nipkow *) section \Interval Trees\ theory Interval_Tree imports "HOL-Data_Structures.Cmp" "HOL-Data_Structures.List_Ins_Del" "HOL-Data_Structures.Isin2" "HOL-Data_Structures.Set_Specs" begin subsection \Intervals\ text \The following definition of intervals uses the \<^bold>\typedef\ command to define the type of non-empty intervals as a subset of the type of pairs \p\ where @{prop "fst p \ snd p"}:\ typedef (overloaded) 'a::linorder ivl = "{p :: 'a \ 'a. fst p \ snd p}" by auto text \More precisely, @{typ "'a ivl"} is isomorphic with that subset via the function @{const Rep_ivl}. Hence the basic interval properties are not immediate but need simple proofs:\ definition low :: "'a::linorder ivl \ 'a" where "low p = fst (Rep_ivl p)" definition high :: "'a::linorder ivl \ 'a" where "high p = snd (Rep_ivl p)" lemma ivl_is_interval: "low p \ high p" by (metis Rep_ivl high_def low_def mem_Collect_eq) lemma ivl_inj: "low p = low q \ high p = high q \ p = q" by (metis Rep_ivl_inverse high_def low_def prod_eqI) text \Now we can forget how exactly intervals were defined.\ instantiation ivl :: (linorder) linorder begin definition ivl_less: "(x < y) = (low x < low y | (low x = low y \ high x < high y))" definition ivl_less_eq: "(x \ y) = (low x < low y | (low x = low y \ high x \ high y))" instance proof fix x y z :: "'a ivl" show a: "(x < y) = (x \ y \ \ y \ x)" using ivl_less ivl_less_eq by force show b: "x \ x" by (simp add: ivl_less_eq) show c: "x \ y \ y \ z \ x \ z" - by (smt ivl_less_eq dual_order.trans less_trans) + using ivl_less_eq by fastforce show d: "x \ y \ y \ x \ x = y" using ivl_less_eq a ivl_inj ivl_less by fastforce show e: "x \ y \ y \ x" by (meson ivl_less_eq leI not_less_iff_gr_or_eq) qed end definition overlap :: "('a::linorder) ivl \ 'a ivl \ bool" where "overlap x y \ (high x \ low y \ high y \ low x)" definition has_overlap :: "('a::linorder) ivl set \ 'a ivl \ bool" where "has_overlap S y \ (\x\S. overlap x y)" subsection \Interval Trees\ type_synonym 'a ivl_tree = "('a ivl * 'a) tree" fun max_hi :: "('a::order_bot) ivl_tree \ 'a" where "max_hi Leaf = bot" | "max_hi (Node _ (_,m) _) = m" definition max3 :: "('a::linorder) ivl \ 'a \ 'a \ 'a" where "max3 a m n = max (high a) (max m n)" fun inv_max_hi :: "('a::{linorder,order_bot}) ivl_tree \ bool" where "inv_max_hi Leaf \ True" | "inv_max_hi (Node l (a, m) r) \ (m = max3 a (max_hi l) (max_hi r) \ inv_max_hi l \ inv_max_hi r)" lemma max_hi_is_max: "inv_max_hi t \ a \ set_tree t \ high a \ max_hi t" by (induct t, auto simp add: max3_def max_def) lemma max_hi_exists: "inv_max_hi t \ t \ Leaf \ \a\set_tree t. high a = max_hi t" proof (induction t rule: tree2_induct) case Leaf then show ?case by auto next case N: (Node l v m r) then show ?case proof (cases l rule: tree2_cases) case Leaf then show ?thesis using N.prems(1) N.IH(2) by (cases r, auto simp add: max3_def max_def le_bot) next case Nl: Node then show ?thesis proof (cases r rule: tree2_cases) case Leaf then show ?thesis using N.prems(1) N.IH(1) Nl by (auto simp add: max3_def max_def le_bot) next case Nr: Node obtain p1 where p1: "p1 \ set_tree l" "high p1 = max_hi l" using N.IH(1) N.prems(1) Nl by auto obtain p2 where p2: "p2 \ set_tree r" "high p2 = max_hi r" using N.IH(2) N.prems(1) Nr by auto then show ?thesis using p1 p2 N.prems(1) by (auto simp add: max3_def max_def) qed qed qed subsection \Insertion and Deletion\ definition node where [simp]: "node l a r = Node l (a, max3 a (max_hi l) (max_hi r)) r" fun insert :: "'a::{linorder,order_bot} ivl \ 'a ivl_tree \ 'a ivl_tree" where "insert x Leaf = Node Leaf (x, high x) Leaf" | "insert x (Node l (a, m) r) = (case cmp x a of EQ \ Node l (a, m) r | LT \ node (insert x l) a r | GT \ node l a (insert x r))" lemma inorder_insert: "sorted (inorder t) \ inorder (insert x t) = ins_list x (inorder t)" by (induct t rule: tree2_induct) (auto simp: ins_list_simps) lemma inv_max_hi_insert: "inv_max_hi t \ inv_max_hi (insert x t)" by (induct t rule: tree2_induct) (auto simp add: max3_def) fun split_min :: "'a::{linorder,order_bot} ivl_tree \ 'a ivl \ 'a ivl_tree" where "split_min (Node l (a, m) r) = (if l = Leaf then (a, r) else let (x,l') = split_min l in (x, node l' a r))" fun delete :: "'a::{linorder,order_bot} ivl \ 'a ivl_tree \ 'a ivl_tree" where "delete x Leaf = Leaf" | "delete x (Node l (a, m) r) = (case cmp x a of LT \ node (delete x l) a r | GT \ node l a (delete x r) | EQ \ if r = Leaf then l else let (a', r') = split_min r in node l a' r')" lemma split_minD: "split_min t = (x,t') \ t \ Leaf \ x # inorder t' = inorder t" by (induct t arbitrary: t' rule: split_min.induct) (auto simp: sorted_lems split: prod.splits if_splits) lemma inorder_delete: "sorted (inorder t) \ inorder (delete x t) = del_list x (inorder t)" by (induct t) (auto simp: del_list_simps split_minD Let_def split: prod.splits) lemma inv_max_hi_split_min: "\ t \ Leaf; inv_max_hi t \ \ inv_max_hi (snd (split_min t))" by (induct t) (auto split: prod.splits) lemma inv_max_hi_delete: "inv_max_hi t \ inv_max_hi (delete x t)" apply (induct t) apply simp using inv_max_hi_split_min by (fastforce simp add: Let_def split: prod.splits) subsection \Search\ text \Does interval \x\ overlap with any interval in the tree?\ fun search :: "'a::{linorder,order_bot} ivl_tree \ 'a ivl \ bool" where "search Leaf x = False" | "search (Node l (a, m) r) x = (if overlap x a then True else if l \ Leaf \ max_hi l \ low x then search l x else search r x)" lemma search_correct: "inv_max_hi t \ sorted (inorder t) \ search t x = has_overlap (set_tree t) x" proof (induction t rule: tree2_induct) case Leaf then show ?case by (auto simp add: has_overlap_def) next case (Node l a m r) have search_l: "search l x = has_overlap (set_tree l) x" using Node.IH(1) Node.prems by (auto simp: sorted_wrt_append) have search_r: "search r x = has_overlap (set_tree r) x" using Node.IH(2) Node.prems by (auto simp: sorted_wrt_append) show ?case proof (cases "overlap a x") case True thus ?thesis by (auto simp: overlap_def has_overlap_def) next case a_disjoint: False then show ?thesis proof cases assume [simp]: "l = Leaf" have search_eval: "search (Node l (a, m) r) x = search r x" using a_disjoint overlap_def by auto show ?thesis unfolding search_eval search_r by (auto simp add: has_overlap_def a_disjoint) next assume "l \ Leaf" then show ?thesis proof (cases "max_hi l \ low x") case max_hi_l_ge: True have "inv_max_hi l" using Node.prems(1) by auto then obtain p where p: "p \ set_tree l" "high p = max_hi l" using \l \ Leaf\ max_hi_exists by auto have search_eval: "search (Node l (a, m) r) x = search l x" using a_disjoint \l \ Leaf\ max_hi_l_ge by (auto simp: overlap_def) show ?thesis proof (cases "low p \ high x") case True have "overlap p x" unfolding overlap_def using True p(2) max_hi_l_ge by auto then show ?thesis unfolding search_eval search_l using p(1) by(auto simp: has_overlap_def overlap_def) next case False have "\overlap x rp" if asm: "rp \ set_tree r" for rp proof - have "low p \ low rp" using asm p(1) Node(4) by(fastforce simp: sorted_wrt_append ivl_less) then show ?thesis using False by (auto simp: overlap_def) qed then show ?thesis unfolding search_eval search_l using a_disjoint by (auto simp: has_overlap_def overlap_def) qed next case False have search_eval: "search (Node l (a, m) r) x = search r x" using a_disjoint False by (auto simp: overlap_def) have "\overlap x lp" if asm: "lp \ set_tree l" for lp using asm False Node.prems(1) max_hi_is_max by (fastforce simp: overlap_def) then show ?thesis unfolding search_eval search_r using a_disjoint by (auto simp: has_overlap_def overlap_def) qed qed qed qed definition empty :: "'a ivl_tree" where "empty = Leaf" subsection \Specification\ locale Interval_Set = Set + fixes has_overlap :: "'t \ 'a::linorder ivl \ bool" assumes set_overlap: "invar s \ has_overlap s x = Interval_Tree.has_overlap (set s) x" fun invar :: "('a::{linorder,order_bot}) ivl_tree \ bool" where "invar t = (inv_max_hi t \ sorted(inorder t))" interpretation S: Interval_Set where empty = Leaf and insert = insert and delete = delete and has_overlap = search and isin = isin and set = set_tree and invar = invar proof (standard, goal_cases) case 1 then show ?case by auto next case 2 then show ?case by (simp add: isin_set_inorder) next case 3 then show ?case by(simp add: inorder_insert set_ins_list flip: set_inorder) next case 4 then show ?case by(simp add: inorder_delete set_del_list flip: set_inorder) next case 5 then show ?case by auto next case 6 then show ?case by (simp add: inorder_insert inv_max_hi_insert sorted_ins_list) next case 7 then show ?case by (simp add: inorder_delete inv_max_hi_delete sorted_del_list) next case 8 then show ?case by (simp add: search_correct) qed end diff --git a/src/HOL/Datatype_Examples/FAE_Sequence.thy b/src/HOL/Datatype_Examples/FAE_Sequence.thy --- a/src/HOL/Datatype_Examples/FAE_Sequence.thy +++ b/src/HOL/Datatype_Examples/FAE_Sequence.thy @@ -1,133 +1,136 @@ theory FAE_Sequence imports "HOL-Library.Confluent_Quotient" begin type_synonym 'a seq = "nat \ 'a" abbreviation finite_range :: "('a \ 'b) \ bool" where "finite_range f \ finite (range f)" lemma finite_range_pair: assumes 1: "finite_range (\x. fst (f x))" and 2: "finite_range (\x. snd (f x))" shows "finite_range f" proof - have "range f \ range (\x. fst (f x)) \ range (\x. snd (f x))" by(auto 4 3 intro: rev_image_eqI dest: sym) then show ?thesis by(rule finite_subset)(use assms in simp) qed definition seq_at :: "'a seq \ 'a \ nat set" where "seq_at f x = f -` {x}" typedef 'a fseq = "{f :: 'a seq. finite_range f}" morphisms ap_fseq Abs_fseq by(rule exI[where x="\_. undefined"]) simp setup_lifting type_definition_fseq lift_bnf 'a fseq [wits: "\x. (\_ :: nat. x)"] subgoal by(metis finite_imageI fun.set_map mem_Collect_eq) subgoal by(auto intro: finite_range_pair) subgoal by auto subgoal by auto done lemma ap_map_fseq [simp]: "ap_fseq (map_fseq f x) = f \ ap_fseq x" by transfer simp lift_definition fseq_eq :: "'a fseq \ 'a fseq \ bool" is "\f g. finite {n. f n \ g n}" . lemma fseq_eq_unfold: "fseq_eq f g \ finite {n. ap_fseq f n \ ap_fseq g n}" by transfer simp lemma reflp_fseq_eq [simp]: "reflp fseq_eq" by(rule reflpI)(simp add: fseq_eq_unfold) lemma symp_fseq_eq [simp]: "symp fseq_eq" by(rule sympI)(simp add: fseq_eq_unfold eq_commute) lemma transp_fseq_eq [simp]: "transp fseq_eq" apply(rule transpI) subgoal for f g h unfolding fseq_eq_unfold by(rule finite_subset[where B="{n. ap_fseq f n \ ap_fseq g n} \ {n. ap_fseq g n \ ap_fseq h n}"]) auto done lemma equivp_fseq_eq [simp]: "equivp fseq_eq" by(simp add: equivp_reflp_symp_transp) functor map_fseq by(simp_all add: fseq.map_id0 fseq.map_comp fun_eq_iff) quotient_type 'a fae_seq = "'a fseq" / fseq_eq by simp lemma map_fseq_eq: "fseq_eq f g \ fseq_eq (map_fseq h f) (map_fseq h g)" unfolding fseq_eq_unfold by(auto elim!: finite_subset[rotated]) lemma finite_set_fseq [simp]: "finite (set_fseq x)" by transfer interpretation wide_intersection_fseq: wide_intersection_finite fseq_eq map_fseq set_fseq by unfold_locales(simp_all add: map_fseq_eq fseq.map_id[unfolded id_def] fseq.set_map cong: fseq.map_cong) lemma fseq_subdistributivity: assumes "A OO B \ bot" shows "rel_fseq A OO fseq_eq OO rel_fseq B \ fseq_eq OO rel_fseq (A OO B) OO fseq_eq" proof(rule predicate2I; elim relcomppE; transfer fixing: A B) fix f and g g' :: "'c seq" and h assume fg: "rel_fun(=) A f g" and gg': "finite {n. g n \ g' n}" and g'h: "rel_fun (=) B g' h" and f: "finite_range f" and g: "finite_range g" and g': "finite_range g'" and h: "finite_range h" from assms obtain a c b where ac: "A a c" and cb: "B c b" by(auto simp add: fun_eq_iff) let ?X = "{n. g n \ g' n}" let ?f = "\n. if n \ ?X then a else f n" let ?h = "\n. if n \ ?X then b else h n" have "range ?f \ insert a (range f)" by auto then have "finite_range ?f" by(rule finite_subset)(simp add: f) moreover have "range ?h \ insert b (range h)" by auto then have "finite_range ?h" by(rule finite_subset)(simp add: h) moreover have "{n. f n \ ?f n} \ ?X" by auto then have "finite {n. f n \ ?f n}" by(rule finite_subset)(simp add: gg') moreover have "{n. ?h n \ h n} \ ?X" by auto then have "finite {n. ?h n \ h n}" by(rule finite_subset)(simp add: gg') moreover have "rel_fun (=) (A OO B) ?f ?h" using fg g'h ac cb by(auto simp add: rel_fun_def) ultimately show "\ya\{f. finite_range f}. finite {n. f n \ ya n} \ (\yb\{f. finite_range f}. rel_fun (=) (A OO B) ya yb \ finite {n. yb n \ h n})" unfolding mem_Collect_eq Bex_def by blast qed lift_bnf 'a fae_seq subgoal by(rule fseq_subdistributivity) subgoal by(rule wide_intersection_fseq.wide_intersection) done lift_definition fseq_at :: "'a fseq \ 'a \ nat set" is seq_at . lemma fae_vimage_singleton: "finite (f -` {x}) \ finite (g -` {x})" if "finite {n. f n \ g n}" proof assume *: "finite (g -` {x})" have "f -` {x} \ g -` {x} \ {n. f n \ g n}" by auto thus "finite (f -` {x})" by(rule finite_subset)(use that * in auto) next assume *: "finite (f -` {x})" have "g -` {x} \ f -` {x} \ {n. f n \ g n}" by auto thus "finite (g -` {x})" by(rule finite_subset)(use that * in auto) qed lift_definition set_fae_seq_alt :: "'a fae_seq \ 'a set" is "\f. {a. infinite (fseq_at f a)}" by(transfer)(clarsimp simp add: set_eq_iff seq_at_def fae_vimage_singleton) lemma fseq_at_infinite_Inr: "\infinite (fseq_at f x); fseq_eq (map_fseq Inr f) g\ \ \x'\set_fseq g. x \ Basic_BNFs.setr x'" - by transfer(auto simp add: seq_at_def vimage_def; smt (z3) finite_nat_set_iff_bounded_le mem_Collect_eq setr.intros) + apply transfer + apply (auto simp add: seq_at_def vimage_def) + apply (smt (verit, ccfv_SIG) finite_subset mem_Collect_eq setr.simps subsetI) + done lemma fseq_at_Inr_infinite: assumes "\g. fseq_eq (map_fseq Inr f) g \ (\x'\set_fseq g. x \ Basic_BNFs.setr x')" shows "infinite (fseq_at f x)" proof assume fin: "finite (fseq_at f x)" let ?g = "map_fseq (\y. if x = y then Inl undefined else Inr y) f" have "fseq_eq (map_fseq Inr f) ?g" using fin by transfer(simp add: seq_at_def vimage_def eq_commute) with assms[of "?g"] show False by(auto simp add: fseq.set_map) qed lemma set_fae_seq_eq_alt: "set_fae_seq f = set_fae_seq_alt f" by transfer(use fseq_at_Inr_infinite in \force simp add: fseq_at_infinite_Inr\) end diff --git a/src/HOL/Datatype_Examples/Regex_ACIDZ.thy b/src/HOL/Datatype_Examples/Regex_ACIDZ.thy --- a/src/HOL/Datatype_Examples/Regex_ACIDZ.thy +++ b/src/HOL/Datatype_Examples/Regex_ACIDZ.thy @@ -1,481 +1,481 @@ theory Regex_ACIDZ imports "HOL-Library.Confluent_Quotient" begin datatype (discs_sels) 'a rexp = Zero | Eps | Atom 'a | Alt "'a rexp" "'a rexp" | Conc "'a rexp" "'a rexp" | Star "'a rexp" fun elim_zeros where "elim_zeros (Alt r s) = (let r' = elim_zeros r; s' = elim_zeros s in if r' = Zero then s' else if s' = Zero then r' else Alt r' s')" | "elim_zeros (Conc r s) = (let r' = elim_zeros r in if r' = Zero then Zero else Conc r' s)" | "elim_zeros r = r" fun distribute where "distribute t (Alt r s) = Alt (distribute t r) (distribute t s)" | "distribute t (Conc r s) = Conc (distribute s r) t" | "distribute t r = Conc r t" inductive ACIDZ (infix "~" 64) where a1: "Alt (Alt r s) t ~ Alt r (Alt s t)" | a2: "Alt r (Alt s t) ~ Alt (Alt r s) t" | c: "Alt r s ~ Alt s r" | i: "r ~ Alt r r" | z: "r ~ s \ r ~ elim_zeros s" | d: "Conc r t ~ distribute t r" | R: "r ~ r" | A: "r ~ r' \ s ~ s' \ Alt r s ~ Alt r' s'" | C: "r ~ r' \ Conc r s ~ Conc r' s" inductive_cases ACIDZ_AltE[elim]: "Alt r s ~ t" inductive_cases ACIDZ_ConcE[elim]: "Conc r s ~ t" inductive_cases ACIDZ_StarE[elim]: "Star r ~ t" declare ACIDZ.intros[intro] abbreviation ACIDZcl (infix "~~" 64) where "(~~) \ equivclp (~)" lemma map_rexp_distribute[simp]: "map_rexp f (distribute t r) = distribute (map_rexp f t) (map_rexp f r)" by (induct r arbitrary: t) auto lemma set_rexp_distribute[simp]: "set_rexp (distribute t r) = set_rexp r \ set_rexp t" by (induct r arbitrary: t) auto lemma Zero_eq_map_rexp_iff[simp]: "Zero = map_rexp f x \ x = Zero" "map_rexp f x = Zero \ x = Zero" by (cases x; auto)+ lemma Alt_eq_map_rexp_iff: "Alt r s = map_rexp f x \ (\r' s'. x = Alt r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" "map_rexp f x = Alt r s \ (\r' s'. x = Alt r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" by (cases x; auto)+ lemma map_rexp_elim_zeros[simp]: "map_rexp f (elim_zeros r) = elim_zeros (map_rexp f r)" by (induct r rule: elim_zeros.induct) (auto simp: Let_def) lemma set_rexp_elim_zeros: "set_rexp (elim_zeros r) \ set_rexp r" by (induct r rule: elim_zeros.induct) (auto simp: Let_def) lemma ACIDZ_map_respects: fixes f :: "'a \ 'b" and r s :: "'a rexp" assumes "r ~ s" shows "map_rexp f r ~ map_rexp f s" using assms by (induct r s rule: ACIDZ.induct) (auto simp: Let_def) lemma ACIDZcl_map_respects: fixes f :: "'a \ 'b" and r s :: "'a rexp" assumes "r ~~ s" shows "map_rexp f r ~~ map_rexp f s" using assms by (induct rule: equivclp_induct) (auto intro: ACIDZ_map_respects equivclp_trans) lemma finite_set_rexp: "finite (set_rexp r)" by (induct r) auto lemma ACIDZ_set_rexp: "r ~ s \ set_rexp s \ set_rexp r" by (induct r s rule: ACIDZ.induct) (auto dest: set_mp[OF set_rexp_elim_zeros] simp: Let_def) lemma ACIDZ_set_rexp': "(~)\<^sup>*\<^sup>* r s \ set_rexp s \ set_rexp r" by (induct rule: rtranclp.induct) (auto dest: ACIDZ_set_rexp) lemma Conc_eq_map_rexp_iff: "Conc r s = map_rexp f x \ (\r' s'. x = Conc r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" "map_rexp f x = Conc r s \ (\r' s'. x = Conc r' s' \ map_rexp f r' = r \ map_rexp f s' = s)" by (cases x; auto)+ lemma Star_eq_map_rexp_iff: "Star r = map_rexp f x \ (\r'. x = Star r' \ map_rexp f r' = r)" "map_rexp f x = Star r \ (\r'. x = Star r' \ map_rexp f r' = r)" by (cases x; auto)+ lemma AAA: "(~)\<^sup>*\<^sup>* r r' \ (~)\<^sup>*\<^sup>* s s' \ (~)\<^sup>*\<^sup>* (Alt r s) (Alt r' s')" proof (induct r r' rule: rtranclp.induct) case (rtrancl_refl r) then show ?case by (induct s s' rule: rtranclp.induct) (auto elim!: rtranclp.rtrancl_into_rtrancl) next case (rtrancl_into_rtrancl a b c) then show ?case by (auto elim!: rtranclp.rtrancl_into_rtrancl) qed lemma CCC: "(~)\<^sup>*\<^sup>* r r' \ (~)\<^sup>*\<^sup>* (Conc r s) (Conc r' s)" proof (induct r r' rule: rtranclp.induct) case (rtrancl_refl r) then show ?case by simp next case (rtrancl_into_rtrancl a b c) then show ?case by (auto elim!: rtranclp.rtrancl_into_rtrancl) qed lemma map_rexp_ACIDZ_inv: "map_rexp f x ~ y \ \z. (~)\<^sup>*\<^sup>* x z \ y = map_rexp f z" proof (induct "map_rexp f x" y arbitrary: x rule: ACIDZ.induct) case (a1 r s t) then obtain r' s' t' where "x = Alt (Alt r' s') t'" "map_rexp f r' = r" "map_rexp f s' = s" "map_rexp f t' = t" by (auto simp: Alt_eq_map_rexp_iff) then show ?case by (intro exI[of _ "Alt r' (Alt s' t')"]) auto next case (a2 r s t) then obtain r' s' t' where "x = Alt r' (Alt s' t')" "map_rexp f r' = r" "map_rexp f s' = s" "map_rexp f t' = t" by (auto simp: Alt_eq_map_rexp_iff) then show ?case by (intro exI[of _ "Alt (Alt r' s') t'"]) auto next case (c r s) then obtain r' s' where "x = Alt r' s'" "map_rexp f r' = r" "map_rexp f s' = s" by (auto simp: Alt_eq_map_rexp_iff) then show ?case by (intro exI[of _ "Alt s' r'"]) auto next case (i r) then show ?case by (intro exI[of _ "Alt r r"]) auto next case (z r) then show ?case by (metis ACIDZ.z R map_rexp_elim_zeros rtranclp.simps) next case (d r s) then obtain r' s' where "x = Conc r' s'" "map_rexp f r' = r" "map_rexp f s' = s" by (auto simp: Conc_eq_map_rexp_iff) then show ?case by (intro exI[of _ "distribute s' r'"]) auto next case (R r) then show ?case by (auto intro: exI[of _ r]) next case (A r rr s ss) then obtain r' s' where "x = Alt r' s'" "map_rexp f r' = r" "map_rexp f s' = s" by (auto simp: Alt_eq_map_rexp_iff) moreover from A(2)[OF this(2)[symmetric]] A(4)[OF this(3)[symmetric]] obtain rr' ss' where "(~)\<^sup>*\<^sup>* r' rr'" "rr = map_rexp f rr'" "(~)\<^sup>*\<^sup>* s' ss'" "ss = map_rexp f ss'" by blast ultimately show ?case using A(1,3) by (intro exI[of _ "Alt rr' ss'"]) (auto intro!: AAA) next case (C r rr s) then obtain r' s' where "x = Conc r' s'" "map_rexp f r' = r" "map_rexp f s' = s" by (auto simp: Conc_eq_map_rexp_iff) moreover from C(2)[OF this(2)[symmetric]] obtain rr' where "(~)\<^sup>*\<^sup>* r' rr'" "rr = map_rexp f rr'" by blast ultimately show ?case using C(1,3) by (intro exI[of _ "Conc rr' s'"]) (auto intro!: CCC) qed lemma reflclp_ACIDZ: "(~)\<^sup>=\<^sup>= = (~)" unfolding fun_eq_iff by auto lemma elim_zeros_distribute_Zero: "elim_zeros r = Zero \ elim_zeros (distribute t r) = Zero" by (induct r arbitrary: t) (auto simp: Let_def split: if_splits) lemma elim_zeros_ACIDZ_Zero: "r ~ s \ elim_zeros r = Zero \ elim_zeros s = Zero" by (induct r s rule: ACIDZ.induct) (auto simp: Let_def elim_zeros_distribute_Zero split: if_splits) lemma AZZ[simp]: "Alt Zero Zero ~ Zero" by (metis elim_zeros.simps(3) elim_zeros_ACIDZ_Zero i z) lemma elim_zeros_idem[simp]: "elim_zeros (elim_zeros r) = elim_zeros r" by (induct r rule: elim_zeros.induct) (auto simp: Let_def) lemma elim_zeros_distribute_idem: "elim_zeros (distribute s (elim_zeros r)) = elim_zeros (distribute s r)" by (induct r arbitrary: s rule: elim_zeros.induct) (auto simp: Let_def) lemma zz: "r ~ s \ t = elim_zeros s \ r ~ t" by (metis z) lemma elim_zeros_ACIDZ1: "r ~ s \ elim_zeros r ~ elim_zeros s" proof (induct r s rule: ACIDZ.induct) case (c r s) then show ?case by (auto simp: Let_def) next case (d r t) then show ?case - by (smt (z3) ACIDZ.d R distribute.simps(3) elim_zeros.simps(2) elim_zeros.simps(3) elim_zeros_distribute_idem z) + by (smt (verit, ccfv_SIG) ACIDZ.d R elim_zeros.simps(2) elim_zeros_distribute_Zero elim_zeros_distribute_idem z) next case (A r r' s s') then show ?case by (auto 0 2 simp: Let_def dest: elim_zeros_ACIDZ_Zero elim: zz[OF ACIDZ.A]) next case (C r r' s) then show ?case - by (smt (z3) ACIDZ.C elim_zeros.simps(2,3) z elim_zeros_ACIDZ_Zero) + by (smt (verit, best) ACIDZ.C R elim_zeros.simps(2) elim_zeros_ACIDZ_Zero z) qed (auto simp: Let_def) lemma AA': "r ~ r' \ s ~ s' \ t = elim_zeros (Alt r' s') \ Alt r s ~ t" by (auto simp del: elim_zeros.simps) lemma distribute_ACIDZ1: "r ~ r' \ distribute s r ~ elim_zeros (distribute s r')" proof (induct r r' arbitrary: s rule: ACIDZ.induct) case (A r r' s s' u) from A(2,4)[of u] show ?case by (auto simp: Let_def elim: zz[OF ACIDZ.A]) next case (C r r' s) then show ?case by (smt (verit, best) ACIDZ.C distribute.simps(2,3) elim_zeros.simps(2,3) z) qed (auto simp del: elim_zeros.simps simp: elim_zeros_distribute_idem) lemma elim_zeros_ACIDZ: "r ~ s \ (~)\<^sup>*\<^sup>* (elim_zeros r) (elim_zeros s)" using elim_zeros_ACIDZ1 by blast lemma elim_zeros_ACIDZ_rtranclp: "(~)\<^sup>*\<^sup>* r s \ (~)\<^sup>*\<^sup>* (elim_zeros r) (elim_zeros s)" by (induct rule: rtranclp_induct) (auto elim!: rtranclp_trans elim_zeros_ACIDZ) lemma distribute_ACIDZ: "(~) r r' \ (~)\<^sup>*\<^sup>* (distribute s r) (elim_zeros (distribute s r'))" by (metis distribute_ACIDZ1 rtranclp.simps) lemma ACIDZ_elim_zeros: "(~) r s \ elim_zeros r = Zero \ elim_zeros s = Zero" by (meson elim_zeros_ACIDZ1 elim_zeros_ACIDZ_Zero) lemma ACIDZ_elim_zeros_rtranclp: "(~)\<^sup>*\<^sup>* r s \ elim_zeros r = Zero \ elim_zeros s = Zero" by (induct rule: rtranclp_induct) (auto elim: ACIDZ_elim_zeros) lemma Alt_elim_zeros[simp]: "Alt (elim_zeros r) s ~ elim_zeros (Alt r s)" "Alt r (elim_zeros s) ~ elim_zeros (Alt r s)" by (smt (verit, ccfv_threshold) ACIDZ.simps elim_zeros.simps(1) elim_zeros_idem)+ lemma strong_confluentp_ACIDZ: "strong_confluentp (~)" proof (rule strong_confluentpI, unfold reflclp_ACIDZ) fix x y z :: "'a rexp" assume "x ~ y" "x ~ z" then show "\u. (~)\<^sup>*\<^sup>* y u \ z ~ u" proof (induct x y arbitrary: z rule: ACIDZ.induct) case (a1 r s t) then show ?case by (auto intro: AAA converse_rtranclp_into_rtranclp) next case (a2 r s t) then show ?case by (auto intro: AAA converse_rtranclp_into_rtranclp) next case (c r s) then show ?case by (cases rule: ACIDZ.cases) (auto 0 4 intro: exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ R], rotated] converse_rtranclp_into_rtranclp[where r="(~)", OF ACIDZ.c]) next case (i r) then show ?case by (auto intro: AAA) next case (z r s) then show ?case by (meson ACIDZ.z elim_zeros_ACIDZ_rtranclp) next case (d r s t) then show ?case by (induct "Conc r s" t rule: ACIDZ.induct) (force intro: exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ z], rotated] exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ z[OF elim_zeros_ACIDZ1]], rotated] distribute_ACIDZ1 elim!: rtranclp_trans)+ next case (R r) then show ?case by auto next note A1 = ACIDZ.A[OF _ R] note A2 = ACIDZ.A[OF R] case (A r r' s s') from A(5,1-4) show ?case proof (induct "Alt r s" z rule: ACIDZ.induct) case inner: (a1 r'' s'') from A(1,3) show ?case unfolding inner.hyps[symmetric] proof (induct "Alt r'' s''" r' rule: ACIDZ.induct[consumes 1, case_names Xa1 Xa2 Xc Xi Xz Xd XR XA XC]) case Xa1 with A(3) show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) (metis A1 a1 a2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xa2 then show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) (metis a1 A1 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xc then show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A2[OF A2]], rotated]) (metis a1 c A1 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xi then show ?case apply (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ ACIDZ.A[OF i ACIDZ.A[OF i]]], rotated]) apply (rule converse_rtranclp_into_rtranclp, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A1, rule c) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule a2) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) apply blast done next case Xz with A show ?case by (auto elim!: exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ z], rotated] converse_rtranclp_into_rtranclp[rotated, OF elim_zeros_ACIDZ_rtranclp] simp del: elim_zeros.simps) qed auto next case inner: (a2 s'' t'') from A(3,1) show ?case unfolding inner.hyps[symmetric] proof (induct "Alt s'' t''" s' rule: ACIDZ.induct[consumes 1, case_names Xa1 Xa2 Xc Xi Xz Xd XR XA XC]) case Xa1 with A(3) show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) (metis a2 A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xa2 then show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) (metis a1 a2 A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xc then show ?case by (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ A1[OF A1]], rotated]) (metis a2 c A2 r_into_rtranclp rtranclp.rtrancl_into_rtrancl) next case Xi then show ?case apply (elim exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ ACIDZ.A[OF ACIDZ.A[OF _ i] i]], rotated]) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A1, rule c) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule A2, rule a1) apply (rule converse_rtranclp_into_rtranclp, rule A2, rule a2) apply (rule converse_rtranclp_into_rtranclp, rule a2) apply blast done next case Xz then show ?case by (auto elim!: exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ z], rotated] converse_rtranclp_into_rtranclp[rotated, OF elim_zeros_ACIDZ_rtranclp] simp del: elim_zeros.simps) qed auto next case (z rs) with A show ?case by (auto elim!: rtranclp_trans intro!: exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ elim_zeros_ACIDZ1], rotated]) qed (auto 5 0 intro: AAA) next case (C r r' s s') from C(3,1-2) show ?case by (induct "Conc r s" s' rule: ACIDZ.induct) (auto intro: CCC elim!: rtranclp_trans exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ elim_zeros_ACIDZ1], rotated] exI[where P = "\x. _ x \ _ ~ x", OF conjI[OF _ distribute_ACIDZ1], rotated]) qed qed lemma confluent_quotient_ACIDZ: "confluent_quotient (~) (~~) (~~) (~~) (~~) (~~) (map_rexp fst) (map_rexp snd) (map_rexp fst) (map_rexp snd) rel_rexp rel_rexp rel_rexp set_rexp set_rexp" by unfold_locales (auto 4 4 dest: ACIDZ_set_rexp' simp: rexp.in_rel rexp.rel_compp dest: map_rexp_ACIDZ_inv intro: rtranclp_into_equivclp intro: equivpI reflpI sympI transpI ACIDZcl_map_respects strong_confluentp_imp_confluentp[OF strong_confluentp_ACIDZ]) lemma wide_intersection_finite_ACIDZ: "wide_intersection_finite (~~) map_rexp set_rexp" by unfold_locales (auto intro: ACIDZcl_map_respects rexp.map_cong simp: rexp.map_id rexp.set_map finite_set_rexp) inductive ACIDZEQ (infix "\" 64) where a: "Alt (Alt r s) t \ Alt r (Alt s t)" | c: "Alt r s \ Alt s r" | i: "Alt r r \ r" | cz: "Conc Zero r \ Zero" | az: "Alt Zero r \ r" | d: "Conc (Alt r s) t \ Alt (Conc r t) (Conc s t)" | A: "r \ r' \ s \ s' \ Alt r s \ Alt r' s'" | C: "r \ r' \ Conc r s \ Conc r' s" | r: "r \ r" | s: "r \ s \ s \ r" | t: "r \ s \ s \ t \ r \ t" context begin private lemma AA: "r ~~ r' \ s ~~ s' \ Alt r s ~~ Alt r' s'" proof (induct rule: equivclp_induct) case base then show ?case by (induct rule: equivclp_induct) (auto elim!: equivclp_trans) next case (step s t) then show ?case by (auto elim!: equivclp_trans) qed private lemma CC: "r ~~ r' \ Conc r s ~~ Conc r' s" proof (induct rule: equivclp_induct) case base then show ?case by simp next case (step s t) then show ?case by (auto elim!: equivclp_trans) qed private lemma CZ: "Conc Zero r ~~ Zero" - by (smt (z3) R elim_zeros.simps(2) elim_zeros.simps(3) r_into_equivclp z) + by (metis R distribute.simps(3) elim_zeros.simps(3) elim_zeros_distribute_Zero r_into_equivclp z) private lemma AZ: "Alt Zero r ~~ r" by (smt (verit, ccfv_threshold) converse_r_into_equivclp converse_rtranclp_into_rtranclp elim_zeros.simps(1) elim_zeros.simps(3) equivclp_def symclpI1 z R) private lemma D: "Conc (Alt r s) t ~~ Alt (Conc r t) (Conc s t)" by (smt (verit, ccfv_threshold) AA ACIDZ.d converse_r_into_equivclp distribute.simps(1) equivclp_sym rtranclp.simps rtranclp_equivclp) lemma ACIDZEQ_le_ACIDZcl: "r \ s \ r ~~ s" by (induct r s rule: ACIDZEQ.induct) (auto intro: AA CC AZ CZ D equivclp_sym equivclp_trans) end lemma ACIDZEQ_z[simp]: "r \ elim_zeros r" by (induct r rule: elim_zeros.induct) (auto 0 3 intro: ACIDZEQ.intros simp: Let_def) lemma ACIDZEQ_d[simp]: "Conc r s \ distribute s r" by (induct s r rule: distribute.induct) (auto 0 3 intro: ACIDZEQ.intros) lemma ACIDZ_le_ACIDZEQ: "r ~ s \ r \ s" by (induct r s rule: ACIDZ.induct) (auto 0 2 intro: ACIDZEQ.intros simp: Let_def) lemma ACIDZcl_le_ACIDZEQ: "r ~~ s \ r \ s" by (induct rule: equivclp_induct) (auto 0 3 intro: ACIDZEQ.intros ACIDZ_le_ACIDZEQ) lemma ACIDZEQ_alt: "(\) = (~~)" by (simp add: ACIDZEQ_le_ACIDZcl ACIDZcl_le_ACIDZEQ antisym predicate2I) quotient_type 'a rexp_ACIDZ = "'a rexp" / "(\)" unfolding ACIDZEQ_alt by (auto intro!: equivpI reflpI sympI transpI) lift_bnf 'a rexp_ACIDZ unfolding ACIDZEQ_alt subgoal for A Q by (rule confluent_quotient.subdistributivity[OF confluent_quotient_ACIDZ]) subgoal for Ss by (elim wide_intersection_finite.wide_intersection[OF wide_intersection_finite_ACIDZ]) done datatype ldl = Prop string | And ldl ldl | Not ldl | Match "ldl rexp_ACIDZ" end diff --git a/src/HOL/Hoare/Hoare_Logic_Abort.thy b/src/HOL/Hoare/Hoare_Logic_Abort.thy --- a/src/HOL/Hoare/Hoare_Logic_Abort.thy +++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy @@ -1,236 +1,236 @@ (* Title: HOL/Hoare/Hoare_Logic_Abort.thy Author: Leonor Prensa Nieto & Tobias Nipkow Copyright 2003 TUM Author: Walter Guttmann (extension to total-correctness proofs) *) section \Hoare Logic with an Abort statement for modelling run time errors\ theory Hoare_Logic_Abort imports Hoare_Syntax Hoare_Tac begin type_synonym 'a bexp = "'a set" type_synonym 'a assn = "'a set" type_synonym 'a var = "'a \ nat" datatype 'a com = Basic "'a \ 'a" | Abort | Seq "'a com" "'a com" | Cond "'a bexp" "'a com" "'a com" | While "'a bexp" "'a assn" "'a var" "'a com" abbreviation annskip ("SKIP") where "SKIP == Basic id" type_synonym 'a sem = "'a option => 'a option => bool" inductive Sem :: "'a com \ 'a sem" where "Sem (Basic f) None None" | "Sem (Basic f) (Some s) (Some (f s))" | "Sem Abort s None" | "Sem c1 s s'' \ Sem c2 s'' s' \ Sem (Seq c1 c2) s s'" | "Sem (Cond b c1 c2) None None" | "s \ b \ Sem c1 (Some s) s' \ Sem (Cond b c1 c2) (Some s) s'" | "s \ b \ Sem c2 (Some s) s' \ Sem (Cond b c1 c2) (Some s) s'" | "Sem (While b x y c) None None" | "s \ b \ Sem (While b x y c) (Some s) (Some s)" | "s \ b \ Sem c (Some s) s'' \ Sem (While b x y c) s'' s' \ Sem (While b x y c) (Some s) s'" inductive_cases [elim!]: "Sem (Basic f) s s'" "Sem (Seq c1 c2) s s'" "Sem (Cond b c1 c2) s s'" lemma Sem_deterministic: assumes "Sem c s s1" and "Sem c s s2" shows "s1 = s2" proof - have "Sem c s s1 \ (\s2. Sem c s s2 \ s1 = s2)" by (induct rule: Sem.induct) (subst Sem.simps, blast)+ thus ?thesis using assms by simp qed definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" where "Valid p c q \ \s s'. Sem c s s' \ s \ Some ` p \ s' \ Some ` q" definition ValidTC :: "'a bexp \ 'a com \ 'a bexp \ bool" where "ValidTC p c q \ \s . s \ p \ (\t . Sem c (Some s) (Some t) \ t \ q)" lemma tc_implies_pc: "ValidTC p c q \ Valid p c q" - by (smt Sem_deterministic ValidTC_def Valid_def image_iff) + by (smt (verit) Sem_deterministic ValidTC_def Valid_def image_iff) lemma tc_extract_function: "ValidTC p c q \ \f . \s . s \ p \ f s \ q" by (meson ValidTC_def) text \The proof rules for partial correctness\ lemma SkipRule: "p \ q \ Valid p (Basic id) q" by (auto simp:Valid_def) lemma BasicRule: "p \ {s. f s \ q} \ Valid p (Basic f) q" by (auto simp:Valid_def) lemma SeqRule: "Valid P c1 Q \ Valid Q c2 R \ Valid P (Seq c1 c2) R" by (auto simp:Valid_def) lemma CondRule: "p \ {s. (s \ b \ s \ w) \ (s \ b \ s \ w')} \ Valid w c1 q \ Valid w' c2 q \ Valid p (Cond b c1 c2) q" by (fastforce simp:Valid_def image_def) lemma While_aux: assumes "Sem (While b i v c) s s'" shows "\s s'. Sem c s s' \ s \ Some ` (I \ b) \ s' \ Some ` I \ s \ Some ` I \ s' \ Some ` (I \ -b)" using assms by (induct "While b i v c" s s') auto lemma WhileRule: "p \ i \ Valid (i \ b) c i \ i \ (-b) \ q \ Valid p (While b i v c) q" apply (clarsimp simp:Valid_def) apply(drule While_aux) apply assumption apply blast apply blast done lemma AbortRule: "p \ {s. False} \ Valid p Abort q" by(auto simp:Valid_def) text \The proof rules for total correctness\ lemma SkipRuleTC: assumes "p \ q" shows "ValidTC p (Basic id) q" by (metis Sem.intros(2) ValidTC_def assms id_def subsetD) lemma BasicRuleTC: assumes "p \ {s. f s \ q}" shows "ValidTC p (Basic f) q" by (metis Ball_Collect Sem.intros(2) ValidTC_def assms) lemma SeqRuleTC: assumes "ValidTC p c1 q" and "ValidTC q c2 r" shows "ValidTC p (Seq c1 c2) r" by (meson assms Sem.intros(4) ValidTC_def) lemma CondRuleTC: assumes "p \ {s. (s \ b \ s \ w) \ (s \ b \ s \ w')}" and "ValidTC w c1 q" and "ValidTC w' c2 q" shows "ValidTC p (Cond b c1 c2) q" proof (unfold ValidTC_def, rule allI) fix s show "s \ p \ (\t . Sem (Cond b c1 c2) (Some s) (Some t) \ t \ q)" apply (cases "s \ b") apply (metis (mono_tags, lifting) Ball_Collect Sem.intros(6) ValidTC_def assms(1,2)) by (metis (mono_tags, lifting) Ball_Collect Sem.intros(7) ValidTC_def assms(1,3)) qed lemma WhileRuleTC: assumes "p \ i" and "\n::nat . ValidTC (i \ b \ {s . v s = n}) c (i \ {s . v s < n})" and "i \ uminus b \ q" shows "ValidTC p (While b i v c) q" proof - { fix s n have "s \ i \ v s = n \ (\t . Sem (While b i v c) (Some s) (Some t) \ t \ q)" proof (induction "n" arbitrary: s rule: less_induct) fix n :: nat fix s :: 'a assume 1: "\(m::nat) s::'a . m < n \ s \ i \ v s = m \ (\t . Sem (While b i v c) (Some s) (Some t) \ t \ q)" show "s \ i \ v s = n \ (\t . Sem (While b i v c) (Some s) (Some t) \ t \ q)" proof (rule impI, cases "s \ b") assume 2: "s \ b" and "s \ i \ v s = n" hence "s \ i \ b \ {s . v s = n}" using assms(1) by auto hence "\t . Sem c (Some s) (Some t) \ t \ i \ {s . v s < n}" by (metis assms(2) ValidTC_def) from this obtain t where 3: "Sem c (Some s) (Some t) \ t \ i \ {s . v s < n}" by auto hence "\u . Sem (While b i v c) (Some t) (Some u) \ u \ q" using 1 by auto thus "\t . Sem (While b i v c) (Some s) (Some t) \ t \ q" using 2 3 Sem.intros(10) by force next assume "s \ b" and "s \ i \ v s = n" thus "\t . Sem (While b i v c) (Some s) (Some t) \ t \ q" using Sem.intros(9) assms(3) by fastforce qed qed } thus ?thesis using assms(1) ValidTC_def by force qed subsection \Concrete syntax\ setup \ Hoare_Syntax.setup {Basic = \<^const_syntax>\Basic\, Skip = \<^const_syntax>\annskip\, Seq = \<^const_syntax>\Seq\, Cond = \<^const_syntax>\Cond\, While = \<^const_syntax>\While\, Valid = \<^const_syntax>\Valid\, ValidTC = \<^const_syntax>\ValidTC\} \ \ \Special syntax for guarded statements and guarded array updates:\ syntax "_guarded_com" :: "bool \ 'a com \ 'a com" ("(2_ \/ _)" 71) "_array_update" :: "'a list \ nat \ 'a \ 'a com" ("(2_[_] :=/ _)" [70, 65] 61) translations "P \ c" \ "IF P THEN c ELSE CONST Abort FI" "a[i] := v" \ "(i < CONST length a) \ (a := CONST list_update a i v)" \ \reverse translation not possible because of duplicate \a\\ text \ Note: there is no special syntax for guarded array access. Thus you must write \j < length a \ a[i] := a!j\. \ subsection \Proof methods: VCG\ declare BasicRule [Hoare_Tac.BasicRule] and SkipRule [Hoare_Tac.SkipRule] and AbortRule [Hoare_Tac.AbortRule] and SeqRule [Hoare_Tac.SeqRule] and CondRule [Hoare_Tac.CondRule] and WhileRule [Hoare_Tac.WhileRule] declare BasicRuleTC [Hoare_Tac.BasicRuleTC] and SkipRuleTC [Hoare_Tac.SkipRuleTC] and SeqRuleTC [Hoare_Tac.SeqRuleTC] and CondRuleTC [Hoare_Tac.CondRuleTC] and WhileRuleTC [Hoare_Tac.WhileRuleTC] method_setup vcg = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare_Tac.hoare_tac ctxt (K all_tac)))\ "verification condition generator" method_setup vcg_simp = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare_Tac.hoare_tac ctxt (asm_full_simp_tac ctxt)))\ "verification condition generator plus simplification" method_setup vcg_tc = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare_Tac.hoare_tc_tac ctxt (K all_tac)))\ "verification condition generator" method_setup vcg_tc_simp = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Hoare_Tac.hoare_tc_tac ctxt (asm_full_simp_tac ctxt)))\ "verification condition generator plus simplification" end diff --git a/src/HOL/Homology/Simplices.thy b/src/HOL/Homology/Simplices.thy --- a/src/HOL/Homology/Simplices.thy +++ b/src/HOL/Homology/Simplices.thy @@ -1,3270 +1,3271 @@ section\Homology, I: Simplices\ theory "Simplices" imports "HOL-Analysis.Function_Metric" "HOL-Analysis.Abstract_Euclidean_Space" "HOL-Algebra.Free_Abelian_Groups" begin subsection\Standard simplices, all of which are topological subspaces of @{text"R^n"}. \ type_synonym 'a chain = "((nat \ real) \ 'a) \\<^sub>0 int" definition standard_simplex :: "nat \ (nat \ real) set" where "standard_simplex p \ {x. (\i. 0 \ x i \ x i \ 1) \ (\i>p. x i = 0) \ (\i\p. x i) = 1}" lemma topspace_standard_simplex: "topspace(subtopology (powertop_real UNIV) (standard_simplex p)) = standard_simplex p" by simp lemma basis_in_standard_simplex [simp]: "(\j. if j = i then 1 else 0) \ standard_simplex p \ i \ p" by (auto simp: standard_simplex_def) lemma nonempty_standard_simplex: "standard_simplex p \ {}" using basis_in_standard_simplex by blast lemma standard_simplex_0: "standard_simplex 0 = {(\j. if j = 0 then 1 else 0)}" by (auto simp: standard_simplex_def) lemma standard_simplex_mono: assumes "p \ q" shows "standard_simplex p \ standard_simplex q" using assms proof (clarsimp simp: standard_simplex_def) fix x :: "nat \ real" assume "\i. 0 \ x i \ x i \ 1" and "\i>p. x i = 0" and "sum x {..p} = 1" then show "sum x {..q} = 1" using sum.mono_neutral_left [of "{..q}" "{..p}" x] assms by auto qed lemma closedin_standard_simplex: "closedin (powertop_real UNIV) (standard_simplex p)" (is "closedin ?X ?S") proof - have eq: "standard_simplex p = (\i. {x. x \ topspace ?X \ x i \ {0..1}}) \ (\i \ {p<..}. {x \ topspace ?X. x i \ {0}}) \ {x \ topspace ?X. (\i\p. x i) \ {1}}" by (auto simp: standard_simplex_def topspace_product_topology) show ?thesis unfolding eq by (rule closedin_Int closedin_Inter continuous_map_sum continuous_map_product_projection closedin_continuous_map_preimage | force | clarify)+ qed lemma standard_simplex_01: "standard_simplex p \ UNIV \\<^sub>E {0..1}" using standard_simplex_def by auto lemma compactin_standard_simplex: "compactin (powertop_real UNIV) (standard_simplex p)" proof (rule closed_compactin) show "compactin (powertop_real UNIV) (UNIV \\<^sub>E {0..1})" by (simp add: compactin_PiE) show "standard_simplex p \ UNIV \\<^sub>E {0..1}" by (simp add: standard_simplex_01) show "closedin (powertop_real UNIV) (standard_simplex p)" by (simp add: closedin_standard_simplex) qed lemma convex_standard_simplex: "\x \ standard_simplex p; y \ standard_simplex p; 0 \ u; u \ 1\ \ (\i. (1 - u) * x i + u * y i) \ standard_simplex p" by (simp add: standard_simplex_def sum.distrib convex_bound_le flip: sum_distrib_left) lemma path_connectedin_standard_simplex: "path_connectedin (powertop_real UNIV) (standard_simplex p)" proof - define g where "g \ \x y::nat\real. \u i. (1 - u) * x i + u * y i" have "continuous_map (subtopology euclideanreal {0..1}) (powertop_real UNIV) (g x y)" if "x \ standard_simplex p" "y \ standard_simplex p" for x y unfolding g_def continuous_map_componentwise by (force intro: continuous_intros) moreover have "g x y ` {0..1} \ standard_simplex p" "g x y 0 = x" "g x y 1 = y" if "x \ standard_simplex p" "y \ standard_simplex p" for x y using that by (auto simp: convex_standard_simplex g_def) ultimately show ?thesis unfolding path_connectedin_def path_connected_space_def pathin_def by (metis continuous_map_in_subtopology euclidean_product_topology top_greatest topspace_euclidean topspace_euclidean_subtopology) qed lemma connectedin_standard_simplex: "connectedin (powertop_real UNIV) (standard_simplex p)" by (simp add: path_connectedin_imp_connectedin path_connectedin_standard_simplex) subsection\Face map\ definition simplical_face :: "nat \ (nat \ 'a) \ nat \ 'a::comm_monoid_add" where "simplical_face k x \ \i. if i < k then x i else if i = k then 0 else x(i -1)" lemma simplical_face_in_standard_simplex: assumes "1 \ p" "k \ p" "x \ standard_simplex (p - Suc 0)" shows "(simplical_face k x) \ standard_simplex p" proof - have x01: "\i. 0 \ x i \ x i \ 1" and sumx: "sum x {..p - Suc 0} = 1" using assms by (auto simp: standard_simplex_def simplical_face_def) have gg: "\g. sum g {..p} = sum g {..k \ p\ sum.union_disjoint [of "{..i\p. if i < k then x i else if i = k then 0 else x (i -1)) = (\i < k. x i) + (\i \ {k..p}. if i = k then 0 else x (i -1))" by (simp add: gg) consider "k \ p - Suc 0" | "k = p" using \k \ p\ by linarith then have "(\i\p. if i < k then x i else if i = k then 0 else x (i -1)) = 1" proof cases case 1 have [simp]: "Suc (p - Suc 0) = p" using \1 \ p\ by auto have "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k+1..p. if i = k then 0 else x (i -1))" by (rule sum.mono_neutral_right) auto also have "\ = (\i = k+1..p. x (i -1))" by simp also have "\ = (\i = k..p-1. x i)" using sum.atLeastAtMost_reindex [of Suc k "p-1" "\i. x (i - Suc 0)"] 1 by simp finally have eq2: "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k..p-1. x i)" . with 1 show ?thesis by (metis (no_types, lifting) One_nat_def eq finite_atLeastAtMost finite_lessThan ivl_disj_int(4) ivl_disj_un(10) sum.union_disjoint sumx) next case 2 have [simp]: "({..p} \ {x. x < p}) = {..p - Suc 0}" using assms by auto have "(\i\p. if i < p then x i else if i = k then 0 else x (i -1)) = (\i\p. if i < p then x i else 0)" by (rule sum.cong) (auto simp: 2) also have "\ = sum x {..p-1}" by (simp add: sum.If_cases) also have "\ = 1" by (simp add: sumx) finally show ?thesis using 2 by simp qed then show ?thesis using assms by (auto simp: standard_simplex_def simplical_face_def) qed subsection\Singular simplices, forcing canonicity outside the intended domain\ definition singular_simplex :: "nat \ 'a topology \ ((nat \ real) \ 'a) \ bool" where "singular_simplex p X f \ continuous_map(subtopology (powertop_real UNIV) (standard_simplex p)) X f \ f \ extensional (standard_simplex p)" abbreviation singular_simplex_set :: "nat \ 'a topology \ ((nat \ real) \ 'a) set" where "singular_simplex_set p X \ Collect (singular_simplex p X)" lemma singular_simplex_empty: "topspace X = {} \ \ singular_simplex p X f" by (simp add: singular_simplex_def continuous_map nonempty_standard_simplex) lemma singular_simplex_mono: "\singular_simplex p (subtopology X T) f; T \ S\ \ singular_simplex p (subtopology X S) f" by (auto simp: singular_simplex_def continuous_map_in_subtopology) lemma singular_simplex_subtopology: "singular_simplex p (subtopology X S) f \ singular_simplex p X f \ f ` (standard_simplex p) \ S" by (auto simp: singular_simplex_def continuous_map_in_subtopology) subsubsection\Singular face\ definition singular_face :: "nat \ nat \ ((nat \ real) \ 'a) \ (nat \ real) \ 'a" where "singular_face p k f \ restrict (f \ simplical_face k) (standard_simplex (p - Suc 0))" lemma singular_simplex_singular_face: assumes f: "singular_simplex p X f" and "1 \ p" "k \ p" shows "singular_simplex (p - Suc 0) X (singular_face p k f)" proof - let ?PT = "(powertop_real UNIV)" have 0: "simplical_face k ` standard_simplex (p - Suc 0) \ standard_simplex p" using assms simplical_face_in_standard_simplex by auto have 1: "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) (subtopology ?PT (standard_simplex p)) (simplical_face k)" proof (clarsimp simp add: continuous_map_in_subtopology simplical_face_in_standard_simplex continuous_map_componentwise 0) fix i have "continuous_map ?PT euclideanreal (\x. if i < k then x i else if i = k then 0 else x (i -1))" by (auto intro: continuous_map_product_projection) then show "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) euclideanreal (\x. simplical_face k x i)" by (simp add: simplical_face_def continuous_map_from_subtopology) qed have 2: "continuous_map (subtopology ?PT (standard_simplex p)) X f" using assms(1) singular_simplex_def by blast show ?thesis by (simp add: singular_simplex_def singular_face_def continuous_map_compose [OF 1 2]) qed subsection\Singular chains\ definition singular_chain :: "[nat, 'a topology, 'a chain] \ bool" where "singular_chain p X c \ Poly_Mapping.keys c \ singular_simplex_set p X" abbreviation singular_chain_set :: "[nat, 'a topology] \ ('a chain) set" where "singular_chain_set p X \ Collect (singular_chain p X)" lemma singular_chain_empty: "topspace X = {} \ singular_chain p X c \ c = 0" by (auto simp: singular_chain_def singular_simplex_empty subset_eq poly_mapping_eqI) lemma singular_chain_mono: "\singular_chain p (subtopology X T) c; T \ S\ \ singular_chain p (subtopology X S) c" unfolding singular_chain_def using singular_simplex_mono by blast lemma singular_chain_subtopology: "singular_chain p (subtopology X S) c \ singular_chain p X c \ (\f \ Poly_Mapping.keys c. f ` (standard_simplex p) \ S)" unfolding singular_chain_def by (fastforce simp add: singular_simplex_subtopology subset_eq) lemma singular_chain_0 [iff]: "singular_chain p X 0" by (auto simp: singular_chain_def) lemma singular_chain_of: "singular_chain p X (frag_of c) \ singular_simplex p X c" by (auto simp: singular_chain_def) lemma singular_chain_cmul: "singular_chain p X c \ singular_chain p X (frag_cmul a c)" by (auto simp: singular_chain_def) lemma singular_chain_minus: "singular_chain p X (-c) \ singular_chain p X c" by (auto simp: singular_chain_def) lemma singular_chain_add: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a+b)" unfolding singular_chain_def using keys_add [of a b] by blast lemma singular_chain_diff: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a-b)" unfolding singular_chain_def using keys_diff [of a b] by blast lemma singular_chain_sum: "(\i. i \ I \ singular_chain p X (f i)) \ singular_chain p X (\i\I. f i)" unfolding singular_chain_def using keys_sum [of f I] by blast lemma singular_chain_extend: "(\c. c \ Poly_Mapping.keys x \ singular_chain p X (f c)) \ singular_chain p X (frag_extend f x)" by (simp add: frag_extend_def singular_chain_cmul singular_chain_sum) subsection\Boundary homomorphism for singular chains\ definition chain_boundary :: "nat \ ('a chain) \ 'a chain" where "chain_boundary p c \ (if p = 0 then 0 else frag_extend (\f. (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f)))) c)" lemma singular_chain_boundary: assumes "singular_chain p X c" shows "singular_chain (p - Suc 0) X (chain_boundary p c)" unfolding chain_boundary_def proof (clarsimp intro!: singular_chain_extend singular_chain_sum singular_chain_cmul) show "\d k. \0 < p; d \ Poly_Mapping.keys c; k \ p\ \ singular_chain (p - Suc 0) X (frag_of (singular_face p k d))" using assms by (auto simp: singular_chain_def intro: singular_simplex_singular_face) qed lemma singular_chain_boundary_alt: "singular_chain (Suc p) X c \ singular_chain p X (chain_boundary (Suc p) c)" using singular_chain_boundary by force lemma chain_boundary_0 [simp]: "chain_boundary p 0 = 0" by (simp add: chain_boundary_def) lemma chain_boundary_cmul: "chain_boundary p (frag_cmul k c) = frag_cmul k (chain_boundary p c)" by (auto simp: chain_boundary_def frag_extend_cmul) lemma chain_boundary_minus: "chain_boundary p (- c) = - (chain_boundary p c)" by (metis chain_boundary_cmul frag_cmul_minus_one) lemma chain_boundary_add: "chain_boundary p (a+b) = chain_boundary p a + chain_boundary p b" by (simp add: chain_boundary_def frag_extend_add) lemma chain_boundary_diff: "chain_boundary p (a-b) = chain_boundary p a - chain_boundary p b" using chain_boundary_add [of p a "-b"] by (simp add: chain_boundary_minus) lemma chain_boundary_sum: "chain_boundary p (sum g I) = sum (chain_boundary p \ g) I" by (induction I rule: infinite_finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_sum': "finite I \ chain_boundary p (sum' g I) = sum' (chain_boundary p \ g) I" by (induction I rule: finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_of: "chain_boundary p (frag_of f) = (if p = 0 then 0 else (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f))))" by (simp add: chain_boundary_def) subsection\Factoring out chains in a subtopology for relative homology\ definition mod_subset where "mod_subset p X \ {(a,b). singular_chain p X (a - b)}" lemma mod_subset_empty [simp]: "(a,b) \ (mod_subset p (subtopology X {})) \ a = b" by (simp add: mod_subset_def singular_chain_empty) lemma mod_subset_refl [simp]: "(c,c) \ mod_subset p X" by (auto simp: mod_subset_def) lemma mod_subset_cmul: assumes "(a,b) \ mod_subset p X" shows "(frag_cmul k a, frag_cmul k b) \ mod_subset p X" using assms by (simp add: mod_subset_def) (metis (no_types, lifting) add_diff_cancel diff_add_cancel frag_cmul_distrib2 singular_chain_cmul) lemma mod_subset_add: "\(c1,c2) \ mod_subset p X; (d1,d2) \ mod_subset p X\ \ (c1+d1, c2+d2) \ mod_subset p X" by (simp add: mod_subset_def add_diff_add singular_chain_add) subsection\Relative cycles $Z_pX (S)$ where $X$ is a topology and $S$ a subset \ definition singular_relcycle :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relcycle p X S \ \c. singular_chain p X c \ (chain_boundary p c, 0) \ mod_subset (p-1) (subtopology X S)" abbreviation singular_relcycle_set where "singular_relcycle_set p X S \ Collect (singular_relcycle p X S)" lemma singular_relcycle_restrict [simp]: "singular_relcycle p X (topspace X \ S) = singular_relcycle p X S" proof - have eq: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) show ?thesis by (force simp: singular_relcycle_def eq) qed lemma singular_relcycle: "singular_relcycle p X S c \ singular_chain p X c \ singular_chain (p-1) (subtopology X S) (chain_boundary p c)" by (simp add: singular_relcycle_def mod_subset_def) lemma singular_relcycle_0 [simp]: "singular_relcycle p X S 0" by (auto simp: singular_relcycle_def) lemma singular_relcycle_cmul: "singular_relcycle p X S c \ singular_relcycle p X S (frag_cmul k c)" by (auto simp: singular_relcycle_def chain_boundary_cmul dest: singular_chain_cmul mod_subset_cmul) lemma singular_relcycle_minus: "singular_relcycle p X S (-c) \ singular_relcycle p X S c" by (simp add: chain_boundary_minus singular_chain_minus singular_relcycle) lemma singular_relcycle_add: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a+b)" by (simp add: singular_relcycle_def chain_boundary_add mod_subset_def singular_chain_add) lemma singular_relcycle_sum: "\\i. i \ I \ singular_relcycle p X S (f i)\ \ singular_relcycle p X S (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: singular_relcycle_add) lemma singular_relcycle_diff: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a-b)" by (metis singular_relcycle_add singular_relcycle_minus uminus_add_conv_diff) lemma singular_cycle: "singular_relcycle p X {} c \ singular_chain p X c \ chain_boundary p c = 0" by (simp add: singular_relcycle_def) lemma singular_cycle_mono: "\singular_relcycle p (subtopology X T) {} c; T \ S\ \ singular_relcycle p (subtopology X S) {} c" by (auto simp: singular_cycle elim: singular_chain_mono) subsection\Relative boundaries $B_p X S$, where $X$ is a topology and $S$ a subset.\ definition singular_relboundary :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relboundary p X S \ \c. \d. singular_chain (Suc p) X d \ (chain_boundary (Suc p) d, c) \ (mod_subset p (subtopology X S))" abbreviation singular_relboundary_set :: "nat \ 'a topology \ 'a set \ ('a chain) set" where "singular_relboundary_set p X S \ Collect (singular_relboundary p X S)" lemma singular_relboundary_restrict [simp]: "singular_relboundary p X (topspace X \ S) = singular_relboundary p X S" unfolding singular_relboundary_def by (metis (no_types, hide_lams) subtopology_subtopology subtopology_topspace) lemma singular_relboundary_alt: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ chain_boundary (Suc p) d = c + e)" unfolding singular_relboundary_def mod_subset_def by fastforce lemma singular_relboundary: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ (chain_boundary (Suc p) d) + e = c)" using singular_chain_minus by (fastforce simp add: singular_relboundary_alt) lemma singular_boundary: "singular_relboundary p X {} c \ (\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c)" by (simp add: singular_relboundary_def) lemma singular_boundary_imp_chain: "singular_relboundary p X {} c \ singular_chain p X c" by (auto simp: singular_relboundary singular_chain_boundary_alt singular_chain_empty topspace_subtopology) lemma singular_boundary_mono: "\T \ S; singular_relboundary p (subtopology X T) {} c\ \ singular_relboundary p (subtopology X S) {} c" by (metis mod_subset_empty singular_chain_mono singular_relboundary_def) lemma singular_relboundary_imp_chain: "singular_relboundary p X S c \ singular_chain p X c" unfolding singular_relboundary singular_chain_subtopology by (blast intro: singular_chain_add singular_chain_boundary_alt) lemma singular_chain_imp_relboundary: "singular_chain p (subtopology X S) c \ singular_relboundary p X S c" unfolding singular_relboundary_def using mod_subset_def singular_chain_minus by fastforce lemma singular_relboundary_0 [simp]: "singular_relboundary p X S 0" unfolding singular_relboundary_def by (rule_tac x=0 in exI) auto lemma singular_relboundary_cmul: "singular_relboundary p X S c \ singular_relboundary p X S (frag_cmul a c)" unfolding singular_relboundary_def by (metis chain_boundary_cmul mod_subset_cmul singular_chain_cmul) lemma singular_relboundary_minus: "singular_relboundary p X S (-c) \ singular_relboundary p X S c" using singular_relboundary_cmul by (metis add.inverse_inverse frag_cmul_minus_one) lemma singular_relboundary_add: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a+b)" unfolding singular_relboundary_def by (metis chain_boundary_add mod_subset_add singular_chain_add) lemma singular_relboundary_diff: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a-b)" by (metis uminus_add_conv_diff singular_relboundary_minus singular_relboundary_add) subsection\The (relative) homology relation\ definition homologous_rel :: "[nat,'a topology,'a set,'a chain,'a chain] \ bool" where "homologous_rel p X S \ \a b. singular_relboundary p X S (a-b)" abbreviation homologous_rel_set where "homologous_rel_set p X S a \ Collect (homologous_rel p X S a)" lemma homologous_rel_restrict [simp]: "homologous_rel p X (topspace X \ S) = homologous_rel p X S" unfolding homologous_rel_def by (metis singular_relboundary_restrict) lemma homologous_rel_refl [simp]: "homologous_rel p X S c c" unfolding homologous_rel_def by auto lemma homologous_rel_sym: "homologous_rel p X S a b = homologous_rel p X S b a" unfolding homologous_rel_def using singular_relboundary_minus by fastforce lemma homologous_rel_trans: assumes "homologous_rel p X S b c" "homologous_rel p X S a b" shows "homologous_rel p X S a c" using homologous_rel_def proof - have "singular_relboundary p X S (b - c)" using assms unfolding homologous_rel_def by blast moreover have "singular_relboundary p X S (b - a)" using assms by (meson homologous_rel_def homologous_rel_sym) ultimately have "singular_relboundary p X S (c - a)" using singular_relboundary_diff by fastforce then show ?thesis by (meson homologous_rel_def homologous_rel_sym) qed lemma homologous_rel_eq: "homologous_rel p X S a = homologous_rel p X S b \ homologous_rel p X S a b" using homologous_rel_sym homologous_rel_trans by fastforce lemma homologous_rel_set_eq: "homologous_rel_set p X S a = homologous_rel_set p X S b \ homologous_rel p X S a b" by (metis homologous_rel_eq mem_Collect_eq) lemma homologous_rel_singular_chain: "homologous_rel p X S a b \ (singular_chain p X a \ singular_chain p X b)" unfolding homologous_rel_def using singular_chain_diff singular_chain_add by (fastforce dest: singular_relboundary_imp_chain) lemma homologous_rel_add: "\homologous_rel p X S a a'; homologous_rel p X S b b'\ \ homologous_rel p X S (a+b) (a'+b')" unfolding homologous_rel_def by (simp add: add_diff_add singular_relboundary_add) lemma homologous_rel_diff: assumes "homologous_rel p X S a a'" "homologous_rel p X S b b'" shows "homologous_rel p X S (a - b) (a' - b')" proof - have "singular_relboundary p X S ((a - a') - (b - b'))" using assms singular_relboundary_diff unfolding homologous_rel_def by blast then show ?thesis by (simp add: homologous_rel_def algebra_simps) qed lemma homologous_rel_sum: assumes f: "finite {i \ I. f i \ 0}" and g: "finite {i \ I. g i \ 0}" and h: "\i. i \ I \ homologous_rel p X S (f i) (g i)" shows "homologous_rel p X S (sum f I) (sum g I)" proof (cases "finite I") case True let ?L = "{i \ I. f i \ 0} \ {i \ I. g i \ 0}" have L: "finite ?L" "?L \ I" using f g by blast+ have "sum f I = sum f ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have "sum g I = sum g ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have *: "homologous_rel p X S (f i) (g i)" if "i \ ?L" for i using h that by auto have "homologous_rel p X S (sum f ?L) (sum g ?L)" using L proof induction case (insert j J) then show ?case by (simp add: h homologous_rel_add) qed auto ultimately show ?thesis by simp qed auto lemma chain_homotopic_imp_homologous_rel: assumes "\c. singular_chain p X c \ singular_chain (Suc p) X' (h c)" "\c. singular_chain (p -1) (subtopology X S) c \ singular_chain p (subtopology X' T) (h' c)" "\c. singular_chain p X c \ (chain_boundary (Suc p) (h c)) + (h'(chain_boundary p c)) = f c - g c" "singular_relcycle p X S c" shows "homologous_rel p X' T (f c) (g c)" proof - have "singular_chain p (subtopology X' T) (chain_boundary (Suc p) (h c) - (f c - g c))" using assms by (metis (no_types, lifting) add_diff_cancel_left' minus_diff_eq singular_chain_minus singular_relcycle) then show ?thesis using assms by (metis homologous_rel_def singular_relboundary singular_relcycle) qed subsection\Show that all boundaries are cycles, the key "chain complex" property.\ lemma chain_boundary_boundary: assumes "singular_chain p X c" shows "chain_boundary (p - Suc 0) (chain_boundary p c) = 0" proof (cases "p -1 = 0") case False then have "2 \ p" by auto show ?thesis using assms unfolding singular_chain_def proof (induction rule: frag_induction) case (one g) then have ss: "singular_simplex p X g" by simp have eql: "{..p} \ {..p - Suc 0} \ {(x, y). y < x} = (\(j,i). (Suc i, j)) ` {(i,j). i \ j \ j \ p -1}" using False by (auto simp: image_def) (metis One_nat_def diff_Suc_1 diff_le_mono le_refl lessE less_imp_le_nat) have eqr: "{..p} \ {..p - Suc 0} - {(x, y). y < x} = {(i,j). i \ j \ j \ p -1}" by auto have eqf: "singular_face (p - Suc 0) i (singular_face p (Suc j) g) = singular_face (p - Suc 0) j (singular_face p i g)" if "i \ j" "j \ p - Suc 0" for i j proof (rule ext) fix t show "singular_face (p - Suc 0) i (singular_face p (Suc j) g) t = singular_face (p - Suc 0) j (singular_face p i g) t" proof (cases "t \ standard_simplex (p -1 -1)") case True have fi: "simplical_face i t \ standard_simplex (p - Suc 0)" using False True simplical_face_in_standard_simplex that by force have fj: "simplical_face j t \ standard_simplex (p - Suc 0)" by (metis False One_nat_def True simplical_face_in_standard_simplex less_one not_less that(2)) have eq: "simplical_face (Suc j) (simplical_face i t) = simplical_face i (simplical_face j t)" using True that ss unfolding standard_simplex_def simplical_face_def by fastforce show ?thesis by (simp add: singular_face_def fi fj eq) qed (simp add: singular_face_def) qed show ?case proof (cases "p = 1") case False have eq0: "frag_cmul (-1) a = b \ a + b = 0" for a b by (simp add: neg_eq_iff_add_eq_0) have *: "(\x\p. \i\p - Suc 0. frag_cmul ((-1) ^ (x + i)) (frag_of (singular_face (p - Suc 0) i (singular_face p x g)))) = 0" apply (simp add: sum.cartesian_product sum.Int_Diff [of "_ \ _" _ "{(x,y). y < x}"]) apply (rule eq0) unfolding frag_cmul_sum prod.case_distrib [of "frag_cmul (-1)"] frag_cmul_cmul eql eqr apply (force simp: inj_on_def sum.reindex add.commute eqf intro: sum.cong) done show ?thesis using False by (simp add: chain_boundary_of chain_boundary_sum chain_boundary_cmul frag_cmul_sum * flip: power_add) qed (simp add: chain_boundary_def) next case (diff a b) then show ?case by (simp add: chain_boundary_diff) qed auto qed (simp add: chain_boundary_def) lemma chain_boundary_boundary_alt: "singular_chain (Suc p) X c \ chain_boundary p (chain_boundary (Suc p) c) = 0" using chain_boundary_boundary by force lemma singular_relboundary_imp_relcycle: assumes "singular_relboundary p X S c" shows "singular_relcycle p X S c" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary singular_relcycle) have 1: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p (chain_boundary (Suc p) d))" using d chain_boundary_boundary_alt by fastforce have 2: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p e)" using \singular_chain p (subtopology X S) e\ singular_chain_boundary by auto have "singular_chain p X c" using assms singular_relboundary_imp_chain by auto moreover have "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p c)" by (simp add: c chain_boundary_add singular_chain_add 1 2) ultimately show ?thesis by (simp add: singular_relcycle) qed lemma homologous_rel_singular_relcycle_1: assumes "homologous_rel p X S c1 c2" "singular_relcycle p X S c1" shows "singular_relcycle p X S c2" using assms by (metis diff_add_cancel homologous_rel_def homologous_rel_sym singular_relboundary_imp_relcycle singular_relcycle_add) lemma homologous_rel_singular_relcycle: assumes "homologous_rel p X S c1 c2" shows "singular_relcycle p X S c1 = singular_relcycle p X S c2" using assms homologous_rel_singular_relcycle_1 using homologous_rel_sym by blast subsection\Operations induced by a continuous map g between topological spaces\ definition simplex_map :: "nat \ ('b \ 'a) \ ((nat \ real) \ 'b) \ (nat \ real) \ 'a" where "simplex_map p g c \ restrict (g \ c) (standard_simplex p)" lemma singular_simplex_simplex_map: "\singular_simplex p X f; continuous_map X X' g\ \ singular_simplex p X' (simplex_map p g f)" unfolding singular_simplex_def simplex_map_def by (auto simp: continuous_map_compose) lemma simplex_map_eq: "\singular_simplex p X c; \x. x \ topspace X \ f x = g x\ \ simplex_map p f c = simplex_map p g c" by (auto simp: singular_simplex_def simplex_map_def continuous_map_def) lemma simplex_map_id_gen: "\singular_simplex p X c; \x. x \ topspace X \ f x = x\ \ simplex_map p f c = c" unfolding singular_simplex_def simplex_map_def continuous_map_def using extensional_arb by fastforce lemma simplex_map_id [simp]: "simplex_map p id = (\c. restrict c (standard_simplex p))" by (auto simp: simplex_map_def) lemma simplex_map_compose: "simplex_map p (h \ g) = simplex_map p h \ simplex_map p g" unfolding simplex_map_def by force lemma singular_face_simplex_map: "\1 \ p; k \ p\ \ singular_face p k (simplex_map p f c) = simplex_map (p - Suc 0) f (c \ simplical_face k)" unfolding simplex_map_def singular_face_def by (force simp: simplical_face_in_standard_simplex) lemma singular_face_restrict [simp]: assumes "p > 0" "i \ p" shows "singular_face p i (restrict f (standard_simplex p)) = singular_face p i f" by (metis assms One_nat_def Suc_leI simplex_map_id singular_face_def singular_face_simplex_map) definition chain_map :: "nat \ ('b \ 'a) \ (((nat \ real) \ 'b) \\<^sub>0 int) \ 'a chain" where "chain_map p g c \ frag_extend (frag_of \ simplex_map p g) c" lemma singular_chain_chain_map: "\singular_chain p X c; continuous_map X X' g\ \ singular_chain p X' (chain_map p g c)" unfolding chain_map_def by (force simp add: singular_chain_def subset_iff intro!: singular_chain_extend singular_simplex_simplex_map) lemma chain_map_0 [simp]: "chain_map p g 0 = 0" by (auto simp: chain_map_def) lemma chain_map_of [simp]: "chain_map p g (frag_of f) = frag_of (simplex_map p g f)" by (simp add: chain_map_def) lemma chain_map_cmul [simp]: "chain_map p g (frag_cmul a c) = frag_cmul a (chain_map p g c)" by (simp add: frag_extend_cmul chain_map_def) lemma chain_map_minus: "chain_map p g (-c) = - (chain_map p g c)" by (simp add: frag_extend_minus chain_map_def) lemma chain_map_add: "chain_map p g (a+b) = chain_map p g a + chain_map p g b" by (simp add: frag_extend_add chain_map_def) lemma chain_map_diff: "chain_map p g (a-b) = chain_map p g a - chain_map p g b" by (simp add: frag_extend_diff chain_map_def) lemma chain_map_sum: "finite I \ chain_map p g (sum f I) = sum (chain_map p g \ f) I" by (simp add: frag_extend_sum chain_map_def) lemma chain_map_eq: "\singular_chain p X c; \x. x \ topspace X \ f x = g x\ \ chain_map p f c = chain_map p g c" unfolding singular_chain_def proof (induction rule: frag_induction) case (one x) then show ?case by (metis (no_types, lifting) chain_map_of mem_Collect_eq simplex_map_eq) qed (auto simp: chain_map_diff) lemma chain_map_id_gen: "\singular_chain p X c; \x. x \ topspace X \ f x = x\ \ chain_map p f c = c" unfolding singular_chain_def by (erule frag_induction) (auto simp: chain_map_diff simplex_map_id_gen) lemma chain_map_ident: "singular_chain p X c \ chain_map p id c = c" by (simp add: chain_map_id_gen) lemma chain_map_id: "chain_map p id = frag_extend (frag_of \ (\f. restrict f (standard_simplex p)))" by (auto simp: chain_map_def) lemma chain_map_compose: "chain_map p (h \ g) = chain_map p h \ chain_map p g" proof show "chain_map p (h \ g) c = (chain_map p h \ chain_map p g) c" for c using subset_UNIV proof (induction c rule: frag_induction) case (one x) then show ?case by simp (metis (mono_tags, lifting) comp_eq_dest_lhs restrict_apply simplex_map_def) next case (diff a b) then show ?case by (simp add: chain_map_diff) qed auto qed lemma singular_simplex_chain_map_id: assumes "singular_simplex p X f" shows "chain_map p f (frag_of (restrict id (standard_simplex p))) = frag_of f" proof - have "(restrict (f \ restrict id (standard_simplex p)) (standard_simplex p)) = f" by (rule ext) (metis assms comp_apply extensional_arb id_apply restrict_apply singular_simplex_def) then show ?thesis by (simp add: simplex_map_def) qed lemma chain_boundary_chain_map: assumes "singular_chain p X c" shows "chain_boundary p (chain_map p g c) = chain_map (p - Suc 0) g (chain_boundary p c)" using assms unfolding singular_chain_def proof (induction c rule: frag_induction) case (one x) then have "singular_face p i (simplex_map p g x) = simplex_map (p - Suc 0) g (singular_face p i x)" if "0 \ i" "i \ p" "p \ 0" for i using that by (fastforce simp add: singular_face_def simplex_map_def simplical_face_in_standard_simplex) then show ?case by (auto simp: chain_boundary_of chain_map_sum) next case (diff a b) then show ?case by (simp add: chain_boundary_diff chain_map_diff) qed auto lemma singular_relcycle_chain_map: assumes "singular_relcycle p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relcycle p X' T (chain_map p g c)" proof - have "continuous_map (subtopology X S) (subtopology X' T) g" using assms using continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce then show ?thesis using chain_boundary_chain_map [of p X c g] by (metis One_nat_def assms(1) assms(2) singular_chain_chain_map singular_relcycle) qed lemma singular_relboundary_chain_map: assumes "singular_relboundary p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relboundary p X' T (chain_map p g c)" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary) have "singular_chain (Suc p) X' (chain_map (Suc p) g d)" using assms(2) d singular_chain_chain_map by blast moreover have "singular_chain p (subtopology X' T) (chain_map p g e)" proof - have "\t. g ` topspace (subtopology t S) \ T" by (metis assms(3) closure_of_subset_subtopology closure_of_topspace dual_order.trans image_mono) then show ?thesis by (meson assms(2) continuous_map_from_subtopology continuous_map_in_subtopology e singular_chain_chain_map) qed moreover have "chain_boundary (Suc p) (chain_map (Suc p) g d) + chain_map p g e = chain_map p g (chain_boundary (Suc p) d + e)" by (metis One_nat_def chain_boundary_chain_map chain_map_add d diff_Suc_1) ultimately show ?thesis unfolding singular_relboundary using c by blast qed subsection\Homology of one-point spaces degenerates except for $p = 0$.\ lemma singular_simplex_singleton: assumes "topspace X = {a}" shows "singular_simplex p X f \ f = restrict (\x. a) (standard_simplex p)" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs proof - have "continuous_map (subtopology (product_topology (\n. euclideanreal) UNIV) (standard_simplex p)) X f" using \singular_simplex p X f\ singular_simplex_def by blast then have "\c. c \ standard_simplex p \ f c = a" by (simp add: assms continuous_map_def) then show ?thesis by (metis (no_types) L extensional_restrict restrict_ext singular_simplex_def) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_simplex_def topspace_subtopology) qed lemma singular_chain_singleton: assumes "topspace X = {a}" shows "singular_chain p X c \ (\b. c = frag_cmul b (frag_of(restrict (\x. a) (standard_simplex p))))" (is "?lhs = ?rhs") proof let ?f = "restrict (\x. a) (standard_simplex p)" assume L: ?lhs with assms have "Poly_Mapping.keys c \ {?f}" by (auto simp: singular_chain_def singular_simplex_singleton) then consider "Poly_Mapping.keys c = {}" | "Poly_Mapping.keys c = {?f}" by blast then show ?rhs proof cases case 1 with L show ?thesis by (metis frag_cmul_zero keys_eq_empty) next case 2 then have "\b. frag_extend frag_of c = frag_cmul b (frag_of (\x\standard_simplex p. a))" by (force simp: frag_extend_def) then show ?thesis by (metis frag_expansion) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_chain_def singular_simplex_singleton) qed lemma chain_boundary_of_singleton: assumes tX: "topspace X = {a}" and sc: "singular_chain p X c" shows "chain_boundary p c = (if p = 0 \ odd p then 0 else frag_extend (\f. frag_of(restrict (\x. a) (standard_simplex (p -1)))) c)" (is "?lhs = ?rhs") proof (cases "p = 0") case False have "?lhs = frag_extend (\f. if odd p then 0 else frag_of(restrict (\x. a) (standard_simplex (p -1)))) c" proof (simp only: chain_boundary_def False if_False, rule frag_extend_eq) fix f assume "f \ Poly_Mapping.keys c" with assms have "singular_simplex p X f" by (auto simp: singular_chain_def) then have *: "\k. k \ p \ singular_face p k f = (\x\standard_simplex (p -1). a)" using False singular_simplex_singular_face by (fastforce simp flip: singular_simplex_singleton [OF tX]) define c where "c \ frag_of (\x\standard_simplex (p -1). a)" have "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (\k\p. frag_cmul ((-1) ^ k) c)" by (auto simp: c_def * intro: sum.cong) also have "\ = (if odd p then 0 else c)" by (induction p) (auto simp: c_def restrict_def) finally show "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (if odd p then 0 else frag_of (\x\standard_simplex (p -1). a))" unfolding c_def . qed also have "\ = ?rhs" by (auto simp: False frag_extend_eq_0) finally show ?thesis . qed (simp add: chain_boundary_def) lemma singular_cycle_singleton: assumes "topspace X = {a}" shows "singular_relcycle p X {} c \ singular_chain p X c \ (p = 0 \ odd p \ c = 0)" proof - have "c = 0" if "singular_chain p X c" and "chain_boundary p c = 0" and "even p" and "p \ 0" using that assms singular_chain_singleton [of X a p c] chain_boundary_of_singleton [OF assms] by (auto simp: frag_extend_cmul) moreover have "chain_boundary p c = 0" if sc: "singular_chain p X c" and "odd p" by (simp add: chain_boundary_of_singleton [OF assms sc] that) moreover have "chain_boundary 0 c = 0" if "singular_chain 0 X c" and "p = 0" by (simp add: chain_boundary_def) ultimately show ?thesis using assms by (auto simp: singular_cycle) qed lemma singular_boundary_singleton: assumes "topspace X = {a}" shows "singular_relboundary p X {} c \ singular_chain p X c \ (odd p \ c = 0)" proof (cases "singular_chain p X c") case True have "\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c" if "singular_chain p X c" and "odd p" proof - obtain b where b: "c = frag_cmul b (frag_of(restrict (\x. a) (standard_simplex p)))" by (metis True assms singular_chain_singleton) let ?d = "frag_cmul b (frag_of (\x\standard_simplex (Suc p). a))" have scd: "singular_chain (Suc p) X ?d" by (metis assms singular_chain_singleton) moreover have "chain_boundary (Suc p) ?d = c" by (simp add: assms scd chain_boundary_of_singleton [of X a "Suc p"] b frag_extend_cmul \odd p\) ultimately show ?thesis by metis qed with True assms show ?thesis by (auto simp: singular_boundary chain_boundary_of_singleton) next case False with assms singular_boundary_imp_chain show ?thesis by metis qed lemma singular_boundary_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary p X {} c \ singular_relcycle p X {} c" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: singular_relboundary_imp_relcycle) show "?rhs \ ?lhs" by (metis assms not_one_le_zero singular_boundary_singleton singular_cycle_singleton) qed lemma singular_boundary_set_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary_set p X {} = singular_relcycle_set p X {}" using singular_boundary_eq_cycle_singleton [OF assms] by blast subsection\Simplicial chains\ text\Simplicial chains, effectively those resulting from linear maps. We still allow the map to be singular, so the name is questionable. These are intended as building-blocks for singular subdivision, rather than as a axis for 1 simplicial homology.\ definition oriented_simplex where "oriented_simplex p l \ (\x\standard_simplex p. \i. (\j\p. l j i * x j))" definition simplicial_simplex where "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f \ (\l. f = oriented_simplex p l)" lemma simplicial_simplex: "simplicial_simplex p S f \ f ` (standard_simplex p) \ S \ (\l. f = oriented_simplex p l)" (is "?lhs = ?rhs") proof assume R: ?rhs have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) (powertop_real UNIV) (\x i. \j\p. l j i * x j)" for l :: " nat \ 'a \ real" unfolding continuous_map_componentwise by (force intro: continuous_intros continuous_map_from_subtopology continuous_map_product_projection) with R show ?lhs unfolding simplicial_simplex_def singular_simplex_subtopology by (auto simp add: singular_simplex_def oriented_simplex_def) qed (simp add: simplicial_simplex_def singular_simplex_subtopology) lemma simplicial_simplex_empty [simp]: "\ simplicial_simplex p {} f" by (simp add: nonempty_standard_simplex simplicial_simplex) definition simplicial_chain where "simplicial_chain p S c \ Poly_Mapping.keys c \ Collect (simplicial_simplex p S)" lemma simplicial_chain_0 [simp]: "simplicial_chain p S 0" by (simp add: simplicial_chain_def) lemma simplicial_chain_of [simp]: "simplicial_chain p S (frag_of c) \ simplicial_simplex p S c" by (simp add: simplicial_chain_def) lemma simplicial_chain_cmul: "simplicial_chain p S c \ simplicial_chain p S (frag_cmul a c)" by (auto simp: simplicial_chain_def) lemma simplicial_chain_diff: "\simplicial_chain p S c1; simplicial_chain p S c2\ \ simplicial_chain p S (c1 - c2)" unfolding simplicial_chain_def by (meson UnE keys_diff subset_iff) lemma simplicial_chain_sum: "(\i. i \ I \ simplicial_chain p S (f i)) \ simplicial_chain p S (sum f I)" unfolding simplicial_chain_def using order_trans [OF keys_sum [of f I]] by (simp add: UN_least) lemma simplicial_simplex_oriented_simplex: "simplicial_simplex p S (oriented_simplex p l) \ ((\x i. \j\p. l j i * x j) ` standard_simplex p \ S)" by (auto simp: simplicial_simplex oriented_simplex_def) lemma simplicial_imp_singular_simplex: "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f" by (simp add: simplicial_simplex_def) lemma simplicial_imp_singular_chain: "simplicial_chain p S c \ singular_chain p (subtopology (powertop_real UNIV) S) c" unfolding simplicial_chain_def singular_chain_def by (auto intro: simplicial_imp_singular_simplex) lemma oriented_simplex_eq: "oriented_simplex p l = oriented_simplex p l' \ (\i. i \ p \ l i = l' i)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix i assume "i \ p" let ?fi = "(\j. if j = i then 1 else 0)" have "(\j\p. l j k * ?fi j) = (\j\p. l' j k * ?fi j)" for k using L \i \ p\ by (simp add: fun_eq_iff oriented_simplex_def split: if_split_asm) with \i \ p\ show "l i = l' i" by (simp add: if_distrib ext cong: if_cong) qed qed (auto simp: oriented_simplex_def) lemma singular_face_oriented_simplex: assumes "1 \ p" "k \ p" shows "singular_face p k (oriented_simplex p l) = oriented_simplex (p -1) (\j. if j < k then l j else l (Suc j))" proof - have "(\j\p. l j i * simplical_face k x j) = (\j\p - Suc 0. (if j < k then l j else l (Suc j)) i * x j)" if "x \ standard_simplex (p - Suc 0)" for i x proof - show ?thesis unfolding simplical_face_def using sum.zero_middle [OF assms, where 'a=real, symmetric] by (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f i * _"] atLeast0AtMost cong: if_cong) qed then show ?thesis using simplical_face_in_standard_simplex assms by (auto simp: singular_face_def oriented_simplex_def restrict_def) qed lemma simplicial_simplex_singular_face: fixes f :: "(nat \ real) \ nat \ real" assumes ss: "simplicial_simplex p S f" and p: "1 \ p" "k \ p" shows "simplicial_simplex (p - Suc 0) S (singular_face p k f)" proof - let ?X = "subtopology (powertop_real UNIV) S" obtain m where l: "singular_simplex p ?X (oriented_simplex p m)" and feq: "f = oriented_simplex p m" using assms by (force simp: simplicial_simplex_def) moreover have "singular_face p k f = oriented_simplex (p - Suc 0) (\i. if i < k then m i else m (Suc i))" unfolding feq singular_face_def oriented_simplex_def apply (simp add: simplical_face_in_standard_simplex [OF p] restrict_compose_left subset_eq) using sum.zero_middle [OF p, where 'a=real, symmetric] unfolding simplical_face_def o_def apply (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f _ * _"] atLeast0AtMost cong: if_cong) done ultimately show ?thesis using p simplicial_simplex_def singular_simplex_singular_face by blast qed lemma simplicial_chain_boundary: "simplicial_chain p S c \ simplicial_chain (p -1) S (chain_boundary p c)" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have "simplicial_simplex p S f" by simp have "simplicial_chain (p - Suc 0) S (frag_of (singular_face p i f))" if "0 < p" "i \ p" for i using that one by (force simp: simplicial_simplex_def singular_simplex_singular_face singular_face_oriented_simplex) then have "simplicial_chain (p - Suc 0) S (chain_boundary p (frag_of f))" unfolding chain_boundary_def frag_extend_of by (auto intro!: simplicial_chain_cmul simplicial_chain_sum) then show ?case by (simp add: simplicial_chain_def [symmetric]) next case (diff a b) then show ?case by (metis chain_boundary_diff simplicial_chain_def simplicial_chain_diff) qed auto subsection\The cone construction on simplicial simplices.\ consts simplex_cone :: "[nat, nat \ real, [nat \ real, nat] \ real, nat \ real, nat] \ real" specification (simplex_cone) simplex_cone: "\p v l. simplex_cone p v (oriented_simplex p l) = oriented_simplex (Suc p) (\i. if i = 0 then v else l(i -1))" proof - have *: "\x. \y. \v. (\l. oriented_simplex (Suc x) (\i. if i = 0 then v else l (i -1))) = (y v \ (oriented_simplex x))" apply (subst choice_iff [symmetric]) by (simp add: oriented_simplex_eq choice_iff [symmetric] function_factors_left [symmetric]) then show ?thesis unfolding o_def by (metis(no_types)) qed lemma simplicial_simplex_simplex_cone: assumes f: "simplicial_simplex p S f" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_simplex (Suc p) T (simplex_cone p v f)" proof - obtain l where l: "\x. x \ standard_simplex p \ oriented_simplex p l x \ S" and feq: "f = oriented_simplex p l" using f by (auto simp: simplicial_simplex) have "oriented_simplex p l x \ S" if "x \ standard_simplex p" for x using f that by (auto simp: simplicial_simplex feq) then have S: "\x. \\i. 0 \ x i \ x i \ 1; \i. i>p \ x i = 0; sum x {..p} = 1\ \ (\i. \j\p. l j i * x j) \ S" by (simp add: oriented_simplex_def standard_simplex_def) have "oriented_simplex (Suc p) (\i. if i = 0 then v else l (i -1)) x \ T" if "x \ standard_simplex (Suc p)" for x proof (simp add: that oriented_simplex_def sum.atMost_Suc_shift del: sum.atMost_Suc) have x01: "\i. 0 \ x i \ x i \ 1" and x0: "\i. i > Suc p \ x i = 0" and x1: "sum x {..Suc p} = 1" using that by (auto simp: oriented_simplex_def standard_simplex_def) obtain a where "a \ S" using f by force show "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) \ T" proof (cases "x 0 = 1") case True then have "sum x {Suc 0..Suc p} = 0" using x1 by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) then have [simp]: "x (Suc j) = 0" if "j\p" for j unfolding sum.atLeast_Suc_atMost_Suc_shift using x01 that by (simp add: sum_nonneg_eq_0_iff) then show ?thesis using T [of 0 a] \a \ S\ by (auto simp: True) next case False then have "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) = (\i. (1 - (1 - x 0)) * v i + (1 - x 0) * (inverse (1 - x 0) * (\j\p. l j i * x (Suc j))))" by (force simp: field_simps) also have "\ \ T" proof (rule T) have "x 0 < 1" by (simp add: False less_le x01) have xle: "x (Suc i) \ (1 - x 0)" for i proof (cases "i \ p") case True have "sum x {0, Suc i} \ sum x {..Suc p}" by (rule sum_mono2) (auto simp: True x01) then show ?thesis using x1 x01 by (simp add: algebra_simps not_less) qed (simp add: x0 x01) have "(\i. (\j\p. l j i * (x (Suc j) * inverse (1 - x 0)))) \ S" proof (rule S) have "x 0 + (\j\p. x (Suc j)) = sum x {..Suc p}" by (metis sum.atMost_Suc_shift) with x1 have "(\j\p. x (Suc j)) = 1 - x 0" by simp with False show "(\j\p. x (Suc j) * inverse (1 - x 0)) = 1" by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero right_inverse sum_distrib_right) qed (use x01 x0 xle \x 0 < 1\ in \auto simp: field_split_simps\) then show "(\i. inverse (1 - x 0) * (\j\p. l j i * x (Suc j))) \ S" by (simp add: field_simps sum_divide_distrib) qed (use x01 in auto) finally show ?thesis . qed qed then show ?thesis by (auto simp: simplicial_simplex feq simplex_cone) qed definition simplicial_cone where "simplicial_cone p v \ frag_extend (frag_of \ simplex_cone p v)" lemma simplicial_chain_simplicial_cone: assumes c: "simplicial_chain p S c" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_chain (Suc p) T (simplicial_cone p v c)" using c unfolding simplicial_chain_def simplicial_cone_def proof (induction rule: frag_induction) case (one x) then show ?case by (simp add: T simplicial_simplex_simplex_cone) next case (diff a b) then show ?case by (metis frag_extend_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma chain_boundary_simplicial_cone_of': assumes "f = oriented_simplex p l" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" proof (simp, intro impI conjI) assume "p = 0" have eq: "(oriented_simplex 0 (\j. if j = 0 then v else l j)) = (\u\standard_simplex 0. v)" by (force simp: oriented_simplex_def standard_simplex_def) show "chain_boundary (Suc 0) (simplicial_cone 0 v (frag_of f)) = frag_of f - frag_of (\u\standard_simplex 0. v)" by (simp add: assms simplicial_cone_def chain_boundary_of \p = 0\ simplex_cone singular_face_oriented_simplex eq cong: if_cong) next assume "0 < p" have 0: "simplex_cone (p - Suc 0) v (singular_face p x (oriented_simplex p l)) = oriented_simplex p (\j. if j < Suc x then if j = 0 then v else l (j -1) else if Suc j = 0 then v else l (Suc j -1))" if "x \ p" for x using \0 < p\ that by (auto simp: Suc_leI singular_face_oriented_simplex simplex_cone oriented_simplex_eq) have 1: "frag_extend (frag_of \ simplex_cone (p - Suc 0) v) (\k = 0..p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k (oriented_simplex p l)))) = - (\k = Suc 0..Suc p. frag_cmul ((-1) ^ k) (frag_of (singular_face (Suc p) k (simplex_cone p v (oriented_simplex p l)))))" unfolding sum.atLeast_Suc_atMost_Suc_shift by (auto simp: 0 simplex_cone singular_face_oriented_simplex frag_extend_sum frag_extend_cmul simp flip: sum_negf) moreover have 2: "singular_face (Suc p) 0 (simplex_cone p v (oriented_simplex p l)) = oriented_simplex p l" by (simp add: simplex_cone singular_face_oriented_simplex) show "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - simplicial_cone (p - Suc 0) v (chain_boundary p (frag_of f))" using \p > 0\ apply (simp add: assms simplicial_cone_def chain_boundary_of atMost_atLeast0 del: sum.atMost_Suc) apply (subst sum.atLeast_Suc_atMost [of 0]) apply (simp_all add: 1 2 del: sum.atMost_Suc) done qed lemma chain_boundary_simplicial_cone_of: assumes "simplicial_simplex p S f" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" using chain_boundary_simplicial_cone_of' assms unfolding simplicial_simplex_def by blast lemma chain_boundary_simplicial_cone: "simplicial_chain p S c \ chain_boundary (Suc p) (simplicial_cone p v c) = c - (if p = 0 then frag_extend (\f. frag_of (\u\standard_simplex p. v)) c else simplicial_cone (p -1) v (chain_boundary p c))" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case by (auto simp: chain_boundary_simplicial_cone_of) qed (auto simp: chain_boundary_diff simplicial_cone_def frag_extend_diff) lemma simplex_map_oriented_simplex: assumes l: "simplicial_simplex p (standard_simplex q) (oriented_simplex p l)" and g: "simplicial_simplex r S g" and "q \ r" shows "simplex_map p g (oriented_simplex p l) = oriented_simplex p (g \ l)" proof - obtain m where geq: "g = oriented_simplex r m" using g by (auto simp: simplicial_simplex_def) have "g (\i. \j\p. l j i * x j) i = (\j\p. g (l j) i * x j)" if "x \ standard_simplex p" for x i proof - have ssr: "(\i. \j\p. l j i * x j) \ standard_simplex r" using l that standard_simplex_mono [OF \q \ r\] unfolding simplicial_simplex_oriented_simplex by auto have lss: "l j \ standard_simplex r" if "j\p" for j proof - have q: "(\x i. \j\p. l j i * x j) ` standard_simplex p \ standard_simplex q" using l by (simp add: simplicial_simplex_oriented_simplex) let ?x = "(\i. if i = j then 1 else 0)" have p: "l j \ (\x i. \j\p. l j i * x j) ` standard_simplex p" proof show "l j = (\i. \j\p. l j i * ?x j)" using \j\p\ by (force simp: if_distrib cong: if_cong) show "?x \ standard_simplex p" by (simp add: that) qed show ?thesis using standard_simplex_mono [OF \q \ r\] q p by blast qed have "g (\i. \j\p. l j i * x j) i = (\j\r. \n\p. m j i * (l n j * x n))" by (simp add: geq oriented_simplex_def sum_distrib_left ssr) also have "... = (\j\p. \n\r. m n i * (l j n * x j))" by (rule sum.swap) also have "... = (\j\p. g (l j) i * x j)" by (simp add: geq oriented_simplex_def sum_distrib_right mult.assoc lss) finally show ?thesis . qed then show ?thesis by (force simp: oriented_simplex_def simplex_map_def o_def) qed lemma chain_map_simplicial_cone: assumes g: "simplicial_simplex r S g" and c: "simplicial_chain p (standard_simplex q) c" and v: "v \ standard_simplex q" and "q \ r" shows "chain_map (Suc p) g (simplicial_cone p v c) = simplicial_cone p (g v) (chain_map p g c)" proof - have *: "simplex_map (Suc p) g (simplex_cone p v f) = simplex_cone p (g v) (simplex_map p g f)" if "f \ Poly_Mapping.keys c" for f proof - have "simplicial_simplex p (standard_simplex q) f" using c that by (auto simp: simplicial_chain_def) then obtain m where feq: "f = oriented_simplex p m" by (auto simp: simplicial_simplex) have 0: "simplicial_simplex p (standard_simplex q) (oriented_simplex p m)" using \simplicial_simplex p (standard_simplex q) f\ feq by blast then have 1: "simplicial_simplex (Suc p) (standard_simplex q) (oriented_simplex (Suc p) (\i. if i = 0 then v else m (i -1)))" using convex_standard_simplex v by (simp flip: simplex_cone add: simplicial_simplex_simplex_cone) show ?thesis using simplex_map_oriented_simplex [OF 1 g \q \ r\] simplex_map_oriented_simplex [of p q m r S g, OF 0 g \q \ r\] by (simp add: feq oriented_simplex_eq simplex_cone) qed show ?thesis by (auto simp: chain_map_def simplicial_cone_def frag_extend_compose * intro: frag_extend_eq) qed subsection\Barycentric subdivision of a linear ("simplicial") simplex's image\ definition simplicial_vertex where "simplicial_vertex i f = f(\j. if j = i then 1 else 0)" lemma simplicial_vertex_oriented_simplex: "simplicial_vertex i (oriented_simplex p l) = (if i \ p then l i else undefined)" by (simp add: simplicial_vertex_def oriented_simplex_def if_distrib cong: if_cong) primrec simplicial_subdivision where "simplicial_subdivision 0 = id" | "simplicial_subdivision (Suc p) = frag_extend (\f. simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (p + 2)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" lemma simplicial_subdivision_0 [simp]: "simplicial_subdivision p 0 = 0" by (induction p) auto lemma simplicial_subdivision_diff: "simplicial_subdivision p (c1-c2) = simplicial_subdivision p c1 - simplicial_subdivision p c2" by (induction p) (auto simp: frag_extend_diff) lemma simplicial_subdivision_of: "simplicial_subdivision p (frag_of f) = (if p = 0 then frag_of f else simplicial_cone (p -1) (\i. (\j\p. simplicial_vertex j f i) / (Suc p)) (simplicial_subdivision (p -1) (chain_boundary p (frag_of f))))" by (induction p) (auto simp: add.commute) lemma simplicial_chain_simplicial_subdivision: "simplicial_chain p S c \ simplicial_chain p S (simplicial_subdivision p c)" proof (induction p arbitrary: S c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by auto then have "simplicial_chain p (f ` standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis Suc.IH diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_simplex subsetI) moreover obtain l where l: "\x. x \ standard_simplex (Suc p) \ (\i. (\j\Suc p. l j i * x j)) \ S" and feq: "f = oriented_simplex (Suc p) l" using f by (fastforce simp: simplicial_simplex oriented_simplex_def simp del: sum.atMost_Suc) have "(\i. (1 - u) * ((\j\Suc p. simplicial_vertex j f i) / (real p + 2)) + u * y i) \ S" if "0 \ u" "u \ 1" and y: "y \ f ` standard_simplex (Suc p)" for y u proof - obtain x where x: "x \ standard_simplex (Suc p)" and yeq: "y = oriented_simplex (Suc p) l x" using y feq by blast have "(\i. \j\Suc p. l j i * ((if j \ Suc p then (1 - u) * inverse (p + 2) + u * x j else 0))) \ S" proof (rule l) have "inverse (2 + real p) \ 1" "(2 + real p) * ((1 - u) * inverse (2 + real p)) + u = 1" by (auto simp add: field_split_simps) then show "(\j. if j \ Suc p then (1 - u) * inverse (real (p + 2)) + u * x j else 0) \ standard_simplex (Suc p)" using x \0 \ u\ \u \ 1\ by (simp add: sum.distrib standard_simplex_def linepath_le_1 flip: sum_distrib_left del: sum.atMost_Suc) qed moreover have "(\i. \j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\i. (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j))" proof fix i have "(\j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\j\Suc p. (1 - u) * l j i / (real p + 2) + u * l j i * x j)" (is "?lhs = _") by (simp add: field_simps cong: sum.cong) also have "\ = (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j)" (is "_ = ?rhs") by (simp add: sum_distrib_left sum.distrib sum_divide_distrib mult.assoc del: sum.atMost_Suc) finally show "?lhs = ?rhs" . qed ultimately show ?thesis using feq x yeq by (simp add: simplicial_vertex_oriented_simplex) (simp add: oriented_simplex_def) qed ultimately show ?case by (simp add: simplicial_chain_simplicial_cone) next case (diff a b) then show ?case by (metis simplicial_chain_diff simplicial_subdivision_diff) qed auto qed auto lemma chain_boundary_simplicial_subdivision: "simplicial_chain p S c \ chain_boundary p (simplicial_subdivision p c) = simplicial_subdivision (p -1) (chain_boundary p c)" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by simp then have "simplicial_chain p S (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision) moreover have "simplicial_chain p S (chain_boundary (Suc p) (frag_of f))" using one simplicial_chain_boundary simplicial_chain_of by fastforce moreover have "simplicial_subdivision (p - Suc 0) (chain_boundary p (chain_boundary (Suc p) (frag_of f))) = 0" by (metis f chain_boundary_boundary_alt simplicial_simplex_def simplicial_subdivision_0 singular_chain_of) ultimately show ?case using chain_boundary_simplicial_cone Suc by (auto simp: chain_boundary_of frag_extend_diff simplicial_cone_def) next case (diff a b) then show ?case by (simp add: simplicial_subdivision_diff chain_boundary_diff frag_extend_diff) qed auto qed auto text \A MESS AND USED ONLY ONCE\ lemma simplicial_subdivision_shrinks: "\simplicial_chain p S c; \f x y. \f \ Poly_Mapping.keys c; x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ d; f \ Poly_Mapping.keys(simplicial_subdivision p c); x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ (p / (Suc p)) * d" proof (induction p arbitrary: d c f x y) case (Suc p) define Sigp where "Sigp \ \f:: (nat \ real) \ nat \ real. \i. (\j\Suc p. simplicial_vertex j f i) / real (p + 2)" let ?CB = "\f. chain_boundary (Suc p) (frag_of f)" have *: "Poly_Mapping.keys (simplicial_cone p (Sigp f) (simplicial_subdivision p (?CB f))) \ {f. \x\standard_simplex (Suc p). \y\standard_simplex (Suc p). \f x k - f y k\ \ real (Suc p) / (real p + 2) * d}" (is "?lhs \ ?rhs") if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "simplicial_simplex (Suc p) S f" using Suc.prems(1) simplicial_chain_def that by auto have 2: "\x y. \x \ standard_simplex (Suc p); y \ standard_simplex (Suc p)\ \ \f x k - f y k\ \ d" by (meson Suc.prems(2) f subsetD le_Suc_eq order_refl standard_simplex_mono) have sub: "Poly_Mapping.keys ((frag_of \ simplex_cone p (Sigp f)) g) \ ?rhs" if "g \ Poly_Mapping.keys (simplicial_subdivision p (?CB f))" for g proof - have 1: "simplicial_chain p S (?CB f)" using ssf simplicial_chain_boundary simplicial_chain_of by fastforce have "simplicial_chain (Suc p) (f ` standard_simplex(Suc p)) (frag_of f)" by (metis simplicial_chain_of simplicial_simplex ssf subset_refl) then have sc_sub: "Poly_Mapping.keys (?CB f) \ Collect (simplicial_simplex p (f ` standard_simplex (Suc p)))" by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def) have led: "\h x y. \h \ Poly_Mapping.keys (chain_boundary (Suc p) (frag_of f)); x \ standard_simplex p; y \ standard_simplex p\ \ \h x k - h y k\ \ d" using Suc.prems(2) f sc_sub by (simp add: simplicial_simplex subset_iff image_iff) metis have "\f' x y. \f' \ Poly_Mapping.keys (simplicial_subdivision p (?CB f)); x \ standard_simplex p; y \ standard_simplex p\ \ \f' x k - f' y k\ \ (p / (Suc p)) * d" by (blast intro: led Suc.IH [of "chain_boundary (Suc p) (frag_of f)", OF 1]) then have g: "\x y. \x \ standard_simplex p; y \ standard_simplex p\ \ \g x k - g y k\ \ (p / (Suc p)) * d" using that by blast have "d \ 0" using Suc.prems(2)[OF f] \x \ standard_simplex (Suc p)\ by force have 3: "simplex_cone p (Sigp f) g \ ?rhs" proof - have "simplicial_simplex p (f ` standard_simplex(Suc p)) g" by (metis (mono_tags, hide_lams) sc_sub mem_Collect_eq simplicial_chain_def simplicial_chain_simplicial_subdivision subsetD that) then obtain m where m: "g ` standard_simplex p \ f ` standard_simplex (Suc p)" and geq: "g = oriented_simplex p m" using ssf by (auto simp: simplicial_simplex) have m_in_gim: "m i \ g ` standard_simplex p" if "i \ p" for i proof show "m i = g (\j. if j = i then 1 else 0)" by (simp add: geq oriented_simplex_def that if_distrib cong: if_cong) show "(\j. if j = i then 1 else 0) \ standard_simplex p" by (simp add: oriented_simplex_def that) qed obtain l where l: "f ` standard_simplex (Suc p) \ S" and feq: "f = oriented_simplex (Suc p) l" using ssf by (auto simp: simplicial_simplex) show ?thesis proof (clarsimp simp add: geq simp del: sum.atMost_Suc) fix x y assume x: "x \ standard_simplex (Suc p)" and y: "y \ standard_simplex (Suc p)" then have x': "(\i. 0 \ x i \ x i \ 1) \ (\i>Suc p. x i = 0) \ (\i\Suc p. x i) = 1" and y': "(\i. 0 \ y i \ y i \ 1) \ (\i>Suc p. y i = 0) \ (\i\Suc p. y i) = 1" by (auto simp: standard_simplex_def) have "\(\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * x j) - (\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * y j)\ \ (1 + real p) * d / (2 + real p)" proof - have zero: "\m (s - Suc 0) k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" if "0 < s" and "s \ Suc p" for s proof - have "m (s - Suc 0) \ f ` standard_simplex (Suc p)" using m m_in_gim that(2) by auto then obtain z where eq: "m (s - Suc 0) = (\i. \j\Suc p. l j i * z j)" and z: "z \ standard_simplex (Suc p)" using feq unfolding oriented_simplex_def by auto show ?thesis unfolding eq proof (rule convex_sum_bound_le) fix i assume i: "i \ {..Suc p}" then have [simp]: "card ({..Suc p} - {i}) = Suc p" by (simp add: card_Suc_Diff1) have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) = (\j\Suc p. \l i k - l j k\ / (p + 2))" by (rule sum.cong) (simp_all add: flip: diff_divide_distrib) also have "\ = (\j \ {..Suc p} - {i}. \l i k - l j k\ / (p + 2))" by (rule sum.mono_neutral_right) auto also have "\ \ (1 + real p) * d / (p + 2)" proof (rule sum_bounded_above_divide) fix i' :: "nat" assume i': "i' \ {..Suc p} - {i}" have lf: "l r \ f ` standard_simplex(Suc p)" if "r \ Suc p" for r proof show "l r = f (\j. if j = r then 1 else 0)" using that by (simp add: feq oriented_simplex_def if_distrib cong: if_cong) show "(\j. if j = r then 1 else 0) \ standard_simplex (Suc p)" by (auto simp: oriented_simplex_def that) qed show "\l i k - l i' k\ / real (p + 2) \ (1 + real p) * d / real (p + 2) / real (card ({..Suc p} - {i}))" using i i' lf [of i] lf [of i'] 2 by (auto simp: image_iff divide_simps) qed auto finally have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) \ (1 + real p) * d / (p + 2)" . then have "\\j\Suc p. l i k / (p + 2) - l j k / (p + 2)\ \ (1 + real p) * d / (p + 2)" by (rule order_trans [OF sum_abs]) then show "\l i k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" by (simp add: sum_subtractf sum_divide_distrib del: sum.atMost_Suc) qed (use standard_simplex_def z in auto) qed have nonz: "\m (s - Suc 0) k - m (r - Suc 0) k\ \ (1 + real p) * d / (2 + real p)" (is "?lhs \ ?rhs") if "r < s" and "0 < r" and "r \ Suc p" and "s \ Suc p" for r s proof - have "?lhs \ (p / (Suc p)) * d" using m_in_gim [of "r - Suc 0"] m_in_gim [of "s - Suc 0"] that g by fastforce also have "\ \ ?rhs" by (simp add: field_simps \0 \ d\) finally show ?thesis . qed have jj: "j \ Suc p \ j' \ Suc p \ \(if j' = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j' -1)) k - (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k\ \ (1 + real p) * d / (2 + real p)" for j j' using \0 \ d\ by (rule_tac a=j and b = "j'" in linorder_less_wlog; force simp: zero nonz simp del: sum.atMost_Suc) show ?thesis apply (rule convex_sum_bound_le) using x' apply blast using x' apply blast apply (subst abs_minus_commute) apply (rule convex_sum_bound_le) using y' apply blast using y' apply blast using jj by blast qed then show "\simplex_cone p (Sigp f) (oriented_simplex p m) x k - simplex_cone p (Sigp f) (oriented_simplex p m) y k\ \ (1 + real p) * d / (real p + 2)" apply (simp add: feq Sigp_def simplicial_vertex_oriented_simplex simplex_cone del: sum.atMost_Suc) apply (simp add: oriented_simplex_def algebra_simps x y del: sum.atMost_Suc) done qed qed show ?thesis using Suc.IH [OF 1, where f=g] 2 3 by simp qed then show ?thesis unfolding simplicial_chain_def simplicial_cone_def by (simp add: order_trans [OF keys_frag_extend] sub UN_subset_iff) qed show ?case using Suc apply (simp del: sum.atMost_Suc) apply (drule subsetD [OF keys_frag_extend]) apply (simp del: sum.atMost_Suc) apply clarify (*OBTAIN?*) apply (rename_tac FFF) using * apply (simp add: add.commute Sigp_def subset_iff) done qed (auto simp: standard_simplex_0) subsection\Singular subdivision\ definition singular_subdivision where "singular_subdivision p \ frag_extend (\f. chain_map p f (simplicial_subdivision p (frag_of(restrict id (standard_simplex p)))))" lemma singular_subdivision_0 [simp]: "singular_subdivision p 0 = 0" by (simp add: singular_subdivision_def) lemma singular_subdivision_add: "singular_subdivision p (a + b) = singular_subdivision p a + singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_add) lemma singular_subdivision_diff: "singular_subdivision p (a - b) = singular_subdivision p a - singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_diff) lemma simplicial_simplex_id [simp]: "simplicial_simplex p S (restrict id (standard_simplex p)) \ standard_simplex p \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: simplicial_simplex_def singular_simplex_def continuous_map_in_subtopology set_mp) next assume R: ?rhs then have cm: "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) (subtopology (powertop_real UNIV) S) id" using continuous_map_from_subtopology_mono continuous_map_id by blast moreover have "\l. restrict id (standard_simplex p) = oriented_simplex p l" apply (rule_tac x="\i j. if i = j then 1 else 0" in exI) apply (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\u. u * _"] cong: if_cong) done ultimately show ?lhs by (simp add: simplicial_simplex_def singular_simplex_def) qed lemma singular_chain_singular_subdivision: "singular_chain p X c \ singular_chain p X (singular_subdivision p c)" unfolding singular_subdivision_def apply (rule singular_chain_extend) apply (rule singular_chain_chain_map [where X = "subtopology (powertop_real UNIV) (standard_simplex p)"]) apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain) by (simp add: singular_chain_def singular_simplex_def subset_iff) lemma naturality_singular_subdivision: "singular_chain p X c \ singular_subdivision p (chain_map p g c) = chain_map p g (singular_subdivision p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "singular_simplex p X f" by auto have "\simplicial_chain p (standard_simplex p) d\ \ chain_map p (simplex_map p g f) d = chain_map p g (chain_map p f d)" for d unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then have "simplex_map p (simplex_map p g f) x = simplex_map p g (simplex_map p f x)" by (force simp: simplex_map_def restrict_compose_left simplicial_simplex) then show ?case by auto qed (auto simp: chain_map_diff) then show ?case using simplicial_chain_simplicial_subdivision [of p "standard_simplex p" "frag_of (restrict id (standard_simplex p))"] by (simp add: singular_subdivision_def) next case (diff a b) then show ?case by (simp add: chain_map_diff singular_subdivision_diff) qed auto lemma simplicial_chain_chain_map: assumes f: "simplicial_simplex q X f" and c: "simplicial_chain p (standard_simplex q) c" shows "simplicial_chain p X (chain_map p f c)" using c unfolding simplicial_chain_def proof (induction c rule: frag_induction) case (one g) have "\n. simplex_map p (oriented_simplex q l) (oriented_simplex p m) = oriented_simplex p n" if m: "singular_simplex p (subtopology (powertop_real UNIV) (standard_simplex q)) (oriented_simplex p m)" for l m proof - have "(\i. \j\p. m j i * x j) \ standard_simplex q" if "x \ standard_simplex p" for x using that m unfolding oriented_simplex_def singular_simplex_def by (auto simp: continuous_map_in_subtopology image_subset_iff) then show ?thesis unfolding oriented_simplex_def simplex_map_def apply (rule_tac x="\j k. (\i\q. l i k * m j i)" in exI) apply (force simp: sum_distrib_left sum_distrib_right mult.assoc intro: sum.swap) done qed then show ?case using f one apply (auto simp: simplicial_simplex_def) apply (rule singular_simplex_simplex_map [where X = "subtopology (powertop_real UNIV) (standard_simplex q)"]) unfolding singular_simplex_def apply (fastforce simp add:)+ done next case (diff a b) then show ?case by (metis chain_map_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma singular_subdivision_simplicial_simplex: "simplicial_chain p S c \ singular_subdivision p c = simplicial_subdivision p c" proof (induction p arbitrary: S c) case 0 then show ?case unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case using singular_simplex_chain_map_id simplicial_imp_singular_simplex by (fastforce simp: singular_subdivision_def simplicial_subdivision_def) qed (auto simp: singular_subdivision_diff) next case (Suc p) show ?case using Suc.prems unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "simplicial_simplex (Suc p) S f" by (auto simp: simplicial_simplex) then have 1: "simplicial_chain p (standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_simplex_id) have 2: "(\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: simplicial_vertex_def standard_simplex_def del: sum.atMost_Suc) have ss_Sp: "(\i. (if i \ Suc p then 1 else 0) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: standard_simplex_def field_split_simps) obtain l where feq: "f = oriented_simplex (Suc p) l" using one unfolding simplicial_simplex by blast then have 3: "f (\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) = (\i. (\j\Suc p. simplicial_vertex j f i) / (real p + 2))" unfolding simplicial_vertex_def oriented_simplex_def by (simp add: ss_Sp if_distrib [of "\x. _ * x"] sum_divide_distrib del: sum.atMost_Suc cong: if_cong) have scp: "singular_chain (Suc p) (subtopology (powertop_real UNIV) (standard_simplex (Suc p))) (frag_of (restrict id (standard_simplex (Suc p))))" by (simp add: simplicial_imp_singular_chain) have scps: "simplicial_chain p (standard_simplex (Suc p)) (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_simplex_id) have scpf: "simplicial_chain p S (chain_map p f (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" using scps simplicial_chain_chain_map ssf by blast have 4: "chain_map p f (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))) = simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))" apply (simp add: chain_boundary_chain_map [OF scp] del: chain_map_of flip: singular_simplex_chain_map_id [OF simplicial_imp_singular_simplex [OF ssf]]) by (metis (no_types) scp singular_chain_boundary_alt Suc.IH [OF scps] Suc.IH [OF scpf] naturality_singular_subdivision) show ?case apply (simp add: singular_subdivision_def del: sum.atMost_Suc) apply (simp only: ssf 1 2 3 4 chain_map_simplicial_cone [of "Suc p" S _ p "Suc p"]) done qed (auto simp: frag_extend_diff singular_subdivision_diff) qed lemma naturality_simplicial_subdivision: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q S g\ \ simplicial_subdivision p (chain_map p g c) = chain_map p g (simplicial_subdivision p c)" apply (simp flip: singular_subdivision_simplicial_simplex) by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain singular_subdivision_simplicial_simplex) lemma chain_boundary_singular_subdivision: "singular_chain p X c \ chain_boundary p (singular_subdivision p c) = singular_subdivision (p - Suc 0) (chain_boundary p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "singular_simplex p X f" by (auto simp: singular_simplex_def) then have scp: "simplicial_chain p (standard_simplex p) (frag_of (restrict id (standard_simplex p)))" by simp have scp1: "simplicial_chain (p - Suc 0) (standard_simplex p) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using simplicial_chain_boundary by force have sgp1: "singular_chain (p - Suc 0) (subtopology (powertop_real UNIV) (standard_simplex p)) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using scp1 simplicial_imp_singular_chain by blast have scpp: "singular_chain p (subtopology (powertop_real UNIV) (standard_simplex p)) (frag_of (restrict id (standard_simplex p)))" using scp simplicial_imp_singular_chain by blast then show ?case unfolding singular_subdivision_def using chain_boundary_chain_map [of p "subtopology (powertop_real UNIV) (standard_simplex p)" _ f] apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain chain_boundary_simplicial_subdivision [OF scp] flip: singular_subdivision_simplicial_simplex [OF scp1] naturality_singular_subdivision [OF sgp1]) by (metis (full_types) singular_subdivision_def chain_boundary_chain_map [OF scpp] singular_simplex_chain_map_id [OF ssf]) qed (auto simp: singular_subdivision_def frag_extend_diff chain_boundary_diff) lemma singular_subdivision_zero: "singular_chain 0 X c \ singular_subdivision 0 c = c" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "restrict (f \ restrict id (standard_simplex 0)) (standard_simplex 0) = f" by (simp add: extensional_restrict restrict_compose_right singular_simplex_def) then show ?case by (auto simp: singular_subdivision_def simplex_map_def) qed (auto simp: singular_subdivision_def frag_extend_diff) primrec subd where "subd 0 = (\x. 0)" | "subd (Suc p) = frag_extend (\f. simplicial_cone (Suc p) (\i. (\j\Suc p. simplicial_vertex j f i) / real (Suc p + 1)) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))))" lemma subd_0 [simp]: "subd p 0 = 0" by (induction p) auto lemma subd_diff [simp]: "subd p (c1 - c2) = subd p c1 - subd p c2" by (induction p) (auto simp: frag_extend_diff) lemma subd_uminus [simp]: "subd p (-c) = - subd p c" by (metis diff_0 subd_0 subd_diff) lemma subd_power_uminus: "subd p (frag_cmul ((-1) ^ k) c) = frag_cmul ((-1) ^ k) (subd p c)" apply (induction k, simp_all) by (metis minus_frag_cmul subd_uminus) lemma subd_power_sum: "subd p (sum f I) = sum (subd p \ f) I" apply (induction I rule: infinite_finite_induct) by auto (metis add_diff_cancel_left' diff_add_cancel subd_diff) lemma subd: "simplicial_chain p (standard_simplex s) c \ (\r g. simplicial_simplex s (standard_simplex r) g \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)) \ simplicial_chain (Suc p) (standard_simplex s) (subd p c) \ (chain_boundary (Suc p) (subd p c)) + (subd (p - Suc 0) (chain_boundary p c)) = (simplicial_subdivision p c) - c" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction rule: frag_induction) case (one f) then obtain l where l: "(\x i. \j\Suc p. l j i * x j) ` standard_simplex (Suc p) \ standard_simplex s" and feq: "f = oriented_simplex (Suc p) l" by (metis (mono_tags) mem_Collect_eq simplicial_simplex simplicial_simplex_oriented_simplex) have scf: "simplicial_chain (Suc p) (standard_simplex s) (frag_of f)" using one by simp have lss: "l i \ standard_simplex s" if "i \ Suc p" for i proof - have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) \ standard_simplex s" using subsetD [OF l] basis_in_standard_simplex that by blast moreover have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) = l i" using that by (simp add: if_distrib [of "\x. _ * x"] del: sum.atMost_Suc cong: if_cong) ultimately show ?thesis by simp qed have *: "(\i. i \ n \ l i \ standard_simplex s) \ (\i. (\j\n. l j i) / (Suc n)) \ standard_simplex s" for n proof (induction n) case (Suc n) let ?x = "\i. (1 - inverse (n + 2)) * ((\j\n. l j i) / (Suc n)) + inverse (n + 2) * l (Suc n) i" have "?x \ standard_simplex s" proof (rule convex_standard_simplex) show "(\i. (\j\n. l j i) / real (Suc n)) \ standard_simplex s" using Suc by simp qed (auto simp: lss Suc inverse_le_1_iff) moreover have "?x = (\i. (\j\Suc n. l j i) / real (Suc (Suc n)))" by (force simp: divide_simps) ultimately show ?case by simp qed auto have **: "(\i. (\j\Suc p. simplicial_vertex j f i) / (2 + real p)) \ standard_simplex s" using * [of "Suc p"] lss by (simp add: simplicial_vertex_oriented_simplex feq) show ?case proof (intro conjI impI allI) fix r g assume g: "simplicial_simplex s (standard_simplex r) g" then obtain m where geq: "g = oriented_simplex s m" using simplicial_simplex by blast have 1: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_subdivision (Suc p) (frag_of f))" by (metis mem_Collect_eq one.hyps simplicial_chain_of simplicial_chain_simplicial_subdivision) have 2: "(\j\Suc p. \i\s. m i k * simplicial_vertex j f i) = (\j\Suc p. simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k)" for k proof (rule sum.cong [OF refl]) fix j assume j: "j \ {..Suc p}" have eq: "simplex_map (Suc p) (oriented_simplex s m) (oriented_simplex (Suc p) l) = oriented_simplex (Suc p) (oriented_simplex s m \ l)" proof (rule simplex_map_oriented_simplex) show "simplicial_simplex (Suc p) (standard_simplex s) (oriented_simplex (Suc p) l)" using one by (simp add: feq flip: oriented_simplex_def) show "simplicial_simplex s (standard_simplex r) (oriented_simplex s m)" using g by (simp add: geq) qed auto show "(\i\s. m i k * simplicial_vertex j f i) = simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k" using one j apply (simp add: feq eq simplicial_vertex_oriented_simplex simplicial_simplex_oriented_simplex image_subset_iff) apply (drule_tac x="(\i. if i = j then 1 else 0)" in bspec) apply (auto simp: oriented_simplex_def lss) done qed have 4: "chain_map (Suc p) g (subd p (chain_boundary (Suc p) (frag_of f))) = subd p (chain_boundary (Suc p) (frag_of (simplex_map (Suc p) g f)))" by (metis (no_types) One_nat_def scf Suc.IH chain_boundary_chain_map chain_map_of diff_Suc_Suc diff_zero g simplicial_chain_boundary simplicial_imp_singular_chain) show "chain_map (Suc (Suc p)) g (subd (Suc p) (frag_of f)) = subd (Suc p) (chain_map (Suc p) g (frag_of f))" using g apply (simp only: subd.simps frag_extend_of) apply (subst chain_map_simplicial_cone [of s "standard_simplex r" _ "Suc p" s], assumption) apply (intro simplicial_chain_diff) using "1" apply auto[1] using one.hyps apply auto[1] apply (metis Suc.IH diff_Suc_1 mem_Collect_eq one.hyps simplicial_chain_boundary simplicial_chain_of) using "**" apply auto[1] apply (rule order_refl) apply (simp only: chain_map_of frag_extend_of) apply (rule arg_cong2 [where f = "simplicial_cone (Suc p)"]) apply (simp add: geq sum_distrib_left oriented_simplex_def ** del: sum.atMost_Suc flip: sum_divide_distrib) using 2 apply (simp only: oriented_simplex_def sum.swap [where A = "{..s}"]) using naturality_simplicial_subdivision scf apply (fastforce simp add: 4 chain_map_diff) done next have sc: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (Suc (Suc p))) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" by (metis diff_Suc_1 nat.simps(3) simplicial_subdivision_of scf simplicial_chain_simplicial_subdivision) have ff: "simplicial_chain (Suc p) (standard_simplex s) (subd p (chain_boundary (Suc p) (frag_of f)))" by (metis (no_types) Suc.IH diff_Suc_1 scf simplicial_chain_boundary) show "simplicial_chain (Suc (Suc p)) (standard_simplex s) (subd (Suc p) (frag_of f))" using one apply (simp only: subd.simps frag_extend_of) apply (rule_tac S="standard_simplex s" in simplicial_chain_simplicial_cone) apply (intro simplicial_chain_diff ff) using sc apply (simp add: algebra_simps) using "**" convex_standard_simplex apply force+ done have "simplicial_chain p (standard_simplex s) (chain_boundary (Suc p) (frag_of f))" using scf simplicial_chain_boundary by fastforce then have "chain_boundary (Suc p) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))) = 0" apply (simp only: chain_boundary_diff) using Suc.IH chain_boundary_boundary [of "Suc p" "subtopology (powertop_real UNIV) (standard_simplex s)" "frag_of f"] by (metis One_nat_def add_diff_cancel_left' subd_0 chain_boundary_simplicial_subdivision plus_1_eq_Suc scf simplicial_imp_singular_chain) then show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f)) + subd (Suc p - Suc 0) (chain_boundary (Suc p) (frag_of f)) = simplicial_subdivision (Suc p) (frag_of f) - frag_of f" apply (simp only: subd.simps frag_extend_of) apply (subst chain_boundary_simplicial_cone [of "Suc p" "standard_simplex s"]) apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision) apply (simp add: simplicial_cone_def del: sum.atMost_Suc simplicial_subdivision.simps) done qed next case (diff a b) then show ?case apply safe apply (metis chain_map_diff subd_diff) apply (metis simplicial_chain_diff subd_diff) apply (auto simp: simplicial_subdivision_diff chain_boundary_diff simp del: simplicial_subdivision.simps subd.simps) by (metis (no_types, lifting) add_diff_add add_uminus_conv_diff diff_0 diff_diff_add) qed auto qed simp lemma chain_homotopic_simplicial_subdivision1: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q (standard_simplex r) g\ \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision2: "simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (subd p c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision3: "simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (subd p c) = (simplicial_subdivision p c) - c - subd (p - Suc 0) (chain_boundary p c)" by (simp add: subd algebra_simps) lemma chain_homotopic_simplicial_subdivision: "\h. (\p. h p 0 = 0) \ (\p c1 c2. h p (c1-c2) = h p c1 - h p c2) \ (\p q r g c. simplicial_chain p (standard_simplex q) c \ simplicial_simplex q (standard_simplex r) g \ chain_map (Suc p) g (h p c) = h p (chain_map p g c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (h p c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (simplicial_subdivision p c) - c)" by (rule_tac x=subd in exI) (fastforce simp: subd) lemma chain_homotopic_singular_subdivision: obtains h where "\p. h p 0 = 0" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" proof - define k where "k \ \p. frag_extend (\f:: (nat \ real) \ 'a. chain_map (Suc p) f (subd p (frag_of(restrict id (standard_simplex p)))))" show ?thesis proof fix p X and c :: "'a chain" assume c: "singular_chain p X c" have "singular_chain (Suc p) X (k p c) \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" using c [unfolded singular_chain_def] proof (induction rule: frag_induction) case (one f) let ?X = "subtopology (powertop_real UNIV) (standard_simplex p)" show ?case proof (simp add: k_def, intro conjI) show "singular_chain (Suc p) X (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p)))))" proof (rule singular_chain_chain_map) show "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) show "continuous_map ?X X f" using one.hyps singular_simplex_def by auto qed next have scp: "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) have feqf: "frag_of (simplex_map p f (restrict id (standard_simplex p))) = frag_of f" using one.hyps singular_simplex_chain_map_id by auto have *: "chain_map p f (subd (p - Suc 0) (\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k id)))) = (\x\p. frag_cmul ((-1) ^ x) (chain_map p (singular_face p x f) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" (is "?lhs = ?rhs") if "p > 0" proof - have eqc: "subd (p - Suc 0) (frag_of (singular_face p i id)) = chain_map p (singular_face p i id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" if "i \ p" for i proof - have 1: "simplicial_chain (p - Suc 0) (standard_simplex (p - Suc 0)) (frag_of (restrict id (standard_simplex (p - Suc 0))))" by simp have 2: "simplicial_simplex (p - Suc 0) (standard_simplex p) (singular_face p i id)" by (metis One_nat_def Suc_leI \0 < p\ simplicial_simplex_id simplicial_simplex_singular_face singular_face_restrict subsetI that) have 3: "simplex_map (p - Suc 0) (singular_face p i id) (restrict id (standard_simplex (p - Suc 0))) = singular_face p i id" by (force simp: simplex_map_def singular_face_def) show ?thesis using chain_homotopic_simplicial_subdivision1 [OF 1 2] that \p > 0\ by (simp add: 3) qed have xx: "simplicial_chain p (standard_simplex(p - Suc 0)) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" by (metis Suc_pred chain_homotopic_simplicial_subdivision2 order_refl simplicial_chain_of simplicial_simplex_id that) have yy: "\k. k \ p \ chain_map p f (chain_map p (singular_face p k id) h) = chain_map p (singular_face p k f) h" if "simplicial_chain p (standard_simplex(p - Suc 0)) h" for h using that unfolding simplicial_chain_def proof (induction h rule: frag_induction) case (one x) then show ?case using one apply (simp add: chain_map_of singular_simplex_def simplicial_simplex_def, auto) apply (rule_tac f=frag_of in arg_cong, rule) apply (simp add: simplex_map_def) by (simp add: continuous_map_in_subtopology image_subset_iff singular_face_def) qed (auto simp: chain_map_diff) have "?lhs = chain_map p f (\k\p. frag_cmul ((-1) ^ k) (chain_map p (singular_face p k id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" by (simp add: subd_power_sum subd_power_uminus eqc) also have "\ = ?rhs" by (simp add: chain_map_sum xx yy) finally show ?thesis . qed have "chain_map p f (simplicial_subdivision p (frag_of (restrict id (standard_simplex p))) - subd (p - Suc 0) (chain_boundary p (frag_of (restrict id (standard_simplex p))))) = singular_subdivision p (frag_of f) - frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f))" apply (simp add: singular_subdivision_def chain_map_diff) apply (clarsimp simp add: chain_boundary_def) apply (simp add: frag_extend_sum frag_extend_cmul *) done then show "chain_boundary (Suc p) (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p))))) + frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f)) = singular_subdivision p (frag_of f) - frag_of f" by (simp add: chain_boundary_chain_map [OF scp] chain_homotopic_simplicial_subdivision3 [where q=p] chain_map_diff feqf) qed next case (diff a b) then show ?case apply (simp only: k_def singular_chain_diff chain_boundary_diff frag_extend_diff singular_subdivision_diff) by (metis (no_types, lifting) add_diff_add diff_add_cancel) qed (auto simp: k_def) then show "singular_chain (Suc p) X (k p c)" "chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by auto qed (auto simp: k_def frag_extend_diff) qed lemma homologous_rel_singular_subdivision: assumes "singular_relcycle p X T c" shows "homologous_rel p X T (singular_subdivision p c) c" proof (cases "p = 0") case True with assms show ?thesis by (auto simp: singular_relcycle_def singular_subdivision_zero) next case False with assms show ?thesis unfolding homologous_rel_def singular_relboundary singular_relcycle by (metis One_nat_def Suc_diff_1 chain_homotopic_singular_subdivision gr_zeroI) qed subsection\Excision argument that we keep doing singular subdivision\ lemma singular_subdivision_power_0 [simp]: "(singular_subdivision p ^^ n) 0 = 0" by (induction n) auto lemma singular_subdivision_power_diff: "(singular_subdivision p ^^ n) (a - b) = (singular_subdivision p ^^ n) a - (singular_subdivision p ^^ n) b" by (induction n) (auto simp: singular_subdivision_diff) lemma iterated_singular_subdivision: "singular_chain p X c \ (singular_subdivision p ^^ n) c = frag_extend (\f. chain_map p f ((simplicial_subdivision p ^^ n) (frag_of(restrict id (standard_simplex p))))) c" proof (induction n arbitrary: c) case 0 then show ?case unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "restrict f (standard_simplex p) = f" by (simp add: extensional_restrict singular_simplex_def) then show ?case by (auto simp: simplex_map_def cong: restrict_cong) qed (auto simp: frag_extend_diff) next case (Suc n) show ?case using Suc.prems unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "singular_simplex p X f" by simp have scp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (restrict id (standard_simplex p))))" proof (induction n) case 0 then show ?case by (metis funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have scnp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (\x\standard_simplex p. x)))" proof (induction n) case 0 then show ?case by (metis eq_id_iff funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have sff: "singular_chain p X (frag_of f)" by (simp add: \singular_simplex p X f\ singular_chain_of) then show ?case using Suc.IH [OF sff] naturality_singular_subdivision [OF simplicial_imp_singular_chain [OF scp], of f] singular_subdivision_simplicial_simplex [OF scnp] by (simp add: singular_chain_of id_def del: restrict_apply) qed (auto simp: singular_subdivision_power_diff singular_subdivision_diff frag_extend_diff) qed lemma chain_homotopic_iterated_singular_subdivision: obtains h where "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" proof (induction n arbitrary: thesis) case 0 show ?case by (rule 0 [of "(\p x. 0)"]) auto next case (Suc n) then obtain k where k: "\p. k p 0 = (0 :: 'a chain)" "\p c1 c2. k p (c1-c2) = k p c1 - k p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (k p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" by metis obtain h where h: "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by (blast intro: chain_homotopic_singular_subdivision) let ?h = "(\p c. singular_subdivision (Suc p) (k p c) + h p c)" show ?case proof (rule Suc.prems) fix p X and c :: "'a chain" assume "singular_chain p X c" then show "singular_chain (Suc p) X (?h p c)" by (simp add: h k singular_chain_add singular_chain_singular_subdivision) next fix p :: "nat" and X :: "'a topology" and c :: "'a chain" assume sc: "singular_chain p X c" have f5: "chain_boundary (Suc p) (singular_subdivision (Suc p) (k p c)) = singular_subdivision p (chain_boundary (Suc p) (k p c))" using chain_boundary_singular_subdivision k(3) sc by fastforce have [simp]: "singular_subdivision (Suc (p - Suc 0)) (k (p - Suc 0) (chain_boundary p c)) = singular_subdivision p (k (p - Suc 0) (chain_boundary p c))" proof (cases p) case 0 then show ?thesis by (simp add: k chain_boundary_def) qed auto show "chain_boundary (Suc p) (?h p c) + ?h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ Suc n) c - c" using chain_boundary_singular_subdivision [of "Suc p" X] apply (simp add: chain_boundary_add f5 h k algebra_simps) - by (smt add.assoc add.commute diff_add_cancel h(4) k(4) sc singular_subdivision_add) + apply (smt (verit, ccfv_threshold) add.commute add_diff_eq diff_add_cancel diff_diff_eq2 h(4) k(4) sc singular_subdivision_add) + done qed (auto simp: k h singular_subdivision_diff) qed lemma llemma: assumes p: "standard_simplex p \ \\" and \: "\U. U \ \ \ openin (powertop_real UNIV) U" obtains d where "0 < d" "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ \ \ K \ U" proof - have "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" if x: "x \ standard_simplex p" for x proof- obtain U where U: "U \ \" "x \ U" using x p by blast then obtain V where finV: "finite {i. V i \ UNIV}" and openV: "\i. open (V i)" and xV: "x \ Pi\<^sub>E UNIV V" and UV: "Pi\<^sub>E UNIV V \ U" using \ unfolding openin_product_topology_alt by force have xVi: "x i \ V i" for i using PiE_mem [OF xV] by simp have "\i. \e>0. \x'. \x' - x i\ < e \ x' \ V i" by (rule openV [unfolded open_real, rule_format, OF xVi]) then obtain d where d: "\i. d i > 0" and dV: "\i x'. \x' - x i\ < d i \ x' \ V i" by metis define e where "e \ Inf (insert 1 (d ` {i. V i \ UNIV})) / 3" have ed3: "e \ d i / 3" if "V i \ UNIV" for i using that finV by (auto simp: e_def intro: cInf_le_finite) show "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" proof (intro exI conjI allI impI) show "e > 0" using d finV by (simp add: e_def finite_less_Inf_iff) fix y assume y: "(\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0)" have "y \ Pi\<^sub>E UNIV V" proof show "y i \ V i" for i proof (cases "p < i") case True then show ?thesis by (metis (mono_tags, lifting) y x mem_Collect_eq standard_simplex_def xVi) next case False show ?thesis proof (cases "V i = UNIV") case False show ?thesis proof (rule dV) have "\y i - x i\ \ 2 * e" using y \\ p < i\ by simp also have "\ < d i" using ed3 [OF False] \e > 0\ by simp finally show "\y i - x i\ < d i" . qed qed auto qed qed auto with UV show "y \ U" by blast qed (use U in auto) qed then obtain e U where eU: "\x. x \ standard_simplex p \ 0 < e x \ U x \ \ \ x \ U x" and UI: "\x y. \x \ standard_simplex p; \i. i \ p \ \y i - x i\ \ 2 * e x; \i. i > p \ y i = 0\ \ y \ U x" by metis define F where "F \ \x. Pi\<^sub>E UNIV (\i. if i \ p then {x i - e x<..S \ F ` standard_simplex p. openin (powertop_real UNIV) S" by (simp add: F_def openin_PiE_gen) moreover have pF: "standard_simplex p \ \(F ` standard_simplex p)" by (force simp: F_def PiE_iff eU) ultimately have "\\. finite \ \ \ \ F ` standard_simplex p \ standard_simplex p \ \\" using compactin_standard_simplex [of p] unfolding compactin_def by force then obtain S where "finite S" and ssp: "S \ standard_simplex p" "standard_simplex p \ \(F ` S)" unfolding ex_finite_subset_image by (auto simp: ex_finite_subset_image) then have "S \ {}" by (auto simp: nonempty_standard_simplex) show ?thesis proof show "Inf (e ` S) > 0" using \finite S\ \S \ {}\ ssp eU by (auto simp: finite_less_Inf_iff) fix k :: "(nat \ real) set" assume k: "k \ standard_simplex p" and kle: "\x y i. \i \ p; x \ k; y \ k\ \ \x i - y i\ \ Inf (e ` S)" show "\U. U \ \ \ k \ U" proof (cases "k = {}") case True then show ?thesis using \S \ {}\ eU equals0I ssp(1) subset_eq p by auto next case False with k ssp obtain x a where "x \ k" "x \ standard_simplex p" and a: "a \ S" and Fa: "x \ F a" by blast then have le_ea: "\i. i \ p \ abs (x i - a i) < e a" by (simp add: F_def PiE_iff if_distrib abs_diff_less_iff cong: if_cong) show ?thesis proof (intro exI conjI) show "U a \ \" using a eU ssp(1) by auto show "k \ U a" proof clarify fix y assume "y \ k" with k have y: "y \ standard_simplex p" by blast show "y \ U a" proof (rule UI) show "a \ standard_simplex p" using a ssp(1) by auto fix i :: "nat" assume "i \ p" then have "\x i - y i\ \ e a" by (meson kle [OF \i \ p\] a \finite S\ \x \ k\ \y \ k\ cInf_le_finite finite_imageI imageI order_trans) then show "\y i - a i\ \ 2 * e a" using le_ea [OF \i \ p\] by linarith next fix i assume "p < i" then show "y i = 0" using standard_simplex_def y by auto qed qed qed qed qed qed proposition sufficient_iterated_singular_subdivision_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_chain p X c" obtains n where "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" proof (cases "c = 0") case False then show ?thesis proof (cases "topspace X = {}") case True show ?thesis using p that by (force simp: singular_chain_empty True) next case False show ?thesis proof (cases "\ = {}") case True then show ?thesis using False X by blast next case False have "\e. 0 < e \ (\K. K \ standard_simplex p \ (\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ e) \ (\V. V \ \ \ f ` K \ V))" if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "singular_simplex p X f" using f p by (auto simp: singular_chain_def) then have fp: "\x. x \ standard_simplex p \ f x \ topspace X" by (auto simp: singular_simplex_def image_subset_iff dest: continuous_map_image_subset_topspace) have "\T. openin (powertop_real UNIV) T \ standard_simplex p \ f -` V = T \ standard_simplex p" if V: "V \ \" for V proof - have "singular_simplex p X f" using p f unfolding singular_chain_def by blast then have "openin (subtopology (powertop_real UNIV) (standard_simplex p)) {x \ standard_simplex p. f x \ V}" using \ [OF \V \ \\] by (simp add: singular_simplex_def continuous_map_def) moreover have "standard_simplex p \ f -` V = {x \ standard_simplex p. f x \ V}" by blast ultimately show ?thesis by (simp add: openin_subtopology) qed then obtain g where gope: "\V. V \ \ \ openin (powertop_real UNIV) (g V)" and geq: "\V. V \ \ \ standard_simplex p \ f -` V = g V \ standard_simplex p" by metis obtain d where "0 < d" and d: "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ g ` \ \ K \ U" proof (rule llemma [of p "g ` \"]) show "standard_simplex p \ \(g ` \)" using geq X fp by (fastforce simp add:) show "openin (powertop_real UNIV) U" if "U \ g ` \" for U :: "(nat \ real) set" using gope that by blast qed auto show ?thesis proof (rule exI, intro allI conjI impI) fix K :: "(nat \ real) set" assume K: "K \ standard_simplex p" and Kd: "\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ d" then have "\U. U \ g ` \ \ K \ U" using d [OF K] by auto then show "\V. V \ \ \ f ` K \ V" using K geq by fastforce qed (rule \d > 0\) qed then obtain \ where epos: "\f \ Poly_Mapping.keys c. 0 < \ f" and e: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ \ f\ \ \V. V \ \ \ f ` K \ V" by metis obtain d where "0 < d" and d: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ d\ \ \V. V \ \ \ f ` K \ V" proof show "Inf (\ ` Poly_Mapping.keys c) > 0" by (simp add: finite_less_Inf_iff \c \ 0\ epos) fix f K assume fK: "f \ Poly_Mapping.keys c" "K \ standard_simplex p" and le: "\x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ Inf (\ ` Poly_Mapping.keys c)" then have lef: "Inf (\ ` Poly_Mapping.keys c) \ \ f" by (auto intro: cInf_le_finite) show "\V. V \ \ \ f ` K \ V" using le lef by (blast intro: dual_order.trans e [OF fK]) qed let ?d = "\m. (simplicial_subdivision p ^^ m) (frag_of (restrict id (standard_simplex p)))" obtain n where n: "(p / (Suc p)) ^ n < d" using real_arch_pow_inv \0 < d\ by fastforce show ?thesis proof fix m h assume "n \ m" and "h \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)" then obtain f where "f \ Poly_Mapping.keys c" "h \ Poly_Mapping.keys (chain_map p f (?d m))" using subsetD [OF keys_frag_extend] iterated_singular_subdivision [OF p, of m] by force then obtain g where g: "g \ Poly_Mapping.keys (?d m)" and heq: "h = restrict (f \ g) (standard_simplex p)" using keys_frag_extend by (force simp: chain_map_def simplex_map_def) have xx: "simplicial_chain p (standard_simplex p) (?d n) \ (\f \ Poly_Mapping.keys(?d n). \x \ standard_simplex p. \y \ standard_simplex p. \f x i - f y i\ \ (p / (Suc p)) ^ n)" for n i proof (induction n) case 0 have "simplicial_simplex p (standard_simplex p) (\a\standard_simplex p. a)" by (metis eq_id_iff order_refl simplicial_simplex_id) moreover have "(\x\standard_simplex p. \y\standard_simplex p. \x i - y i\ \ 1)" unfolding standard_simplex_def by (auto simp: abs_if dest!: spec [where x=i]) ultimately show ?case unfolding power_0 funpow_0 by simp next case (Suc n) show ?case unfolding power_Suc funpow.simps o_def proof (intro conjI ballI) show "simplicial_chain p (standard_simplex p) (simplicial_subdivision p (?d n))" by (simp add: Suc simplicial_chain_simplicial_subdivision) show "\f x i - f y i\ \ real p / real (Suc p) * (real p / real (Suc p)) ^ n" if "f \ Poly_Mapping.keys (simplicial_subdivision p (?d n))" and "x \ standard_simplex p" and "y \ standard_simplex p" for f x y using Suc that by (blast intro: simplicial_subdivision_shrinks) qed qed have "g ` standard_simplex p \ standard_simplex p" using g xx [of m] unfolding simplicial_chain_def simplicial_simplex by auto moreover have "\g x i - g y i\ \ d" if "i \ p" "x \ standard_simplex p" "y \ standard_simplex p" for x y i proof - have "\g x i - g y i\ \ (p / (Suc p)) ^ m" using g xx [of m] that by blast also have "\ \ (p / (Suc p)) ^ n" by (auto intro: power_decreasing [OF \n \ m\]) finally show ?thesis using n by simp qed then have "\x i - y i\ \ d" if "x \ g ` (standard_simplex p)" "y \ g ` (standard_simplex p)" "i \ p" for i x y using that by blast ultimately show "\V\\. h ` standard_simplex p \ V" using \f \ Poly_Mapping.keys c\ d [of f "g ` standard_simplex p"] by (simp add: Bex_def heq image_image) qed qed qed qed force lemma small_homologous_rel_relcycle_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_relcycle p X S c" obtains c' where "singular_relcycle p X S c'" "homologous_rel p X S c c'" "\f. f \ Poly_Mapping.keys c' \ \V \ \. f ` (standard_simplex p) \ V" proof - have "singular_chain p X c" "(chain_boundary p c, 0) \ (mod_subset (p - Suc 0) (subtopology X S))" using p unfolding singular_relcycle_def by auto then obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" by (blast intro: sufficient_iterated_singular_subdivision_exists [OF \ X]) let ?c' = "(singular_subdivision p ^^ n) c" show ?thesis proof show "homologous_rel p X S c ?c'" apply (induction n, simp_all) by (metis p homologous_rel_singular_subdivision homologous_rel_singular_relcycle homologous_rel_trans homologous_rel_sym) then show "singular_relcycle p X S ?c'" by (metis homologous_rel_singular_relcycle p) next fix f :: "(nat \ real) \ 'a" assume "f \ Poly_Mapping.keys ?c'" then show "\V\\. f ` standard_simplex p \ V" by (rule n [OF order_refl]) qed qed lemma excised_chain_exists: fixes S :: "'a set" assumes "X closure_of U \ X interior_of T" "T \ S" "singular_chain p (subtopology X S) c" obtains n d e where "singular_chain p (subtopology X (S - U)) d" "singular_chain p (subtopology X T) e" "(singular_subdivision p ^^ n) c = d + e" proof - have *: "\n d e. singular_chain p (subtopology X (S - U)) d \ singular_chain p (subtopology X T) e \ (singular_subdivision p ^^ n) c = d + e" if c: "singular_chain p (subtopology X S) c" and X: "X closure_of U \ X interior_of T" "U \ topspace X" and S: "T \ S" "S \ topspace X" for p X c S and T U :: "'a set" proof - obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ {S \ X interior_of T, S - X closure_of U}. f ` standard_simplex p \ V" apply (rule sufficient_iterated_singular_subdivision_exists [of "{S \ X interior_of T, S - X closure_of U}"]) using X S c by (auto simp: topspace_subtopology openin_subtopology_Int2 openin_subtopology_diff_closed) let ?c' = "\n. (singular_subdivision p ^^ n) c" have "singular_chain p (subtopology X S) (?c' m)" for m by (induction m) (auto simp: singular_chain_singular_subdivision c) then have scp: "singular_chain p (subtopology X S) (?c' n)" . have SS: "Poly_Mapping.keys (?c' n) \ singular_simplex_set p (subtopology X (S - U)) \ singular_simplex_set p (subtopology X T)" proof (clarsimp) fix f assume f: "f \ Poly_Mapping.keys ((singular_subdivision p ^^ n) c)" and non: "\ singular_simplex p (subtopology X T) f" show "singular_simplex p (subtopology X (S - U)) f" using n [OF order_refl f] scp f non closure_of_subset [OF \U \ topspace X\] interior_of_subset [of X T] by (fastforce simp: image_subset_iff singular_simplex_subtopology singular_chain_def) qed show ?thesis unfolding singular_chain_def using frag_split [OF SS] by metis qed have "(subtopology X (topspace X \ S)) = (subtopology X S)" by (metis subtopology_subtopology subtopology_topspace) with assms have c: "singular_chain p (subtopology X (topspace X \ S)) c" by simp have Xsub: "X closure_of (topspace X \ U) \ X interior_of (topspace X \ T)" using assms closure_of_restrict interior_of_restrict by fastforce obtain n d e where d: "singular_chain p (subtopology X (topspace X \ S - topspace X \ U)) d" and e: "singular_chain p (subtopology X (topspace X \ T)) e" and de: "(singular_subdivision p ^^ n) c = d + e" using *[OF c Xsub, simplified] assms by force show thesis proof show "singular_chain p (subtopology X (S - U)) d" by (metis d Diff_Int_distrib inf.cobounded2 singular_chain_mono) show "singular_chain p (subtopology X T) e" by (metis e inf.cobounded2 singular_chain_mono) show "(singular_subdivision p ^^ n) c = d + e" by (rule de) qed qed lemma excised_relcycle_exists: fixes S :: "'a set" assumes X: "X closure_of U \ X interior_of T" and "T \ S" and c: "singular_relcycle p (subtopology X S) T c" obtains c' where "singular_relcycle p (subtopology X (S - U)) (T - U) c'" "homologous_rel p (subtopology X S) T c c'" proof - have [simp]: "(S - U) \ (T - U) = T - U" "S \ T = T" using \T \ S\ by auto have scc: "singular_chain p (subtopology X S) c" and scp1: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p c)" using c by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology) obtain n d e where d: "singular_chain p (subtopology X (S - U)) d" and e: "singular_chain p (subtopology X T) e" and de: "(singular_subdivision p ^^ n) c = d + e" using excised_chain_exists [OF X \T \ S\ scc] . have scSUd: "singular_chain (p - Suc 0) (subtopology X (S - U)) (chain_boundary p d)" by (simp add: singular_chain_boundary d) have sccn: "singular_chain p (subtopology X S) ((singular_subdivision p ^^ n) c)" for n by (induction n) (auto simp: singular_chain_singular_subdivision scc) have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c))" proof (induction n) case (Suc n) then show ?case by (simp add: singular_chain_singular_subdivision chain_boundary_singular_subdivision [OF sccn]) qed (auto simp: scp1) then have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c - e))" by (simp add: chain_boundary_diff singular_chain_diff singular_chain_boundary e) with de have scTd: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p d)" by simp show thesis proof have "singular_chain (p - Suc 0) X (chain_boundary p d)" using scTd singular_chain_subtopology by blast with scSUd scTd have "singular_chain (p - Suc 0) (subtopology X (T - U)) (chain_boundary p d)" by (fastforce simp add: singular_chain_subtopology) then show "singular_relcycle p (subtopology X (S - U)) (T - U) d" by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology d) have "homologous_rel p (subtopology X S) T (c-0) ((singular_subdivision p ^^ n) c - e)" proof (rule homologous_rel_diff) show "homologous_rel p (subtopology X S) T c ((singular_subdivision p ^^ n) c)" proof (induction n) case (Suc n) then show ?case apply simp apply (rule homologous_rel_trans) using c homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision homologous_rel_sym by blast qed auto show "homologous_rel p (subtopology X S) T 0 e" unfolding homologous_rel_def using e by (intro singular_relboundary_diff singular_chain_imp_relboundary; simp add: subtopology_subtopology) qed with de show "homologous_rel p (subtopology X S) T c d" by simp qed qed subsection\Homotopy invariance\ theorem homotopic_imp_homologous_rel_chain_maps: assumes hom: "homotopic_with (\h. h ` T \ V) S U f g" and c: "singular_relcycle p S T c" shows "homologous_rel p U V (chain_map p f c) (chain_map p g c)" proof - note sum.atMost_Suc [simp del] have contf: "continuous_map S U f" and contg: "continuous_map S U g" using homotopic_with_imp_continuous_maps [OF hom] by metis+ obtain h where conth: "continuous_map (prod_topology (top_of_set {0..1::real}) S) U h" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = g x" and hV: "\t x. \0 \ t; t \ 1; x \ T\ \ h(t,x) \ V" using hom by (fastforce simp: homotopic_with_def) define vv where "vv \ \j i. if i = Suc j then 1 else (0::real)" define ww where "ww \ \j i. if i=0 \ i = Suc j then 1 else (0::real)" define simp where "simp \ \q i. oriented_simplex (Suc q) (\j. if j \ i then vv j else ww(j -1))" define pr where "pr \ \q c. \i\q. frag_cmul ((-1) ^ i) (frag_of (simplex_map (Suc q) (\z. h(z 0, c(z \ Suc))) (simp q i)))" have ss_ss: "simplicial_simplex (Suc q) ({x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}) (simp q i)" if "i \ q" for q i proof - have "(\j\Suc q. (if j \ i then vv j 0 else ww (j -1) 0) * x j) \ {0..1}" if "x \ standard_simplex (Suc q)" for x proof - have "(\j\Suc q. if j \ i then 0 else x j) \ sum x {..Suc q}" using that unfolding standard_simplex_def by (force intro!: sum_mono) with \i \ q\ that show ?thesis by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg cong: if_cong) qed moreover have "(\k. \j\Suc q. (if j \ i then vv j k else ww (j -1) k) * x j) \ Suc \ standard_simplex q" if "x \ standard_simplex (Suc q)" for x proof - have card: "({..q} \ {k. Suc k = j}) = {j-1}" if "0 < j" "j \ Suc q" for j using that by auto have eq: "(\j\Suc q. \k\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) = (\j\Suc q. x j)" by (rule sum.cong [OF refl]) (use \i \ q\ in \simp add: sum.If_cases card\) have "(\j\Suc q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) \ sum x {..Suc q}" for k using that unfolding standard_simplex_def by (force intro!: sum_mono) then show ?thesis using \i \ q\ that by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg sum.swap [where A = "atMost q"] eq cong: if_cong) qed ultimately show ?thesis by (simp add: that simplicial_simplex_oriented_simplex simp_def image_subset_iff if_distribR) qed obtain prism where prism: "\q. prism q 0 = 0" "\q c. singular_chain q S c \ singular_chain (Suc q) U (prism q c)" "\q c. singular_chain q (subtopology S T) c \ singular_chain (Suc q) (subtopology U V) (prism q c)" "\q c. singular_chain q S c \ chain_boundary (Suc q) (prism q c) = chain_map q g c - chain_map q f c - prism (q -1) (chain_boundary q c)" proof show "(frag_extend \ pr) q 0 = 0" for q by (simp add: pr_def) next show "singular_chain (Suc q) U ((frag_extend \ pr) q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) U (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X S (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) qed (use one in \simp add: singular_simplex_def\) show "continuous_map X U (\z. h (z 0, m (z \ Suc)))" apply (rule continuous_map_compose [unfolded o_def, OF _ conth]) using 0 1 by (simp add: continuous_map_pairwise o_def) qed qed next case (diff a b) then show ?case apply (simp add: frag_extend_diff keys_diff) using singular_chain_def singular_chain_diff by blast qed auto next show "singular_chain (Suc q) (subtopology U V) ((frag_extend \ pr) q c)" if "singular_chain q (subtopology S T) c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) (subtopology U V) (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X (subtopology S T) (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) show "continuous_map (subtopology (powertop_real UNIV) (standard_simplex q)) (subtopology S T) m" using one continuous_map_into_fulltopology by (auto simp: singular_simplex_def) qed have "continuous_map X (subtopology U V) (h \ (\z. (z 0, m (z \ Suc))))" proof (rule continuous_map_compose) show "continuous_map X (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (\z. (z 0, m (z \ Suc)))" using 0 1 by (simp add: continuous_map_pairwise o_def) have "continuous_map (subtopology (prod_topology euclideanreal S) ({0..1} \ T)) U h" by (metis conth continuous_map_from_subtopology subtopology_Times subtopology_topspace) with hV show "continuous_map (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (subtopology U V) h" by (force simp: topspace_subtopology continuous_map_in_subtopology subtopology_restrict subtopology_Times) qed then show "continuous_map X (subtopology U V) (\z. h (z 0, m (z \ Suc)))" by (simp add: o_def) qed qed next case (diff a b) then show ?case by (metis comp_apply frag_extend_diff singular_chain_diff) qed auto next show "chain_boundary (Suc q) ((frag_extend \ pr) q c) = chain_map q g c - chain_map q f c - (frag_extend \ pr) (q -1) (chain_boundary q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) have eq2: "Sigma S T = (\i. (i,i)) ` {i \ S. i \ T i} \ (Sigma S (\i. T i - {i}))" for S :: "nat set" and T by force have 1: "(\(i,j)\(\i. (i, i)) ` {i. i \ q \ i \ Suc q}. frag_cmul (((-1) ^ i) * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(\i. (i, i)) ` {i. i \ q}. frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = frag_of (simplex_map q g m) - frag_of (simplex_map q f m)" proof - have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) (standard_simplex q) = restrict (g \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "(\j\Suc q. if j = 0 then 0 else x (j - Suc 0)) = (\j\q. x j)" by (simp add: sum.atMost_Suc_shift) with x have "simp q 0 (simplical_face 0 x) 0 = 1" apply (simp add: oriented_simplex_def simp_def simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib ww_def standard_simplex_def cong: if_cong) done moreover have "(\n. if n \ q then x n else 0) = x" using standard_simplex_def x by auto then have "(\n. simp q 0 (simplical_face 0 x) (Suc n)) = x" unfolding oriented_simplex_def simp_def ww_def using x apply (simp add: simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib) apply (simp add: if_distribR if_distrib cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) x = (g \ m) x" by (simp add: o_def h1) qed then have a: "frag_of (singular_face (Suc q) 0 (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q 0))) = frag_of (simplex_map q g m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) (standard_simplex q) = restrict (f \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" then have "simp q q (simplical_face (Suc q) x) 0 = 0" unfolding oriented_simplex_def simp_def by (simp add: simplical_face_in_standard_simplex sum.atMost_Suc) (simp add: simplical_face_def vv_def) moreover have "(\n. simp q q (simplical_face (Suc q) x) (Suc n)) = x" unfolding oriented_simplex_def simp_def vv_def using x apply (simp add: simplical_face_in_standard_simplex) apply (force simp: standard_simplex_def simplical_face_def if_distribR if_distrib [of "\x. x * _"] sum.atMost_Suc cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) x = (f \ m) x" by (simp add: o_def h0) qed then have b: "frag_of (singular_face (Suc q) (Suc q) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q q))) = frag_of (simplex_map q f m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have sfeq: "simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q (Suc i) \ simplical_face (Suc i)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face (Suc i))" if "i < q" for i unfolding simplex_map_def proof (rule restrict_ext) fix x assume "x \ standard_simplex q" then have "(simp q (Suc i) \ simplical_face (Suc i)) x = (simp q i \ simplical_face (Suc i)) x" unfolding oriented_simplex_def simp_def simplical_face_def by (force intro: sum.cong) then show "((\z. h (z 0, m (z \ Suc))) \ (simp q (Suc i) \ simplical_face (Suc i))) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face (Suc i))) x" by simp qed have eqq: "{i. i \ q \ i \ Suc q} = {..q}" by force have qeq: "{..q} = insert 0 ((\i. Suc i) ` {i. i < q})" "{i. i \ q} = insert q {i. i < q}" using le_imp_less_Suc less_Suc_eq_0_disj by auto show ?thesis using a b apply (simp add: sum.reindex inj_on_def eqq) apply (simp add: qeq sum.insert_if sum.reindex sum_negf singular_face_simplex_map sfeq) done qed have 2: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ i * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = - frag_extend (pr (q - Suc 0)) (chain_boundary q (frag_of m))" proof (cases "q=0") case True then show ?thesis by (simp add: chain_boundary_def flip: sum.Sigma) next case False have eq: "{..q - Suc 0} \ {..q} = Sigma {..q - Suc 0} (\i. {0..min q i}) \ Sigma {..q} (\i. {i<..q})" by force have I: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ (i + j)) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q - Suc 0}. {0..min q i}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" proof - have seq: "simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) (i - Suc 0)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face j)" if ij: "i \ q" "j \ i" "j \ i" for i j unfolding simplex_map_def proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "i > 0" using that by force then have iq: "i - Suc 0 \ q - Suc 0" using \i \ q\ False by simp have q0_eq: "{..Suc q} = insert 0 (Suc ` {..q})" by (auto simp: image_def gr0_conv_Suc) have \: "simp (q - Suc 0) (i - Suc 0) x 0 = simp q i (simplical_face j x) 0" using False x ij unfolding oriented_simplex_def simp_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex) apply (force simp: simplical_face_def q0_eq sum.reindex intro!: sum.cong) done have \: "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) = simp q i (simplical_face j x) \ Suc" proof fix k show "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) k = (simp q i (simplical_face j x) \ Suc) k" using False x ij unfolding oriented_simplex_def simp_def o_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex if_distribR) apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] cong: if_cong) apply (intro impI conjI) apply (force simp: sum.atMost_Suc intro: sum.cong) apply (force simp: q0_eq sum.reindex intro!: sum.cong) done qed have "simp (q - Suc 0) (i - Suc 0) x \ Suc \ standard_simplex (q - Suc 0)" using ss_ss [OF iq] \i \ q\ False \i > 0\ apply (simp add: simplicial_simplex image_subset_iff) using \x \ standard_simplex q\ by blast then show "((\z. h (z 0, singular_face q j m (z \ Suc))) \ simp (q - Suc 0) (i - Suc 0)) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face j)) x" by (simp add: singular_face_def \ \) qed have [simp]: "(-1::int) ^ (i + j - Suc 0) = - ((-1) ^ (i + j))" if "i \ j" for i j::nat proof - have "i + j > 0" using that by blast then show ?thesis by (metis (no_types, hide_lams) One_nat_def Suc_diff_1 add.inverse_inverse mult.left_neutral mult_minus_left power_Suc) qed show ?thesis apply (rule sum.eq_general_inverses [where h = "\(a,b). (a-1,b)" and k = "\(a,b). (Suc a,b)"]) using False apply (auto simp: singular_face_simplex_map seq add.commute) done qed have *: "singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)) = simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i)" if ij: "i < j" "j \ q" for i j proof - have iq: "i \ q - Suc 0" using that by auto have sf_eqh: "singular_face (Suc q) (Suc j) (\x. if x \ standard_simplex (Suc q) then ((\z. h (z 0, m (z \ Suc))) \ simp q i) x else undefined) x = h (simp (q - Suc 0) i x 0, singular_face q j m (\xa. simp (q - Suc 0) i x (Suc xa)))" if x: "x \ standard_simplex q" for x proof - let ?f = "\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0" have fm: "simplical_face (Suc j) x \ standard_simplex (Suc q)" using ss_ss [OF iq] that ij by (simp add: simplical_face_in_standard_simplex) have ss: "?f \ standard_simplex (q - Suc 0)" unfolding standard_simplex_def proof (intro CollectI conjI impI allI) fix k show "0 \ ?f k" using that by (simp add: sum_nonneg standard_simplex_def) show "?f k \ 1" using x sum_le_included [of "{..q}" "{..q}" x "id"] by (simp add: standard_simplex_def) assume k: "q - Suc 0 < k" show "?f k = 0" by (rule sum.neutral) (use that x iq k standard_simplex_def in auto) next have "(\k\q - Suc 0. ?f k) = (\(k,j) \ ({..q - Suc 0} \ {..q}) \ {(k,j). if j \ i then k = j else Suc k = j}. x j)" apply (simp add: sum.Sigma) by (rule sum.mono_neutral_cong) (auto simp: split: if_split_asm) also have "\ = sum x {..q}" apply (rule sum.eq_general_inverses [where h = "\(k,j). if j\i \ k=j \ j>i \ Suc k = j then j else Suc q" and k = "\j. if j \ i then (j,j) else (j - Suc 0, j)"]) using ij by auto also have "\ = 1" using x by (simp add: standard_simplex_def) finally show "(\k\q - Suc 0. ?f k) = 1" by (simp add: standard_simplex_def) qed let ?g = "\k. if k \ i then 0 else if k < Suc j then x k else if k = Suc j then 0 else x (k - Suc 0)" have eq: "{..Suc q} = {..j} \ {Suc j} \ Suc`{j<..q}" "{..q} = {..j} \ {j<..q}" using ij image_iff less_Suc_eq_0_disj less_Suc_eq_le by (force simp: image_iff)+ then have "(\k\Suc q. ?g k) = (\k\{..j} \ {Suc j} \ Suc`{j<..q}. ?g k)" by simp also have "\ = (\k\{..j} \ Suc`{j<..q}. ?g k)" by (rule sum.mono_neutral_right) auto also have "\ = (\k\{..j}. ?g k) + (\k\Suc`{j<..q}. ?g k)" by (rule sum.union_disjoint) auto also have "\ = (\k\{..j}. ?g k) + (\k\{j<..q}. ?g (Suc k))" by (auto simp: sum.reindex) also have "\ = (\k\{..j}. if k \ i then 0 else x k) + (\k\{j<..q}. if k \ i then 0 else x k)" by (intro sum.cong arg_cong2 [of concl: "(+)"]) (use ij in auto) also have "\ = (\k\q. if k \ i then 0 else x k)" unfolding eq by (subst sum.union_disjoint) auto finally have "(\k\Suc q. ?g k) = (\k\q. if k \ i then 0 else x k)" . then have QQ: "(\l\Suc q. if l \ i then 0 else simplical_face (Suc j) x l) = (\j\q. if j \ i then 0 else x j)" by (simp add: simplical_face_def cong: if_cong) have WW: "(\k. \l\Suc q. if l \ i then if k = l then simplical_face (Suc j) x l else 0 else if Suc k = l then simplical_face (Suc j) x l else 0) = simplical_face j (\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0)" proof - have *: "(\l\q. if l \ i then 0 else if Suc k = l then x (l - Suc 0) else 0) = (\l\q. if l \ i then if k - Suc 0 = l then x l else 0 else if k = l then x l else 0)" (is "?lhs = ?rhs") if "k \ q" "k > j" for k proof (cases "k \ q") case True have "?lhs = sum (\l. x (l - Suc 0)) {Suc k}" "?rhs = sum x {k}" by (rule sum.mono_neutral_cong_right; use True ij that in auto)+ then show ?thesis by simp next case False have "?lhs = 0" "?rhs = 0" by (rule sum.neutral; use False ij in auto)+ then show ?thesis by simp qed show ?thesis apply (rule ext) unfolding simplical_face_def using ij apply (auto simp: sum.atMost_Suc cong: if_cong) apply (force simp flip: ivl_disj_un(2) intro: sum.neutral) apply (auto simp: *) done qed show ?thesis using False that iq unfolding oriented_simplex_def simp_def vv_def ww_def apply (simp add: if_distribR cong: if_cong) apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] o_def cong: if_cong) apply (simp add: singular_face_def fm ss QQ WW) done qed show ?thesis unfolding simplex_map_def restrict_def apply (simp add: simplicial_simplex image_subset_iff o_def sf_eqh fun_eq_iff) apply (simp add: singular_face_def) done qed have sgeq: "(SIGMA i:{..q}. {i..q} - {i}) = (SIGMA i:{..q}. {i<..q})" by force have II: "(\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ (i + j))) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q}. {i<..q}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" by (force simp: * sgeq add.commute intro: sum.cong) show ?thesis using False apply (simp add: chain_boundary_def frag_extend_sum frag_extend_cmul frag_cmul_sum pr_def flip: sum_negf power_add) apply (subst sum.swap [where A = "{..q}"]) apply (simp add: sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal I II) done qed have *: "\a+b = w; c+d = -z\ \ (a + c) + (b+d) = w-z" for a b w c d z :: "'c \\<^sub>0 int" by (auto simp: algebra_simps) have eq: "{..q} \ {..Suc q} = Sigma {..q} (\i. {0..min (Suc q) i}) \ Sigma {..q} (\i. {Suc i..Suc q})" by force show ?case apply (subst pr_def) apply (simp add: chain_boundary_sum chain_boundary_cmul) apply (subst chain_boundary_def) apply (simp add: frag_cmul_sum sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal sum.atLeast_Suc_atMost_Suc_shift del: sum.cl_ivl_Suc flip: comm_monoid_add_class.sum.Sigma) apply (simp add: comm_monoid_add_class.sum.Sigma eq2 [of _ "\i. {_ i.._ i}"]) apply (simp add: sum.union_disjoint disjoint_iff_not_equal * [OF 1 2]) done next case (diff a b) then show ?case by (simp add: chain_boundary_diff frag_extend_diff chain_map_diff) qed auto qed have *: "singular_chain p (subtopology U V) (prism (p - Suc 0) (chain_boundary p c))" if "singular_chain p S c" "singular_chain (p - Suc 0) (subtopology S T) (chain_boundary p c)" proof (cases "p") case 0 then show ?thesis by (simp add: chain_boundary_def prism) next case (Suc p') with prism that show ?thesis by auto qed then show ?thesis using c unfolding singular_relcycle_def homologous_rel_def singular_relboundary_def mod_subset_def apply (rule_tac x="- prism p c" in exI) by (simp add: chain_boundary_minus prism(2) prism(4) singular_chain_minus) qed end diff --git a/src/HOL/Library/Float.thy b/src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy +++ b/src/HOL/Library/Float.thy @@ -1,2446 +1,2446 @@ (* Title: HOL/Library/Float.thy Author: Johannes Hölzl, Fabian Immler Copyright 2012 TU München *) section \Floating-Point Numbers\ theory Float imports Log_Nat Lattice_Algebras begin definition "float = {m * 2 powr e | (m :: int) (e :: int). True}" typedef float = float morphisms real_of_float float_of unfolding float_def by auto setup_lifting type_definition_float declare real_of_float [code_unfold] lemmas float_of_inject[simp] declare [[coercion "real_of_float :: float \ real"]] lemma real_of_float_eq: "f1 = f2 \ real_of_float f1 = real_of_float f2" for f1 f2 :: float unfolding real_of_float_inject .. declare real_of_float_inverse[simp] float_of_inverse [simp] declare real_of_float [simp] subsection \Real operations preserving the representation as floating point number\ lemma floatI: "m * 2 powr e = x \ x \ float" for m e :: int by (auto simp: float_def) lemma zero_float[simp]: "0 \ float" by (auto simp: float_def) lemma one_float[simp]: "1 \ float" by (intro floatI[of 1 0]) simp lemma numeral_float[simp]: "numeral i \ float" by (intro floatI[of "numeral i" 0]) simp lemma neg_numeral_float[simp]: "- numeral i \ float" by (intro floatI[of "- numeral i" 0]) simp lemma real_of_int_float[simp]: "real_of_int x \ float" for x :: int by (intro floatI[of x 0]) simp lemma real_of_nat_float[simp]: "real x \ float" for x :: nat by (intro floatI[of x 0]) simp lemma two_powr_int_float[simp]: "2 powr (real_of_int i) \ float" for i :: int by (intro floatI[of 1 i]) simp lemma two_powr_nat_float[simp]: "2 powr (real i) \ float" for i :: nat by (intro floatI[of 1 i]) simp lemma two_powr_minus_int_float[simp]: "2 powr - (real_of_int i) \ float" for i :: int by (intro floatI[of 1 "-i"]) simp lemma two_powr_minus_nat_float[simp]: "2 powr - (real i) \ float" for i :: nat by (intro floatI[of 1 "-i"]) simp lemma two_powr_numeral_float[simp]: "2 powr numeral i \ float" by (intro floatI[of 1 "numeral i"]) simp lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \ float" by (intro floatI[of 1 "- numeral i"]) simp lemma two_pow_float[simp]: "2 ^ n \ float" by (intro floatI[of 1 n]) (simp add: powr_realpow) lemma plus_float[simp]: "r \ float \ p \ float \ r + p \ float" unfolding float_def proof (safe, simp) have *: "\(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" if "e1 \ e2" for e1 m1 e2 m2 :: int proof - from that have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1" by (simp add: powr_diff field_simps flip: powr_realpow) then show ?thesis by blast qed fix e1 m1 e2 m2 :: int consider "e2 \ e1" | "e1 \ e2" by (rule linorder_le_cases) then show "\(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e" proof cases case 1 from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps) next case 2 then show ?thesis by (rule *) qed qed lemma uminus_float[simp]: "x \ float \ -x \ float" apply (auto simp: float_def) apply hypsubst_thin apply (rename_tac m e) apply (rule_tac x="-m" in exI) apply (rule_tac x="e" in exI) apply (simp add: field_simps) done lemma times_float[simp]: "x \ float \ y \ float \ x * y \ float" apply (auto simp: float_def) apply hypsubst_thin apply (rename_tac mx my ex ey) apply (rule_tac x="mx * my" in exI) apply (rule_tac x="ex + ey" in exI) apply (simp add: powr_add) done lemma minus_float[simp]: "x \ float \ y \ float \ x - y \ float" using plus_float [of x "- y"] by simp lemma abs_float[simp]: "x \ float \ \x\ \ float" by (cases x rule: linorder_cases[of 0]) auto lemma sgn_of_float[simp]: "x \ float \ sgn x \ float" by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float) lemma div_power_2_float[simp]: "x \ float \ x / 2^d \ float" apply (auto simp add: float_def) apply hypsubst_thin apply (rename_tac m e) apply (rule_tac x="m" in exI) apply (rule_tac x="e - d" in exI) apply (simp flip: powr_realpow powr_add add: field_simps) done lemma div_power_2_int_float[simp]: "x \ float \ x / (2::int)^d \ float" apply (auto simp add: float_def) apply hypsubst_thin apply (rename_tac m e) apply (rule_tac x="m" in exI) apply (rule_tac x="e - d" in exI) apply (simp flip: powr_realpow powr_add add: field_simps) done lemma div_numeral_Bit0_float[simp]: assumes "x / numeral n \ float" shows "x / (numeral (Num.Bit0 n)) \ float" proof - have "(x / numeral n) / 2^1 \ float" by (intro assms div_power_2_float) also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))" by (induct n) auto finally show ?thesis . qed lemma div_neg_numeral_Bit0_float[simp]: assumes "x / numeral n \ float" shows "x / (- numeral (Num.Bit0 n)) \ float" proof - have "- (x / numeral (Num.Bit0 n)) \ float" using assms by simp also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)" by simp finally show ?thesis . qed lemma power_float[simp]: assumes "a \ float" shows "a ^ b \ float" proof - from assms obtain m e :: int where "a = m * 2 powr e" by (auto simp: float_def) then show ?thesis by (auto intro!: floatI[where m="m^b" and e = "e*b"] simp: power_mult_distrib powr_realpow[symmetric] powr_powr) qed lift_definition Float :: "int \ int \ float" is "\(m::int) (e::int). m * 2 powr e" by simp declare Float.rep_eq[simp] code_datatype Float lemma compute_real_of_float[code]: "real_of_float (Float m e) = (if e \ 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))" by (simp add: powr_int) subsection \Arithmetic operations on floating point numbers\ instantiation float :: "{ring_1,linorder,linordered_ring,linordered_idom,numeral,equal}" begin lift_definition zero_float :: float is 0 by simp declare zero_float.rep_eq[simp] lift_definition one_float :: float is 1 by simp declare one_float.rep_eq[simp] lift_definition plus_float :: "float \ float \ float" is "(+)" by simp declare plus_float.rep_eq[simp] lift_definition times_float :: "float \ float \ float" is "(*)" by simp declare times_float.rep_eq[simp] lift_definition minus_float :: "float \ float \ float" is "(-)" by simp declare minus_float.rep_eq[simp] lift_definition uminus_float :: "float \ float" is "uminus" by simp declare uminus_float.rep_eq[simp] lift_definition abs_float :: "float \ float" is abs by simp declare abs_float.rep_eq[simp] lift_definition sgn_float :: "float \ float" is sgn by simp declare sgn_float.rep_eq[simp] lift_definition equal_float :: "float \ float \ bool" is "(=) :: real \ real \ bool" . lift_definition less_eq_float :: "float \ float \ bool" is "(\)" . declare less_eq_float.rep_eq[simp] lift_definition less_float :: "float \ float \ bool" is "(<)" . declare less_float.rep_eq[simp] instance by standard (transfer; fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+ end lemma real_of_float [simp]: "real_of_float (of_nat n) = of_nat n" by (induct n) simp_all lemma real_of_float_of_int_eq [simp]: "real_of_float (of_int z) = of_int z" by (cases z rule: int_diff_cases) (simp_all add: of_rat_diff) lemma Float_0_eq_0[simp]: "Float 0 e = 0" by transfer simp lemma real_of_float_power[simp]: "real_of_float (f^n) = real_of_float f^n" for f :: float by (induct n) simp_all lemma real_of_float_min: "real_of_float (min x y) = min (real_of_float x) (real_of_float y)" and real_of_float_max: "real_of_float (max x y) = max (real_of_float x) (real_of_float y)" for x y :: float by (simp_all add: min_def max_def) instance float :: unbounded_dense_linorder proof fix a b :: float show "\c. a < c" apply (intro exI[of _ "a + 1"]) apply transfer apply simp done show "\c. c < a" apply (intro exI[of _ "a - 1"]) apply transfer apply simp done show "\c. a < c \ c < b" if "a < b" apply (rule exI[of _ "(a + b) * Float 1 (- 1)"]) using that apply transfer apply (simp add: powr_minus) done qed instantiation float :: lattice_ab_group_add begin definition inf_float :: "float \ float \ float" where "inf_float a b = min a b" definition sup_float :: "float \ float \ float" where "sup_float a b = max a b" instance by standard (transfer; simp add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+ end lemma float_numeral[simp]: "real_of_float (numeral x :: float) = numeral x" apply (induct x) apply simp apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq float_of_inverse plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float) done lemma transfer_numeral [transfer_rule]: "rel_fun (=) pcr_float (numeral :: _ \ real) (numeral :: _ \ float)" by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def) lemma float_neg_numeral[simp]: "real_of_float (- numeral x :: float) = - numeral x" by simp lemma transfer_neg_numeral [transfer_rule]: "rel_fun (=) pcr_float (- numeral :: _ \ real) (- numeral :: _ \ float)" by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def) lemma float_of_numeral: "numeral k = float_of (numeral k)" and float_of_neg_numeral: "- numeral k = float_of (- numeral k)" unfolding real_of_float_eq by simp_all subsection \Quickcheck\ instantiation float :: exhaustive begin definition exhaustive_float where "exhaustive_float f d = Quickcheck_Exhaustive.exhaustive (\x. Quickcheck_Exhaustive.exhaustive (\y. f (Float x y)) d) d" instance .. end context includes term_syntax begin definition [code_unfold]: "valtermify_float x y = Code_Evaluation.valtermify Float {\} x {\} y" end instantiation float :: full_exhaustive begin definition "full_exhaustive_float f d = Quickcheck_Exhaustive.full_exhaustive (\x. Quickcheck_Exhaustive.full_exhaustive (\y. f (valtermify_float x y)) d) d" instance .. end instantiation float :: random begin definition "Quickcheck_Random.random i = scomp (Quickcheck_Random.random (2 ^ nat_of_natural i)) (\man. scomp (Quickcheck_Random.random i) (\exp. Pair (valtermify_float man exp)))" instance .. end subsection \Represent floats as unique mantissa and exponent\ lemma int_induct_abs[case_names less]: fixes j :: int assumes H: "\n. (\i. \i\ < \n\ \ P i) \ P n" shows "P j" proof (induct "nat \j\" arbitrary: j rule: less_induct) case less show ?case by (rule H[OF less]) simp qed lemma int_cancel_factors: fixes n :: int assumes "1 < r" shows "n = 0 \ (\k i. n = k * r ^ i \ \ r dvd k)" proof (induct n rule: int_induct_abs) case (less n) have "\k i. n = k * r ^ Suc i \ \ r dvd k" if "n \ 0" "n = m * r" for m proof - from that have "\m \ < \n\" using \1 < r\ by (simp add: abs_mult) from less[OF this] that show ?thesis by auto qed then show ?case by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0) qed lemma mult_powr_eq_mult_powr_iff_asym: fixes m1 m2 e1 e2 :: int assumes m1: "\ 2 dvd m1" and "e1 \ e2" shows "m1 * 2 powr e1 = m2 * 2 powr e2 \ m1 = m2 \ e1 = e2" (is "?lhs \ ?rhs") proof show ?rhs if eq: ?lhs proof - have "m1 \ 0" using m1 unfolding dvd_def by auto from \e1 \ e2\ eq have "m1 = m2 * 2 powr nat (e2 - e1)" by (simp add: powr_diff field_simps) also have "\ = m2 * 2^nat (e2 - e1)" by (simp add: powr_realpow) finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)" by linarith with m1 have "m1 = m2" by (cases "nat (e2 - e1)") (auto simp add: dvd_def) then show ?thesis using eq \m1 \ 0\ by (simp add: powr_inj) qed show ?lhs if ?rhs using that by simp qed lemma mult_powr_eq_mult_powr_iff: "\ 2 dvd m1 \ \ 2 dvd m2 \ m1 * 2 powr e1 = m2 * 2 powr e2 \ m1 = m2 \ e1 = e2" for m1 m2 e1 e2 :: int using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2] using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1] by (cases e1 e2 rule: linorder_le_cases) auto lemma floatE_normed: assumes x: "x \ float" obtains (zero) "x = 0" | (powr) m e :: int where "x = m * 2 powr e" "\ 2 dvd m" "x \ 0" proof - have "\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m" if "x \ 0" proof - from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def) with \x \ 0\ int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\ 2 dvd k" by auto with \\ 2 dvd k\ x show ?thesis apply (rule_tac exI[of _ "k"]) apply (rule_tac exI[of _ "e + int i"]) apply (simp add: powr_add powr_realpow) done qed with that show thesis by blast qed lemma float_normed_cases: fixes f :: float obtains (zero) "f = 0" | (powr) m e :: int where "real_of_float f = m * 2 powr e" "\ 2 dvd m" "f \ 0" proof (atomize_elim, induct f) case (float_of y) then show ?case by (cases rule: floatE_normed) (auto simp: zero_float_def) qed definition mantissa :: "float \ int" where "mantissa f = fst (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) \ (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p))" definition exponent :: "float \ int" where "exponent f = snd (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) \ (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p))" lemma exponent_0[simp]: "exponent 0 = 0" (is ?E) and mantissa_0[simp]: "mantissa 0 = 0" (is ?M) proof - have "\p::int \ int. fst p = 0 \ snd p = 0 \ p = (0, 0)" by auto then show ?E ?M by (auto simp add: mantissa_def exponent_def zero_float_def) qed lemma mantissa_exponent: "real_of_float f = mantissa f * 2 powr exponent f" (is ?E) and mantissa_not_dvd: "f \ 0 \ \ 2 dvd mantissa f" (is "_ \ ?D") proof cases assume [simp]: "f \ 0" have "f = mantissa f * 2 powr exponent f \ \ 2 dvd mantissa f" proof (cases f rule: float_normed_cases) case zero then show ?thesis by simp next case (powr m e) then have "\p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) \ (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p)" by auto then show ?thesis unfolding exponent_def mantissa_def by (rule someI2_ex) simp qed then show ?E ?D by auto qed simp lemma mantissa_noteq_0: "f \ 0 \ mantissa f \ 0" using mantissa_not_dvd[of f] by auto lemma mantissa_eq_zero_iff: "mantissa x = 0 \ x = 0" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - from that have z: "0 = real_of_float x" using mantissa_exponent by simp show ?thesis by (simp add: zero_float_def z) qed show ?lhs if ?rhs using that by simp qed lemma mantissa_pos_iff: "0 < mantissa x \ 0 < x" by (auto simp: mantissa_exponent algebra_split_simps) lemma mantissa_nonneg_iff: "0 \ mantissa x \ 0 \ x" by (auto simp: mantissa_exponent algebra_split_simps) lemma mantissa_neg_iff: "0 > mantissa x \ 0 > x" by (auto simp: mantissa_exponent algebra_split_simps) lemma fixes m e :: int defines "f \ float_of (m * 2 powr e)" assumes dvd: "\ 2 dvd m" shows mantissa_float: "mantissa f = m" (is "?M") and exponent_float: "m \ 0 \ exponent f = e" (is "_ \ ?E") proof cases assume "m = 0" with dvd show "mantissa f = m" by auto next assume "m \ 0" then have f_not_0: "f \ 0" by (simp add: f_def zero_float_def) from mantissa_exponent[of f] have "m * 2 powr e = mantissa f * 2 powr exponent f" by (auto simp add: f_def) then show ?M ?E using mantissa_not_dvd[OF f_not_0] dvd by (auto simp: mult_powr_eq_mult_powr_iff) qed subsection \Compute arithmetic operations\ lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f" unfolding real_of_float_eq mantissa_exponent[of f] by simp lemma Float_cases [cases type: float]: fixes f :: float obtains (Float) m e :: int where "f = Float m e" using Float_mantissa_exponent[symmetric] by (atomize_elim) auto lemma denormalize_shift: assumes f_def: "f = Float m e" and not_0: "f \ 0" obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i" proof from mantissa_exponent[of f] f_def have "m * 2 powr e = mantissa f * 2 powr exponent f" by simp then have eq: "m = mantissa f * 2 powr (exponent f - e)" by (simp add: powr_diff field_simps) moreover have "e \ exponent f" proof (rule ccontr) assume "\ e \ exponent f" then have pos: "exponent f < e" by simp then have "2 powr (exponent f - e) = 2 powr - real_of_int (e - exponent f)" by simp also have "\ = 1 / 2^nat (e - exponent f)" using pos by (simp flip: powr_realpow add: powr_diff) finally have "m * 2^nat (e - exponent f) = real_of_int (mantissa f)" using eq by simp then have "mantissa f = m * 2^nat (e - exponent f)" by linarith with \exponent f < e\ have "2 dvd mantissa f" apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"]) apply (cases "nat (e - exponent f)") apply auto done then show False using mantissa_not_dvd[OF not_0] by simp qed ultimately have "real_of_int m = mantissa f * 2^nat (exponent f - e)" by (simp flip: powr_realpow) with \e \ exponent f\ show "m = mantissa f * 2 ^ nat (exponent f - e)" by linarith show "e = exponent f - nat (exponent f - e)" using \e \ exponent f\ by auto qed context begin qualified lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0" by transfer simp qualified lemma compute_float_one[code_unfold, code]: "1 = Float 1 0" by transfer simp lift_definition normfloat :: "float \ float" is "\x. x" . lemma normloat_id[simp]: "normfloat x = x" by transfer rule qualified lemma compute_normfloat[code]: "normfloat (Float m e) = (if m mod 2 = 0 \ m \ 0 then normfloat (Float (m div 2) (e + 1)) else if m = 0 then 0 else Float m e)" by transfer (auto simp add: powr_add zmod_eq_0_iff) qualified lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k" by transfer simp qualified lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k" by transfer simp qualified lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1" by transfer simp qualified lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)" by transfer (simp add: field_simps powr_add) qualified lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 = (if m1 = 0 then Float m2 e2 else if m2 = 0 then Float m1 e1 else if e1 \ e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1 else Float (m2 + m1 * 2^nat (e1 - e2)) e2)" by transfer (simp add: field_simps powr_realpow[symmetric] powr_diff) qualified lemma compute_float_minus[code]: "f - g = f + (-g)" for f g :: float by simp qualified lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)" by transfer (simp add: sgn_mult) lift_definition is_float_pos :: "float \ bool" is "(<) 0 :: real \ bool" . qualified lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \ 0 < m" by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0]) lift_definition is_float_nonneg :: "float \ bool" is "(\) 0 :: real \ bool" . qualified lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \ 0 \ m" by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0]) lift_definition is_float_zero :: "float \ bool" is "(=) 0 :: real \ bool" . qualified lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \ 0 = m" by transfer (auto simp add: is_float_zero_def) qualified lemma compute_float_abs[code]: "\Float m e\ = Float \m\ e" by transfer (simp add: abs_mult) qualified lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)" by transfer simp end subsection \Lemmas for types \<^typ>\real\, \<^typ>\nat\, \<^typ>\int\\ lemmas real_of_ints = of_int_add of_int_minus of_int_diff of_int_mult of_int_power of_int_numeral of_int_neg_numeral lemmas int_of_reals = real_of_ints[symmetric] subsection \Rounding Real Numbers\ definition round_down :: "int \ real \ real" where "round_down prec x = \x * 2 powr prec\ * 2 powr -prec" definition round_up :: "int \ real \ real" where "round_up prec x = \x * 2 powr prec\ * 2 powr -prec" lemma round_down_float[simp]: "round_down prec x \ float" unfolding round_down_def by (auto intro!: times_float simp flip: of_int_minus) lemma round_up_float[simp]: "round_up prec x \ float" unfolding round_up_def by (auto intro!: times_float simp flip: of_int_minus) lemma round_up: "x \ round_up prec x" by (simp add: powr_minus_divide le_divide_eq round_up_def ceiling_correct) lemma round_down: "round_down prec x \ x" by (simp add: powr_minus_divide divide_le_eq round_down_def) lemma round_up_0[simp]: "round_up p 0 = 0" unfolding round_up_def by simp lemma round_down_0[simp]: "round_down p 0 = 0" unfolding round_down_def by simp lemma round_up_diff_round_down: "round_up prec x - round_down prec x \ 2 powr -prec" proof - have "round_up prec x - round_down prec x = (\x * 2 powr prec\ - \x * 2 powr prec\) * 2 powr -prec" by (simp add: round_up_def round_down_def field_simps) also have "\ \ 1 * 2 powr -prec" by (rule mult_mono) (auto simp flip: of_int_diff simp: ceiling_diff_floor_le_1) finally show ?thesis by simp qed lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x" unfolding round_down_def by (simp add: powr_add powr_mult field_simps powr_diff) (simp flip: powr_add) lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x" unfolding round_up_def by (simp add: powr_add powr_mult field_simps powr_diff) (simp flip: powr_add) lemma round_up_uminus_eq: "round_up p (-x) = - round_down p x" and round_down_uminus_eq: "round_down p (-x) = - round_up p x" by (auto simp: round_up_def round_down_def ceiling_def) lemma round_up_mono: "x \ y \ round_up p x \ round_up p y" by (auto intro!: ceiling_mono simp: round_up_def) lemma round_up_le1: assumes "x \ 1" "prec \ 0" shows "round_up prec x \ 1" proof - have "real_of_int \x * 2 powr prec\ \ real_of_int \2 powr real_of_int prec\" using assms by (auto intro!: ceiling_mono) also have "\ = 2 powr prec" using assms by (auto simp: powr_int intro!: exI[where x="2^nat prec"]) finally show ?thesis by (simp add: round_up_def) (simp add: powr_minus inverse_eq_divide) qed lemma round_up_less1: assumes "x < 1 / 2" "p > 0" shows "round_up p x < 1" proof - have "x * 2 powr p < 1 / 2 * 2 powr p" using assms by simp also have "\ \ 2 powr p - 1" using \p > 0\ by (auto simp: powr_diff powr_int field_simps self_le_power) finally show ?thesis using \p > 0\ by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_iff) qed lemma round_down_ge1: assumes x: "x \ 1" assumes prec: "p \ - log 2 x" shows "1 \ round_down p x" proof cases assume nonneg: "0 \ p" have "2 powr p = real_of_int \2 powr real_of_int p\" using nonneg by (auto simp: powr_int) also have "\ \ real_of_int \x * 2 powr p\" using assms by (auto intro!: floor_mono) finally show ?thesis by (simp add: round_down_def) (simp add: powr_minus inverse_eq_divide) next assume neg: "\ 0 \ p" have "x = 2 powr (log 2 x)" using x by simp also have "2 powr (log 2 x) \ 2 powr - p" using prec by auto finally have x_le: "x \ 2 powr -p" . from neg have "2 powr real_of_int p \ 2 powr 0" by (intro powr_mono) auto also have "\ \ \2 powr 0::real\" by simp also have "\ \ \x * 2 powr (real_of_int p)\" unfolding of_int_le_iff using x x_le by (intro floor_mono) (simp add: powr_minus_divide field_simps) finally show ?thesis using prec x by (simp add: round_down_def powr_minus_divide pos_le_divide_eq) qed lemma round_up_le0: "x \ 0 \ round_up p x \ 0" unfolding round_up_def by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff) subsection \Rounding Floats\ definition div_twopow :: "int \ nat \ int" where [simp]: "div_twopow x n = x div (2 ^ n)" definition mod_twopow :: "int \ nat \ int" where [simp]: "mod_twopow x n = x mod (2 ^ n)" lemma compute_div_twopow[code]: "div_twopow x n = (if x = 0 \ x = -1 \ n = 0 then x else div_twopow (x div 2) (n - 1))" by (cases n) (auto simp: zdiv_zmult2_eq div_eq_minus1) lemma compute_mod_twopow[code]: "mod_twopow x n = (if n = 0 then 0 else x mod 2 + 2 * mod_twopow (x div 2) (n - 1))" by (cases n) (auto simp: zmod_zmult2_eq) lift_definition float_up :: "int \ float \ float" is round_up by simp declare float_up.rep_eq[simp] lemma round_up_correct: "round_up e f - f \ {0..2 powr -e}" unfolding atLeastAtMost_iff proof have "round_up e f - f \ round_up e f - round_down e f" using round_down by simp also have "\ \ 2 powr -e" using round_up_diff_round_down by simp finally show "round_up e f - f \ 2 powr - (real_of_int e)" by simp qed (simp add: algebra_simps round_up) lemma float_up_correct: "real_of_float (float_up e f) - real_of_float f \ {0..2 powr -e}" by transfer (rule round_up_correct) lift_definition float_down :: "int \ float \ float" is round_down by simp declare float_down.rep_eq[simp] lemma round_down_correct: "f - (round_down e f) \ {0..2 powr -e}" unfolding atLeastAtMost_iff proof have "f - round_down e f \ round_up e f - round_down e f" using round_up by simp also have "\ \ 2 powr -e" using round_up_diff_round_down by simp finally show "f - round_down e f \ 2 powr - (real_of_int e)" by simp qed (simp add: algebra_simps round_down) lemma float_down_correct: "real_of_float f - real_of_float (float_down e f) \ {0..2 powr -e}" by transfer (rule round_down_correct) context begin qualified lemma compute_float_down[code]: "float_down p (Float m e) = (if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)" proof (cases "p + e < 0") case True then have "real_of_int ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" using powr_realpow[of 2 "nat (-(p + e))"] by simp also have "\ = 1 / 2 powr p / 2 powr e" unfolding powr_minus_divide of_int_minus by (simp add: powr_add) finally show ?thesis using \p + e < 0\ apply transfer apply (simp add: round_down_def field_simps flip: floor_divide_of_int_eq) apply (metis (no_types, hide_lams) Float.rep_eq add.inverse_inverse compute_real_of_float diff_minus_eq_add floor_divide_of_int_eq int_of_reals(1) linorder_not_le minus_add_distrib of_int_eq_numeral_power_cancel_iff powr_add) done next case False then have r: "real_of_int e + real_of_int p = real (nat (e + p))" by simp have r: "\(m * 2 powr e) * 2 powr real_of_int p\ = (m * 2 powr e) * 2 powr real_of_int p" by (auto intro: exI[where x="m*2^nat (e+p)"] simp add: ac_simps powr_add[symmetric] r powr_realpow) with \\ p + e < 0\ show ?thesis by transfer (auto simp add: round_down_def field_simps powr_add powr_minus) qed lemma abs_round_down_le: "\f - (round_down e f)\ \ 2 powr -e" using round_down_correct[of f e] by simp lemma abs_round_up_le: "\f - (round_up e f)\ \ 2 powr -e" using round_up_correct[of e f] by simp lemma round_down_nonneg: "0 \ s \ 0 \ round_down p s" by (auto simp: round_down_def) lemma ceil_divide_floor_conv: assumes "b \ 0" shows "\real_of_int a / real_of_int b\ = (if b dvd a then a div b else \real_of_int a / real_of_int b\ + 1)" proof (cases "b dvd a") case True then show ?thesis by (simp add: ceiling_def floor_divide_of_int_eq dvd_neg_div flip: of_int_minus divide_minus_left) next case False then have "a mod b \ 0" by auto then have ne: "real_of_int (a mod b) / real_of_int b \ 0" using \b \ 0\ by auto have "\real_of_int a / real_of_int b\ = \real_of_int a / real_of_int b\ + 1" apply (rule ceiling_eq) apply (auto simp flip: floor_divide_of_int_eq) proof - have "real_of_int \real_of_int a / real_of_int b\ \ real_of_int a / real_of_int b" by simp moreover have "real_of_int \real_of_int a / real_of_int b\ \ real_of_int a / real_of_int b" apply (subst (2) real_of_int_div_aux) unfolding floor_divide_of_int_eq using ne \b \ 0\ apply auto done ultimately show "real_of_int \real_of_int a / real_of_int b\ < real_of_int a / real_of_int b" by arith qed then show ?thesis using \\ b dvd a\ by simp qed qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)" by transfer (simp add: round_down_uminus_eq) end lemma bitlen_Float: fixes m e defines [THEN meta_eq_to_obj_eq]: "f \ Float m e" shows "bitlen \mantissa f\ + exponent f = (if m = 0 then 0 else bitlen \m\ + e)" proof (cases "m = 0") case True then show ?thesis by (simp add: f_def bitlen_alt_def) next case False then have "f \ 0" unfolding real_of_float_eq by (simp add: f_def) then have "mantissa f \ 0" by (simp add: mantissa_eq_zero_iff) moreover obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i" by (rule f_def[THEN denormalize_shift, OF \f \ 0\]) ultimately show ?thesis by (simp add: abs_mult) qed lemma float_gt1_scale: assumes "1 \ Float m e" shows "0 \ e + (bitlen m - 1)" proof - have "0 < Float m e" using assms by auto then have "0 < m" using powr_gt_zero[of 2 e] by (auto simp: zero_less_mult_iff) then have "m \ 0" by auto show ?thesis proof (cases "0 \ e") case True then show ?thesis using \0 < m\ by (simp add: bitlen_alt_def) next case False have "(1::int) < 2" by simp let ?S = "2^(nat (-e))" have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"] by (auto simp: powr_minus field_simps) then have "1 \ real_of_int m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"] by (auto simp: powr_minus) then have "1 * ?S \ real_of_int m * inverse ?S * ?S" by (rule mult_right_mono) auto then have "?S \ real_of_int m" unfolding mult.assoc by auto then have "?S \ m" unfolding of_int_le_iff[symmetric] by auto from this bitlen_bounds[OF \0 < m\, THEN conjunct2] have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF \1 < 2\, symmetric] by (rule order_le_less_trans) then have "-e < bitlen m" using False by auto then show ?thesis by auto qed qed subsection \Truncating Real Numbers\ definition truncate_down::"nat \ real \ real" where "truncate_down prec x = round_down (prec - \log 2 \x\\) x" lemma truncate_down: "truncate_down prec x \ x" using round_down by (simp add: truncate_down_def) lemma truncate_down_le: "x \ y \ truncate_down prec x \ y" by (rule order_trans[OF truncate_down]) lemma truncate_down_zero[simp]: "truncate_down prec 0 = 0" by (simp add: truncate_down_def) lemma truncate_down_float[simp]: "truncate_down p x \ float" by (auto simp: truncate_down_def) definition truncate_up::"nat \ real \ real" where "truncate_up prec x = round_up (prec - \log 2 \x\\) x" lemma truncate_up: "x \ truncate_up prec x" using round_up by (simp add: truncate_up_def) lemma truncate_up_le: "x \ y \ x \ truncate_up prec y" by (rule order_trans[OF _ truncate_up]) lemma truncate_up_zero[simp]: "truncate_up prec 0 = 0" by (simp add: truncate_up_def) lemma truncate_up_uminus_eq: "truncate_up prec (-x) = - truncate_down prec x" and truncate_down_uminus_eq: "truncate_down prec (-x) = - truncate_up prec x" by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def) lemma truncate_up_float[simp]: "truncate_up p x \ float" by (auto simp: truncate_up_def) lemma mult_powr_eq: "0 < b \ b \ 1 \ 0 < x \ x * b powr y = b powr (y + log b x)" by (simp_all add: powr_add) lemma truncate_down_pos: assumes "x > 0" shows "truncate_down p x > 0" proof - have "0 \ log 2 x - real_of_int \log 2 x\" by (simp add: algebra_simps) with assms show ?thesis apply (auto simp: truncate_down_def round_down_def mult_powr_eq intro!: ge_one_powr_ge_zero mult_pos_pos) by linarith qed lemma truncate_down_nonneg: "0 \ y \ 0 \ truncate_down prec y" by (auto simp: truncate_down_def round_down_def) lemma truncate_down_ge1: "1 \ x \ 1 \ truncate_down p x" apply (auto simp: truncate_down_def algebra_simps intro!: round_down_ge1) apply linarith done lemma truncate_up_nonpos: "x \ 0 \ truncate_up prec x \ 0" by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg) lemma truncate_up_le1: assumes "x \ 1" shows "truncate_up p x \ 1" proof - consider "x \ 0" | "x > 0" by arith then show ?thesis proof cases case 1 with truncate_up_nonpos[OF this, of p] show ?thesis by simp next case 2 then have le: "\log 2 \x\\ \ 0" using assms by (auto simp: log_less_iff) from assms have "0 \ int p" by simp from add_mono[OF this le] show ?thesis using assms by (simp add: truncate_up_def round_up_le1 add_mono) qed qed lemma truncate_down_shift_int: "truncate_down p (x * 2 powr real_of_int k) = truncate_down p x * 2 powr k" by (cases "x = 0") (simp_all add: algebra_simps abs_mult log_mult truncate_down_def round_down_shift[of _ _ k, simplified]) lemma truncate_down_shift_nat: "truncate_down p (x * 2 powr real k) = truncate_down p x * 2 powr k" by (metis of_int_of_nat_eq truncate_down_shift_int) lemma truncate_up_shift_int: "truncate_up p (x * 2 powr real_of_int k) = truncate_up p x * 2 powr k" by (cases "x = 0") (simp_all add: algebra_simps abs_mult log_mult truncate_up_def round_up_shift[of _ _ k, simplified]) lemma truncate_up_shift_nat: "truncate_up p (x * 2 powr real k) = truncate_up p x * 2 powr k" by (metis of_int_of_nat_eq truncate_up_shift_int) subsection \Truncating Floats\ lift_definition float_round_up :: "nat \ float \ float" is truncate_up by (simp add: truncate_up_def) lemma float_round_up: "real_of_float x \ real_of_float (float_round_up prec x)" using truncate_up by transfer simp lemma float_round_up_zero[simp]: "float_round_up prec 0 = 0" by transfer simp lift_definition float_round_down :: "nat \ float \ float" is truncate_down by (simp add: truncate_down_def) lemma float_round_down: "real_of_float (float_round_down prec x) \ real_of_float x" using truncate_down by transfer simp lemma float_round_down_zero[simp]: "float_round_down prec 0 = 0" by transfer simp lemmas float_round_up_le = order_trans[OF _ float_round_up] and float_round_down_le = order_trans[OF float_round_down] lemma minus_float_round_up_eq: "- float_round_up prec x = float_round_down prec (- x)" and minus_float_round_down_eq: "- float_round_down prec x = float_round_up prec (- x)" by (transfer; simp add: truncate_down_uminus_eq truncate_up_uminus_eq)+ context begin qualified lemma compute_float_round_down[code]: "float_round_down prec (Float m e) = (let d = bitlen \m\ - int prec - 1 in if 0 < d then Float (div_twopow m (nat d)) (e + d) else Float m e)" using Float.compute_float_down[of "Suc prec - bitlen \m\ - e" m e, symmetric] by transfer (simp add: field_simps abs_mult log_mult bitlen_alt_def truncate_down_def cong del: if_weak_cong) qualified lemma compute_float_round_up[code]: "float_round_up prec x = - float_round_down prec (-x)" by transfer (simp add: truncate_down_uminus_eq) end lemma truncate_up_nonneg_mono: assumes "0 \ x" "x \ y" shows "truncate_up prec x \ truncate_up prec y" proof - consider "\log 2 x\ = \log 2 y\" | "\log 2 x\ \ \log 2 y\" "0 < x" | "x \ 0" by arith then show ?thesis proof cases case 1 then show ?thesis using assms by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono) next case 2 from assms \0 < x\ have "log 2 x \ log 2 y" by auto with \\log 2 x\ \ \log 2 y\\ have logless: "log 2 x < log 2 y" by linarith have flogless: "\log 2 x\ < \log 2 y\" using \\log 2 x\ \ \log 2 y\\ \log 2 x \ log 2 y\ by linarith have "truncate_up prec x = real_of_int \x * 2 powr real_of_int (int prec - \log 2 x\ )\ * 2 powr - real_of_int (int prec - \log 2 x\)" using assms by (simp add: truncate_up_def round_up_def) also have "\x * 2 powr real_of_int (int prec - \log 2 x\)\ \ (2 ^ (Suc prec))" proof (simp only: ceiling_le_iff) have "x * 2 powr real_of_int (int prec - \log 2 x\) \ x * (2 powr real (Suc prec) / (2 powr log 2 x))" using real_of_int_floor_add_one_ge[of "log 2 x"] assms by (auto simp: algebra_simps simp flip: powr_diff intro!: mult_left_mono) then show "x * 2 powr real_of_int (int prec - \log 2 x\) \ real_of_int ((2::int) ^ (Suc prec))" using \0 < x\ by (simp add: powr_realpow powr_add) qed then have "real_of_int \x * 2 powr real_of_int (int prec - \log 2 x\)\ \ 2 powr int (Suc prec)" by (auto simp: powr_realpow powr_add) (metis power_Suc of_int_le_numeral_power_cancel_iff) also have "2 powr - real_of_int (int prec - \log 2 x\) \ 2 powr - real_of_int (int prec - \log 2 y\ + 1)" using logless flogless by (auto intro!: floor_mono) also have "2 powr real_of_int (int (Suc prec)) \ 2 powr (log 2 y + real_of_int (int prec - \log 2 y\ + 1))" using assms \0 < x\ by (auto simp: algebra_simps) finally have "truncate_up prec x \ 2 powr (log 2 y + real_of_int (int prec - \log 2 y\ + 1)) * 2 powr - real_of_int (int prec - \log 2 y\ + 1)" by simp also have "\ = 2 powr (log 2 y + real_of_int (int prec - \log 2 y\) - real_of_int (int prec - \log 2 y\))" by (subst powr_add[symmetric]) simp also have "\ = y" using \0 < x\ assms by (simp add: powr_add) also have "\ \ truncate_up prec y" by (rule truncate_up) finally show ?thesis . next case 3 then show ?thesis using assms by (auto intro!: truncate_up_le) qed qed lemma truncate_up_switch_sign_mono: assumes "x \ 0" "0 \ y" shows "truncate_up prec x \ truncate_up prec y" proof - note truncate_up_nonpos[OF \x \ 0\] also note truncate_up_le[OF \0 \ y\] finally show ?thesis . qed lemma truncate_down_switch_sign_mono: assumes "x \ 0" and "0 \ y" and "x \ y" shows "truncate_down prec x \ truncate_down prec y" proof - note truncate_down_le[OF \x \ 0\] also note truncate_down_nonneg[OF \0 \ y\] finally show ?thesis . qed lemma truncate_down_nonneg_mono: assumes "0 \ x" "x \ y" shows "truncate_down prec x \ truncate_down prec y" proof - consider "x \ 0" | "\log 2 \x\\ = \log 2 \y\\" | "0 < x" "\log 2 \x\\ \ \log 2 \y\\" by arith then show ?thesis proof cases case 1 with assms have "x = 0" "0 \ y" by simp_all then show ?thesis by (auto intro!: truncate_down_nonneg) next case 2 then show ?thesis using assms by (auto simp: truncate_down_def round_down_def intro!: floor_mono) next case 3 from \0 < x\ have "log 2 x \ log 2 y" "0 < y" "0 \ y" using assms by auto with \\log 2 \x\\ \ \log 2 \y\\\ have logless: "log 2 x < log 2 y" and flogless: "\log 2 x\ < \log 2 y\" unfolding atomize_conj abs_of_pos[OF \0 < x\] abs_of_pos[OF \0 < y\] by (metis floor_less_cancel linorder_cases not_le) have "2 powr prec \ y * 2 powr real prec / (2 powr log 2 y)" using \0 < y\ by simp also have "\ \ y * 2 powr real (Suc prec) / (2 powr (real_of_int \log 2 y\ + 1))" using \0 \ y\ \0 \ x\ assms(2) by (auto intro!: powr_mono divide_left_mono simp: of_nat_diff powr_add powr_diff) also have "\ = y * 2 powr real (Suc prec) / (2 powr real_of_int \log 2 y\ * 2)" by (auto simp: powr_add) finally have "(2 ^ prec) \ \y * 2 powr real_of_int (int (Suc prec) - \log 2 \y\\ - 1)\" using \0 \ y\ by (auto simp: powr_diff le_floor_iff powr_realpow powr_add) then have "(2 ^ (prec)) * 2 powr - real_of_int (int prec - \log 2 \y\\) \ truncate_down prec y" by (auto simp: truncate_down_def round_down_def) moreover have "x \ (2 ^ prec) * 2 powr - real_of_int (int prec - \log 2 \y\\)" proof - have "x = 2 powr (log 2 \x\)" using \0 < x\ by simp also have "\ \ (2 ^ (Suc prec )) * 2 powr - real_of_int (int prec - \log 2 \x\\)" using real_of_int_floor_add_one_ge[of "log 2 \x\"] \0 < x\ by (auto simp flip: powr_realpow powr_add simp: algebra_simps powr_mult_base le_powr_iff) also have "2 powr - real_of_int (int prec - \log 2 \x\\) \ 2 powr - real_of_int (int prec - \log 2 \y\\ + 1)" using logless flogless \x > 0\ \y > 0\ by (auto intro!: floor_mono) finally show ?thesis by (auto simp flip: powr_realpow simp: powr_diff assms of_nat_diff) qed ultimately show ?thesis by (metis dual_order.trans truncate_down) qed qed lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)" and truncate_up_eq_truncate_down: "truncate_up p x = - truncate_down p (-x)" by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq) lemma truncate_down_mono: "x \ y \ truncate_down p x \ truncate_down p y" apply (cases "0 \ x") apply (rule truncate_down_nonneg_mono, assumption+) apply (simp add: truncate_down_eq_truncate_up) apply (cases "0 \ y") apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono) done lemma truncate_up_mono: "x \ y \ truncate_up p x \ truncate_up p y" by (simp add: truncate_up_eq_truncate_down truncate_down_mono) lemma truncate_up_nonneg: "0 \ truncate_up p x" if "0 \ x" by (simp add: that truncate_up_le) lemma truncate_up_pos: "0 < truncate_up p x" if "0 < x" by (meson less_le_trans that truncate_up) lemma truncate_up_less_zero_iff[simp]: "truncate_up p x < 0 \ x < 0" proof - have f1: "\n r. truncate_up n r + truncate_down n (- 1 * r) = 0" by (simp add: truncate_down_uminus_eq) have f2: "(\v0 v1. truncate_up v0 v1 + truncate_down v0 (- 1 * v1) = 0) = (\v0 v1. truncate_up v0 v1 = - 1 * truncate_down v0 (- 1 * v1))" by (auto simp: truncate_up_eq_truncate_down) have f3: "\x1. ((0::real) < x1) = (\ x1 \ 0)" by fastforce have "(- 1 * x \ 0) = (0 \ x)" by force then have "0 \ x \ \ truncate_down p (- 1 * x) \ 0" using f3 by (meson truncate_down_pos) then have "(0 \ truncate_up p x) \ (\ 0 \ x)" using f2 f1 truncate_up_nonneg by force then show ?thesis by linarith qed lemma truncate_up_nonneg_iff[simp]: "truncate_up p x \ 0 \ x \ 0" using truncate_up_less_zero_iff[of p x] truncate_up_nonneg[of x] by linarith lemma truncate_down_less_zero_iff[simp]: "truncate_down p x < 0 \ x < 0" by (metis le_less_trans not_less_iff_gr_or_eq truncate_down truncate_down_pos truncate_down_zero) lemma truncate_down_nonneg_iff[simp]: "truncate_down p x \ 0 \ x \ 0" using truncate_down_less_zero_iff[of p x] truncate_down_nonneg[of x p] by linarith lemma truncate_down_eq_zero_iff[simp]: "truncate_down prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_down_less_zero_iff truncate_down_pos truncate_down_zero) lemma truncate_up_eq_zero_iff[simp]: "truncate_up prec x = 0 \ x = 0" by (metis not_less_iff_gr_or_eq truncate_up_less_zero_iff truncate_up_pos truncate_up_zero) subsection \Approximation of positive rationals\ lemma div_mult_twopow_eq: "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" for a b :: nat by (cases "b = 0") (simp_all add: div_mult2_eq[symmetric] ac_simps) lemma real_div_nat_eq_floor_of_divide: "a div b = real_of_int \a / b\" for a b :: nat by (simp add: floor_divide_of_nat_eq [of a b]) definition "rat_precision prec x y = (let d = bitlen x - bitlen y in int prec - d + (if Float (abs x) 0 < Float (abs y) d then 1 else 0))" lemma floor_log_divide_eq: assumes "i > 0" "j > 0" "p > 1" shows "\log p (i / j)\ = floor (log p i) - floor (log p j) - (if i \ j * p powr (floor (log p i) - floor (log p j)) then 0 else 1)" proof - let ?l = "log p" let ?fl = "\x. floor (?l x)" have "\?l (i / j)\ = \?l i - ?l j\" using assms by (auto simp: log_divide) also have "\ = floor (real_of_int (?fl i - ?fl j) + (?l i - ?fl i - (?l j - ?fl j)))" (is "_ = floor (_ + ?r)") by (simp add: algebra_simps) also note floor_add2 also note \p > 1\ note powr = powr_le_cancel_iff[symmetric, OF \1 < p\, THEN iffD2] note powr_strict = powr_less_cancel_iff[symmetric, OF \1 < p\, THEN iffD2] have "floor ?r = (if i \ j * p powr (?fl i - ?fl j) then 0 else -1)" (is "_ = ?if") using assms by (linarith | auto intro!: floor_eq2 intro: powr_strict powr simp: powr_diff powr_add field_split_simps algebra_simps)+ finally show ?thesis by simp qed lemma truncate_down_rat_precision: "truncate_down prec (real x / real y) = round_down (rat_precision prec x y) (real x / real y)" and truncate_up_rat_precision: "truncate_up prec (real x / real y) = round_up (rat_precision prec x y) (real x / real y)" unfolding truncate_down_def truncate_up_def rat_precision_def by (cases x; cases y) (auto simp: floor_log_divide_eq algebra_simps bitlen_alt_def) lift_definition lapprox_posrat :: "nat \ nat \ nat \ float" is "\prec (x::nat) (y::nat). truncate_down prec (x / y)" by simp context begin qualified lemma compute_lapprox_posrat[code]: "lapprox_posrat prec x y = (let l = rat_precision prec x y; d = if 0 \ l then x * 2^nat l div y else x div 2^nat (- l) div y in normfloat (Float d (- l)))" unfolding div_mult_twopow_eq by transfer (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def truncate_down_rat_precision del: two_powr_minus_int_float) end lift_definition rapprox_posrat :: "nat \ nat \ nat \ float" is "\prec (x::nat) (y::nat). truncate_up prec (x / y)" by simp context begin qualified lemma compute_rapprox_posrat[code]: fixes prec x y defines "l \ rat_precision prec x y" shows "rapprox_posrat prec x y = (let l = l; (r, s) = if 0 \ l then (x * 2^nat l, y) else (x, y * 2^nat(-l)); d = r div s; m = r mod s in normfloat (Float (d + (if m = 0 \ y = 0 then 0 else 1)) (- l)))" proof (cases "y = 0") assume "y = 0" then show ?thesis by transfer simp next assume "y \ 0" show ?thesis proof (cases "0 \ l") case True define x' where "x' = x * 2 ^ nat l" have "int x * 2 ^ nat l = x'" by (simp add: x'_def) moreover have "real x * 2 powr l = real x'" by (simp flip: powr_realpow add: \0 \ l\ x'_def) ultimately show ?thesis using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] \0 \ l\ \y \ 0\ l_def[symmetric, THEN meta_eq_to_obj_eq] apply transfer apply (auto simp add: round_up_def truncate_up_rat_precision) apply (metis dvd_triv_left of_nat_dvd_iff) apply (metis floor_divide_of_int_eq of_int_of_nat_eq) done next case False define y' where "y' = y * 2 ^ nat (- l)" from \y \ 0\ have "y' \ 0" by (simp add: y'_def) have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def) moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'" using \\ 0 \ l\ by (simp flip: powr_realpow add: powr_minus y'_def field_simps) ultimately show ?thesis using ceil_divide_floor_conv[of y' x] \\ 0 \ l\ \y' \ 0\ \y \ 0\ l_def[symmetric, THEN meta_eq_to_obj_eq] apply transfer apply (auto simp add: round_up_def ceil_divide_floor_conv truncate_up_rat_precision) apply (metis dvd_triv_left of_nat_dvd_iff) apply (metis floor_divide_of_int_eq of_int_of_nat_eq) done qed qed end lemma rat_precision_pos: assumes "0 \ x" and "0 < y" and "2 * x < y" shows "rat_precision n (int x) (int y) > 0" proof - have "0 < x \ log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) then have "bitlen (int x) < bitlen (int y)" using assms by (simp add: bitlen_alt_def) (auto intro!: floor_mono simp add: one_add_floor) then show ?thesis using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def) qed lemma rapprox_posrat_less1: "0 \ x \ 0 < y \ 2 * x < y \ real_of_float (rapprox_posrat n x y) < 1" by transfer (simp add: rat_precision_pos round_up_less1 truncate_up_rat_precision) lift_definition lapprox_rat :: "nat \ int \ int \ float" is "\prec (x::int) (y::int). truncate_down prec (x / y)" by simp context begin qualified lemma compute_lapprox_rat[code]: "lapprox_rat prec x y = (if y = 0 then 0 else if 0 \ x then (if 0 < y then lapprox_posrat prec (nat x) (nat y) else - (rapprox_posrat prec (nat x) (nat (-y)))) else (if 0 < y then - (rapprox_posrat prec (nat (-x)) (nat y)) else lapprox_posrat prec (nat (-x)) (nat (-y))))" by transfer (simp add: truncate_up_uminus_eq) lift_definition rapprox_rat :: "nat \ int \ int \ float" is "\prec (x::int) (y::int). truncate_up prec (x / y)" by simp lemma "rapprox_rat = rapprox_posrat" by transfer auto lemma "lapprox_rat = lapprox_posrat" by transfer auto qualified lemma compute_rapprox_rat[code]: "rapprox_rat prec x y = - lapprox_rat prec (-x) y" by transfer (simp add: truncate_down_uminus_eq) qualified lemma compute_truncate_down[code]: "truncate_down p (Ratreal r) = (let (a, b) = quotient_of r in lapprox_rat p a b)" by transfer (auto split: prod.split simp: of_rat_divide dest!: quotient_of_div) qualified lemma compute_truncate_up[code]: "truncate_up p (Ratreal r) = (let (a, b) = quotient_of r in rapprox_rat p a b)" by transfer (auto split: prod.split simp: of_rat_divide dest!: quotient_of_div) end subsection \Division\ definition "real_divl prec a b = truncate_down prec (a / b)" definition "real_divr prec a b = truncate_up prec (a / b)" lift_definition float_divl :: "nat \ float \ float \ float" is real_divl by (simp add: real_divl_def) context begin qualified lemma compute_float_divl[code]: "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)" apply transfer unfolding real_divl_def of_int_1 mult_1 truncate_down_shift_int[symmetric] apply (simp add: powr_diff powr_minus) done lift_definition float_divr :: "nat \ float \ float \ float" is real_divr by (simp add: real_divr_def) qualified lemma compute_float_divr[code]: "float_divr prec x y = - float_divl prec (-x) y" by transfer (simp add: real_divr_def real_divl_def truncate_down_uminus_eq) end subsection \Approximate Addition\ definition "plus_down prec x y = truncate_down prec (x + y)" definition "plus_up prec x y = truncate_up prec (x + y)" lemma float_plus_down_float[intro, simp]: "x \ float \ y \ float \ plus_down p x y \ float" by (simp add: plus_down_def) lemma float_plus_up_float[intro, simp]: "x \ float \ y \ float \ plus_up p x y \ float" by (simp add: plus_up_def) lift_definition float_plus_down :: "nat \ float \ float \ float" is plus_down .. lift_definition float_plus_up :: "nat \ float \ float \ float" is plus_up .. lemma plus_down: "plus_down prec x y \ x + y" and plus_up: "x + y \ plus_up prec x y" by (auto simp: plus_down_def truncate_down plus_up_def truncate_up) lemma float_plus_down: "real_of_float (float_plus_down prec x y) \ x + y" and float_plus_up: "x + y \ real_of_float (float_plus_up prec x y)" by (transfer; rule plus_down plus_up)+ lemmas plus_down_le = order_trans[OF plus_down] and plus_up_le = order_trans[OF _ plus_up] and float_plus_down_le = order_trans[OF float_plus_down] and float_plus_up_le = order_trans[OF _ float_plus_up] lemma compute_plus_up[code]: "plus_up p x y = - plus_down p (-x) (-y)" using truncate_down_uminus_eq[of p "x + y"] by (auto simp: plus_down_def plus_up_def) lemma truncate_down_log2_eqI: assumes "\log 2 \x\\ = \log 2 \y\\" assumes "\x * 2 powr (p - \log 2 \x\\)\ = \y * 2 powr (p - \log 2 \x\\)\" shows "truncate_down p x = truncate_down p y" using assms by (auto simp: truncate_down_def round_down_def) lemma sum_neq_zeroI: "\a\ \ k \ \b\ < k \ a + b \ 0" "\a\ > k \ \b\ \ k \ a + b \ 0" for a k :: real by auto lemma abs_real_le_2_powr_bitlen[simp]: "\real_of_int m2\ < 2 powr real_of_int (bitlen \m2\)" proof (cases "m2 = 0") case True then show ?thesis by simp next case False then have "\m2\ < 2 ^ nat (bitlen \m2\)" using bitlen_bounds[of "\m2\"] by (auto simp: powr_add bitlen_nonneg) then show ?thesis by (metis bitlen_nonneg powr_int of_int_abs of_int_less_numeral_power_cancel_iff zero_less_numeral) qed lemma floor_sum_times_2_powr_sgn_eq: fixes ai p q :: int and a b :: real assumes "a * 2 powr p = ai" and b_le_1: "\b * 2 powr (p + 1)\ \ 1" and leqp: "q \ p" shows "\(a + b) * 2 powr q\ = \(2 * ai + sgn b) * 2 powr (q - p - 1)\" proof - consider "b = 0" | "b > 0" | "b < 0" by arith then show ?thesis proof cases case 1 then show ?thesis by (simp flip: assms(1) powr_add add: algebra_simps powr_mult_base) next case 2 then have "b * 2 powr p < \b * 2 powr (p + 1)\" by simp also note b_le_1 finally have b_less_1: "b * 2 powr real_of_int p < 1" . from b_less_1 \b > 0\ have floor_eq: "\b * 2 powr real_of_int p\ = 0" "\sgn b / 2\ = 0" by (simp_all add: floor_eq_iff) have "\(a + b) * 2 powr q\ = \(a + b) * 2 powr p * 2 powr (q - p)\" by (simp add: algebra_simps flip: powr_realpow powr_add) also have "\ = \(ai + b * 2 powr p) * 2 powr (q - p)\" by (simp add: assms algebra_simps) also have "\ = \(ai + b * 2 powr p) / real_of_int ((2::int) ^ nat (p - q))\" using assms by (simp add: algebra_simps divide_powr_uminus flip: powr_realpow powr_add) also have "\ = \ai / real_of_int ((2::int) ^ nat (p - q))\" by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) finally have "\(a + b) * 2 powr real_of_int q\ = \real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\" . moreover have "\(2 * ai + (sgn b)) * 2 powr (real_of_int (q - p) - 1)\ = \real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\" proof - have "\(2 * ai + sgn b) * 2 powr (real_of_int (q - p) - 1)\ = \(ai + sgn b / 2) * 2 powr (q - p)\" by (subst powr_diff) (simp add: field_simps) also have "\ = \(ai + sgn b / 2) / real_of_int ((2::int) ^ nat (p - q))\" using leqp by (simp flip: powr_realpow add: powr_diff) also have "\ = \ai / real_of_int ((2::int) ^ nat (p - q))\" by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) finally show ?thesis . qed ultimately show ?thesis by simp next case 3 then have floor_eq: "\b * 2 powr (real_of_int p + 1)\ = -1" using b_le_1 by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus intro!: mult_neg_pos split: if_split_asm) have "\(a + b) * 2 powr q\ = \(2*a + 2*b) * 2 powr p * 2 powr (q - p - 1)\" by (simp add: algebra_simps powr_mult_base flip: powr_realpow powr_add) also have "\ = \(2 * (a * 2 powr p) + 2 * b * 2 powr p) * 2 powr (q - p - 1)\" by (simp add: algebra_simps) also have "\ = \(2 * ai + b * 2 powr (p + 1)) / 2 powr (1 - q + p)\" using assms by (simp add: algebra_simps powr_mult_base divide_powr_uminus) also have "\ = \(2 * ai + b * 2 powr (p + 1)) / real_of_int ((2::int) ^ nat (p - q + 1))\" using assms by (simp add: algebra_simps flip: powr_realpow) also have "\ = \(2 * ai - 1) / real_of_int ((2::int) ^ nat (p - q + 1))\" using \b < 0\ assms by (simp add: floor_divide_of_int_eq floor_eq floor_divide_real_eq_div del: of_int_mult of_int_power of_int_diff) also have "\ = \(2 * ai - 1) * 2 powr (q - p - 1)\" using assms by (simp add: algebra_simps divide_powr_uminus flip: powr_realpow) finally show ?thesis using \b < 0\ by simp qed qed lemma log2_abs_int_add_less_half_sgn_eq: fixes ai :: int and b :: real assumes "\b\ \ 1/2" and "ai \ 0" shows "\log 2 \real_of_int ai + b\\ = \log 2 \ai + sgn b / 2\\" proof (cases "b = 0") case True then show ?thesis by simp next case False define k where "k = \log 2 \ai\\" then have "\log 2 \ai\\ = k" by simp then have k: "2 powr k \ \ai\" "\ai\ < 2 powr (k + 1)" by (simp_all add: floor_log_eq_powr_iff \ai \ 0\) have "k \ 0" using assms by (auto simp: k_def) define r where "r = \ai\ - 2 ^ nat k" have r: "0 \ r" "r < 2 powr k" using \k \ 0\ k by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int) then have "r \ (2::int) ^ nat k - 1" using \k \ 0\ by (auto simp: powr_int) from this[simplified of_int_le_iff[symmetric]] \0 \ k\ have r_le: "r \ 2 powr k - 1" by (auto simp: algebra_simps powr_int) (metis of_int_1 of_int_add of_int_le_numeral_power_cancel_iff) have "\ai\ = 2 powr k + r" using \k \ 0\ by (auto simp: k_def r_def simp flip: powr_realpow) have pos: "\b\ < 1 \ 0 < 2 powr k + (r + b)" for b :: real using \0 \ k\ \ai \ 0\ by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps split: if_split_asm) have less: "\sgn ai * b\ < 1" and less': "\sgn (sgn ai * b) / 2\ < 1" using \\b\ \ _\ by (auto simp: abs_if sgn_if split: if_split_asm) have floor_eq: "\b::real. \b\ \ 1 / 2 \ \log 2 (1 + (r + b) / 2 powr k)\ = (if r = 0 \ b < 0 then -1 else 0)" using \k \ 0\ r r_le by (auto simp: floor_log_eq_powr_iff powr_minus_divide field_simps sgn_if) from \real_of_int \ai\ = _\ have "\ai + b\ = 2 powr k + (r + sgn ai * b)" using \\b\ \ _\ \0 \ k\ r by (auto simp add: sgn_if abs_if) also have "\log 2 \\ = \log 2 (2 powr k + r + sgn (sgn ai * b) / 2)\" proof - have "2 powr k + (r + (sgn ai) * b) = 2 powr k * (1 + (r + sgn ai * b) / 2 powr k)" by (simp add: field_simps) also have "\log 2 \\ = k + \log 2 (1 + (r + sgn ai * b) / 2 powr k)\" using pos[OF less] by (subst log_mult) (simp_all add: log_mult powr_mult field_simps) also let ?if = "if r = 0 \ sgn ai * b < 0 then -1 else 0" have "\log 2 (1 + (r + sgn ai * b) / 2 powr k)\ = ?if" using \\b\ \ _\ by (intro floor_eq) (auto simp: abs_mult sgn_if) also have "\ = \log 2 (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k)\" by (subst floor_eq) (auto simp: sgn_if) also have "k + \ = \log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\" unfolding int_add_floor using pos[OF less'] \\b\ \ _\ by (simp add: field_simps add_log_eq_powr del: floor_add2) also have "2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k) = 2 powr k + r + sgn (sgn ai * b) / 2" by (simp add: sgn_if field_simps) finally show ?thesis . qed also have "2 powr k + r + sgn (sgn ai * b) / 2 = \ai + sgn b / 2\" unfolding \real_of_int \ai\ = _\[symmetric] using \ai \ 0\ by (auto simp: abs_if sgn_if algebra_simps) finally show ?thesis . qed context begin qualified lemma compute_far_float_plus_down: fixes m1 e1 m2 e2 :: int and p :: nat defines "k1 \ Suc p - nat (bitlen \m1\)" assumes H: "bitlen \m2\ \ e1 - e2 - k1 - 2" "m1 \ 0" "m2 \ 0" "e1 \ e2" shows "float_plus_down p (Float m1 e1) (Float m2 e2) = float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))" proof - let ?a = "real_of_float (Float m1 e1)" let ?b = "real_of_float (Float m2 e2)" let ?sum = "?a + ?b" let ?shift = "real_of_int e2 - real_of_int e1 + real k1 + 1" let ?m1 = "m1 * 2 ^ Suc k1" let ?m2 = "m2 * 2 powr ?shift" let ?m2' = "sgn m2 / 2" let ?e = "e1 - int k1 - 1" have sum_eq: "?sum = (?m1 + ?m2) * 2 powr ?e" by (auto simp flip: powr_add powr_mult powr_realpow simp: powr_mult_base algebra_simps) have "\?m2\ * 2 < 2 powr (bitlen \m2\ + ?shift + 1)" by (auto simp: field_simps powr_add powr_mult_base powr_diff abs_mult) also have "\ \ 2 powr 0" using H by (intro powr_mono) auto finally have abs_m2_less_half: "\?m2\ < 1 / 2" by simp then have "\real_of_int m2\ < 2 powr -(?shift + 1)" unfolding powr_minus_divide by (auto simp: bitlen_alt_def field_simps powr_mult_base abs_mult) also have "\ \ 2 powr real_of_int (e1 - e2 - 2)" by simp finally have b_less_quarter: "\?b\ < 1/4 * 2 powr real_of_int e1" by (simp add: powr_add field_simps powr_diff abs_mult) also have "1/4 < \real_of_int m1\ / 2" using \m1 \ 0\ by simp finally have b_less_half_a: "\?b\ < 1/2 * \?a\" by (simp add: algebra_simps powr_mult_base abs_mult) then have a_half_less_sum: "\?a\ / 2 < \?sum\" by (auto simp: field_simps abs_if split: if_split_asm) from b_less_half_a have "\?b\ < \?a\" "\?b\ \ \?a\" by simp_all have "\real_of_float (Float m1 e1)\ \ 1/4 * 2 powr real_of_int e1" using \m1 \ 0\ by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult) then have "?sum \ 0" using b_less_quarter by (rule sum_neq_zeroI) then have "?m1 + ?m2 \ 0" unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff) have "\real_of_int ?m1\ \ 2 ^ Suc k1" "\?m2'\ < 2 ^ Suc k1" using \m1 \ 0\ \m2 \ 0\ by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps) then have sum'_nz: "?m1 + ?m2' \ 0" by (intro sum_neq_zeroI) have "\log 2 \real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\\ = \log 2 \?m1 + ?m2\\ + ?e" using \?m1 + ?m2 \ 0\ unfolding floor_add[symmetric] sum_eq by (simp add: abs_mult log_mult) linarith also have "\log 2 \?m1 + ?m2\\ = \log 2 \?m1 + sgn (real_of_int m2 * 2 powr ?shift) / 2\\" using abs_m2_less_half \m1 \ 0\ by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult) also have "sgn (real_of_int m2 * 2 powr ?shift) = sgn m2" by (auto simp: sgn_if zero_less_mult_iff less_not_sym) also have "\?m1 + ?m2'\ * 2 powr ?e = \?m1 * 2 + sgn m2\ * 2 powr (?e - 1)" by (auto simp: field_simps powr_minus[symmetric] powr_diff powr_mult_base) then have "\log 2 \?m1 + ?m2'\\ + ?e = \log 2 \real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1))\\" using \?m1 + ?m2' \ 0\ unfolding floor_add_int by (simp add: log_add_eq_powr abs_mult_pos del: floor_add2) finally have "\log 2 \?sum\\ = \log 2 \real_of_float (Float (?m1*2 + sgn m2) (?e - 1))\\" . then have "plus_down p (Float m1 e1) (Float m2 e2) = truncate_down p (Float (?m1*2 + sgn m2) (?e - 1))" unfolding plus_down_def proof (rule truncate_down_log2_eqI) let ?f = "(int p - \log 2 \real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\\)" let ?ai = "m1 * 2 ^ (Suc k1)" have "\(?a + ?b) * 2 powr real_of_int ?f\ = \(real_of_int (2 * ?ai) + sgn ?b) * 2 powr real_of_int (?f - - ?e - 1)\" proof (rule floor_sum_times_2_powr_sgn_eq) show "?a * 2 powr real_of_int (-?e) = real_of_int ?ai" by (simp add: powr_add powr_realpow[symmetric] powr_diff) show "\?b * 2 powr real_of_int (-?e + 1)\ \ 1" using abs_m2_less_half by (simp add: abs_mult powr_add[symmetric] algebra_simps powr_mult_base) next have "e1 + \log 2 \real_of_int m1\\ - 1 = \log 2 \?a\\ - 1" using \m1 \ 0\ by (simp add: int_add_floor algebra_simps log_mult abs_mult del: floor_add2) also have "\ \ \log 2 \?a + ?b\\" using a_half_less_sum \m1 \ 0\ \?sum \ 0\ unfolding floor_diff_of_int[symmetric] by (auto simp add: log_minus_eq_powr powr_minus_divide intro!: floor_mono) finally have "int p - \log 2 \?a + ?b\\ \ p - (bitlen \m1\) - e1 + 2" by (auto simp: algebra_simps bitlen_alt_def \m1 \ 0\) also have "\ \ - ?e" using bitlen_nonneg[of "\m1\"] by (simp add: k1_def) finally show "?f \ - ?e" by simp qed also have "sgn ?b = sgn m2" using powr_gt_zero[of 2 e2] by (auto simp add: sgn_if zero_less_mult_iff simp del: powr_gt_zero) also have "\(real_of_int (2 * ?m1) + real_of_int (sgn m2)) * 2 powr real_of_int (?f - - ?e - 1)\ = \Float (?m1 * 2 + sgn m2) (?e - 1) * 2 powr ?f\" by (simp flip: powr_add powr_realpow add: algebra_simps) finally show "\(?a + ?b) * 2 powr ?f\ = \real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\" . qed then show ?thesis by transfer (simp add: plus_down_def ac_simps Let_def) qed lemma compute_float_plus_down_naive[code]: "float_plus_down p x y = float_round_down p (x + y)" by transfer (auto simp: plus_down_def) qualified lemma compute_float_plus_down[code]: fixes p::nat and m1 e1 m2 e2::int shows "float_plus_down p (Float m1 e1) (Float m2 e2) = (if m1 = 0 then float_round_down p (Float m2 e2) else if m2 = 0 then float_round_down p (Float m1 e1) else (if e1 \ e2 then (let k1 = Suc p - nat (bitlen \m1\) in if bitlen \m2\ > e1 - e2 - k1 - 2 then float_round_down p ((Float m1 e1) + (Float m2 e2)) else float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))) else float_plus_down p (Float m2 e2) (Float m1 e1)))" proof - { assume "bitlen \m2\ \ e1 - e2 - (Suc p - nat (bitlen \m1\)) - 2" "m1 \ 0" "m2 \ 0" "e1 \ e2" note compute_far_float_plus_down[OF this] } then show ?thesis by transfer (simp add: Let_def plus_down_def ac_simps) qed qualified lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)" using truncate_down_uminus_eq[of p "x + y"] by transfer (simp add: plus_down_def plus_up_def ac_simps) lemma mantissa_zero: "mantissa 0 = 0" by (fact mantissa_0) qualified lemma compute_float_less[code]: "a < b \ is_float_pos (float_plus_down 0 b (- a))" using truncate_down[of 0 "b - a"] truncate_down_pos[of "b - a" 0] by transfer (auto simp: plus_down_def) qualified lemma compute_float_le[code]: "a \ b \ is_float_nonneg (float_plus_down 0 b (- a))" using truncate_down[of 0 "b - a"] truncate_down_nonneg[of "b - a" 0] by transfer (auto simp: plus_down_def) end lemma plus_down_mono: "plus_down p a b \ plus_down p c d" if "a + b \ c + d" by (auto simp: plus_down_def intro!: truncate_down_mono that) lemma plus_up_mono: "plus_up p a b \ plus_up p c d" if "a + b \ c + d" by (auto simp: plus_up_def intro!: truncate_up_mono that) subsection \Approximate Multiplication\ lemma mult_mono_nonpos_nonneg: "a * b \ c * d" if "a \ c" "a \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" by (meson dual_order.trans mult_left_mono_neg mult_right_mono that) lemma mult_mono_nonneg_nonpos: "b * a \ d * c" if "a \ c" "c \ 0" "0 \ d" "d \ b" for a b c d::"'a::ordered_ring" by (meson dual_order.trans mult_right_mono_neg mult_left_mono that) lemma mult_mono_nonpos_nonpos: "a * b \ c * d" if "a \ c" "a \ 0" "b \ d" "d \ 0" for a b c d::real by (meson dual_order.trans mult_left_mono_neg mult_right_mono_neg that) lemma mult_float_mono1: notes mono_rules = plus_down_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_down prec (nprt aa * pprt bc) (plus_down prec (nprt ba * nprt bc) (plus_down prec (pprt aa * pprt ac) (pprt ba * nprt ac))) \ plus_down prec (nprt a * pprt bb) (plus_down prec (nprt b * nprt bb) (plus_down prec (pprt a * pprt ab) (pprt b * nprt ab)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ lemma mult_float_mono2: notes mono_rules = plus_up_mono add_mono nprt_mono nprt_le_zero zero_le_pprt pprt_mono shows "a \ b \ ab \ bb \ aa \ a \ b \ ba \ ac \ ab \ bb \ bc \ plus_up prec (pprt b * pprt bb) (plus_up prec (pprt a * nprt bb) (plus_up prec (nprt b * pprt ab) (nprt a * nprt ab))) \ plus_up prec (pprt ba * pprt bc) (plus_up prec (pprt aa * nprt bc) (plus_up prec (nprt ba * pprt ac) (nprt aa * nprt ac)))" apply (rule order_trans) apply (rule mono_rules | assumption)+ apply (rule mult_mono) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonneg_nonpos) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonneg) apply (rule mono_rules | assumption)+ apply (rule mult_mono_nonpos_nonpos) apply (rule mono_rules | assumption)+ by (rule order_refl)+ subsection \Approximate Power\ lemma div2_less_self[termination_simp]: "odd n \ n div 2 < n" for n :: nat by (simp add: odd_pos) fun power_down :: "nat \ real \ nat \ real" where "power_down p x 0 = 1" | "power_down p x (Suc n) = (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2) else truncate_down (Suc p) (x * power_down p x n))" fun power_up :: "nat \ real \ nat \ real" where "power_up p x 0 = 1" | "power_up p x (Suc n) = (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2) else truncate_up p (x * power_up p x n))" lift_definition power_up_fl :: "nat \ float \ nat \ float" is power_up by (induct_tac rule: power_up.induct) simp_all lift_definition power_down_fl :: "nat \ float \ nat \ float" is power_down by (induct_tac rule: power_down.induct) simp_all lemma power_float_transfer[transfer_rule]: "(rel_fun pcr_float (rel_fun (=) pcr_float)) (^) (^)" unfolding power_def by transfer_prover lemma compute_power_up_fl[code]: "power_up_fl p x 0 = 1" "power_up_fl p x (Suc n) = (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2) else float_round_up p (x * power_up_fl p x n))" and compute_power_down_fl[code]: "power_down_fl p x 0 = 1" "power_down_fl p x (Suc n) = (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2) else float_round_down (Suc p) (x * power_down_fl p x n))" unfolding atomize_conj by transfer simp lemma power_down_pos: "0 < x \ 0 < power_down p x n" by (induct p x n rule: power_down.induct) (auto simp del: odd_Suc_div_two intro!: truncate_down_pos) lemma power_down_nonneg: "0 \ x \ 0 \ power_down p x n" by (induct p x n rule: power_down.induct) (auto simp del: odd_Suc_div_two intro!: truncate_down_nonneg mult_nonneg_nonneg) lemma power_down: "0 \ x \ power_down p x n \ x ^ n" proof (induct p x n rule: power_down.induct) case (2 p x n) have ?case if "odd n" proof - from that 2 have "(power_down p x (Suc n div 2)) ^ 2 \ (x ^ (Suc n div 2)) ^ 2" by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two) also have "\ = x ^ (Suc n div 2 * 2)" by (simp flip: power_mult) also have "Suc n div 2 * 2 = Suc n" using \odd n\ by presburger finally show ?thesis using that by (auto intro!: truncate_down_le simp del: odd_Suc_div_two) qed then show ?case by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg) qed simp lemma power_up: "0 \ x \ x ^ n \ power_up p x n" proof (induct p x n rule: power_up.induct) case (2 p x n) have ?case if "odd n" proof - from that even_Suc have "Suc n = Suc n div 2 * 2" by presburger then have "x ^ Suc n \ (x ^ (Suc n div 2))\<^sup>2" by (simp flip: power_mult) also from that 2 have "\ \ (power_up p x (Suc n div 2))\<^sup>2" by (auto intro: power_mono simp del: odd_Suc_div_two) finally show ?thesis using that by (auto intro!: truncate_up_le simp del: odd_Suc_div_two) qed then show ?case by (auto intro!: truncate_up_le mult_left_mono 2) qed simp lemmas power_up_le = order_trans[OF _ power_up] and power_up_less = less_le_trans[OF _ power_up] and power_down_le = order_trans[OF power_down] lemma power_down_fl: "0 \ x \ power_down_fl p x n \ x ^ n" by transfer (rule power_down) lemma power_up_fl: "0 \ x \ x ^ n \ power_up_fl p x n" by transfer (rule power_up) lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n" by transfer simp lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n" by transfer simp lemmas [simp del] = power_down.simps(2) power_up.simps(2) lemmas power_down_simp = power_down.simps(2) lemmas power_up_simp = power_up.simps(2) lemma power_down_even_nonneg: "even n \ 0 \ power_down p x n" by (induct p x n rule: power_down.induct) (auto simp: power_down_simp simp del: odd_Suc_div_two intro!: truncate_down_nonneg ) lemma power_down_eq_zero_iff[simp]: "power_down prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_down_simp[of _ _ "x - 1"] by (cases x) (auto simp add: div2_less_self) qed lemma power_down_nonneg_iff[simp]: "power_down prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_down_simp[of _ _ "x - 1"] by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff) qed lemma power_down_neg_iff[simp]: "power_down prec b n < 0 \ b < 0 \ odd n" using power_down_nonneg_iff[of prec b n] by (auto simp del: power_down_nonneg_iff) lemma power_down_nonpos_iff[simp]: notes [simp del] = power_down_neg_iff power_down_eq_zero_iff shows "power_down prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_down_neg_iff[of prec b n] power_down_eq_zero_iff[of prec b n] by auto lemma power_down_mono: "power_down prec a n \ power_down prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_down.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_down prec a j \ b * power_down prec b j" - by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_down_even_nonneg) + by (metis IH(1) IH(2) \even j\ lessI linear mult_mono mult_mono' mult_mono_nonpos_nonneg power_down_even_nonneg) then have "truncate_down (Suc prec) (a * power_down prec a j) \ truncate_down (Suc prec) (b * power_down prec b j)" by (auto intro!: truncate_down_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_down_simp) next assume [simp]: "odd j" have "power_down prec 0 (Suc (j div 2)) \ - power_down prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_down (Suc prec) ((power_down prec a (Suc (j div 2)))\<^sup>2) \ truncate_down (Suc prec) ((power_down prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_down_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_down_simp) qed qed simp qed lemma power_up_even_nonneg: "even n \ 0 \ power_up p x n" by (induct p x n rule: power_up.induct) (auto simp: power_up.simps simp del: odd_Suc_div_two intro!: ) lemma power_up_eq_zero_iff[simp]: "power_up prec b n = 0 \ b = 0 \ n \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) then show ?case using power_up_simp[of _ _ "x - 1"] by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff div2_less_self) qed lemma power_up_nonneg_iff[simp]: "power_up prec b n \ 0 \ even n \ b \ 0" proof (induction n arbitrary: b rule: less_induct) case (less x) show ?case using less(1)[of "x - 1" b] power_up_simp[of _ _ "x - 1"] by (cases x) (auto simp: algebra_split_simps zero_le_mult_iff) qed lemma power_up_neg_iff[simp]: "power_up prec b n < 0 \ b < 0 \ odd n" using power_up_nonneg_iff[of prec b n] by (auto simp del: power_up_nonneg_iff) lemma power_up_nonpos_iff[simp]: notes [simp del] = power_up_neg_iff power_up_eq_zero_iff shows "power_up prec b n \ 0 \ b < 0 \ odd n \ b = 0 \ n \ 0" using power_up_neg_iff[of prec b n] power_up_eq_zero_iff[of prec b n] by auto lemma power_up_mono: "power_up prec a n \ power_up prec b n" if "((0 \ a \ a \ b)\(odd n \ a \ b) \ (even n \ a \ 0 \ b \ a))" using that proof (induction n arbitrary: a b rule: less_induct) case (less i) show ?case proof (cases i) case j: (Suc j) note IH = less[unfolded j even_Suc not_not] note [simp del] = power_up.simps show ?thesis proof cases assume [simp]: "even j" have "a * power_up prec a j \ b * power_up prec b j" - by (smt IH(1) IH(2) \even j\ lessI mult_mono' mult_mono_nonpos_nonneg power_up_even_nonneg) + by (metis IH(1) IH(2) \even j\ lessI linear mult_mono mult_mono' mult_mono_nonpos_nonneg power_up_even_nonneg) then have "truncate_up prec (a * power_up prec a j) \ truncate_up prec (b * power_up prec b j)" by (auto intro!: truncate_up_mono simp: abs_le_square_iff[symmetric] abs_real_def) then show ?thesis unfolding j by (simp add: power_up_simp) next assume [simp]: "odd j" have "power_up prec 0 (Suc (j div 2)) \ - power_up prec b (Suc (j div 2))" if "b < 0" "even (j div 2)" apply (rule order_trans[where y=0]) using IH that by (auto simp: div2_less_self) then have "truncate_up prec ((power_up prec a (Suc (j div 2)))\<^sup>2) \ truncate_up prec ((power_up prec b (Suc (j div 2)))\<^sup>2)" using IH by (auto intro!: truncate_up_mono intro: order_trans[where y=0] simp: abs_le_square_iff[symmetric] abs_real_def div2_less_self) then show ?thesis unfolding j by (simp add: power_up_simp) qed qed simp qed subsection \Lemmas needed by Approximate\ lemma Float_num[simp]: "real_of_float (Float 1 0) = 1" "real_of_float (Float 1 1) = 2" "real_of_float (Float 1 2) = 4" "real_of_float (Float 1 (- 1)) = 1/2" "real_of_float (Float 1 (- 2)) = 1/4" "real_of_float (Float 1 (- 3)) = 1/8" "real_of_float (Float (- 1) 0) = -1" "real_of_float (Float (numeral n) 0) = numeral n" "real_of_float (Float (- numeral n) 0) = - numeral n" using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"] using powr_realpow[of 2 2] powr_realpow[of 2 3] using powr_minus[of "2::real" 1] powr_minus[of "2::real" 2] powr_minus[of "2::real" 3] by auto lemma real_of_Float_int[simp]: "real_of_float (Float n 0) = real n" by simp lemma float_zero[simp]: "real_of_float (Float 0 e) = 0" by simp lemma abs_div_2_less: "a \ 0 \ a \ -1 \ \(a::int) div 2\ < \a\" by arith lemma lapprox_rat: "real_of_float (lapprox_rat prec x y) \ real_of_int x / real_of_int y" by (simp add: lapprox_rat.rep_eq truncate_down) lemma mult_div_le: fixes a b :: int assumes "b > 0" shows "a \ b * (a div b)" proof - from minus_div_mult_eq_mod [symmetric, of a b] have "a = b * (a div b) + a mod b" by simp also have "\ \ b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign) using assms apply simp done finally show ?thesis by simp qed lemma lapprox_rat_nonneg: assumes "0 \ x" and "0 \ y" shows "0 \ real_of_float (lapprox_rat n x y)" using assms by transfer (simp add: truncate_down_nonneg) lemma rapprox_rat: "real_of_int x / real_of_int y \ real_of_float (rapprox_rat prec x y)" by transfer (simp add: truncate_up) lemma rapprox_rat_le1: assumes "0 \ x" "0 < y" "x \ y" shows "real_of_float (rapprox_rat n x y) \ 1" using assms by transfer (simp add: truncate_up_le1) lemma rapprox_rat_nonneg_nonpos: "0 \ x \ y \ 0 \ real_of_float (rapprox_rat n x y) \ 0" by transfer (simp add: truncate_up_nonpos divide_nonneg_nonpos) lemma rapprox_rat_nonpos_nonneg: "x \ 0 \ 0 \ y \ real_of_float (rapprox_rat n x y) \ 0" by transfer (simp add: truncate_up_nonpos divide_nonpos_nonneg) lemma real_divl: "real_divl prec x y \ x / y" by (simp add: real_divl_def truncate_down) lemma real_divr: "x / y \ real_divr prec x y" by (simp add: real_divr_def truncate_up) lemma float_divl: "real_of_float (float_divl prec x y) \ x / y" by transfer (rule real_divl) lemma real_divl_lower_bound: "0 \ x \ 0 \ y \ 0 \ real_divl prec x y" by (simp add: real_divl_def truncate_down_nonneg) lemma float_divl_lower_bound: "0 \ x \ 0 \ y \ 0 \ real_of_float (float_divl prec x y)" by transfer (rule real_divl_lower_bound) lemma exponent_1: "exponent 1 = 0" using exponent_float[of 1 0] by (simp add: one_float_def) lemma mantissa_1: "mantissa 1 = 1" using mantissa_float[of 1 0] by (simp add: one_float_def) lemma bitlen_1: "bitlen 1 = 1" by (simp add: bitlen_alt_def) lemma float_upper_bound: "x \ 2 powr (bitlen \mantissa x\ + exponent x)" proof (cases "x = 0") case True then show ?thesis by simp next case False then have "mantissa x \ 0" using mantissa_eq_zero_iff by auto have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent) also have "mantissa x \ \mantissa x\" by simp also have "\ \ 2 powr (bitlen \mantissa x\)" using bitlen_bounds[of "\mantissa x\"] bitlen_nonneg \mantissa x \ 0\ by (auto simp del: of_int_abs simp add: powr_int) finally show ?thesis by (simp add: powr_add) qed lemma real_divl_pos_less1_bound: assumes "0 < x" "x \ 1" shows "1 \ real_divl prec 1 x" using assms by (auto intro!: truncate_down_ge1 simp: real_divl_def) lemma float_divl_pos_less1_bound: "0 < real_of_float x \ real_of_float x \ 1 \ prec \ 1 \ 1 \ real_of_float (float_divl prec 1 x)" by transfer (rule real_divl_pos_less1_bound) lemma float_divr: "real_of_float x / real_of_float y \ real_of_float (float_divr prec x y)" by transfer (rule real_divr) lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x \ 1" shows "1 \ real_divr prec 1 x" proof - have "1 \ 1 / x" using \0 < x\ and \x \ 1\ by auto also have "\ \ real_divr prec 1 x" using real_divr[where x = 1 and y = x] by auto finally show ?thesis by auto qed lemma float_divr_pos_less1_lower_bound: "0 < x \ x \ 1 \ 1 \ float_divr prec 1 x" by transfer (rule real_divr_pos_less1_lower_bound) lemma real_divr_nonpos_pos_upper_bound: "x \ 0 \ 0 \ y \ real_divr prec x y \ 0" by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff) lemma float_divr_nonpos_pos_upper_bound: "real_of_float x \ 0 \ 0 \ real_of_float y \ real_of_float (float_divr prec x y) \ 0" by transfer (rule real_divr_nonpos_pos_upper_bound) lemma real_divr_nonneg_neg_upper_bound: "0 \ x \ y \ 0 \ real_divr prec x y \ 0" by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff) lemma float_divr_nonneg_neg_upper_bound: "0 \ real_of_float x \ real_of_float y \ 0 \ real_of_float (float_divr prec x y) \ 0" by transfer (rule real_divr_nonneg_neg_upper_bound) lemma Float_le_zero_iff: "Float a b \ 0 \ a \ 0" by (auto simp: zero_float_def mult_le_0_iff) lemma real_of_float_pprt[simp]: fixes a :: float shows "real_of_float (pprt a) = pprt (real_of_float a)" unfolding pprt_def sup_float_def max_def sup_real_def by auto lemma real_of_float_nprt[simp]: fixes a :: float shows "real_of_float (nprt a) = nprt (real_of_float a)" unfolding nprt_def inf_float_def min_def inf_real_def by auto context begin lift_definition int_floor_fl :: "float \ int" is floor . qualified lemma compute_int_floor_fl[code]: "int_floor_fl (Float m e) = (if 0 \ e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))" apply transfer apply (simp add: powr_int floor_divide_of_int_eq) apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power floor_of_int of_int_mult) done lift_definition floor_fl :: "float \ float" is "\x. real_of_int \x\" by simp qualified lemma compute_floor_fl[code]: "floor_fl (Float m e) = (if 0 \ e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" apply transfer apply (simp add: powr_int floor_divide_of_int_eq) apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power of_int_mult) done end lemma floor_fl: "real_of_float (floor_fl x) \ real_of_float x" by transfer simp lemma int_floor_fl: "real_of_int (int_floor_fl x) \ real_of_float x" by transfer simp lemma floor_pos_exp: "exponent (floor_fl x) \ 0" proof (cases "floor_fl x = 0") case True then show ?thesis by (simp add: floor_fl_def) next case False have eq: "floor_fl x = Float \real_of_float x\ 0" by transfer simp obtain i where "\real_of_float x\ = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i" by (rule denormalize_shift[OF eq False]) then show ?thesis by simp qed lemma compute_mantissa[code]: "mantissa (Float m e) = (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)" by (auto simp: mantissa_float Float.abs_eq simp flip: zero_float_def) lemma compute_exponent[code]: "exponent (Float m e) = (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)" by (auto simp: exponent_float Float.abs_eq simp flip: zero_float_def) lifting_update Float.float.lifting lifting_forget Float.float.lifting end diff --git a/src/HOL/Library/Interval_Float.thy b/src/HOL/Library/Interval_Float.thy --- a/src/HOL/Library/Interval_Float.thy +++ b/src/HOL/Library/Interval_Float.thy @@ -1,361 +1,361 @@ section \Approximate Operations on Intervals of Floating Point Numbers\ theory Interval_Float imports Interval Float begin definition mid :: "float interval \ float" where "mid i = (lower i + upper i) * Float 1 (-1)" lemma mid_in_interval: "mid i \\<^sub>i i" using lower_le_upper[of i] by (auto simp: mid_def set_of_eq powr_minus) lemma mid_le: "lower i \ mid i" "mid i \ upper i" using mid_in_interval by (auto simp: set_of_eq) definition centered :: "float interval \ float interval" where "centered i = i - interval_of (mid i)" definition "split_float_interval x = split_interval x ((lower x + upper x) * Float 1 (-1))" lemma split_float_intervalD: "split_float_interval X = (A, B) \ set_of X \ set_of A \ set_of B" by (auto dest!: split_intervalD simp: split_float_interval_def) lemma split_float_interval_bounds: shows lower_split_float_interval1: "lower (fst (split_float_interval X)) = lower X" and lower_split_float_interval2: "lower (snd (split_float_interval X)) = mid X" and upper_split_float_interval1: "upper (fst (split_float_interval X)) = mid X" and upper_split_float_interval2: "upper (snd (split_float_interval X)) = upper X" using mid_le[of X] by (auto simp: split_float_interval_def mid_def[symmetric] min_def max_def real_of_float_eq lower_split_interval1 lower_split_interval2 upper_split_interval1 upper_split_interval2) lemmas float_round_down_le[intro] = order_trans[OF float_round_down] and float_round_up_ge[intro] = order_trans[OF _ float_round_up] text \TODO: many of the lemmas should move to theories Float or Approximation (the latter should be based on type @{type interval}.\ subsection "Intervals with Floating Point Bounds" context includes interval.lifting begin lift_definition round_interval :: "nat \ float interval \ float interval" is "\p. \(l, u). (float_round_down p l, float_round_up p u)" by (auto simp: intro!: float_round_down_le float_round_up_le) lemma lower_round_ivl[simp]: "lower (round_interval p x) = float_round_down p (lower x)" by transfer auto lemma upper_round_ivl[simp]: "upper (round_interval p x) = float_round_up p (upper x)" by transfer auto lemma round_ivl_correct: "set_of A \ set_of (round_interval prec A)" by (auto simp: set_of_eq float_round_down_le float_round_up_le) lift_definition truncate_ivl :: "nat \ real interval \ real interval" is "\p. \(l, u). (truncate_down p l, truncate_up p u)" by (auto intro!: truncate_down_le truncate_up_le) lemma lower_truncate_ivl[simp]: "lower (truncate_ivl p x) = truncate_down p (lower x)" by transfer auto lemma upper_truncate_ivl[simp]: "upper (truncate_ivl p x) = truncate_up p (upper x)" by transfer auto lemma truncate_ivl_correct: "set_of A \ set_of (truncate_ivl prec A)" by (auto simp: set_of_eq intro!: truncate_down_le truncate_up_le) lift_definition real_interval::"float interval \ real interval" is "\(l, u). (real_of_float l, real_of_float u)" by auto lemma lower_real_interval[simp]: "lower (real_interval x) = lower x" by transfer auto lemma upper_real_interval[simp]: "upper (real_interval x) = upper x" by transfer auto definition "set_of' x = (case x of None \ UNIV | Some i \ set_of (real_interval i))" lemma real_interval_min_interval[simp]: "real_interval (min_interval a b) = min_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_min) lemma real_interval_max_interval[simp]: "real_interval (max_interval a b) = max_interval (real_interval a) (real_interval b)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max) lemma in_intervalI: "x \\<^sub>i X" if "lower X \ x" "x \ upper X" using that by (auto simp: set_of_eq) abbreviation in_real_interval ("(_/ \\<^sub>r _)" [51, 51] 50) where "x \\<^sub>r X \ x \\<^sub>i real_interval X" lemma in_real_intervalI: "x \\<^sub>r X" if "lower X \ x" "x \ upper X" for x::real and X::"float interval" using that by (intro in_intervalI) auto subsection \intros for \real_interval\\ lemma in_round_intervalI: "x \\<^sub>r A \ x \\<^sub>r (round_interval prec A)" by (auto simp: set_of_eq float_round_down_le float_round_up_le) lemma zero_in_float_intervalI: "0 \\<^sub>r 0" by (auto simp: set_of_eq) lemma plus_in_float_intervalI: "a + b \\<^sub>r A + B" if "a \\<^sub>r A" "b \\<^sub>r B" using that by (auto simp: set_of_eq) lemma minus_in_float_intervalI: "a - b \\<^sub>r A - B" if "a \\<^sub>r A" "b \\<^sub>r B" using that by (auto simp: set_of_eq) lemma uminus_in_float_intervalI: "-a \\<^sub>r -A" if "a \\<^sub>r A" using that by (auto simp: set_of_eq) lemma real_interval_times: "real_interval (A * B) = real_interval A * real_interval B" by (auto simp: interval_eq_iff lower_times upper_times min_def max_def) lemma times_in_float_intervalI: "a * b \\<^sub>r A * B" if "a \\<^sub>r A" "b \\<^sub>r B" using times_in_intervalI[OF that] by (auto simp: real_interval_times) lemma real_interval_abs: "real_interval (abs_interval A) = abs_interval (real_interval A)" by (auto simp: interval_eq_iff min_def max_def) lemma abs_in_float_intervalI: "abs a \\<^sub>r abs_interval A" if "a \\<^sub>r A" by (auto simp: set_of_abs_interval real_interval_abs intro!: imageI that) lemma interval_of[intro,simp]: "x \\<^sub>r interval_of x" by (auto simp: set_of_eq) lemma split_float_interval_realD: "split_float_interval X = (A, B) \ x \\<^sub>r X \ x \\<^sub>r A \ x \\<^sub>r B" by (auto simp: set_of_eq prod_eq_iff split_float_interval_bounds) subsection \bounds for lists\ lemma lower_Interval: "lower (Interval x) = fst x" and upper_Interval: "upper (Interval x) = snd x" if "fst x \ snd x" using that by (auto simp: lower_def upper_def Interval_inverse split_beta') definition all_in_i :: "'a::preorder list \ 'a interval list \ bool" (infix "(all'_in\<^sub>i)" 50) where "x all_in\<^sub>i I = (length x = length I \ (\i < length I. x ! i \\<^sub>i I ! i))" definition all_in :: "real list \ float interval list \ bool" (infix "(all'_in)" 50) where "x all_in I = (length x = length I \ (\i < length I. x ! i \\<^sub>r I ! i))" definition all_subset :: "'a::order interval list \ 'a interval list \ bool" (infix "(all'_subset)" 50) where "I all_subset J = (length I = length J \ (\i < length I. set_of (I!i) \ set_of (J!i)))" lemmas [simp] = all_in_def all_subset_def lemma all_subsetD: assumes "I all_subset J" assumes "x all_in I" shows "x all_in J" using assms by (auto simp: set_of_eq; fastforce) lemma round_interval_mono: "set_of (round_interval prec X) \ set_of (round_interval prec Y)" if "set_of X \ set_of Y" using that by transfer (auto simp: float_round_down.rep_eq float_round_up.rep_eq truncate_down_mono truncate_up_mono) lemma Ivl_simps[simp]: "lower (Ivl a b) = min a b" "upper (Ivl a b) = b" subgoal by transfer simp subgoal by transfer simp done lemma set_of_subset_iff: "set_of X \ set_of Y \ lower Y \ lower X \ upper X \ upper Y" for X Y::"'a::linorder interval" by (auto simp: set_of_eq subset_iff) lemma set_of_subset_iff': "set_of a \ set_of (b :: 'a :: linorder interval) \ a \ b" unfolding less_eq_interval_def set_of_subset_iff .. lemma bounds_of_interval_eq_lower_upper: "bounds_of_interval ivl = (lower ivl, upper ivl)" if "lower ivl \ upper ivl" using that by (auto simp: lower.rep_eq upper.rep_eq) lemma real_interval_Ivl: "real_interval (Ivl a b) = Ivl a b" by transfer (auto simp: min_def) lemma set_of_mul_contains_real_zero: "0 \\<^sub>r (A * B)" if "0 \\<^sub>r A \ 0 \\<^sub>r B" using that set_of_mul_contains_zero[of A B] by (auto simp: set_of_eq) fun subdivide_interval :: "nat \ float interval \ float interval list" where "subdivide_interval 0 I = [I]" | "subdivide_interval (Suc n) I = ( let m = mid I in (subdivide_interval n (Ivl (lower I) m)) @ (subdivide_interval n (Ivl m (upper I))) )" lemma subdivide_interval_length: shows "length (subdivide_interval n I) = 2^n" by(induction n arbitrary: I, simp_all add: Let_def) lemma lower_le_mid: "lower x \ mid x" "real_of_float (lower x) \ mid x" and mid_le_upper: "mid x \ upper x" "real_of_float (mid x) \ upper x" unfolding mid_def subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) subgoal by transfer (auto simp: powr_neg_one) done lemma subdivide_interval_correct: "list_ex (\i. x \\<^sub>r i) (subdivide_interval n I)" if "x \\<^sub>r I" for x::real using that proof(induction n arbitrary: x I) case 0 then show ?case by simp next case (Suc n) from \x \\<^sub>r I\ consider "x \\<^sub>r Ivl (lower I) (mid I)" | "x \\<^sub>r Ivl (mid I) (upper I)" by (cases "x \ real_of_float (mid I)") (auto simp: set_of_eq min_def lower_le_mid mid_le_upper) from this[case_names lower upper] show ?case by cases (use Suc.IH in \auto simp: Let_def\) qed fun interval_list_union :: "'a::lattice interval list \ 'a interval" where "interval_list_union [] = undefined" | "interval_list_union [I] = I" | "interval_list_union (I#Is) = sup I (interval_list_union Is)" lemma interval_list_union_correct: assumes "S \ []" assumes "i < length S" shows "set_of (S!i) \ set_of (interval_list_union S)" using assms proof(induction S arbitrary: i) case (Cons a S i) thus ?case proof(cases S) fix b S' assume "S = b # S'" hence "S \ []" by simp show ?thesis proof(cases i) case 0 show ?thesis apply(cases S) using interval_union_mono1 by (auto simp add: 0) next case (Suc i_prev) hence "i_prev < length S" using Cons(3) by simp from Cons(1)[OF \S \ []\ this] Cons(1) have "set_of ((a # S) ! i) \ set_of (interval_list_union S)" by (simp add: \i = Suc i_prev\) also have "... \ set_of (interval_list_union (a # S))" using \S \ []\ apply(cases S) using interval_union_mono2 by auto finally show ?thesis . qed qed simp qed simp lemma split_domain_correct: fixes x :: "real list" assumes "x all_in I" assumes split_correct: "\x a I. x \\<^sub>r I \ list_ex (\i::float interval. x \\<^sub>r i) (split I)" shows "list_ex (\s. x all_in s) (split_domain split I)" using assms(1) proof(induction I arbitrary: x) case (Cons I Is x) have "x \ []" using Cons(2) by auto obtain x' xs where x_decomp: "x = x' # xs" using \x \ []\ list.exhaust by auto hence "x' \\<^sub>r I" "xs all_in Is" using Cons(2) by auto show ?case using Cons(1)[OF \xs all_in Is\] split_correct[OF \x' \\<^sub>r I\] apply (auto simp add: list_ex_iff set_of_eq) - by (smt length_Cons less_Suc_eq_0_disj nth_Cons_0 nth_Cons_Suc x_decomp) + by (smt (verit, ccfv_SIG) One_nat_def Suc_pred \x \ []\ le_simps(3) length_greater_0_conv length_tl linorder_not_less list.sel(3) neq0_conv nth_Cons' x_decomp) qed simp lift_definition(code_dt) inverse_float_interval::"nat \ float interval \ float interval option" is "\prec (l, u). if (0 < l \ u < 0) then Some (float_divl prec 1 u, float_divr prec 1 l) else None" by (auto intro!: order_trans[OF float_divl] order_trans[OF _ float_divr] simp: divide_simps) lemma inverse_float_interval_eq_Some_conv: defines "one \ (1::float)" shows "inverse_float_interval p X = Some R \ (lower X > 0 \ upper X < 0) \ lower R = float_divl p one (upper X) \ upper R = float_divr p one (lower X)" by clarsimp (transfer fixing: one, force simp: one_def split: if_splits) lemma inverse_float_interval: "inverse ` set_of (real_interval X) \ set_of (real_interval Y)" if "inverse_float_interval p X = Some Y" using that apply (clarsimp simp: set_of_eq inverse_float_interval_eq_Some_conv) by (intro order_trans[OF float_divl] order_trans[OF _ float_divr] conjI) (auto simp: divide_simps) lemma inverse_float_intervalI: "x \\<^sub>r X \ inverse x \ set_of' (inverse_float_interval p X)" using inverse_float_interval[of p X] by (auto simp: set_of'_def split: option.splits) lemma inverse_float_interval_eqI: "inverse_float_interval p X = Some IVL \ x \\<^sub>r X \ inverse x \\<^sub>r IVL" using inverse_float_intervalI[of x X p] by (auto simp: set_of'_def) lemma real_interval_abs_interval[simp]: "real_interval (abs_interval x) = abs_interval (real_interval x)" by (auto simp: interval_eq_set_of_iff set_of_eq real_of_float_max real_of_float_min) lift_definition floor_float_interval::"float interval \ float interval" is "\(l, u). (floor_fl l, floor_fl u)" by (auto intro!: floor_mono simp: floor_fl.rep_eq) lemma lower_floor_float_interval[simp]: "lower (floor_float_interval x) = floor_fl (lower x)" by transfer auto lemma upper_floor_float_interval[simp]: "upper (floor_float_interval x) = floor_fl (upper x)" by transfer auto lemma floor_float_intervalI: "\x\ \\<^sub>r floor_float_interval X" if "x \\<^sub>r X" using that by (auto simp: set_of_eq floor_fl_def floor_mono) end subsection \constants for code generation\ definition lowerF::"float interval \ float" where "lowerF = lower" definition upperF::"float interval \ float" where "upperF = upper" end \ No newline at end of file diff --git a/src/HOL/Library/Poly_Mapping.thy b/src/HOL/Library/Poly_Mapping.thy --- a/src/HOL/Library/Poly_Mapping.thy +++ b/src/HOL/Library/Poly_Mapping.thy @@ -1,1878 +1,1878 @@ (* Author: Andreas Lochbihler, ETH Zurich Author: Florian Haftmann, TU Muenchen with some material ported from HOL Light by LCP *) section \Polynomial mapping: combination of almost everywhere zero functions with an algebraic view\ theory Poly_Mapping imports Groups_Big_Fun Fun_Lexorder More_List begin subsection \Preliminary: auxiliary operations for \emph{almost everywhere zero}\ text \ A central notion for polynomials are functions being \emph{almost everywhere zero}. For these we provide some auxiliary definitions and lemmas. \ lemma finite_mult_not_eq_zero_leftI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. g a * f a \ 0}" proof - have "{a. g a * f a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_rightI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. f a * g a \ 0}" proof - have "{a. f a * g a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_prodI: fixes f g :: "'a \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {(a, b). f a * g b \ 0}" proof - from assms have "finite (?F \ ?G)" by blast then have "finite {(a, b). f a \ 0 \ g b \ 0}" by simp then show ?thesis by (rule rev_finite_subset) auto qed lemma finite_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {a + b | a b. f a \ 0 \ g b \ 0}" (is "finite ?FG") proof - from assms have "finite (?F \ ?G)" by (simp add: finite_cartesian_product_iff) then have "finite (case_prod plus ` (?F \ ?G))" by (rule finite_imageI) also have "case_prod plus ` (?F \ ?G) = ?FG" by auto finally show ?thesis by simp qed lemma finite_mult_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" assumes "finite {b. g b \ 0}" shows "finite {a + b | a b. f a * g b \ 0}" proof - from assms have "finite {a + b | a b. f a \ 0 \ g b \ 0}" by (rule finite_not_eq_zero_sumI) then show ?thesis by (rule rev_finite_subset) (auto dest: mult_not_zero) qed lemma finite_Sum_any_not_eq_zero_weakenI: assumes "finite {a. \b. f a b \ 0}" shows "finite {a. Sum_any (f a) \ 0}" proof - have "{a. Sum_any (f a) \ 0} \ {a. \b. f a b \ 0}" by (auto elim: Sum_any.not_neutral_obtains_not_neutral) then show ?thesis using assms by (rule finite_subset) qed context zero begin definition "when" :: "'a \ bool \ 'a" (infixl "when" 20) where "(a when P) = (if P then a else 0)" text \ Case distinctions always complicate matters, particularly when nested. The @{const "when"} operation allows to minimise these if @{term 0} is the false-case value and makes proof obligations much more readable. \ lemma "when" [simp]: "P \ (a when P) = a" "\ P \ (a when P) = 0" by (simp_all add: when_def) lemma when_simps [simp]: "(a when True) = a" "(a when False) = 0" by simp_all lemma when_cong: assumes "P \ Q" and "Q \ a = b" shows "(a when P) = (b when Q)" using assms by (simp add: when_def) lemma zero_when [simp]: "(0 when P) = 0" by (simp add: when_def) lemma when_when: "(a when P when Q) = (a when P \ Q)" by (cases Q) simp_all lemma when_commute: "(a when Q when P) = (a when P when Q)" by (simp add: when_when conj_commute) lemma when_neq_zero [simp]: "(a when P) \ 0 \ P \ a \ 0" by (cases P) simp_all end context monoid_add begin lemma when_add_distrib: "(a + b when P) = (a when P) + (b when P)" by (simp add: when_def) end context semiring_1 begin lemma zero_power_eq: "0 ^ n = (1 when n = 0)" by (simp add: power_0_left) end context comm_monoid_add begin lemma Sum_any_when_equal [simp]: "(\a. (f a when a = b)) = f b" by (simp add: when_def) lemma Sum_any_when_equal' [simp]: "(\a. (f a when b = a)) = f b" by (simp add: when_def) lemma Sum_any_when_independent: "(\a. g a when P) = ((\a. g a) when P)" by (cases P) simp_all lemma Sum_any_when_dependent_prod_right: "(\(a, b). g a when b = h a) = (\a. g a)" proof - have "inj_on (\a. (a, h a)) {a. g a \ 0}" by (rule inj_onI) auto then show ?thesis unfolding Sum_any.expand_set by (rule sum.reindex_cong) auto qed lemma Sum_any_when_dependent_prod_left: "(\(a, b). g b when a = h b) = (\b. g b)" proof - have "(\(a, b). g b when a = h b) = (\(b, a). g b when a = h b)" by (rule Sum_any.reindex_cong [of prod.swap]) (simp_all add: fun_eq_iff) then show ?thesis by (simp add: Sum_any_when_dependent_prod_right) qed end context cancel_comm_monoid_add begin lemma when_diff_distrib: "(a - b when P) = (a when P) - (b when P)" by (simp add: when_def) end context group_add begin lemma when_uminus_distrib: "(- a when P) = - (a when P)" by (simp add: when_def) end context mult_zero begin lemma mult_when: "a * (b when P) = (a * b when P)" by (cases P) simp_all lemma when_mult: "(a when P) * b = (a * b when P)" by (cases P) simp_all end subsection \Type definition\ text \ The following type is of central importance: \ typedef (overloaded) ('a, 'b) poly_mapping ("(_ \\<^sub>0 /_)" [1, 0] 0) = "{f :: 'a \ 'b::zero. finite {x. f x \ 0}}" morphisms lookup Abs_poly_mapping proof - have "(\_::'a. (0 :: 'b)) \ ?poly_mapping" by simp then show ?thesis by (blast intro!: exI) qed declare lookup_inverse [simp] declare lookup_inject [simp] lemma lookup_Abs_poly_mapping [simp]: "finite {x. f x \ 0} \ lookup (Abs_poly_mapping f) = f" using Abs_poly_mapping_inverse [of f] by simp lemma finite_lookup [simp]: "finite {k. lookup f k \ 0}" using poly_mapping.lookup [of f] by simp lemma finite_lookup_nat [simp]: fixes f :: "'a \\<^sub>0 nat" shows "finite {k. 0 < lookup f k}" using poly_mapping.lookup [of f] by simp lemma poly_mapping_eqI: assumes "\k. lookup f k = lookup g k" shows "f = g" using assms unfolding poly_mapping.lookup_inject [symmetric] by blast lemma poly_mapping_eq_iff: "a = b \ lookup a = lookup b" by auto text \ We model the universe of functions being \emph{almost everywhere zero} by means of a separate type @{typ "('a, 'b) poly_mapping"}. For convenience we provide a suggestive infix syntax which is a variant of the usual function space syntax. Conversion between both types happens through the morphisms \begin{quote} @{term_type lookup} \end{quote} \begin{quote} @{term_type Abs_poly_mapping} \end{quote} satisfying \begin{quote} @{thm lookup_inverse} \end{quote} \begin{quote} @{thm lookup_Abs_poly_mapping} \end{quote} Luckily, we have rarely to deal with those low-level morphisms explicitly but rely on Isabelle's \emph{lifting} package with its method \transfer\ and its specification tool \lift_definition\. \ setup_lifting type_definition_poly_mapping code_datatype Abs_poly_mapping\\FIXME? workaround for preventing \code_abstype\ setup\ text \ @{typ "'a \\<^sub>0 'b"} serves distinctive purposes: \begin{enumerate} \item A clever nesting as @{typ "(nat \\<^sub>0 nat) \\<^sub>0 'a"} later in theory \MPoly\ gives a suitable representation type for polynomials \emph{almost for free}: Interpreting @{typ "nat \\<^sub>0 nat"} as a mapping from variable identifiers to exponents yields monomials, and the whole type maps monomials to coefficients. Lets call this the \emph{ultimate interpretation}. \item A further more specialised type isomorphic to @{typ "nat \\<^sub>0 'a"} is apt to direct implementation using code generation \cite{Haftmann-Nipkow:2010:code}. \end{enumerate} Note that despite the names \emph{mapping} and \emph{lookup} suggest something implementation-near, it is best to keep @{typ "'a \\<^sub>0 'b"} as an abstract \emph{algebraic} type providing operations like \emph{addition}, \emph{multiplication} without any notion of key-order, data structures etc. This implementations-specific notions are easily introduced later for particular implementations but do not provide any gain for specifying logical properties of polynomials. \ subsection \Additive structure\ text \ The additive structure covers the usual operations \0\, \+\ and (unary and binary) \-\. Recalling the ultimate interpretation, it is obvious that these have just lift the corresponding operations on values to mappings. Isabelle has a rich hierarchy of algebraic type classes, and in such situations of pointwise lifting a typical pattern is to have instantiations for a considerable number of type classes. The operations themselves are specified using \lift_definition\, where the proofs of the \emph{almost everywhere zero} property can be significantly involved. The @{const lookup} operation is supposed to be usable explicitly (unless in other situations where the morphisms between types are somehow internal to the \emph{lifting} package). Hence it is good style to provide explicit rewrite rules how @{const lookup} acts on operations immediately. \ instantiation poly_mapping :: (type, zero) zero begin lift_definition zero_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 0" by simp instance .. end lemma Abs_poly_mapping [simp]: "Abs_poly_mapping (\k. 0) = 0" by (simp add: zero_poly_mapping.abs_eq) lemma lookup_zero [simp]: "lookup 0 k = 0" by transfer rule instantiation poly_mapping :: (type, monoid_add) monoid_add begin lift_definition plus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k + f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x + f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x + f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_add: "lookup (f + g) k = lookup f k + lookup g k" by transfer rule instance poly_mapping :: (type, comm_monoid_add) comm_monoid_add by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ lemma lookup_sum: "lookup (sum pp X) i = sum (\x. lookup (pp x) i) X" by (induction rule: infinite_finite_induct) (auto simp: lookup_add) (*instance poly_mapping :: (type, "{monoid_add, cancel_semigroup_add}") cancel_semigroup_add by intro_classes (transfer, simp add: fun_eq_iff)+*) instantiation poly_mapping :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k - f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x - f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x - f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff diff_diff_add)+ end instantiation poly_mapping :: (type, ab_group_add) ab_group_add begin lift_definition uminus_poly_mapping :: "('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is uminus by simp instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_uminus [simp]: "lookup (- f) k = - lookup f k" by transfer simp lemma lookup_minus: "lookup (f - g) k = lookup f k - lookup g k" by transfer rule subsection \Multiplicative structure\ instantiation poly_mapping :: (zero, zero_neq_one) zero_neq_one begin lift_definition one_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 1 when k = 0" by simp instance by intro_classes (transfer, simp add: fun_eq_iff) end lemma lookup_one: "lookup 1 k = (1 when k = 0)" by transfer rule lemma lookup_one_zero [simp]: "lookup 1 0 = 1" by transfer simp definition prod_fun :: "('a \ 'b) \ ('a \ 'b) \ 'a::monoid_add \ 'b::semiring_0" where "prod_fun f1 f2 k = (\l. f1 l * (\q. (f2 q when k = l + q)))" lemma prod_fun_unfold_prod: fixes f g :: "'a :: monoid_add \ 'b::semiring_0" assumes fin_f: "finite {a. f a \ 0}" assumes fin_g: "finite {b. g b \ 0}" shows "prod_fun f g k = (\(a, b). f a * g b when k = a + b)" proof - let ?C = "{a. f a \ 0} \ {b. g b \ 0}" from fin_f fin_g have "finite ?C" by blast moreover have "{a. \b. (f a * g b when k = a + b) \ 0} \ {b. \a. (f a * g b when k = a + b) \ 0} \ {a. f a \ 0} \ {b. g b \ 0}" by auto ultimately show ?thesis using fin_g by (auto simp add: prod_fun_def Sum_any.cartesian_product [of "{a. f a \ 0} \ {b. g b \ 0}"] Sum_any_right_distrib mult_when) qed lemma finite_prod_fun: fixes f1 f2 :: "'a :: monoid_add \ 'b :: semiring_0" assumes fin1: "finite {l. f1 l \ 0}" and fin2: "finite {q. f2 q \ 0}" shows "finite {k. prod_fun f1 f2 k \ 0}" proof - have *: "finite {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" using assms by simp { fix k l have "{q. (f2 q when k = l + q) \ 0} \ {q. f2 q \ 0 \ k = l + q}" by auto with fin2 have "sum f2 {q. f2 q \ 0 \ k = l + q} = (\q. (f2 q when k = l + q))" by (simp add: Sum_any.expand_superset [of "{q. f2 q \ 0 \ k = l + q}"]) } note aux = this have "{k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0} \ {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto elim!: Sum_any.not_neutral_obtains_not_neutral) also have "\ \ {k. (\l. f1 l \ 0 \ sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto dest: mult_not_zero) also have "\ \ {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" by (auto elim!: sum.not_neutral_contains_not_neutral) finally have "finite {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0}" using * by (rule finite_subset) with aux have "finite {k. (\l. f1 l * (\q. (f2 q when k = l + q))) \ 0}" by simp with fin2 show ?thesis by (simp add: prod_fun_def) qed instantiation poly_mapping :: (monoid_add, semiring_0) semiring_0 begin lift_definition times_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is prod_fun by(rule finite_prod_fun) instance proof fix a b c :: "'a \\<^sub>0 'b" show "a * b * c = a * (b * c)" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" (is "finite ?F") assume fin_g: "finite {b. g b \ 0}" (is "finite ?G") assume fin_h: "finite {c. h c \ 0}" (is "finite ?H") from fin_f fin_g have fin_fg: "finite {(a, b). f a * g b \ 0}" (is "finite ?FG") by (rule finite_mult_not_eq_zero_prodI) from fin_g fin_h have fin_gh: "finite {(b, c). g b * h c \ 0}" (is "finite ?GH") by (rule finite_mult_not_eq_zero_prodI) from fin_f fin_g have fin_fg': "finite {a + b | a b. f a * g b \ 0}" (is "finite ?FG'") by (rule finite_mult_not_eq_zero_sumI) then have fin_fg'': "finite {d. (\(a, b). f a * g b when d = a + b) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) from fin_g fin_h have fin_gh': "finite {b + c | b c. g b * h c \ 0}" (is "finite ?GH'") by (rule finite_mult_not_eq_zero_sumI) then have fin_gh'': "finite {d. (\(b, c). g b * h c when d = b + c) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) show "prod_fun (prod_fun f g) h = prod_fun f (prod_fun g h)" (is "?lhs = ?rhs") proof fix k from fin_f fin_g fin_h fin_fg'' have "?lhs k = (\(ab, c). (\(a, b). f a * g b when ab = a + b) * h c when k = ab + c)" by (simp add: prod_fun_unfold_prod) also have "\ = (\(ab, c). (\(a, b). f a * g b * h c when k = ab + c when ab = a + b))" apply (subst Sum_any_left_distrib) using fin_fg apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def conj_commute) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)" apply (subst Sum_any.cartesian_product2 [of "(?FG' \ ?H) \ ?FG"]) apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)" by (rule Sum_any.cong) (simp add: split_def when_def) also have "\ = (\(ab, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when ab = (case cab of (c, a, b) \ a + b))" by (simp add: split_def) also have "\ = (\(c, a, b). f a * g b * h c when k = a + b + c)" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when bc = (case cab of (c, a, b) \ b + c))" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, c, a, b). f a * g b * h c when k = a + b + c when bc = b + c)" by (simp add: split_def) also have "\ = (\(bc, c, a, b). f a * g b * h c when bc = b + c when k = a + bc)" by (rule Sum_any.cong) (simp add: split_def when_def ac_simps) also have "\ = (\(a, bc, b, c). f a * g b * h c when bc = b + c when k = a + bc)" proof - have "bij (\(a, d, b, c). (d, c, a, b))" by (auto intro!: bijI injI surjI [of _ "\(d, c, a, b). (a, d, b, c)"] simp add: split_def) then show ?thesis by (rule Sum_any.reindex_cong) auto qed also have "\ = (\(a, bc). (\(b, c). f a * g b * h c when bc = b + c when k = a + bc))" apply (subst Sum_any.cartesian_product2 [of "(?F \ ?GH') \ ?GH"]) apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero) done also have "\ = (\(a, bc). f a * (\(b, c). g b * h c when bc = b + c) when k = a + bc)" apply (subst Sum_any_right_distrib) using fin_gh apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def ac_simps) done also from fin_f fin_g fin_h fin_gh'' have "\ = ?rhs k" by (simp add: prod_fun_unfold_prod) finally show "?lhs k = ?rhs k" . qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (rule ext) apply (auto simp add: prod_fun_def algebra_simps) apply (subst Sum_any.distrib) using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI) done qed show "a * (b + c) = a * b + a * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun f (\k. g k + h k) = (\k. prod_fun f g k + prod_fun f h k)" apply (rule ext) apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) apply (auto intro: fin_g fin_h) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_f apply (rule finite_mult_not_eq_zero_rightI) done qed show "0 * a = 0" by transfer (simp add: prod_fun_def [abs_def]) show "a * 0 = 0" by transfer (simp add: prod_fun_def [abs_def]) qed end lemma lookup_mult: "lookup (f * g) k = (\l. lookup f l * (\q. lookup g q when k = l + q))" by transfer (simp add: prod_fun_def) instance poly_mapping :: (comm_monoid_add, comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a \\<^sub>0 'b" show "a * b = b * a" proof transfer fix f g :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" assume fin_g: "finite {b. g b \ 0}" show "prod_fun f g = prod_fun g f" proof fix k have fin1: "\l. finite {a. (f a when k = l + a) \ 0}" using fin_f by auto have fin2: "\l. finite {b. (g b when k = l + b) \ 0}" using fin_g by auto from fin_f fin_g have "finite {(a, b). f a \ 0 \ g b \ 0}" (is "finite ?AB") by simp show "prod_fun f g k = prod_fun g f k" apply (simp add: prod_fun_def) apply (subst Sum_any_right_distrib) apply (rule fin2) apply (subst Sum_any_right_distrib) apply (rule fin1) apply (subst Sum_any.swap [of ?AB]) apply (fact \finite ?AB\) apply (auto simp add: mult_when ac_simps) done qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps) apply (subst Sum_any.distrib) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_g apply (rule finite_mult_not_eq_zero_rightI) apply simp_all done qed qed instance poly_mapping :: (monoid_add, semiring_0_cancel) semiring_0_cancel .. instance poly_mapping :: (comm_monoid_add, comm_semiring_0_cancel) comm_semiring_0_cancel .. instance poly_mapping :: (monoid_add, semiring_1) semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def] when_mult) show "a * 1 = a" apply transfer apply (simp add: prod_fun_def [abs_def] Sum_any_right_distrib Sum_any_left_distrib mult_when) apply (subst when_commute) apply simp done qed instance poly_mapping :: (comm_monoid_add, comm_semiring_1) comm_semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def]) qed instance poly_mapping :: (monoid_add, semiring_1_cancel) semiring_1_cancel .. instance poly_mapping :: (monoid_add, ring) ring .. instance poly_mapping :: (comm_monoid_add, comm_ring) comm_ring .. instance poly_mapping :: (monoid_add, ring_1) ring_1 .. instance poly_mapping :: (comm_monoid_add, comm_ring_1) comm_ring_1 .. subsection \Single-point mappings\ lift_definition single :: "'a \ 'b \ 'a \\<^sub>0 'b::zero" is "\k v k'. (v when k = k')" by simp lemma inj_single [iff]: "inj (single k)" proof (rule injI, transfer) fix k :: 'b and a b :: "'a::zero" assume "(\k'. a when k = k') = (\k'. b when k = k')" then have "(\k'. a when k = k') k = (\k'. b when k = k') k" by (rule arg_cong) then show "a = b" by simp qed lemma lookup_single: "lookup (single k v) k' = (v when k = k')" by transfer rule lemma lookup_single_eq [simp]: "lookup (single k v) k = v" by transfer simp lemma lookup_single_not_eq: "k \ k' \ lookup (single k v) k' = 0" by transfer simp lemma single_zero [simp]: "single k 0 = 0" by transfer simp lemma single_one [simp]: "single 0 1 = 1" by transfer simp lemma single_add: "single k (a + b) = single k a + single k b" by transfer (simp add: fun_eq_iff when_add_distrib) lemma single_uminus: "single k (- a) = - single k a" by transfer (simp add: fun_eq_iff when_uminus_distrib) lemma single_diff: "single k (a - b) = single k a - single k b" by transfer (simp add: fun_eq_iff when_diff_distrib) lemma single_numeral [simp]: "single 0 (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_add single_zero single_one single_add) lemma lookup_numeral: "lookup (numeral n) k = (numeral n when k = 0)" proof - have "lookup (numeral n) k = lookup (single 0 (numeral n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma single_of_nat [simp]: "single 0 (of_nat n) = of_nat n" by (induct n) (simp_all add: single_add) lemma lookup_of_nat: "lookup (of_nat n) k = (of_nat n when k = 0)" proof - have "lookup (of_nat n) k = lookup (single 0 (of_nat n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma of_nat_single: "of_nat = single 0 \ of_nat" by (simp add: fun_eq_iff) lemma mult_single: "single k a * single l b = single (k + l) (a * b)" proof transfer fix k l :: 'a and a b :: 'b show "prod_fun (\k'. a when k = k') (\k'. b when l = k') = (\k'. a * b when k + l = k')" proof fix k' have "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\n. a * b when l = n when k' = k + n)" by (simp add: prod_fun_def Sum_any_right_distrib mult_when when_mult) also have "\ = (\n. a * b when k' = k + n when l = n)" by (simp add: when_when conj_commute) also have "\ = (a * b when k' = k + l)" by simp also have "\ = (a * b when k + l = k')" by (simp add: when_def) finally show "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\k'. a * b when k + l = k') k'" . qed qed instance poly_mapping :: (monoid_add, semiring_char_0) semiring_char_0 by intro_classes (auto intro: inj_compose inj_of_nat simp add: of_nat_single) instance poly_mapping :: (monoid_add, ring_char_0) ring_char_0 .. lemma single_of_int [simp]: "single 0 (of_int k) = of_int k" by (cases k) (simp_all add: single_diff single_uminus) lemma lookup_of_int: "lookup (of_int l) k = (of_int l when k = 0)" proof - have "lookup (of_int l) k = lookup (single 0 (of_int l)) k" by simp then show ?thesis unfolding lookup_single by simp qed subsection \Integral domains\ instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", semiring_no_zero_divisors) semiring_no_zero_divisors text \The @{class "linorder"} constraint is a pragmatic device for the proof --- maybe it can be dropped\ proof fix f g :: "'a \\<^sub>0 'b" assume "f \ 0" and "g \ 0" then show "f * g \ 0" proof transfer fix f g :: "'a \ 'b" define F where "F = {a. f a \ 0}" moreover define G where "G = {a. g a \ 0}" ultimately have [simp]: "\a. f a \ 0 \ a \ F" "\b. g b \ 0 \ b \ G" by simp_all assume "finite {a. f a \ 0}" then have [simp]: "finite F" by simp assume "finite {a. g a \ 0}" then have [simp]: "finite G" by simp assume "f \ (\a. 0)" then obtain a where "f a \ 0" by (auto simp add: fun_eq_iff) assume "g \ (\b. 0)" then obtain b where "g b \ 0" by (auto simp add: fun_eq_iff) from \f a \ 0\ and \g b \ 0\ have "F \ {}" and "G \ {}" by auto note Max_F = \finite F\ \F \ {}\ note Max_G = \finite G\ \G \ {}\ from Max_F and Max_G have [simp]: "Max F \ F" "Max G \ G" by auto from Max_F Max_G have [dest!]: "\a. a \ F \ a \ Max F" "\b. b \ G \ b \ Max G" by auto define q where "q = Max F + Max G" have "(\(a, b). f a * g b when q = a + b) = (\(a, b). f a * g b when q = a + b when a \ F \ b \ G)" by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr) also have "\ = (\(a, b). f a * g b when (Max F, Max G) = (a, b))" proof (rule Sum_any.cong) fix ab :: "'a \ 'a" obtain a b where [simp]: "ab = (a, b)" by (cases ab) simp_all have [dest!]: "a \ Max F \ Max F \ a \ a < Max F" "b \ Max G \ Max G \ b \ b < Max G" by auto show "(case ab of (a, b) \ f a * g b when q = a + b when a \ F \ b \ G) = (case ab of (a, b) \ f a * g b when (Max F, Max G) = (a, b))" by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"]) qed also have "\ = (\ab. (case ab of (a, b) \ f a * g b) when (Max F, Max G) = ab)" unfolding split_def when_def by auto also have "\ \ 0" by simp finally have "prod_fun f g q \ 0" by (simp add: prod_fun_unfold_prod) then show "prod_fun f g \ (\k. 0)" by (auto simp add: fun_eq_iff) qed qed instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_no_zero_divisors) ring_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", idom) idom .. subsection \Mapping order\ instantiation poly_mapping :: (linorder, "{zero, linorder}") linorder begin lift_definition less_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is less_fun . lift_definition less_eq_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is "\f g. less_fun f g \ f = g" . instance proof (rule class.Orderings.linorder.of_class.intro) show "class.linorder (less_eq :: (_ \\<^sub>0 _) \ _) less" proof (rule linorder_strictI, rule order_strictI) fix f g h :: "'a \\<^sub>0 'b" show "f \ g \ f < g \ f = g" by transfer (rule refl) show "\ f < f" by transfer (rule less_fun_irrefl) show "f < g \ f = g \ g < f" proof transfer fix f g :: "'a \ 'b" assume "finite {k. f k \ 0}" and "finite {k. g k \ 0}" then have "finite ({k. f k \ 0} \ {k. g k \ 0})" by simp moreover have "{k. f k \ g k} \ {k. f k \ 0} \ {k. g k \ 0}" by auto ultimately have "finite {k. f k \ g k}" by (rule rev_finite_subset) then show "less_fun f g \ f = g \ less_fun g f" by (rule less_fun_trichotomy) qed assume "f < g" then show "\ g < f" by transfer (rule less_fun_asym) note \f < g\ moreover assume "g < h" ultimately show "f < h" by transfer (rule less_fun_trans) qed qed end instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, linorder}") ordered_ab_semigroup_add proof (intro_classes, transfer) fix f g h :: "'a \ 'b" assume *: "less_fun f g \ f = g" { assume "less_fun f g" then obtain k where "f k < g k" "(\k'. k' < k \ f k' = g k')" by (blast elim!: less_funE) then have "h k + f k < h k + g k" "(\k'. k' < k \ h k' + f k' = h k' + g k')" by simp_all then have "less_fun (\k. h k + f k) (\k. h k + g k)" by (blast intro: less_funI) } with * show "less_fun (\k. h k + f k) (\k. h k + g k) \ (\k. h k + f k) = (\k. h k + g k)" by (auto simp add: fun_eq_iff) qed instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_comm_monoid_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_cancel_comm_monoid_add .. instance poly_mapping :: (linorder, linordered_ab_group_add) linordered_ab_group_add .. text \ For pragmatism we leave out the final elements in the hierarchy: @{class linordered_ring}, @{class linordered_ring_strict}, @{class linordered_idom}; remember that the order instance is a mere technical device, not a deeper algebraic property. \ subsection \Fundamental mapping notions\ lift_definition keys :: "('a \\<^sub>0 'b::zero) \ 'a set" is "\f. {k. f k \ 0}" . lift_definition range :: "('a \\<^sub>0 'b::zero) \ 'b set" is "\f :: 'a \ 'b. Set.range f - {0}" . lemma finite_keys [simp]: "finite (keys f)" by transfer lemma not_in_keys_iff_lookup_eq_zero: "k \ keys f \ lookup f k = 0" by transfer simp lemma lookup_not_eq_zero_eq_in_keys: "lookup f k \ 0 \ k \ keys f" by transfer simp lemma lookup_eq_zero_in_keys_contradict [dest]: "lookup f k = 0 \ \ k \ keys f" by (simp add: not_in_keys_iff_lookup_eq_zero) lemma finite_range [simp]: "finite (Poly_Mapping.range p)" proof transfer fix f :: "'b \ 'a" assume *: "finite {x. f x \ 0}" have "Set.range f - {0} \ f ` {x. f x \ 0}" by auto thus "finite (Set.range f - {0})" by(rule finite_subset)(rule finite_imageI[OF *]) qed lemma in_keys_lookup_in_range [simp]: "k \ keys f \ lookup f k \ range f" by transfer simp lemma in_keys_iff: "x \ (keys s) = (lookup s x \ 0)" by (transfer, simp) lemma keys_zero [simp]: "keys 0 = {}" by transfer simp lemma range_zero [simp]: "range 0 = {}" by transfer auto lemma keys_add: "keys (f + g) \ keys f \ keys g" by transfer auto lemma keys_one [simp]: "keys 1 = {0}" by transfer simp lemma range_one [simp]: "range 1 = {1}" by transfer (auto simp add: when_def) lemma keys_single [simp]: "keys (single k v) = (if v = 0 then {} else {k})" by transfer simp lemma range_single [simp]: "range (single k v) = (if v = 0 then {} else {v})" by transfer (auto simp add: when_def) lemma keys_mult: "keys (f * g) \ {a + b | a b. a \ keys f \ b \ keys g}" apply transfer apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral) apply blast done lemma setsum_keys_plus_distrib: assumes hom_0: "\k. f k 0 = 0" and hom_plus: "\k. k \ Poly_Mapping.keys p \ Poly_Mapping.keys q \ f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k) = f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k)" shows "(\k\Poly_Mapping.keys (p + q). f k (Poly_Mapping.lookup (p + q) k)) = (\k\Poly_Mapping.keys p. f k (Poly_Mapping.lookup p k)) + (\k\Poly_Mapping.keys q. f k (Poly_Mapping.lookup q k))" (is "?lhs = ?p + ?q") proof - let ?A = "Poly_Mapping.keys p \ Poly_Mapping.keys q" have "?lhs = (\k\?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))" apply(rule sum.mono_neutral_cong_left) apply(simp_all add: Poly_Mapping.keys_add) apply(transfer fixing: f) apply(auto simp add: hom_0)[1] apply(transfer fixing: f) apply(auto simp add: hom_0)[1] done also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))" by(rule sum.cong)(simp_all add: hom_plus) also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k)) + (\k\?A. f k (Poly_Mapping.lookup q k))" (is "_ = ?p' + ?q'") by(simp add: sum.distrib) also have "?p' = ?p" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) also have "?q' = ?q" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) finally show ?thesis . qed subsection \Degree\ definition degree :: "(nat \\<^sub>0 'a::zero) \ nat" where "degree f = Max (insert 0 (Suc ` keys f))" lemma degree_zero [simp]: "degree 0 = 0" unfolding degree_def by transfer simp lemma degree_one [simp]: "degree 1 = 1" unfolding degree_def by transfer simp lemma degree_single_zero [simp]: "degree (single k 0) = 0" unfolding degree_def by transfer simp lemma degree_single_not_zero [simp]: "v \ 0 \ degree (single k v) = Suc k" unfolding degree_def by transfer simp lemma degree_zero_iff [simp]: "degree f = 0 \ f = 0" unfolding degree_def proof transfer fix f :: "nat \ 'a" assume "finite {n. f n \ 0}" then have fin: "finite (insert 0 (Suc ` {n. f n \ 0}))" by auto show "Max (insert 0 (Suc ` {n. f n \ 0})) = 0 \ f = (\n. 0)" (is "?P \ ?Q") proof assume ?P have "{n. f n \ 0} = {}" proof (rule ccontr) assume "{n. f n \ 0} \ {}" then obtain n where "n \ {n. f n \ 0}" by blast then have "{n. f n \ 0} = insert n {n. f n \ 0}" by auto then have "Suc ` {n. f n \ 0} = insert (Suc n) (Suc ` {n. f n \ 0})" by auto with \?P\ have "Max (insert 0 (insert (Suc n) (Suc ` {n. f n \ 0}))) = 0" by simp then have "Max (insert (Suc n) (insert 0 (Suc ` {n. f n \ 0}))) = 0" by (simp add: insert_commute) with fin have "max (Suc n) (Max (insert 0 (Suc ` {n. f n \ 0}))) = 0" by simp then show False by simp qed then show ?Q by (simp add: fun_eq_iff) next assume ?Q then show ?P by simp qed qed lemma degree_greater_zero_in_keys: assumes "0 < degree f" shows "degree f - 1 \ keys f" proof - from assms have "keys f \ {}" by (auto simp add: degree_def) then show ?thesis unfolding degree_def by (simp add: mono_Max_commute [symmetric] mono_Suc) qed lemma in_keys_less_degree: "n \ keys f \ n < degree f" unfolding degree_def by transfer (auto simp add: Max_gr_iff) lemma beyond_degree_lookup_zero: "degree f \ n \ lookup f n = 0" unfolding degree_def by transfer auto lemma degree_add: "degree (f + g) \ max (degree f) (Poly_Mapping.degree g)" unfolding degree_def proof transfer fix f g :: "nat \ 'a" assume f: "finite {x. f x \ 0}" assume g: "finite {x. g x \ 0}" let ?f = "Max (insert 0 (Suc ` {k. f k \ 0}))" let ?g = "Max (insert 0 (Suc ` {k. g k \ 0}))" have "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ Max (insert 0 (Suc ` ({k. f k \ 0} \ {k. g k \ 0})))" by (rule Max.subset_imp) (insert f g, auto) also have "\ = max ?f ?g" using f g by (simp_all add: image_Un Max_Un [symmetric]) finally show "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ max (Max (insert 0 (Suc ` {k. f k \ 0}))) (Max (insert 0 (Suc ` {k. g k \ 0})))" . qed lemma sorted_list_of_set_keys: "sorted_list_of_set (keys f) = filter (\k. k \ keys f) [0..Inductive structure\ lift_definition update :: "'a \ 'b \ ('a \\<^sub>0 'b::zero) \ 'a \\<^sub>0 'b" is "\k v f. f(k := v)" proof - fix f :: "'a \ 'b" and k' v assume "finite {k. f k \ 0}" then have "finite (insert k' {k. f k \ 0})" by simp then show "finite {k. (f(k' := v)) k \ 0}" by (rule rev_finite_subset) auto qed lemma update_induct [case_names const update]: assumes const': "P 0" assumes update': "\f a b. a \ keys f \ b \ 0 \ P f \ P (update a b f)" shows "P f" proof - obtain g where "f = Abs_poly_mapping g" and "finite {a. g a \ 0}" by (cases f) simp_all define Q where "Q g = P (Abs_poly_mapping g)" for g from \finite {a. g a \ 0}\ have "Q g" proof (induct g rule: finite_update_induct) case const with const' Q_def show ?case by simp next case (update a b g) from \finite {a. g a \ 0}\ \g a = 0\ have "a \ keys (Abs_poly_mapping g)" by (simp add: Abs_poly_mapping_inverse keys.rep_eq) moreover note \b \ 0\ moreover from \Q g\ have "P (Abs_poly_mapping g)" by (simp add: Q_def) ultimately have "P (update a b (Abs_poly_mapping g))" by (rule update') also from \finite {a. g a \ 0}\ have "update a b (Abs_poly_mapping g) = Abs_poly_mapping (g(a := b))" by (simp add: update.abs_eq eq_onp_same_args) finally show ?case by (simp add: Q_def fun_upd_def) qed then show ?thesis by (simp add: Q_def \f = Abs_poly_mapping g\) qed lemma lookup_update: "lookup (update k v f) k' = (if k = k' then v else lookup f k')" by transfer simp lemma keys_update: "keys (update k v f) = (if v = 0 then keys f - {k} else insert k (keys f))" by transfer auto subsection \Quasi-functorial structure\ lift_definition map :: "('b::zero \ 'c::zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c::zero)" is "\g f k. g (f k) when f k \ 0" by simp context fixes f :: "'b \ 'a" assumes inj_f: "inj f" begin lift_definition map_key :: "('a \\<^sub>0 'c::zero) \ 'b \\<^sub>0 'c" is "\p. p \ f" proof - fix g :: "'c \ 'd" and p :: "'a \ 'c" assume "finite {x. p x \ 0}" hence "finite (f ` {y. p (f y) \ 0})" by(rule finite_subset[rotated]) auto thus "finite {x. (p \ f) x \ 0}" unfolding o_def by(rule finite_imageD)(rule subset_inj_on[OF inj_f], simp) qed end lemma map_key_compose: assumes [transfer_rule]: "inj f" "inj g" shows "map_key f (map_key g p) = map_key (g \ f) p" proof - from assms have [transfer_rule]: "inj (g \ f)" by(simp add: inj_compose) show ?thesis by transfer(simp add: o_assoc) qed lemma map_key_id: "map_key (\x. x) p = p" proof - have [transfer_rule]: "inj (\x. x)" by simp show ?thesis by transfer(simp add: o_def) qed context fixes f :: "'a \ 'b" assumes inj_f [transfer_rule]: "inj f" begin lemma map_key_map: "map_key f (map g p) = map g (map_key f p)" by transfer (simp add: fun_eq_iff) lemma map_key_plus: "map_key f (p + q) = map_key f p + map_key f q" by transfer (simp add: fun_eq_iff) lemma keys_map_key: "keys (map_key f p) = f -` keys p" by transfer auto lemma map_key_zero [simp]: "map_key f 0 = 0" by transfer (simp add: fun_eq_iff) lemma map_key_single [simp]: "map_key f (single (f k) v) = single k v" by transfer (simp add: fun_eq_iff inj_onD [OF inj_f] when_def) end lemma mult_map_scale_conv_mult: "map ((*) s) p = single 0 s * p" proof(transfer fixing: s) fix p :: "'a \ 'b" assume *: "finite {x. p x \ 0}" { fix x have "prod_fun (\k'. s when 0 = k') p x = (\l :: 'a. if l = 0 then s * (\q. p q when x = q) else 0)" by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta) also have "\ = (\k. s * p k when p k \ 0) x" by(simp add: when_def) also note calculation } then show "(\k. s * p k when p k \ 0) = prod_fun (\k'. s when 0 = k') p" by(simp add: fun_eq_iff) qed lemma map_single [simp]: "(c = 0 \ f 0 = 0) \ map f (single x c) = single x (f c)" by transfer(auto simp add: fun_eq_iff when_def) lemma map_eq_zero_iff: "map f p = 0 \ (\k \ keys p. f (lookup p k) = 0)" by transfer(auto simp add: fun_eq_iff when_def) subsection \Canonical dense representation of @{typ "nat \\<^sub>0 'a"}\ abbreviation no_trailing_zeros :: "'a :: zero list \ bool" where "no_trailing_zeros \ no_trailing ((=) 0)" lift_definition "nth" :: "'a list \ (nat \\<^sub>0 'a::zero)" is "nth_default 0" by (fact finite_nth_default_neq_default) text \ The opposite direction is directly specified on (later) type \nat_mapping\. \ lemma nth_Nil [simp]: "nth [] = 0" by transfer (simp add: fun_eq_iff) lemma nth_singleton [simp]: "nth [v] = single 0 v" proof (transfer, rule ext) fix n :: nat and v :: 'a show "nth_default 0 [v] n = (v when 0 = n)" by (auto simp add: nth_default_def nth_append) qed lemma nth_replicate [simp]: "nth (replicate n 0 @ [v]) = single n v" proof (transfer, rule ext) fix m n :: nat and v :: 'a show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)" by (auto simp add: nth_default_def nth_append) qed lemma nth_strip_while [simp]: "nth (strip_while ((=) 0) xs) = nth xs" by transfer (fact nth_default_strip_while_dflt) lemma nth_strip_while' [simp]: "nth (strip_while (\k. k = 0) xs) = nth xs" by (subst eq_commute) (fact nth_strip_while) lemma nth_eq_iff: "nth xs = nth ys \ strip_while (HOL.eq 0) xs = strip_while (HOL.eq 0) ys" by transfer (simp add: nth_default_eq_iff) lemma lookup_nth [simp]: "lookup (nth xs) = nth_default 0 xs" by (fact nth.rep_eq) lemma keys_nth [simp]: "keys (nth xs) = fst ` {(n, v) \ set (enumerate 0 xs). v \ 0}" proof transfer fix xs :: "'a list" { fix n assume "nth_default 0 xs n \ 0" then have "n < length xs" and "xs ! n \ 0" by (auto simp add: nth_default_def split: if_splits) then have "(n, xs ! n) \ {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" (is "?x \ ?A") by (auto simp add: in_set_conv_nth enumerate_eq_zip) then have "fst ?x \ fst ` ?A" by blast then have "n \ fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by simp } then show "{k. nth_default 0 xs k \ 0} = fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by (auto simp add: in_enumerate_iff_nth_default_eq) qed lemma range_nth [simp]: "range (nth xs) = set xs - {0}" by transfer simp lemma degree_nth: "no_trailing_zeros xs \ degree (nth xs) = length xs" unfolding degree_def proof transfer fix xs :: "'a list" assume *: "no_trailing_zeros xs" let ?A = "{n. nth_default 0 xs n \ 0}" let ?f = "nth_default 0 xs" let ?bound = "Max (insert 0 (Suc ` {n. ?f n \ 0}))" show "?bound = length xs" proof (cases "xs = []") case False with * obtain n where n: "n < length xs" "xs ! n \ 0" by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv) then have "?bound = Max (Suc ` {k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs})" by (subst Max_insert) (auto simp add: nth_default_def) also let ?A = "{k. k < length xs \ xs ! k \ 0}" have "{k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs} = ?A" by auto also have "Max (Suc ` ?A) = Suc (Max ?A)" using n by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc) also { have "Max ?A \ ?A" using n Max_in [of ?A] by fastforce hence "Suc (Max ?A) \ length xs" by simp moreover from * False have "length xs - 1 \ ?A" by(auto simp add: no_trailing_unfold last_conv_nth) hence "length xs - 1 \ Max ?A" using Max_ge[of ?A "length xs - 1"] by auto hence "length xs \ Suc (Max ?A)" by simp ultimately have "Suc (Max ?A) = length xs" by simp } finally show ?thesis . qed simp qed lemma nth_trailing_zeros [simp]: "nth (xs @ replicate n 0) = nth xs" by transfer simp lemma nth_idem: "nth (List.map (lookup f) [0.. n" shows "nth (List.map (lookup f) [0..Canonical sparse representation of @{typ "'a \\<^sub>0 'b"}\ lift_definition the_value :: "('a \ 'b) list \ 'a \\<^sub>0 'b::zero" is "\xs k. case map_of xs k of None \ 0 | Some v \ v" proof - fix xs :: "('a \ 'b) list" have fin: "finite {k. \v. map_of xs k = Some v}" using finite_dom_map_of [of xs] unfolding dom_def by auto then show "finite {k. (case map_of xs k of None \ 0 | Some v \ v) \ 0}" using fin by (simp split: option.split) qed definition items :: "('a::linorder \\<^sub>0 'b::zero) \ ('a \ 'b) list" where "items f = List.map (\k. (k, lookup f k)) (sorted_list_of_set (keys f))" text \ For the canonical sparse representation we provide both directions of morphisms since the specification of ordered association lists in theory \OAList\ will support arbitrary linear orders @{class linorder} as keys, not just natural numbers @{typ nat}. \ lemma the_value_items [simp]: "the_value (items f) = f" unfolding items_def by transfer (simp add: fun_eq_iff map_of_map_restrict restrict_map_def) lemma lookup_the_value: "lookup (the_value xs) k = (case map_of xs k of None \ 0 | Some v \ v)" by transfer rule lemma items_the_value: assumes "sorted (List.map fst xs)" and "distinct (List.map fst xs)" and "0 \ snd ` set xs" shows "items (the_value xs) = xs" proof - from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs" unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sorted_sort_id) moreover from assms have "keys (the_value xs) = fst ` set xs" by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr) ultimately show ?thesis unfolding items_def using assms by (auto simp add: lookup_the_value intro: map_idI) qed lemma the_value_Nil [simp]: "the_value [] = 0" by transfer (simp add: fun_eq_iff) lemma the_value_Cons [simp]: "the_value (x # xs) = update (fst x) (snd x) (the_value xs)" by transfer (simp add: fun_eq_iff) lemma items_zero [simp]: "items 0 = []" unfolding items_def by simp lemma items_one [simp]: "items 1 = [(0, 1)]" unfolding items_def by transfer simp lemma items_single [simp]: "items (single k v) = (if v = 0 then [] else [(k, v)])" unfolding items_def by simp lemma in_set_items_iff [simp]: "(k, v) \ set (items f) \ k \ keys f \ lookup f k = v" unfolding items_def by transfer auto subsection \Size estimation\ context fixes f :: "'a \ nat" and g :: "'b :: zero \ nat" begin definition poly_mapping_size :: "('a \\<^sub>0 'b) \ nat" where "poly_mapping_size m = g 0 + (\k \ keys m. Suc (f k + g (lookup m k)))" lemma poly_mapping_size_0 [simp]: "poly_mapping_size 0 = g 0" by (simp add: poly_mapping_size_def) lemma poly_mapping_size_single [simp]: "poly_mapping_size (single k v) = (if v = 0 then g 0 else g 0 + f k + g v + 1)" unfolding poly_mapping_size_def by transfer simp lemma keys_less_poly_mapping_size: "k \ keys m \ f k + g (lookup m k) < poly_mapping_size m" unfolding poly_mapping_size_def proof transfer fix k :: 'a and m :: "'a \ 'b" and f :: "'a \ nat" and g let ?keys = "{k. m k \ 0}" assume *: "finite ?keys" "k \ ?keys" then have "f k + g (m k) = (\k' \ ?keys. f k' + g (m k') when k' = k)" by (simp add: sum.delta when_def) also have "\ < (\k' \ ?keys. Suc (f k' + g (m k')))" using * by (intro sum_strict_mono) (auto simp add: when_def) also have "\ \ g 0 + \" by simp finally have "f k + g (m k) < \" . then show "f k + g (m k) < g 0 + (\k | m k \ 0. Suc (f k + g (m k)))" by simp qed lemma lookup_le_poly_mapping_size: "g (lookup m k) \ poly_mapping_size m" proof (cases "k \ keys m") case True with keys_less_poly_mapping_size [of k m] show ?thesis by simp next case False then show ?thesis by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff) qed lemma poly_mapping_size_estimation: "k \ keys m \ y \ f k + g (lookup m k) \ y < poly_mapping_size m" using keys_less_poly_mapping_size by (auto intro: le_less_trans) lemma poly_mapping_size_estimation2: assumes "v \ range m" and "y \ g v" shows "y < poly_mapping_size m" proof - from assms obtain k where *: "lookup m k = v" "v \ 0" by transfer blast from * have "k \ keys m" by (simp add: in_keys_iff) then show ?thesis proof (rule poly_mapping_size_estimation) from assms * have "y \ g (lookup m k)" by simp then show "y \ f k + g (lookup m k)" by simp qed qed end lemma poly_mapping_size_one [simp]: "poly_mapping_size f g 1 = g 0 + f 0 + g 1 + 1" unfolding poly_mapping_size_def by transfer simp lemma poly_mapping_size_cong [fundef_cong]: "m = m' \ g 0 = g' 0 \ (\k. k \ keys m' \ f k = f' k) \ (\v. v \ range m' \ g v = g' v) \ poly_mapping_size f g m = poly_mapping_size f' g' m'" by (auto simp add: poly_mapping_size_def intro!: sum.cong) instantiation poly_mapping :: (type, zero) size begin definition "size = poly_mapping_size (\_. 0) (\_. 0)" instance .. end subsection \Further mapping operations and properties\ text \It is like in algebra: there are many definitions, some are also used\ lift_definition mapp :: "('a \ 'b :: zero \ 'c :: zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c)" is "\f p k. (if k \ keys p then f k (lookup p k) else 0)" by simp lemma mapp_cong [fundef_cong]: "\ m = m'; \k. k \ keys m' \ f k (lookup m' k) = f' k (lookup m' k) \ \ mapp f m = mapp f' m'" by transfer (auto simp add: fun_eq_iff) lemma lookup_mapp: "lookup (mapp f p) k = (f k (lookup p k) when k \ keys p)" by (simp add: mapp.rep_eq) lemma keys_mapp_subset: "keys (mapp f p) \ keys p" by (meson in_keys_iff mapp.rep_eq subsetI) subsection\Free Abelian Groups Over a Type\ abbreviation frag_of :: "'a \ 'a \\<^sub>0 int" where "frag_of c \ Poly_Mapping.single c (1::int)" lemma lookup_frag_of [simp]: "Poly_Mapping.lookup(frag_of c) = (\x. if x = c then 1 else 0)" by (force simp add: lookup_single_not_eq) lemma frag_of_nonzero [simp]: "frag_of a \ 0" proof - let ?f = "\x. if x = a then 1 else (0::int)" have "?f \ (\x. 0::int)" by (auto simp: fun_eq_iff) then have "Poly_Mapping.lookup (Abs_poly_mapping ?f) \ Poly_Mapping.lookup (Abs_poly_mapping (\x. 0))" by fastforce then show ?thesis by (metis lookup_single_eq lookup_zero) qed definition frag_cmul :: "int \ ('a \\<^sub>0 int) \ ('a \\<^sub>0 int)" where "frag_cmul c a = Abs_poly_mapping (\x. c * Poly_Mapping.lookup a x)" lemma frag_cmul_zero [simp]: "frag_cmul 0 x = 0" by (simp add: frag_cmul_def) lemma frag_cmul_zero2 [simp]: "frag_cmul c 0 = 0" by (simp add: frag_cmul_def) lemma frag_cmul_one [simp]: "frag_cmul 1 x = x" by (auto simp: frag_cmul_def Poly_Mapping.poly_mapping.lookup_inverse) lemma frag_cmul_minus_one [simp]: "frag_cmul (-1) x = -x" by (simp add: frag_cmul_def uminus_poly_mapping_def poly_mapping_eqI) lemma frag_cmul_cmul [simp]: "frag_cmul c (frag_cmul d x) = frag_cmul (c*d) x" by (simp add: frag_cmul_def mult_ac) lemma lookup_frag_cmul [simp]: "poly_mapping.lookup (frag_cmul c x) i = c * poly_mapping.lookup x i" by (simp add: frag_cmul_def) lemma minus_frag_cmul [simp]: "- frag_cmul k x = frag_cmul (-k) x" by (simp add: poly_mapping_eqI) lemma keys_frag_of: "Poly_Mapping.keys(frag_of a) = {a}" by simp lemma finite_cmul_nonzero: "finite {x. c * Poly_Mapping.lookup a x \ (0::int)}" by simp lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \ Poly_Mapping.keys a" using finite_cmul_nonzero [of c a] by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI) lemma keys_cmul_iff [iff]: "i \ Poly_Mapping.keys (frag_cmul c x) \ i \ Poly_Mapping.keys x \ c \ 0" apply (auto simp: ) apply (meson subsetD keys_cmul) by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff) lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a" by (metis (no_types, hide_lams) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym) lemma keys_diff: "Poly_Mapping.keys(a - b) \ Poly_Mapping.keys a \ Poly_Mapping.keys b" by (auto simp add: in_keys_iff lookup_minus) lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \ c = 0" by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI) lemma frag_cmul_eq_0_iff [simp]: "frag_cmul k c = 0 \ k=0 \ c=0" by auto (metis subsetI subset_antisym keys_cmul_iff keys_eq_empty) lemma frag_of_eq: "frag_of x = frag_of y \ x = y" by (metis lookup_single_eq lookup_single_not_eq zero_neq_one) lemma frag_cmul_distrib: "frag_cmul (c+d) a = frag_cmul c a + frag_cmul d a" by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) lemma frag_cmul_distrib2: "frag_cmul c (a+b) = frag_cmul c a + frag_cmul c b" proof - have "finite {x. poly_mapping.lookup a x + poly_mapping.lookup b x \ 0}" using keys_add [of a b] by (metis (no_types, lifting) finite_keys finite_subset keys.rep_eq lookup_add mem_Collect_eq subsetI) then show ?thesis by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) qed lemma frag_cmul_diff_distrib: "frag_cmul (a - b) c = frag_cmul a c - frag_cmul b c" by (auto simp: left_diff_distrib lookup_minus poly_mapping_eqI) lemma frag_cmul_sum: "frag_cmul a (sum b I) = (\i\I. frag_cmul a (b i))" proof (induction rule: infinite_finite_induct) case (insert i I) then show ?case by (auto simp: algebra_simps frag_cmul_distrib2) qed auto lemma keys_sum: "Poly_Mapping.keys(sum b I) \ (\i \I. Poly_Mapping.keys(b i))" proof (induction I rule: infinite_finite_induct) case (insert i I) then show ?case using keys_add [of "b i" "sum b I"] by auto qed auto definition frag_extend :: "('b \ 'a \\<^sub>0 int) \ ('b \\<^sub>0 int) \ 'a \\<^sub>0 int" where "frag_extend b x \ (\i \ Poly_Mapping.keys x. frag_cmul (Poly_Mapping.lookup x i) (b i))" lemma frag_extend_0 [simp]: "frag_extend b 0 = 0" by (simp add: frag_extend_def) lemma frag_extend_of [simp]: "frag_extend f (frag_of a) = f a" by (simp add: frag_extend_def) lemma frag_extend_cmul: "frag_extend f (frag_cmul c x) = frag_cmul c (frag_extend f x)" by (auto simp: frag_extend_def frag_cmul_sum intro: sum.mono_neutral_cong_left) lemma frag_extend_minus: "frag_extend f (- x) = - (frag_extend f x)" using frag_extend_cmul [of f "-1"] by simp lemma frag_extend_add: "frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)" proof - have *: "(\i\Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))" "(\i\Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))" by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left) have "frag_extend f (a+b) = (\i\Poly_Mapping.keys (a + b). frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) " by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib) also have "... = (\i \ Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i))" apply (rule sum.mono_neutral_cong_left) using keys_add [of a b] apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric]) done also have "... = (frag_extend f a) + (frag_extend f b)" by (auto simp: * sum.distrib frag_extend_def) finally show ?thesis . qed lemma frag_extend_diff: "frag_extend f (a-b) = (frag_extend f a) - (frag_extend f b)" by (metis (no_types, hide_lams) add_uminus_conv_diff frag_extend_add frag_extend_minus) lemma frag_extend_sum: "finite I \ frag_extend f (\i\I. g i) = sum (frag_extend f o g) I" by (induction I rule: finite_induct) (simp_all add: frag_extend_add) lemma frag_extend_eq: "(\f. f \ Poly_Mapping.keys c \ g f = h f) \ frag_extend g c = frag_extend h c" by (simp add: frag_extend_def) lemma frag_extend_eq_0: "(\x. x \ Poly_Mapping.keys c \ f x = 0) \ frag_extend f c = 0" by (simp add: frag_extend_def) lemma keys_frag_extend: "Poly_Mapping.keys(frag_extend f c) \ (\x \ Poly_Mapping.keys c. Poly_Mapping.keys(f x))" unfolding frag_extend_def - by (smt SUP_mono' keys_cmul keys_sum order_trans) + using keys_sum by fastforce lemma frag_expansion: "a = frag_extend frag_of a" proof - have *: "finite I \ Poly_Mapping.lookup (\i\I. frag_cmul (Poly_Mapping.lookup a i) (frag_of i)) j = (if j \ I then Poly_Mapping.lookup a j else 0)" for I j by (induction I rule: finite_induct) (auto simp: lookup_single lookup_add) show ?thesis unfolding frag_extend_def by (rule poly_mapping_eqI) (fastforce simp add: in_keys_iff *) qed lemma frag_closure_minus_cmul: assumes "P 0" and P: "\x y. \P x; P y\ \ P(x - y)" "P c" shows "P(frag_cmul k c)" proof - have "P (frag_cmul (int n) c)" for n apply (induction n) apply (simp_all add: assms frag_cmul_distrib) by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P) then show ?thesis by (metis (no_types, hide_lams) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases) qed lemma frag_induction [consumes 1, case_names zero one diff]: assumes supp: "Poly_Mapping.keys c \ S" and 0: "P 0" and sing: "\x. x \ S \ P(frag_of x)" and diff: "\a b. \P a; P b\ \ P(a - b)" shows "P c" proof - have "P (\i\I. frag_cmul (poly_mapping.lookup c i) (frag_of i))" if "I \ Poly_Mapping.keys c" for I using finite_subset [OF that finite_keys [of c]] that supp proof (induction I arbitrary: c rule: finite_induct) case empty then show ?case by (auto simp: 0) next case (insert i I c) have ab: "a+b = a - (0 - b)" for a b :: "'a \\<^sub>0 int" by simp have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))" by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff) show ?case apply (simp add: insert.hyps) apply (subst ab) using insert apply (blast intro: assms Pfrag) done qed then show ?thesis by (subst frag_expansion) (auto simp add: frag_extend_def) qed lemma frag_extend_compose: "frag_extend f (frag_extend (frag_of o g) c) = frag_extend (f o g) c" using subset_UNIV by (induction c rule: frag_induction) (auto simp: frag_extend_diff) lemma frag_split: fixes c :: "'a \\<^sub>0 int" assumes "Poly_Mapping.keys c \ S \ T" obtains d e where "Poly_Mapping.keys d \ S" "Poly_Mapping.keys e \ T" "d + e = c" proof let ?d = "frag_extend (\f. if f \ S then frag_of f else 0) c" let ?e = "frag_extend (\f. if f \ S then 0 else frag_of f) c" show "Poly_Mapping.keys ?d \ S" "Poly_Mapping.keys ?e \ T" using assms by (auto intro!: order_trans [OF keys_frag_extend] split: if_split_asm) show "?d + ?e = c" using assms proof (induction c rule: frag_induction) case (diff a b) then show ?case by (metis (no_types, lifting) frag_extend_diff add_diff_eq diff_add_eq diff_add_eq_diff_diff_swap) qed auto qed hide_const (open) lookup single update keys range map map_key degree nth the_value items foldr mapp end \ No newline at end of file diff --git a/src/HOL/Library/Ramsey.thy b/src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy +++ b/src/HOL/Library/Ramsey.thy @@ -1,999 +1,999 @@ (* Title: HOL/Library/Ramsey.thy Author: Tom Ridge. Full finite version by L C Paulson. *) section \Ramsey's Theorem\ theory Ramsey imports Infinite_Set FuncSet begin subsection \Preliminary definitions\ subsubsection \The $n$-element subsets of a set $A$\ definition nsets :: "['a set, nat] \ 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999) where "nsets A n \ {N. N \ A \ finite N \ card N = n}" lemma nsets_mono: "A \ B \ nsets A n \ nsets B n" by (auto simp: nsets_def) lemma nsets_Pi_contra: "A' \ A \ Pi ([A]\<^bsup>n\<^esup>) B \ Pi ([A']\<^bsup>n\<^esup>) B" by (auto simp: nsets_def) lemma nsets_2_eq: "nsets A 2 = (\x\A. \y\A - {x}. {{x, y}})" by (auto simp: nsets_def card_2_iff) lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})" by (auto simp: nsets_2_eq) lemma doubleton_in_nsets_2 [simp]: "{x,y} \ [A]\<^bsup>2\<^esup> \ x \ A \ y \ A \ x \ y" by (auto simp: nsets_2_eq Set.doubleton_eq_iff) lemma nsets_3_eq: "nsets A 3 = (\x\A. \y\A - {x}. \z\A - {x,y}. {{x,y,z}})" by (simp add: eval_nat_numeral nsets_def card_Suc_eq) blast lemma nsets_4_eq: "[A]\<^bsup>4\<^esup> = (\u\A. \x\A - {u}. \y\A - {u,x}. \z\A - {u,x,y}. {{u,x,y,z}})" (is "_ = ?rhs") proof show "[A]\<^bsup>4\<^esup> \ ?rhs" by (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) blast show "?rhs \ [A]\<^bsup>4\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) by (metis insert_iff singletonD) qed lemma nsets_disjoint_2: "X \ Y = {} \ [X \ Y]\<^bsup>2\<^esup> = [X]\<^bsup>2\<^esup> \ [Y]\<^bsup>2\<^esup> \ (\x\X. \y\Y. {{x,y}})" by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff) lemma ordered_nsets_2_eq: fixes A :: "'a::linorder set" shows "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ x ?rhs" unfolding numeral_nat apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less) by (metis antisym) show "?rhs \ nsets A 2" unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq) qed lemma ordered_nsets_3_eq: fixes A :: "'a::linorder set" shows "nsets A 3 = {{x,y,z} | x y z. x \ A \ y \ A \ z \ A \ x y ?rhs" apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) by (metis insert_commute linorder_cases) show "?rhs \ nsets A 3" apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) by (metis empty_iff insert_iff not_less_iff_gr_or_eq) qed lemma ordered_nsets_4_eq: fixes A :: "'a::linorder set" shows "[A]\<^bsup>4\<^esup> = {U. \u x y z. U = {u,x,y,z} \ u \ A \ x \ A \ y \ A \ z \ A \ u < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - { fix U assume "U \ [A]\<^bsup>4\<^esup>" then obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \ A" by (simp add: nsets_def) (metis finite_set_strict_sorted) then have "?RHS U" unfolding numeral_nat length_Suc_conv by auto blast } moreover have "Collect ?RHS \ [A]\<^bsup>4\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) apply (subst card_insert_disjoint, auto)+ done ultimately show ?thesis by auto qed lemma ordered_nsets_5_eq: fixes A :: "'a::linorder set" shows "[A]\<^bsup>5\<^esup> = {U. \u v x y z. U = {u,v,x,y,z} \ u \ A \ v \ A \ x \ A \ y \ A \ z \ A \ u < v \ v < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - { fix U assume "U \ [A]\<^bsup>5\<^esup>" then obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \ A" apply (simp add: nsets_def) by (metis finite_set_strict_sorted) then have "?RHS U" unfolding numeral_nat length_Suc_conv by auto blast } moreover have "Collect ?RHS \ [A]\<^bsup>5\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) apply (subst card_insert_disjoint, auto)+ done ultimately show ?thesis by auto qed lemma binomial_eq_nsets: "n choose k = card (nsets {0.. finite A \ card A < r" unfolding nsets_def proof (intro iffI conjI) assume that: "{N. N \ A \ finite N \ card N = r} = {}" show "finite A" using infinite_arbitrarily_large that by auto then have "\ r \ card A" using that by (simp add: set_eq_iff) (metis obtain_subset_with_card_n) then show "card A < r" using not_less by blast next show "{N. N \ A \ finite N \ card N = r} = {}" if "finite A \ card A < r" using that card_mono leD by auto qed lemma nsets_eq_empty: "\finite A; card A < r\ \ nsets A r = {}" by (simp add: nsets_eq_empty_iff) lemma nsets_empty_iff: "nsets {} r = (if r=0 then {{}} else {})" by (auto simp: nsets_def) lemma nsets_singleton_iff: "nsets {a} r = (if r=0 then {{}} else if r=1 then {{a}} else {})" by (auto simp: nsets_def card_gt_0_iff subset_singleton_iff) lemma nsets_self [simp]: "nsets {..x. {x}) ` A" using card_eq_SucD by (force simp: nsets_def) lemma inj_on_nsets: assumes "inj_on f A" shows "inj_on (\X. f ` X) ([A]\<^bsup>n\<^esup>)" using assms unfolding nsets_def by (metis (no_types, lifting) inj_on_inverseI inv_into_image_cancel mem_Collect_eq) lemma bij_betw_nsets: assumes "bij_betw f A B" shows "bij_betw (\X. f ` X) ([A]\<^bsup>n\<^esup>) ([B]\<^bsup>n\<^esup>)" proof - have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>" using assms apply (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset) by (metis card_image inj_on_finite order_refl subset_image_inj) with assms show ?thesis by (auto simp: bij_betw_def inj_on_nsets) qed lemma nset_image_obtains: assumes "X \ [f`A]\<^bsup>k\<^esup>" "inj_on f A" obtains Y where "Y \ [A]\<^bsup>k\<^esup>" "X = f ` Y" using assms apply (clarsimp simp add: nsets_def subset_image_iff) by (metis card_image finite_imageD inj_on_subset) lemma nsets_image_funcset: assumes "g \ S \ T" and "inj_on g S" shows "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" using assms by (fastforce simp add: nsets_def card_image inj_on_subset subset_iff simp flip: image_subset_iff_funcset) lemma nsets_compose_image_funcset: assumes f: "f \ [T]\<^bsup>k\<^esup> \ D" and "g \ S \ T" and "inj_on g S" shows "f \ (\X. g ` X) \ [S]\<^bsup>k\<^esup> \ D" proof - have "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" using assms by (simp add: nsets_image_funcset) then show ?thesis using f by fastforce qed subsubsection \Partition predicates\ definition partn :: "'a set \ nat \ nat \ 'b set \ bool" where "partn \ \ \ \ \ \f \ nsets \ \ \ \. \H \ nsets \ \. \\\\. f ` (nsets H \) \ {\}" definition partn_lst :: "'a set \ nat list \ nat \ bool" where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. \i < length \. \H \ nsets \ (\!i). f ` (nsets H \) \ {i}" lemma partn_lst_greater_resource: fixes M::nat assumes M: "partn_lst {.. \" and "M \ N" shows "partn_lst {.. \" proof (clarsimp simp: partn_lst_def) fix f assume "f \ nsets {.. \ {..}" then have "f \ nsets {.. \ {..}" by (meson Pi_anti_mono \M \ N\ lessThan_subset_iff nsets_mono subsetD) then obtain i H where i: "i < length \" and H: "H \ nsets {.. ! i)" and subi: "f ` nsets H \ \ {i}" using M partn_lst_def by blast have "H \ nsets {.. ! i)" using \M \ N\ H by (auto simp: nsets_def subset_iff) then show "\i. \H\nsets {.. ! i). f ` nsets H \ \ {i}" using i subi by blast qed lemma partn_lst_fewer_colours: assumes major: "partn_lst \ (n#\) \" and "n \ \" shows "partn_lst \ \ \" proof (clarsimp simp: partn_lst_def) fix f :: "'a set \ nat" assume f: "f \ [\]\<^bsup>\\<^esup> \ {..}" then obtain i H where i: "i < Suc (length \)" and H: "H \ [\]\<^bsup>((n # \) ! i)\<^esup>" and hom: "\x\[H]\<^bsup>\\<^esup>. Suc (f x) = i" using \n \ \\ major [unfolded partn_lst_def, rule_format, of "Suc o f"] by (fastforce simp: image_subset_iff nsets_eq_empty_iff) show "\i. \H\nsets \ (\ ! i). f ` [H]\<^bsup>\\<^esup> \ {i}" proof (cases i) case 0 then have "[H]\<^bsup>\\<^esup> = {}" using hom by blast then show ?thesis using 0 H \n \ \\ by (simp add: nsets_eq_empty_iff) (simp add: nsets_def) next case (Suc i') then show ?thesis using i H hom by auto qed qed lemma partn_lst_eq_partn: "partn_lst {..Finite versions of Ramsey's theorem\ text \ To distinguish the finite and infinite ones, lower and upper case names are used. \ subsubsection \Trivial cases\ text \Vacuous, since we are dealing with 0-sets!\ lemma ramsey0: "\N::nat. partn_lst {..Just the pigeon hole principle, since we are dealing with 1-sets\ lemma ramsey1: "\N::nat. partn_lst {..iH\nsets {.. {i}" if "f \ nsets {.. {.. \i. {q. q \ q0+q1 \ f {q} = i}" have "A 0 \ A 1 = {..q0 + q1}" using that by (auto simp: A_def PiE_iff nsets_one lessThan_Suc_atMost le_Suc_eq) moreover have "A 0 \ A 1 = {}" by (auto simp: A_def) ultimately have "q0 + q1 \ card (A 0) + card (A 1)" by (metis card_Un_le card_atMost eq_imp_le le_SucI le_trans) then consider "card (A 0) \ q0" | "card (A 1) \ q1" by linarith then obtain i where "i < Suc (Suc 0)" "card (A i) \ [q0, q1] ! i" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) then obtain B where "B \ A i" "card B = [q0, q1] ! i" "finite B" by (meson obtain_subset_with_card_n) then have "B \ nsets {.. f ` nsets B (Suc 0) \ {i}" by (auto simp: A_def nsets_def card_1_singleton_iff) then show ?thesis using \i < Suc (Suc 0)\ by auto qed then show ?thesis by (clarsimp simp: partn_lst_def) blast qed subsubsection \Ramsey's theorem with two colours and arbitrary exponents (hypergraph version)\ proposition ramsey2_full: "\N::nat. partn_lst {.. 0" by simp show ?thesis using Suc.prems proof (induct k \ "q1 + q2" arbitrary: q1 q2) case 0 show ?case proof show "partn_lst {..<1::nat} [q1, q2] (Suc r)" using nsets_empty_iff subset_insert 0 by (fastforce simp: partn_lst_def funcset_to_empty_iff nsets_eq_empty image_subset_iff) qed next case (Suc k) consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto then show ?case proof cases case 1 then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)" unfolding partn_lst_def using \r > 0\ by (fastforce simp add: nsets_empty_iff nsets_singleton_iff lessThan_Suc) then show ?thesis by blast next case 2 with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto then obtain p1 p2::nat where p1: "partn_lst {..iH\nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \ {i}" if f: "f \ nsets {..p} (Suc r) \ {.. \R. f (insert p R)" have "f (insert p i) \ {.. nsets {.. nsets {.. {.. {i}" and U: "U \ nsets {.. {.. nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q1" unfolding nsets_def using \q1 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q1" by (simp add: izero inq1) ultimately show ?thesis by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) next case jone then have "u ` V \ nsets {..p} q2" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD) ultimately show ?thesis using jone not_less_eq by fastforce qed next case ione then have "U \ nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q2" unfolding nsets_def using \q2 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q2" by (simp add: ione inq1) ultimately show ?thesis by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc) next case jzero then have "u ` V \ nsets {..p} q1" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj apply (clarsimp simp add: h_def image_subset_iff nsets_def) by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj) ultimately show ?thesis using jzero not_less_eq by fastforce qed qed qed then show "partn_lst {..Full Ramsey's theorem with multiple colours and arbitrary exponents\ theorem ramsey_full: "\N::nat. partn_lst {.. "length qs" arbitrary: qs) case 0 then show ?case by (rule_tac x=" r" in exI) (simp add: partn_lst_def) next case (Suc k) note IH = this show ?case proof (cases k) case 0 with Suc obtain q where "qs = [q]" by (metis length_0_conv length_Suc_conv) then show ?thesis by (rule_tac x=q in exI) (auto simp: partn_lst_def funcset_to_empty_iff) next case (Suc k') then obtain q1 q2 l where qs: "qs = q1#q2#l" by (metis Suc.hyps(2) length_Suc_conv) then obtain q::nat where q: "partn_lst {..qs = q1 # q2 # l\ by fastforce have keq: "Suc (length l) = k" using IH qs by auto show ?thesis proof (intro exI conjI) show "partn_lst {.. nsets {.. {.. \X. if f X < Suc (Suc 0) then 0 else f X - Suc 0" have "g \ nsets {.. {.. {i}" and U: "U \ nsets {..iH\nsets {.. {i}" proof (cases "i = 0") case True then have "U \ nsets {.. {0, Suc 0}" using U gi unfolding g_def by (auto simp: image_subset_iff) then obtain u where u: "bij_betw u {.. {..U \ nsets {.. mem_Collect_eq nsets_def) + by (smt (verit) U mem_Collect_eq nsets_def) have u_nsets: "u ` X \ nsets {.. nsets {.. \X. f (u ` X)" have "f (u ` X) < Suc (Suc 0)" if "X \ nsets {.. nsets U r" using u u_nsets that by (auto simp: nsets_def bij_betwE subset_eq) then show ?thesis using f01 by auto qed then have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. (\x. (u ` x)) ` nsets V r" apply (clarsimp simp add: nsets_def image_iff) by (metis card_eq_0_iff card_image image_is_empty subset_image_inj) then have "f ` nsets (u ` V) r \ h ` nsets V r" by (auto simp: h_def) then show "f ` nsets (u ` V) r \ {j}" using hj by auto show "(u ` V) \ nsets {.. {Suc i}" using i gi False apply (auto simp: g_def image_subset_iff) by (metis Suc_lessD Suc_pred g_def gi image_subset_iff not_less_eq singleton_iff) show "U \ nsets {..Simple graph version\ text \This is the most basic version in terms of cliques and independent sets, i.e. the version for graphs and 2 colours. \ definition "clique V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" definition "indep V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" lemma ramsey2: "\r\1. \(V::'a set) (E::'a set set). finite V \ card V \ r \ (\R \ V. card R = m \ clique R E \ card R = n \ indep R E)" proof - obtain N where "N \ Suc 0" and N: "partn_lst {..R\V. card R = m \ clique R E \ card R = n \ indep R E" if "finite V" "N \ card V" for V :: "'a set" and E :: "'a set set" proof - from that obtain v where u: "inj_on v {.. V" by (metis card_le_inj card_lessThan finite_lessThan) define f where "f \ \e. if v ` e \ E then 0 else Suc 0" have f: "f \ nsets {.. {.. {i}" and U: "U \ nsets {.. V" using U u by (auto simp: image_subset_iff nsets_def) show "card (v ` U) = m \ clique (v ` U) E \ card (v ` U) = n \ indep (v ` U) E" using i unfolding numeral_2_eq_2 using gi U u apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq) apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm) done qed qed then show ?thesis using \Suc 0 \ N\ by auto qed subsection \Preliminaries\ subsubsection \``Axiom'' of Dependent Choice\ primrec choice :: "('a \ bool) \ ('a \ 'a) set \ nat \ 'a" where \ \An integer-indexed chain of choices\ choice_0: "choice P r 0 = (SOME x. P x)" | choice_Suc: "choice P r (Suc n) = (SOME y. P y \ (choice P r n, y) \ r)" lemma choice_n: assumes P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" shows "P (choice P r n)" proof (induct n) case 0 show ?case by (force intro: someI P0) next case Suc then show ?case by (auto intro: someI2_ex [OF Pstep]) qed lemma dependent_choice: assumes trans: "trans r" and P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" obtains f :: "nat \ 'a" where "\n. P (f n)" and "\n m. n < m \ (f n, f m) \ r" proof fix n show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) next fix n m :: nat assume "n < m" from Pstep [OF choice_n [OF P0 Pstep]] have "(choice P r k, choice P r (Suc k)) \ r" for k by (auto intro: someI2_ex) then show "(choice P r n, choice P r m) \ r" by (auto intro: less_Suc_induct [OF \n < m\] transD [OF trans]) qed subsubsection \Partition functions\ definition part_fn :: "nat \ nat \ 'a set \ ('a set \ nat) \ bool" \ \the function \<^term>\f\ partitions the \<^term>\r\-subsets of the typically infinite set \<^term>\Y\ into \<^term>\s\ distinct categories.\ where "part_fn r s Y f \ (f \ nsets Y r \ {..For induction, we decrease the value of \<^term>\r\ in partitions.\ lemma part_fn_Suc_imp_part_fn: "\infinite Y; part_fn (Suc r) s Y f; y \ Y\ \ part_fn r s (Y - {y}) (\u. f (insert y u))" by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert) lemma part_fn_subset: "part_fn r s YY f \ Y \ YY \ part_fn r s Y f" unfolding part_fn_def nsets_def by blast subsection \Ramsey's Theorem: Infinitary Version\ lemma Ramsey_induction: fixes s r :: nat and YY :: "'a set" and f :: "'a set \ nat" assumes "infinite YY" "part_fn r s YY f" shows "\Y' t'. Y' \ YY \ infinite Y' \ t' < s \ (\X. X \ Y' \ finite X \ card X = r \ f X = t')" using assms proof (induct r arbitrary: YY f) case 0 then show ?case by (auto simp add: part_fn_def card_eq_0_iff cong: conj_cong) next case (Suc r) show ?case proof - from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \ YY" by blast let ?ramr = "{((y, Y, t), (y', Y', t')). y' \ Y \ Y' \ Y}" let ?propr = "\(y, Y, t). y \ YY \ y \ Y \ Y \ YY \ infinite Y \ t < s \ (\X. X\Y \ finite X \ card X = r \ (f \ insert y) X = t)" from Suc.prems have infYY': "infinite (YY - {yy})" by auto from Suc.prems have partf': "part_fn r s (YY - {yy}) (f \ insert yy)" by (simp add: o_def part_fn_Suc_imp_part_fn yy) have transr: "trans ?ramr" by (force simp add: trans_def) from Suc.hyps [OF infYY' partf'] obtain Y0 and t0 where "Y0 \ YY - {yy}" "infinite Y0" "t0 < s" "X \ Y0 \ finite X \ card X = r \ (f \ insert yy) X = t0" for X by blast with yy have propr0: "?propr(yy, Y0, t0)" by blast have proprstep: "\y. ?propr y \ (x, y) \ ?ramr" if x: "?propr x" for x proof (cases x) case (fields yx Yx tx) with x obtain yx' where yx': "yx' \ Yx" by (blast dest: infinite_imp_nonempty) from fields x have infYx': "infinite (Yx - {yx'})" by auto with fields x yx' Suc.prems have partfx': "part_fn r s (Yx - {yx'}) (f \ insert yx')" by (simp add: o_def part_fn_Suc_imp_part_fn part_fn_subset [where YY=YY and Y=Yx]) from Suc.hyps [OF infYx' partfx'] obtain Y' and t' where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" "X \ Y' \ finite X \ card X = r \ (f \ insert yx') X = t'" for X by blast from fields x Y' yx' have "?propr (yx', Y', t') \ (x, (yx', Y', t')) \ ?ramr" by blast then show ?thesis .. qed from dependent_choice [OF transr propr0 proprstep] obtain g where pg: "?propr (g n)" and rg: "n < m \ (g n, g m) \ ?ramr" for n m :: nat by blast let ?gy = "fst \ g" let ?gt = "snd \ snd \ g" have rangeg: "\k. range ?gt \ {.. range ?gt" then obtain n where "x = ?gt n" .. with pg [of n] show "x \ {.. ?gy m'" by (cases "g m", cases "g m'") auto qed show ?thesis proof (intro exI conjI) from pg show "?gy ` {n. ?gt n = s'} \ YY" by (auto simp add: Let_def split_beta) from infeqs' show "infinite (?gy ` {n. ?gt n = s'})" by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) show "s' < s" by (rule less') show "\X. X \ ?gy ` {n. ?gt n = s'} \ finite X \ card X = Suc r \ f X = s'" proof - have "f X = s'" if X: "X \ ?gy ` {n. ?gt n = s'}" and cardX: "finite X" "card X = Suc r" for X proof - from X obtain AA where AA: "AA \ {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" by (auto simp add: subset_image_iff) with cardX have "AA \ {}" by auto then have AAleast: "(LEAST x. x \ AA) \ AA" by (auto intro: LeastI_ex) show ?thesis proof (cases "g (LEAST x. x \ AA)") case (fields ya Ya ta) with AAleast Xeq have ya: "ya \ X" by (force intro!: rev_image_eqI) then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb) also have "\ = ta" proof - have *: "X - {ya} \ Ya" proof fix x assume x: "x \ X - {ya}" then obtain a' where xeq: "x = ?gy a'" and a': "a' \ AA" by (auto simp add: Xeq) with fields x have "a' \ (LEAST x. x \ AA)" by auto with Least_le [of "\x. x \ AA", OF a'] have "(LEAST x. x \ AA) < a'" by arith from xeq fields rg [OF this] show "x \ Ya" by auto qed have "card (X - {ya}) = r" by (simp add: cardX ya) with pg [of "LEAST x. x \ AA"] fields cardX * show ?thesis by (auto simp del: insert_Diff_single) qed also from AA AAleast fields have "\ = s'" by auto finally show ?thesis . qed qed then show ?thesis by blast qed qed qed qed theorem Ramsey: fixes s r :: nat and Z :: "'a set" and f :: "'a set \ nat" shows "\infinite Z; \X. X \ Z \ finite X \ card X = r \ f X < s\ \ \Y t. Y \ Z \ infinite Y \ t < s \ (\X. X \ Y \ finite X \ card X = r \ f X = t)" by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def]) corollary Ramsey2: fixes s :: nat and Z :: "'a set" and f :: "'a set \ nat" assumes infZ: "infinite Z" and part: "\x\Z. \y\Z. x \ y \ f {x, y} < s" shows "\Y t. Y \ Z \ infinite Y \ t < s \ (\x\Y. \y\Y. x\y \ f {x, y} = t)" proof - from part have part2: "\X. X \ Z \ finite X \ card X = 2 \ f X < s" by (fastforce simp add: eval_nat_numeral card_Suc_eq) obtain Y t where *: "Y \ Z" "infinite Y" "t < s" "(\X. X \ Y \ finite X \ card X = 2 \ f X = t)" by (insert Ramsey [OF infZ part2]) auto then have "\x\Y. \y\Y. x \ y \ f {x, y} = t" by auto with * show ?thesis by iprover qed corollary Ramsey_nsets: fixes f :: "'a set \ nat" assumes "infinite Z" "f ` nsets Z r \ {.. Z" "infinite Y" "t < s" "f ` nsets Y r \ {t}" using Ramsey [of Z r f s] assms by (auto simp: nsets_def image_subset_iff) subsection \Disjunctive Well-Foundedness\ text \ An application of Ramsey's theorem to program termination. See @{cite "Podelski-Rybalchenko"}. \ definition disj_wf :: "('a \ 'a) set \ bool" where "disj_wf r \ (\T. \n::nat. (\i r = (\i 'a) \ (nat \ ('a \ 'a) set) \ nat set \ nat" where "transition_idx s T A = (LEAST k. \i j. A = {i, j} \ i < j \ (s j, s i) \ T k)" lemma transition_idx_less: assumes "i < j" "(s j, s i) \ T k" "k < n" shows "transition_idx s T {i, j} < n" proof - from assms(1,2) have "transition_idx s T {i, j} \ k" by (simp add: transition_idx_def, blast intro: Least_le) with assms(3) show ?thesis by simp qed lemma transition_idx_in: assumes "i < j" "(s j, s i) \ T k" shows "(s j, s i) \ T (transition_idx s T {i, j})" using assms by (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR cong: conj_cong) (erule LeastI) text \To be equal to the union of some well-founded relations is equivalent to being the subset of such a union.\ lemma disj_wf: "disj_wf r \ (\T. \n::nat. (\i r \ (\iT n. \\i \ (T ` {.. \ (\i r)) \ r = (\i r)" by (force simp add: wf_Int1) show ?thesis unfolding disj_wf_def by auto (metis "*") qed theorem trans_disj_wf_implies_wf: assumes "trans r" and "disj_wf r" shows "wf r" proof (simp only: wf_iff_no_infinite_down_chain, rule notI) assume "\s. \i. (s (Suc i), s i) \ r" then obtain s where sSuc: "\i. (s (Suc i), s i) \ r" .. from \disj_wf r\ obtain T and n :: nat where wfT: "\kkk. (s j, s i) \ T k \ ki < j\ have "(s j, s i) \ r" proof (induct rule: less_Suc_induct) case 1 then show ?case by (simp add: sSuc) next case 2 with \trans r\ show ?case unfolding trans_def by blast qed then show ?thesis by (auto simp add: r) qed have "i < j \ transition_idx s T {i, j} < n" for i j using s_in_T transition_idx_less by blast then have trless: "i \ j \ transition_idx s T {i, j} < n" for i j by (metis doubleton_eq_iff less_linear) have "\K k. K \ UNIV \ infinite K \ k < n \ (\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k)" by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat) then obtain K and k where infK: "infinite K" and "k < n" and allk: "\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k" by auto have "(s (enumerate K (Suc m)), s (enumerate K m)) \ T k" for m :: nat proof - let ?j = "enumerate K (Suc m)" let ?i = "enumerate K m" have ij: "?i < ?j" by (simp add: enumerate_step infK) have "?j \ K" "?i \ K" by (simp_all add: enumerate_in_set infK) with ij have k: "k = transition_idx s T {?i, ?j}" by (simp add: allk) from s_in_T [OF ij] obtain k' where "(s ?j, s ?i) \ T k'" "k' T k" by (simp add: k transition_idx_in ij) qed then have "\ wf (T k)" by (meson wf_iff_no_infinite_down_chain) with wfT \k < n\ show False by blast qed end