# HG changeset patch # User wenzelm # Date 1533996175 -7200 # Node ID 682ff0e84387bbba681be62a0abfee4fb0cfb393 # Parent 0c62e3b4f4c03898141091181c7130a7e458db8e# Parent ec0b2833cfc40e127ab34ca4f1fdb33fe80056c6 merged; diff -r 0c62e3b4f4c0 -r 682ff0e84387 CONTRIBUTORS --- a/CONTRIBUTORS Tue Aug 07 11:39:40 2018 +0200 +++ b/CONTRIBUTORS Sat Aug 11 16:02:55 2018 +0200 @@ -3,6 +3,10 @@ listed as an author in one of the source files of this Isabelle distribution. +Contributions to this Isabelle version +-------------------------------------- + + Contributions to Isabelle2018 ----------------------------- diff -r 0c62e3b4f4c0 -r 682ff0e84387 NEWS --- a/NEWS Tue Aug 07 11:39:40 2018 +0200 +++ b/NEWS Sat Aug 11 16:02:55 2018 +0200 @@ -4,6 +4,11 @@ (Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.) +New in this Isabelle version +---------------------------- + + + New in Isabelle2018 (August 2018) --------------------------------- diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/AbelCoset.thy Sat Aug 11 16:02:55 2018 +0200 @@ -269,17 +269,15 @@ by (rule a_comm_group) interpret subgroup "H" "(add_monoid G)" by (rule a_subgroup) - - show "abelian_subgroup H G" - apply unfold_locales - proof (simp add: r_coset_def l_coset_def, clarsimp) - fix x - assume xcarr: "x \ carrier G" - from a_subgroup have Hcarr: "H \ carrier G" - unfolding subgroup_def by simp - from xcarr Hcarr show "(\h\H. {h \\<^bsub>G\<^esub> x}) = (\h\H. {x \\<^bsub>G\<^esub> h})" + have "(\xa\H. {xa \ x}) = (\xa\H. {x \ xa})" if "x \ carrier G" for x + proof - + have "H \ carrier G" + using a_subgroup that unfolding subgroup_def by simp + with that show "(\h\H. {h \\<^bsub>G\<^esub> x}) = (\h\H. {x \\<^bsub>G\<^esub> h})" using m_comm [simplified] by fastforce qed + then show "abelian_subgroup H G" + by unfold_locales (auto simp: r_coset_def l_coset_def) qed lemma abelian_subgroupI3: @@ -304,14 +302,6 @@ by (rule normal.inv_op_closed2 [OF a_normal, folded a_inv_def, simplified monoid_record_simps]) -text\Alternative characterization of normal subgroups\ -lemma (in abelian_group) a_normal_inv_iff: - "(N \ (add_monoid G)) = - (subgroup N (add_monoid G) & (\x \ carrier G. \h \ N. x \ h \ (\ x) \ N))" - (is "_ = ?rhs") -by (rule group.normal_inv_iff [OF a_group, - folded a_inv_def, simplified monoid_record_simps]) - lemma (in abelian_group) a_lcos_m_assoc: "\ M \ carrier G; g \ carrier G; h \ carrier G \ \ g <+ (h <+ M) = (g \ h) <+ M" by (rule group.lcos_m_assoc [OF a_group, @@ -322,13 +312,11 @@ by (rule group.lcos_mult_one [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) - lemma (in abelian_group) a_l_coset_subset_G: "\ H \ carrier G; x \ carrier G \ \ x <+ H \ carrier G" by (rule group.l_coset_subset_G [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) - lemma (in abelian_group) a_l_coset_swap: "\y \ x <+ H; x \ carrier G; subgroup H (add_monoid G)\ \ x \ y <+ H" by (rule group.l_coset_swap [OF a_group, @@ -498,15 +486,15 @@ text \Since the Factorization is based on an \emph{abelian} subgroup, is results in a commutative group\ -theorem (in abelian_subgroup) a_factorgroup_is_comm_group: - "comm_group (G A_Mod H)" -apply (intro comm_group.intro comm_monoid.intro) prefer 3 - apply (rule a_factorgroup_is_group) - apply (rule group.axioms[OF a_factorgroup_is_group]) -apply (rule comm_monoid_axioms.intro) -apply (unfold A_FactGroup_def FactGroup_def RCOSETS_def, fold set_add_def a_r_coset_def, clarsimp) -apply (simp add: a_rcos_sum a_comm) -done +theorem (in abelian_subgroup) a_factorgroup_is_comm_group: "comm_group (G A_Mod H)" +proof - + have "Group.comm_monoid_axioms (G A_Mod H)" + apply (rule comm_monoid_axioms.intro) + apply (auto simp: A_FactGroup_def FactGroup_def RCOSETS_def a_normal add.m_comm normal.rcos_sum) + done + then show ?thesis + by (intro comm_group.intro comm_monoid.intro) (simp_all add: a_factorgroup_is_group group.is_monoid) +qed lemma add_A_FactGroup [simp]: "X \\<^bsub>(G A_Mod H)\<^esub> X' = X <+>\<^bsub>G\<^esub> X'" by (simp add: A_FactGroup_def set_add_def) @@ -552,11 +540,8 @@ interpret G: abelian_group G by fact interpret H: abelian_group H by fact show ?thesis - apply (intro abelian_group_hom.intro abelian_group_hom_axioms.intro) - apply fact - apply fact - apply (rule a_group_hom) - done + by (intro abelian_group_hom.intro abelian_group_hom_axioms.intro + G.abelian_group_axioms H.abelian_group_axioms a_group_hom) qed lemma (in abelian_group_hom) is_abelian_group_hom: @@ -576,8 +561,7 @@ lemma (in abelian_group_hom) zero_closed [simp]: "h \ \ carrier H" -by (rule group_hom.one_closed[OF a_group_hom, - simplified ring_record_simps]) + by simp lemma (in abelian_group_hom) hom_zero [simp]: "h \ = \\<^bsub>H\<^esub>" @@ -586,8 +570,7 @@ lemma (in abelian_group_hom) a_inv_closed [simp]: "x \ carrier G ==> h (\x) \ carrier H" -by (rule group_hom.inv_closed[OF a_group_hom, - folded a_inv_def, simplified ring_record_simps]) + by simp lemma (in abelian_group_hom) hom_a_inv [simp]: "x \ carrier G ==> h (\x) = \\<^bsub>H\<^esub> (h x)" @@ -596,19 +579,15 @@ lemma (in abelian_group_hom) additive_subgroup_a_kernel: "additive_subgroup (a_kernel G H h) G" -apply (rule additive_subgroup.intro) -apply (rule group_hom.subgroup_kernel[OF a_group_hom, - folded a_kernel_def, simplified ring_record_simps]) -done + by (simp add: additive_subgroup.intro a_group_hom a_kernel_def group_hom.subgroup_kernel) text\The kernel of a homomorphism is an abelian subgroup\ lemma (in abelian_group_hom) abelian_subgroup_a_kernel: "abelian_subgroup (a_kernel G H h) G" -apply (rule abelian_subgroupI) -apply (rule group_hom.normal_kernel[OF a_group_hom, - folded a_kernel_def, simplified ring_record_simps]) -apply (simp add: G.a_comm) -done + apply (rule abelian_subgroupI) + apply (simp add: G.abelian_group_axioms abelian_subgroup.a_normal abelian_subgroupI3 additive_subgroup_a_kernel) + apply (simp add: G.a_comm) + done lemma (in abelian_group_hom) A_FactGroup_nonempty: assumes X: "X \ carrier (G A_Mod a_kernel G H h)" @@ -715,48 +694,34 @@ qed lemma (in abelian_subgroup) a_repr_independence': - assumes y: "y \ H +> x" - and xcarr: "x \ carrier G" + assumes "y \ H +> x" "x \ carrier G" shows "H +> x = H +> y" - apply (rule a_repr_independence) - apply (rule y) - apply (rule xcarr) - apply (rule a_subgroup) - done + using a_repr_independence a_subgroup assms by blast lemma (in abelian_subgroup) a_repr_independenceD: - assumes ycarr: "y \ carrier G" - and repr: "H +> x = H +> y" + assumes "y \ carrier G" "H +> x = H +> y" shows "y \ H +> x" -by (rule group.repr_independenceD [OF a_group a_subgroup, - folded a_r_coset_def, simplified monoid_record_simps]) (rule ycarr, rule repr) + by (simp add: a_rcos_self assms) lemma (in abelian_subgroup) a_rcosets_carrier: "X \ a_rcosets H \ X \ carrier G" -by (rule subgroup.rcosets_carrier [OF a_subgroup a_group, - folded A_RCOSETS_def, simplified monoid_record_simps]) + using a_rcosets_part_G by auto subsubsection \Addition of Subgroups\ lemma (in abelian_monoid) set_add_closed: - assumes Acarr: "A \ carrier G" - and Bcarr: "B \ carrier G" + assumes "A \ carrier G" "B \ carrier G" shows "A <+> B \ carrier G" -by (rule monoid.set_mult_closed [OF a_monoid, - folded set_add_def, simplified monoid_record_simps]) (rule Acarr, rule Bcarr) + by (simp add: assms add.set_mult_closed set_add_defs(1)) lemma (in abelian_group) add_additive_subgroups: assumes subH: "additive_subgroup H G" - and subK: "additive_subgroup K G" + and subK: "additive_subgroup K G" shows "additive_subgroup (H <+> K) G" -apply (rule additive_subgroup.intro) -apply (unfold set_add_def) -apply (intro comm_group.mult_subgroups) - apply (rule a_comm_group) - apply (rule additive_subgroup.a_subgroup[OF subH]) -apply (rule additive_subgroup.a_subgroup[OF subK]) -done + unfolding set_add_def + using add.mult_subgroups additive_subgroup_def subH subK + by (blast intro: additive_subgroup.intro) end diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Bij.thy --- a/src/HOL/Algebra/Bij.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Bij.thy Sat Aug 11 16:02:55 2018 +0200 @@ -46,15 +46,11 @@ by (simp add: Bij_def compose_inv_into_id) theorem group_BijGroup: "group (BijGroup S)" -apply (simp add: BijGroup_def) -apply (rule groupI) - apply (simp add: compose_Bij) - apply (simp add: id_Bij) - apply (simp add: compose_Bij) - apply (blast intro: compose_assoc [symmetric] dest: Bij_imp_funcset) - apply (simp add: id_Bij Bij_imp_funcset Bij_imp_extensional, simp) -apply (blast intro: Bij_compose_restrict_eq restrict_inv_into_Bij) -done + apply (simp add: BijGroup_def) + apply (rule groupI) + apply (auto simp: compose_Bij id_Bij Bij_imp_funcset Bij_imp_extensional compose_assoc [symmetric]) + apply (blast intro: Bij_compose_restrict_eq restrict_inv_into_Bij) + done subsection\Automorphisms Form a Group\ @@ -63,13 +59,18 @@ by (simp add: Bij_def bij_betw_def inv_into_into) lemma Bij_inv_into_lemma: - assumes eq: "\x y. \x \ S; y \ S\ \ h(g x y) = g (h x) (h y)" - shows "\h \ Bij S; g \ S \ S \ S; x \ S; y \ S\ - \ inv_into S h (g x y) = g (inv_into S h x) (inv_into S h y)" -apply (simp add: Bij_def bij_betw_def) -apply (subgoal_tac "\x'\S. \y'\S. x = h x' \ y = h y'", clarify) - apply (simp add: eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem], blast) -done + assumes eq: "\x y. \x \ S; y \ S\ \ h(g x y) = g (h x) (h y)" + and hg: "h \ Bij S" "g \ S \ S \ S" and "x \ S" "y \ S" + shows "inv_into S h (g x y) = g (inv_into S h x) (inv_into S h y)" +proof - + have "h ` S = S" + by (metis (no_types) Bij_def Int_iff assms(2) bij_betw_def mem_Collect_eq) + with \x \ S\ \y \ S\ have "\x'\S. \y'\S. x = h x' \ y = h y'" + by auto + then show ?thesis + using assms + by (auto simp add: Bij_def bij_betw_def eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem]) +qed definition @@ -94,8 +95,7 @@ lemma inv_BijGroup: "f \ Bij S \ m_inv (BijGroup S) f = (\x \ S. (inv_into S f) x)" -apply (rule group.inv_equality) -apply (rule group_BijGroup) +apply (rule group.inv_equality [OF group_BijGroup]) apply (simp_all add:BijGroup_def restrict_inv_into_Bij Bij_compose_restrict_eq) done diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Complete_Lattice.thy --- a/src/HOL/Algebra/Complete_Lattice.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Complete_Lattice.thy Sat Aug 11 16:02:55 2018 +0200 @@ -680,22 +680,25 @@ next case False show ?thesis - proof (rule_tac x="\\<^bsub>L\<^esub> A" in exI, rule least_UpperI, simp_all) + proof (intro exI least_UpperI, simp_all) show b:"\ x. x \ A \ x \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A" using a by (auto intro: L.sup_upper, meson L.at_least_at_most_closed L.sup_upper subset_trans) show "\y. y \ Upper (L\carrier := \a..b\\<^bsub>L\<^esub>\) A \ \\<^bsub>L\<^esub>A \\<^bsub>L\<^esub> y" using a L.at_least_at_most_closed by (rule_tac L.sup_least, auto intro: funcset_mem simp add: Upper_def) - from a show "A \ \a..b\\<^bsub>L\<^esub>" - by (auto) - from a show "\\<^bsub>L\<^esub>A \ \a..b\\<^bsub>L\<^esub>" - apply (rule_tac L.at_least_at_most_member) - apply (auto) - apply (meson L.at_least_at_most_closed L.sup_closed subset_trans) - apply (meson False L.at_least_at_most_closed L.at_least_at_most_lower L.le_trans L.sup_closed b all_not_in_conv assms(2) contra_subsetD subset_trans) - apply (rule L.sup_least) - apply (auto simp add: assms) - using L.at_least_at_most_closed apply blast - done + from a show *: "A \ \a..b\\<^bsub>L\<^esub>" + by auto + show "\\<^bsub>L\<^esub>A \ \a..b\\<^bsub>L\<^esub>" + proof (rule_tac L.at_least_at_most_member) + show 1: "\\<^bsub>L\<^esub>A \ carrier L" + by (meson L.at_least_at_most_closed L.sup_closed subset_trans *) + show "a \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A" + by (meson "*" False L.at_least_at_most_closed L.at_least_at_most_lower L.le_trans L.sup_upper 1 all_not_in_conv assms(2) set_mp subset_trans) + show "\\<^bsub>L\<^esub>A \\<^bsub>L\<^esub> b" + proof (rule L.sup_least) + show "A \ carrier L" "\x. x \ A \ x \\<^bsub>L\<^esub> b" + using * L.at_least_at_most_closed by blast+ + qed (simp add: assms) + qed qed qed show "\s. is_glb (L\carrier := \a..b\\<^bsub>L\<^esub>\) s A" @@ -711,15 +714,17 @@ using a L.at_least_at_most_closed by (force intro!: L.inf_lower) show "\y. y \ Lower (L\carrier := \a..b\\<^bsub>L\<^esub>\) A \ y \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A" using a L.at_least_at_most_closed by (rule_tac L.inf_greatest, auto intro: funcset_carrier' simp add: Lower_def) - from a show "A \ \a..b\\<^bsub>L\<^esub>" - by (auto) - from a show "\\<^bsub>L\<^esub>A \ \a..b\\<^bsub>L\<^esub>" - apply (rule_tac L.at_least_at_most_member) - apply (auto) - apply (meson L.at_least_at_most_closed L.inf_closed subset_trans) - apply (meson L.at_least_at_most_closed L.at_least_at_most_lower L.inf_greatest assms(2) set_rev_mp subset_trans) - apply (meson False L.at_least_at_most_closed L.at_least_at_most_upper L.inf_closed L.le_trans b all_not_in_conv assms(3) contra_subsetD subset_trans) - done + from a show *: "A \ \a..b\\<^bsub>L\<^esub>" + by auto + show "\\<^bsub>L\<^esub>A \ \a..b\\<^bsub>L\<^esub>" + proof (rule_tac L.at_least_at_most_member) + show 1: "\\<^bsub>L\<^esub>A \ carrier L" + by (meson "*" L.at_least_at_most_closed L.inf_closed subset_trans) + show "a \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A" + by (meson "*" L.at_least_at_most_closed L.at_least_at_most_lower L.inf_greatest assms(2) subsetD subset_trans) + show "\\<^bsub>L\<^esub>A \\<^bsub>L\<^esub> b" + by (meson * 1 False L.at_least_at_most_closed L.at_least_at_most_upper L.inf_lower L.le_trans all_not_in_conv assms(3) set_mp subset_trans) + qed qed qed qed @@ -731,7 +736,7 @@ text \The set of fixed points of a complete lattice is itself a complete lattice\ theorem Knaster_Tarski: - assumes "weak_complete_lattice L" "f \ carrier L \ carrier L" "isotone L L f" + assumes "weak_complete_lattice L" and f: "f \ carrier L \ carrier L" and "isotone L L f" shows "weak_complete_lattice (fpl L f)" (is "weak_complete_lattice ?L'") proof - interpret L: weak_complete_lattice L @@ -805,15 +810,14 @@ show "is_lub ?L'' (LFP\<^bsub>?L'\<^esub> f) A" proof (rule least_UpperI, simp_all) fix x - assume "x \ Upper ?L'' A" - hence "LFP\<^bsub>?L'\<^esub> f \\<^bsub>?L'\<^esub> x" - apply (rule_tac L'.LFP_lowerbound) - apply (auto simp add: Upper_def) - apply (simp add: A AL L.at_least_at_most_member L.sup_least set_rev_mp) - apply (simp add: Pi_iff assms(2) fps_def, rule_tac L.weak_refl) - apply (auto) - apply (rule funcset_mem[of f "carrier L"], simp_all add: assms(2)) - done + assume x: "x \ Upper ?L'' A" + have "LFP\<^bsub>?L'\<^esub> f \\<^bsub>?L'\<^esub> x" + proof (rule L'.LFP_lowerbound, simp_all) + show "x \ \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub>" + using x by (auto simp add: Upper_def A AL L.at_least_at_most_member L.sup_least set_rev_mp) + with x show "f x \\<^bsub>L\<^esub> x" + by (simp add: Upper_def) (meson L.at_least_at_most_closed L.use_fps L.weak_refl subsetD f_top_chain imageI) + qed thus " LFP\<^bsub>?L'\<^esub> f \\<^bsub>L\<^esub> x" by (simp) next @@ -838,17 +842,13 @@ by (auto simp add: at_least_at_most_def) have "LFP\<^bsub>?L'\<^esub> f .=\<^bsub>?L'\<^esub> f (LFP\<^bsub>?L'\<^esub> f)" proof (rule "L'.LFP_weak_unfold", simp_all) - show "f \ \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub> \ \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub>" - apply (auto simp add: Pi_def at_least_at_most_def) - using assms(2) apply blast - apply (meson AL funcset_mem L.le_trans L.sup_closed assms(2) assms(3) pf_w use_iso2) - using assms(2) apply blast - done - from assms(3) show "Mono\<^bsub>L\carrier := \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub>\\<^esub> f" - apply (auto simp add: isotone_def) - using L'.weak_partial_order_axioms apply blast - apply (meson L.at_least_at_most_closed subsetCE) - done + have "\x. \x \ carrier L; \\<^bsub>L\<^esub>A \\<^bsub>L\<^esub> x\ \ \\<^bsub>L\<^esub>A \\<^bsub>L\<^esub> f x" + by (meson AL funcset_mem L.le_trans L.sup_closed assms(2) assms(3) pf_w use_iso2) + with f show "f \ \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub> \ \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub>" + by (auto simp add: Pi_def at_least_at_most_def) + show "Mono\<^bsub>L\carrier := \\\<^bsub>L\<^esub>A..\\<^bsub>L\<^esub>\\<^bsub>L\<^esub>\\<^esub> f" + using L'.weak_partial_order_axioms assms(3) + by (auto simp add: isotone_def) (meson L.at_least_at_most_closed subsetCE) qed thus "f (LFP\<^bsub>?L'\<^esub> f) .=\<^bsub>L\<^esub> LFP\<^bsub>?L'\<^esub> f" by (simp add: L.equivalence_axioms funcset_carrier' c assms(2) equivalence.sym) @@ -889,7 +889,6 @@ thus ?thesis by (meson AL L.inf_closed L.le_trans assms(3) b(1) b(2) fx use_iso2 w) qed - show "\\<^bsub>L\<^esub> \\<^bsub>L\<^esub> f x" by (simp add: fx) qed @@ -905,12 +904,16 @@ proof (rule greatest_LowerI, simp_all) fix x assume "x \ Lower ?L'' A" - hence "x \\<^bsub>?L'\<^esub> GFP\<^bsub>?L'\<^esub> f" - apply (rule_tac L'.GFP_upperbound) - apply (auto simp add: Lower_def) - apply (meson A AL L.at_least_at_most_member L.bottom_lower L.weak_complete_lattice_axioms fps_carrier subsetCE weak_complete_lattice.inf_greatest) - apply (simp add: funcset_carrier' L.sym assms(2) fps_def) - done + then have x: "\y. y \ A \ y \ fps L f \ x \\<^bsub>L\<^esub> y" "x \ fps L f" + by (auto simp add: Lower_def) + have "x \\<^bsub>?L'\<^esub> GFP\<^bsub>?L'\<^esub> f" + unfolding Lower_def + proof (rule_tac L'.GFP_upperbound; simp) + show "x \ \\\<^bsub>L\<^esub>..\\<^bsub>L\<^esub>A\\<^bsub>L\<^esub>" + by (meson x A AL L.at_least_at_most_member L.bottom_lower L.inf_greatest contra_subsetD fps_carrier) + show "x \\<^bsub>L\<^esub> f x" + using x by (simp add: funcset_carrier' L.sym assms(2) fps_def) + qed thus "x \\<^bsub>L\<^esub> GFP\<^bsub>?L'\<^esub> f" by (simp) next @@ -935,17 +938,14 @@ by (auto simp add: at_least_at_most_def) have "GFP\<^bsub>?L'\<^esub> f .=\<^bsub>?L'\<^esub> f (GFP\<^bsub>?L'\<^esub> f)" proof (rule "L'.GFP_weak_unfold", simp_all) - show "f \ \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub> \ \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub>" - apply (auto simp add: Pi_def at_least_at_most_def) - using assms(2) apply blast - apply (simp add: funcset_carrier' assms(2)) - apply (meson AL funcset_carrier L.inf_closed L.le_trans assms(2) assms(3) pf_w use_iso2) - done - from assms(3) show "Mono\<^bsub>L\carrier := \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub>\\<^esub> f" - apply (auto simp add: isotone_def) - using L'.weak_partial_order_axioms apply blast - using L.at_least_at_most_closed apply (blast intro: funcset_carrier') - done + have "\x. \x \ carrier L; x \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A\ \ f x \\<^bsub>L\<^esub> \\<^bsub>L\<^esub>A" + by (meson AL funcset_carrier L.inf_closed L.le_trans assms(2) assms(3) pf_w use_iso2) + with assms(2) show "f \ \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub> \ \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub>" + by (auto simp add: Pi_def at_least_at_most_def) + have "\x y. \x \ \\\<^bsub>L\<^esub>..\\<^bsub>L\<^esub>A\\<^bsub>L\<^esub>; y \ \\\<^bsub>L\<^esub>..\\<^bsub>L\<^esub>A\\<^bsub>L\<^esub>; x \\<^bsub>L\<^esub> y\ \ f x \\<^bsub>L\<^esub> f y" + by (meson L.at_least_at_most_closed subsetD use_iso1 assms(3)) + with L'.weak_partial_order_axioms show "Mono\<^bsub>L\carrier := \\\<^bsub>L\<^esub>..?w\\<^bsub>L\<^esub>\\<^esub> f" + by (auto simp add: isotone_def) qed thus "f (GFP\<^bsub>?L'\<^esub> f) .=\<^bsub>L\<^esub> GFP\<^bsub>?L'\<^esub> f" by (simp add: L.equivalence_axioms funcset_carrier' c assms(2) equivalence.sym) @@ -1117,17 +1117,16 @@ qed show "\\<^bsub>fpl L f\<^esub>A \\<^bsub>L\<^esub> f (\\<^bsub>L\<^esub>A)" proof - + have *: "\\<^bsub>fpl L f\<^esub>A \ carrier L" + using FA infA by blast have "\x. x \ A \ \\<^bsub>fpl L f\<^esub>A \\<^bsub>fpl L f\<^esub> x" by (rule L'.inf_lower, simp_all add: assms) hence "\\<^bsub>fpl L f\<^esub>A \\<^bsub>L\<^esub> (\\<^bsub>L\<^esub>A)" - apply (rule_tac L.inf_greatest, simp_all add: A) - using FA infA apply blast - done + by (rule_tac L.inf_greatest, simp_all add: A *) hence 1:"f(\\<^bsub>fpl L f\<^esub>A) \\<^bsub>L\<^esub> f(\\<^bsub>L\<^esub>A)" by (metis (no_types, lifting) A FA L.inf_closed assms(2) infA subsetCE use_iso1) have 2:"\\<^bsub>fpl L f\<^esub>A \\<^bsub>L\<^esub> f (\\<^bsub>fpl L f\<^esub>A)" by (metis (no_types, lifting) FA L.sym L.use_fps L.weak_complete_lattice_axioms PiE assms(4) infA subsetCE weak_complete_lattice_def weak_partial_order.weak_refl) - show ?thesis using FA fA infA by (auto intro!: L.le_trans[OF 2 1] ic fc, metis FA PiE assms(4) subsetCE) qed @@ -1189,21 +1188,11 @@ lemma sup_pres_is_join_pres: assumes "weak_sup_pres X Y f" shows "join_pres X Y f" - using assms - apply (simp add: join_pres_def weak_sup_pres_def, safe) - apply (rename_tac x y) - apply (drule_tac x="{x, y}" in spec) - apply (auto simp add: join_def) -done + using assms by (auto simp: join_pres_def weak_sup_pres_def join_def) lemma inf_pres_is_meet_pres: assumes "weak_inf_pres X Y f" shows "meet_pres X Y f" - using assms - apply (simp add: meet_pres_def weak_inf_pres_def, safe) - apply (rename_tac x y) - apply (drule_tac x="{x, y}" in spec) - apply (auto simp add: meet_def) -done + using assms by (auto simp: meet_pres_def weak_inf_pres_def meet_def) end diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Coset.thy Sat Aug 11 16:02:55 2018 +0200 @@ -440,45 +440,36 @@ shows "N \ G" using assms normal_inv_iff by blast -corollary (in group) normal_invE: - assumes "N \ G" - shows "subgroup N G" and "\x h. \ x \ carrier G; h \ N \ \ x \ h \ inv x \ N" - using assms normal_inv_iff apply blast - by (simp add: assms normal.inv_op_closed2) - - -lemma (in group) one_is_normal : - "{\} \ G" -proof(intro normal_invI ) +lemma (in group) one_is_normal: "{\} \ G" +proof(intro normal_invI) show "subgroup {\} G" by (simp add: subgroup_def) - show "\x h. x \ carrier G \ h \ {\} \ x \ h \ inv x \ {\}" by simp -qed +qed simp subsection\More Properties of Left Cosets\ lemma (in group) l_repr_independence: - assumes "y \ x <# H" "x \ carrier G" "subgroup H G" + assumes "y \ x <# H" "x \ carrier G" and HG: "subgroup H G" shows "x <# H = y <# H" proof - obtain h' where h': "h' \ H" "y = x \ h'" using assms(1) unfolding l_coset_def by blast hence "x \ h = y \ ((inv h') \ h)" if "h \ H" for h proof - - have f3: "h' \ carrier G" - by (meson assms(3) h'(1) subgroup.mem_carrier) - have f4: "h \ carrier G" - by (meson assms(3) subgroup.mem_carrier that) - then show ?thesis - by (metis assms(2) f3 h'(2) inv_closed inv_solve_right m_assoc m_closed) + have "h' \ carrier G" + by (meson HG h'(1) subgroup.mem_carrier) + moreover have "h \ carrier G" + by (meson HG subgroup.mem_carrier that) + ultimately show ?thesis + by (metis assms(2) h'(2) inv_closed inv_solve_right m_assoc m_closed) qed - hence "\ xh. xh \ x <# H \ xh \ y <# H" - unfolding l_coset_def by (metis (no_types, lifting) UN_iff assms(3) h'(1) subgroup_def) - moreover have "\ h. h \ H \ y \ h = x \ (h' \ h)" - using h' by (meson assms(2) assms(3) m_assoc subgroup.mem_carrier) - hence "\ yh. yh \ y <# H \ yh \ x <# H" - unfolding l_coset_def using subgroup.m_closed[OF assms(3) h'(1)] by blast + hence "\xh. xh \ x <# H \ xh \ y <# H" + unfolding l_coset_def by (metis (no_types, lifting) UN_iff HG h'(1) subgroup_def) + moreover have "\h. h \ H \ y \ h = x \ (h' \ h)" + using h' by (meson assms(2) HG m_assoc subgroup.mem_carrier) + hence "\yh. yh \ y <# H \ yh \ x <# H" + unfolding l_coset_def using subgroup.m_closed[OF HG h'(1)] by blast ultimately show ?thesis by blast qed @@ -655,8 +646,8 @@ shows "hb \ a \ (\h\H. {h \ b})" proof - interpret subgroup H G by fact - from p show ?thesis apply (rule_tac UN_I [of "hb \ ((inv ha) \ h)"]) - apply blast by (simp add: inv_solve_left m_assoc) + from p show ?thesis + by (rule_tac UN_I [of "hb \ ((inv ha) \ h)"]) (auto simp: inv_solve_left m_assoc) qed lemma (in group) rcos_disjoint: @@ -666,9 +657,8 @@ proof - interpret subgroup H G by fact from p show ?thesis - apply (simp add: RCOSETS_def r_coset_def) - apply (blast intro: rcos_equation assms sym) - done + unfolding RCOSETS_def r_coset_def + by (blast intro: rcos_equation assms sym) qed @@ -761,26 +751,26 @@ proof - interpret subgroup H G by fact show ?thesis - apply (rule equalityI) - apply (force simp add: RCOSETS_def r_coset_def) - apply (auto simp add: RCOSETS_def intro: rcos_self assms) - done + unfolding RCOSETS_def r_coset_def by auto qed lemma (in group) cosets_finite: "\c \ rcosets H; H \ carrier G; finite (carrier G)\ \ finite c" -apply (auto simp add: RCOSETS_def) -apply (simp add: r_coset_subset_G [THEN finite_subset]) -done + unfolding RCOSETS_def + by (auto simp add: r_coset_subset_G [THEN finite_subset]) text\The next two lemmas support the proof of \card_cosets_equal\.\ lemma (in group) inj_on_f: - "\H \ carrier G; a \ carrier G\ \ inj_on (\y. y \ inv a) (H #> a)" -apply (rule inj_onI) -apply (subgoal_tac "x \ carrier G \ y \ carrier G") - prefer 2 apply (blast intro: r_coset_subset_G [THEN subsetD]) -apply (simp add: subsetD) -done + assumes "H \ carrier G" and a: "a \ carrier G" + shows "inj_on (\y. y \ inv a) (H #> a)" +proof + fix x y + assume "x \ H #> a" "y \ H #> a" and xy: "x \ inv a = y \ inv a" + then have "x \ carrier G" "y \ carrier G" + using assms r_coset_subset_G by blast+ + with xy a show "x = y" + by auto +qed lemma (in group) inj_on_g: "\H \ carrier G; a \ carrier G\ \ inj_on (\y. y \ a) H" @@ -827,16 +817,17 @@ using rcosets_part_G by auto proposition (in group) lagrange_finite: - "\finite(carrier G); subgroup H G\ - \ card(rcosets H) * card(H) = order(G)" -apply (simp (no_asm_simp) add: order_def rcosets_part_G [symmetric]) -apply (subst mult.commute) -apply (rule card_partition) - apply (simp add: rcosets_subset_PowG [THEN finite_subset]) - apply (simp add: rcosets_part_G) - apply (simp add: card_rcosets_equal subgroup.subset) -apply (simp add: rcos_disjoint) -done + assumes "finite(carrier G)" and HG: "subgroup H G" + shows "card(rcosets H) * card(H) = order(G)" +proof - + have "card H * card (rcosets H) = card (\(rcosets H))" + proof (rule card_partition) + show "\c1 c2. \c1 \ rcosets H; c2 \ rcosets H; c1 \ c2\ \ c1 \ c2 = {}" + using HG rcos_disjoint by auto + qed (auto simp: assms finite_UnionD rcosets_part_G card_rcosets_equal subgroup.subset) + then show ?thesis + by (simp add: HG mult.commute order_def rcosets_part_G) +qed theorem (in group) lagrange: assumes "subgroup H G" @@ -844,29 +835,29 @@ proof (cases "finite (carrier G)") case True thus ?thesis using lagrange_finite assms by simp next - case False note inf_G = this + case False thus ?thesis proof (cases "finite H") - case False thus ?thesis using inf_G by (simp add: order_def) + case False thus ?thesis using \infinite (carrier G)\ by (simp add: order_def) next - case True note finite_H = this + case True have "infinite (rcosets H)" - proof (rule ccontr) - assume "\ infinite (rcosets H)" + proof + assume "finite (rcosets H)" hence finite_rcos: "finite (rcosets H)" by simp hence "card (\(rcosets H)) = (\R\(rcosets H). card R)" - using card_Union_disjoint[of "rcosets H"] finite_H rcos_disjoint[OF assms(1)] + using card_Union_disjoint[of "rcosets H"] \finite H\ rcos_disjoint[OF assms(1)] rcosets_finite[where ?H = H] by (simp add: assms subgroup.subset) hence "order G = (\R\(rcosets H). card R)" by (simp add: assms order_def rcosets_part_G) hence "order G = (\R\(rcosets H). card H)" using card_rcosets_equal by (simp add: assms subgroup.subset) hence "order G = (card H) * (card (rcosets H))" by simp - hence "order G \ 0" using finite_rcos finite_H assms ex_in_conv + hence "order G \ 0" using finite_rcos \finite H\ assms ex_in_conv rcosets_part_G subgroup.one_closed by fastforce - thus False using inf_G order_gt_0_iff_finite by blast + thus False using \infinite (carrier G)\ order_gt_0_iff_finite by blast qed - thus ?thesis using inf_G by (simp add: order_def) + thus ?thesis using \infinite (carrier G)\ by (simp add: order_def) qed qed @@ -908,8 +899,8 @@ theorem (in normal) factorgroup_is_group: "group (G Mod H)" -apply (simp add: FactGroup_def) -apply (rule groupI) + unfolding FactGroup_def + apply (rule groupI) apply (simp add: setmult_closed) apply (simp add: normal_imp_subgroup subgroup_in_rcosets [OF is_group]) apply (simp add: restrictI setmult_closed rcosets_assoc) @@ -922,10 +913,20 @@ by (simp add: FactGroup_def) lemma (in normal) inv_FactGroup: - "X \ carrier (G Mod H) \ inv\<^bsub>G Mod H\<^esub> X = set_inv X" -apply (rule group.inv_equality [OF factorgroup_is_group]) -apply (simp_all add: FactGroup_def setinv_closed rcosets_inv_mult_group_eq) -done + assumes "X \ carrier (G Mod H)" + shows "inv\<^bsub>G Mod H\<^esub> X = set_inv X" +proof - + have X: "X \ rcosets H" + using assms by (simp add: FactGroup_def) + moreover have "set_inv X <#> X = H" + using X by (simp add: normal.rcosets_inv_mult_group_eq normal_axioms) + moreover have "Group.group (G Mod H)" + using normal.factorgroup_is_group normal_axioms by blast + moreover have "set_inv X \ rcosets H" + by (simp add: \X \ rcosets H\ setinv_closed) + ultimately show ?thesis + by (simp add: FactGroup_def group.inv_equality) +qed text\The coset map is a homomorphism from @{term G} to the quotient group @{term "G Mod H"}\ @@ -945,15 +946,13 @@ where "kernel G H h = {x. x \ carrier G \ h x = \\<^bsub>H\<^esub>}" lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G" -apply (rule subgroup.intro) -apply (auto simp add: kernel_def group.intro is_group) -done + by (auto simp add: kernel_def group.intro is_group intro: subgroup.intro) text\The kernel of a homomorphism is a normal subgroup\ lemma (in group_hom) normal_kernel: "(kernel G H h) \ G" -apply (simp add: G.normal_inv_iff subgroup_kernel) -apply (simp add: kernel_def) -done + apply (simp only: G.normal_inv_iff subgroup_kernel) + apply (simp add: kernel_def) + done lemma (in group_hom) FactGroup_nonempty: assumes X: "X \ carrier (G Mod kernel G H h)" @@ -982,37 +981,40 @@ lemma (in group_hom) FactGroup_hom: "(\X. the_elem (h`X)) \ hom (G Mod (kernel G H h)) H" -apply (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed) -proof (intro ballI) - fix X and X' - assume X: "X \ carrier (G Mod kernel G H h)" - and X': "X' \ carrier (G Mod kernel G H h)" - then - obtain g and g' - where "g \ carrier G" and "g' \ carrier G" - and "X = kernel G H h #> g" and "X' = kernel G H h #> g'" - by (auto simp add: FactGroup_def RCOSETS_def) - hence all: "\x\X. h x = h g" "\x\X'. h x = h g'" - and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" - by (force simp add: kernel_def r_coset_def image_def)+ - hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' - by (auto dest!: FactGroup_nonempty intro!: image_eqI - simp add: set_mult_def - subsetD [OF Xsub] subsetD [OF X'sub]) - then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" - by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) +proof - + have "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + if X: "X \ carrier (G Mod kernel G H h)" and X': "X' \ carrier (G Mod kernel G H h)" for X X' + proof - + obtain g and g' + where "g \ carrier G" and "g' \ carrier G" + and "X = kernel G H h #> g" and "X' = kernel G H h #> g'" + using X X' by (auto simp add: FactGroup_def RCOSETS_def) + hence all: "\x\X. h x = h g" "\x\X'. h x = h g'" + and Xsub: "X \ carrier G" and X'sub: "X' \ carrier G" + by (force simp add: kernel_def r_coset_def image_def)+ + hence "h ` (X <#> X') = {h g \\<^bsub>H\<^esub> h g'}" using X X' + by (auto dest!: FactGroup_nonempty intro!: image_eqI + simp add: set_mult_def + subsetD [OF Xsub] subsetD [OF X'sub]) + then show "the_elem (h ` (X <#> X')) = the_elem (h ` X) \\<^bsub>H\<^esub> the_elem (h ` X')" + by (auto simp add: all FactGroup_nonempty X X' the_elem_image_unique) + qed + then show ?thesis + by (simp add: hom_def FactGroup_the_elem_mem normal.factorgroup_is_group [OF normal_kernel] group.axioms monoid.m_closed) qed text\Lemma for the following injectivity result\ lemma (in group_hom) FactGroup_subset: - "\g \ carrier G; g' \ carrier G; h g = h g'\ - \ kernel G H h #> g \ kernel G H h #> g'" -apply (clarsimp simp add: kernel_def r_coset_def) -apply (rename_tac y) -apply (rule_tac x="y \ g \ inv g'" in exI) -apply (simp add: G.m_assoc) -done + assumes "g \ carrier G" "g' \ carrier G" "h g = h g'" + shows "kernel G H h #> g \ kernel G H h #> g'" + unfolding kernel_def r_coset_def +proof clarsimp + fix y + assume "y \ carrier G" "h y = \\<^bsub>H\<^esub>" + with assms show "\x. x \ carrier G \ h x = \\<^bsub>H\<^esub> \ y \ g = x \ g'" + by (rule_tac x="y \ g \ inv g'" in exI) (auto simp: G.m_assoc) +qed lemma (in group_hom) FactGroup_inj_on: "inj_on (\X. the_elem (h ` X)) (carrier (G Mod kernel G H h))" @@ -1113,13 +1115,13 @@ from hHN obtain h1 h2 where h1h2 : "h1 \ H" "h2 \ N" "h = (h1,h2)" unfolding DirProd_def by fastforce hence h1h2GK : "h1 \ carrier G" "h2 \ carrier K" - using normal_imp_subgroup subgroup.subset assms apply blast+. + using normal_imp_subgroup subgroup.subset assms by blast+ have "inv\<^bsub>G \\ K\<^esub> x = (inv\<^bsub>G\<^esub> x1,inv\<^bsub>K\<^esub> x2)" using inv_DirProd[OF group_axioms assms(1) x1x2(1)x1x2(2)] x1x2 by auto hence "x \\<^bsub>G \\ K\<^esub> h \\<^bsub>G \\ K\<^esub> inv\<^bsub>G \\ K\<^esub> x = (x1 \ h1 \ inv x1,x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)" using h1h2 x1x2 h1h2GK by auto moreover have "x1 \ h1 \ inv x1 \ H" "x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2 \ N" - using normal_invE group.normal_invE[OF assms(1)] assms x1x2 h1h2 apply auto. + using assms x1x2 h1h2 assms by (simp_all add: normal.inv_op_closed2) hence "(x1 \ h1 \ inv x1, x2 \\<^bsub>K\<^esub> h2 \\<^bsub>K\<^esub> inv\<^bsub>K\<^esub> x2)\ H \ N" by auto ultimately show " x \\<^bsub>G \\ K\<^esub> h \\<^bsub>G \\ K\<^esub> inv\<^bsub>G \\ K\<^esub> x \ H \ N" by auto qed @@ -1133,7 +1135,7 @@ proof- have R :"(\(X, Y). X \ Y) \ carrier (G Mod H) \ carrier (K Mod N) \ carrier (G \\ K Mod H \ N)" - unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def apply simp by blast + unfolding r_coset_def Sigma_def DirProd_def FactGroup_def RCOSETS_def by force moreover have "(\x\carrier (G Mod H). \y\carrier (K Mod N). \xa\carrier (G Mod H). \ya\carrier (K Mod N). (x <#> xa) \ (y <#>\<^bsub>K\<^esub> ya) = x \ y <#>\<^bsub>G \\ K\<^esub> xa \ ya)" unfolding set_mult_def by force @@ -1143,8 +1145,14 @@ by (metis assms(2) assms(3) normal_def partial_object.select_convs(1)) moreover have "(\(X, Y). X \ Y) ` (carrier (G Mod H) \ carrier (K Mod N)) = carrier (G \\ K Mod H \ N)" - unfolding image_def apply auto using R apply force - unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force + proof - + have 1: "\x a b. \a \ carrier (G Mod H); b \ carrier (K Mod N)\ \ a \ b \ carrier (G \\ K Mod H \ N)" + using R by force + have 2: "\z. z \ carrier (G \\ K Mod H \ N) \ \x\carrier (G Mod H). \y\carrier (K Mod N). z = x \ y" + unfolding DirProd_def FactGroup_def RCOSETS_def r_coset_def by force + show ?thesis + unfolding image_def by (auto simp: intro: 1 2) + qed ultimately show ?thesis unfolding iso_def hom_def bij_betw_def inj_on_def by simp qed @@ -1262,7 +1270,7 @@ have hH:"h\carrier(GH)" using hHK HK_def GH_def by auto have "(\x\carrier (GH). \h\H1. x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ H1)" - using assms normal_invE GH_def normal.inv_op_closed2 by fastforce + using assms GH_def normal.inv_op_closed2 by fastforce hence INCL_1 : "x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ H1" using xH H1K_def p2 by blast have " x \\<^bsub>GH\<^esub> h \\<^bsub>GH\<^esub> inv\<^bsub>GH\<^esub> x \ HK" @@ -1275,7 +1283,6 @@ qed qed - lemma (in group) normal_inter_subgroup: assumes "subgroup H G" and "N \ G" @@ -1288,8 +1295,8 @@ ultimately have "N \ H \ G\carrier := K \ H\" using normal_inter[of K H N] assms(1) by blast moreover have "K \ H = H" using K_def assms subgroup.subset by blast - ultimately show "normal (N\H) (G\carrier := H\)" by auto + ultimately show "normal (N\H) (G\carrier := H\)" + by auto qed - end diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Divisibility.thy Sat Aug 11 16:02:55 2018 +0200 @@ -547,22 +547,14 @@ using pf by (elim properfactorE) lemma (in monoid) properfactor_trans1 [trans]: - assumes dvds: "a divides b" "properfactor G b c" - and carr: "a \ carrier G" "c \ carrier G" + assumes "a divides b" "properfactor G b c" "a \ carrier G" "c \ carrier G" shows "properfactor G a c" - using dvds carr - apply (elim properfactorE, intro properfactorI) - apply (iprover intro: divides_trans)+ - done + by (meson divides_trans properfactorE properfactorI assms) lemma (in monoid) properfactor_trans2 [trans]: - assumes dvds: "properfactor G a b" "b divides c" - and carr: "a \ carrier G" "b \ carrier G" + assumes "properfactor G a b" "b divides c" "a \ carrier G" "b \ carrier G" shows "properfactor G a c" - using dvds carr - apply (elim properfactorE, intro properfactorI) - apply (iprover intro: divides_trans)+ - done + by (meson divides_trans properfactorE properfactorI assms) lemma properfactor_lless: fixes G (structure) @@ -660,23 +652,20 @@ using assms by (fast elim: irreducibleE) lemma (in monoid_cancel) irreducible_cong [trans]: - assumes irred: "irreducible G a" - and aa': "a \ a'" "a \ carrier G" "a' \ carrier G" + assumes "irreducible G a" "a \ a'" "a \ carrier G" "a' \ carrier G" shows "irreducible G a'" - using assms - apply (auto simp: irreducible_def assoc_unit_l) - apply (metis aa' associated_sym properfactor_cong_r) - done +proof - + have "a' divides a" + by (meson \a \ a'\ associated_def) + then show ?thesis + by (metis (no_types) assms assoc_unit_l irreducibleE irreducibleI monoid.properfactor_trans2 monoid_axioms) +qed lemma (in monoid) irreducible_prod_rI: - assumes airr: "irreducible G a" - and bunit: "b \ Units G" - and carr[simp]: "a \ carrier G" "b \ carrier G" + assumes "irreducible G a" "b \ Units G" "a \ carrier G" "b \ carrier G" shows "irreducible G (a \ b)" - using airr carr bunit - apply (elim irreducibleE, intro irreducibleI) - using prod_unit_r apply blast - using associatedI2' properfactor_cong_r by auto + using assms + by (metis (no_types, lifting) associatedI2' irreducible_def monoid.m_closed monoid_axioms prod_unit_r properfactor_cong_r) lemma (in comm_monoid) irreducible_prod_lI: assumes birr: "irreducible G b" @@ -764,9 +753,7 @@ and pp': "p \ p'" "p \ carrier G" "p' \ carrier G" shows "prime G p'" using assms - apply (auto simp: prime_def assoc_unit_l) - apply (metis pp' associated_sym divides_cong_l) - done + by (auto simp: prime_def assoc_unit_l) (metis pp' associated_sym divides_cong_l) (*by Paulo Emílio de Vilhena*) lemma (in comm_monoid_cancel) prime_irreducible: @@ -849,9 +836,7 @@ and "set as \ carrier G" and "set bs \ carrier G" shows "\a\set bs. irreducible G a" using assms - apply (clarsimp simp add: list_all2_conv_all_nth set_conv_nth) - apply (blast intro: irreducible_cong) - done + by (fastforce simp add: list_all2_conv_all_nth set_conv_nth intro: irreducible_cong) text \Permutations\ @@ -1001,15 +986,7 @@ then have f: "f \ carrier G" by blast show ?case - proof (cases "f = a") - case True - then show ?thesis - using Cons.prems by auto - next - case False - with Cons show ?thesis - by clarsimp (metis f divides_prod_l multlist_closed) - qed + using Cons.IH Cons.prems(1) Cons.prems(2) divides_prod_l f by auto qed auto lemma (in comm_monoid_cancel) multlist_listassoc_cong: @@ -1051,9 +1028,7 @@ and "set fs \ carrier G" and "set fs' \ carrier G" shows "foldr (\) fs \ \ foldr (\) fs' \" using assms - apply (elim essentially_equalE) - apply (simp add: multlist_perm_cong multlist_listassoc_cong perm_closed) - done + by (metis essentially_equal_def multlist_listassoc_cong multlist_perm_cong perm_closed) subsubsection \Factorization in irreducible elements\ @@ -1120,9 +1095,6 @@ and carr[simp]: "set fs \ carrier G" shows "fs = []" proof (cases fs) - case Nil - then show ?thesis . -next case fs: (Cons f fs') from carr have fcarr[simp]: "f \ carrier G" and carr'[simp]: "set fs' \ carrier G" by (simp_all add: fs) @@ -1874,6 +1846,18 @@ qed lemma (in factorial_monoid) properfactor_fmset: + assumes "properfactor G a b" + and "wfactors G as a" + and "wfactors G bs b" + and "a \ carrier G" + and "b \ carrier G" + and "set as \ carrier G" + and "set bs \ carrier G" + shows "fmset G as \# fmset G bs" + using assms + by (meson divides_as_fmsubset properfactor_divides) + +lemma (in factorial_monoid) properfactor_fmset_ne: assumes pf: "properfactor G a b" and "wfactors G as a" and "wfactors G bs b" @@ -1881,11 +1865,8 @@ and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" - shows "fmset G as \# fmset G bs \ fmset G as \ fmset G bs" - using pf - apply safe - apply (meson assms divides_as_fmsubset monoid.properfactor_divides monoid_axioms) - by (meson assms associated_def comm_monoid_cancel.ee_wfactorsD comm_monoid_cancel.fmset_ee factorial_monoid_axioms factorial_monoid_def properfactorE) + shows "fmset G as \ fmset G bs" + using properfactorE [OF pf] assms divides_as_fmsubset by force subsection \Irreducible Elements are Prime\ @@ -2246,75 +2227,70 @@ qed lemma (in gcd_condition_monoid) gcdof_cong_l: - assumes a'a: "a' \ a" - and agcd: "a gcdof b c" - and a'carr: "a' \ carrier G" and carr': "a \ carrier G" "b \ carrier G" "c \ carrier G" + assumes "a' \ a" "a gcdof b c" "a' \ carrier G" and carr': "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "a' gcdof b c" proof - - note carr = a'carr carr' interpret weak_lower_semilattice "division_rel G" by simp have "is_glb (division_rel G) a' {b, c}" - by (subst greatest_Lower_cong_l[of _ a]) (simp_all add: a'a carr gcdof_greatestLower[symmetric] agcd) + by (subst greatest_Lower_cong_l[of _ a]) (simp_all add: assms gcdof_greatestLower[symmetric]) then have "a' \ carrier G \ a' gcdof b c" by (simp add: gcdof_greatestLower carr') then show ?thesis .. qed lemma (in gcd_condition_monoid) gcd_closed [simp]: - assumes carr: "a \ carrier G" "b \ carrier G" + assumes "a \ carrier G" "b \ carrier G" shows "somegcd G a b \ carrier G" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - apply (simp add: somegcd_meet[OF carr]) - apply (rule meet_closed[simplified], fact+) - done + using assms meet_closed by (simp add: somegcd_meet) qed lemma (in gcd_condition_monoid) gcd_isgcd: - assumes carr: "a \ carrier G" "b \ carrier G" + assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) gcdof a b" proof - interpret weak_lower_semilattice "division_rel G" by simp - from carr have "somegcd G a b \ carrier G \ (somegcd G a b) gcdof a b" + from assms have "somegcd G a b \ carrier G \ (somegcd G a b) gcdof a b" by (simp add: gcdof_greatestLower inf_of_two_greatest meet_def somegcd_meet) then show "(somegcd G a b) gcdof a b" by simp qed lemma (in gcd_condition_monoid) gcd_exists: - assumes carr: "a \ carrier G" "b \ carrier G" + assumes "a \ carrier G" "b \ carrier G" shows "\x\carrier G. x = somegcd G a b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - by (metis carr(1) carr(2) gcd_closed) + by (metis assms gcd_closed) qed lemma (in gcd_condition_monoid) gcd_divides_l: - assumes carr: "a \ carrier G" "b \ carrier G" + assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides a" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - by (metis carr(1) carr(2) gcd_isgcd isgcd_def) + by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides_r: - assumes carr: "a \ carrier G" "b \ carrier G" + assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - by (metis carr gcd_isgcd isgcd_def) + by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides: - assumes sub: "z divides x" "z divides y" + assumes "z divides x" "z divides y" and L: "x \ carrier G" "y \ carrier G" "z \ carrier G" shows "z divides (somegcd G x y)" proof - @@ -2325,49 +2301,25 @@ qed lemma (in gcd_condition_monoid) gcd_cong_l: - assumes xx': "x \ x'" - and carr: "x \ carrier G" "x' \ carrier G" "y \ carrier G" + assumes "x \ x'" "x \ carrier G" "x' \ carrier G" "y \ carrier G" shows "somegcd G x y \ somegcd G x' y" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - apply (simp add: somegcd_meet carr) - apply (rule meet_cong_l[simplified], fact+) - done + using somegcd_meet assms + by (metis eq_object.select_convs(1) meet_cong_l partial_object.select_convs(1)) qed lemma (in gcd_condition_monoid) gcd_cong_r: - assumes carr: "x \ carrier G" "y \ carrier G" "y' \ carrier G" - and yy': "y \ y'" + assumes "y \ y'" "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "somegcd G x y \ somegcd G x y'" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - apply (simp add: somegcd_meet carr) - apply (rule meet_cong_r[simplified], fact+) - done + by (meson associated_def assms gcd_closed gcd_divides gcd_divides_l gcd_divides_r monoid.divides_trans monoid_axioms) qed -(* -lemma (in gcd_condition_monoid) asc_cong_gcd_l [intro]: - assumes carr: "b \ carrier G" - shows "asc_cong (\a. somegcd G a b)" -using carr -unfolding CONG_def -by clarsimp (blast intro: gcd_cong_l) - -lemma (in gcd_condition_monoid) asc_cong_gcd_r [intro]: - assumes carr: "a \ carrier G" - shows "asc_cong (\b. somegcd G a b)" -using carr -unfolding CONG_def -by clarsimp (blast intro: gcd_cong_r) - -lemmas (in gcd_condition_monoid) asc_cong_gcd_split [simp] = - assoc_split[OF _ asc_cong_gcd_l] assoc_split[OF _ asc_cong_gcd_r] -*) - lemma (in gcd_condition_monoid) gcdI: assumes dvd: "a divides b" "a divides c" and others: "\y. \y\carrier G; y divides b; y divides c\ \ y divides a" @@ -2390,25 +2342,23 @@ lemma (in gcd_condition_monoid) SomeGcd_ex: assumes "finite A" "A \ carrier G" "A \ {}" - shows "\x\ carrier G. x = SomeGcd G A" + shows "\x \ carrier G. x = SomeGcd G A" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis - apply (simp add: SomeGcd_def) - apply (rule finite_inf_closed[simplified], fact+) - done + using finite_inf_closed by (simp add: assms SomeGcd_def) qed lemma (in gcd_condition_monoid) gcd_assoc: - assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" + assumes "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "somegcd G (somegcd G a b) c \ somegcd G a (somegcd G b c)" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis unfolding associated_def - by (meson carr divides_trans gcd_divides gcd_divides_l gcd_divides_r gcd_exists) + by (meson assms divides_trans gcd_divides gcd_divides_l gcd_divides_r gcd_exists) qed lemma (in gcd_condition_monoid) gcd_mult: @@ -2641,141 +2591,124 @@ using Cons.IH Cons.prems(1) by force qed - -lemma (in primeness_condition_monoid) wfactors_unique__hlp_induct: - "\a as'. a \ carrier G \ set as \ carrier G \ set as' \ carrier G \ - wfactors G as a \ wfactors G as' a \ essentially_equal G as as'" -proof (induct as) +proposition (in primeness_condition_monoid) wfactors_unique: + assumes "wfactors G as a" "wfactors G as' a" + and "a \ carrier G" "set as \ carrier G" "set as' \ carrier G" + shows "essentially_equal G as as'" + using assms +proof (induct as arbitrary: a as') case Nil - show ?case - apply (clarsimp simp: wfactors_def) - by (metis Units_one_closed assoc_unit_r list_update_nonempty unit_wfactors_empty unitfactor_ee wfactorsI) + then have "a \ \" + by (meson Units_one_closed one_closed perm.Nil perm_wfactorsD unit_wfactors) + then have "as' = []" + using Nil.prems assoc_unit_l unit_wfactors_empty by blast + then show ?case + by auto next case (Cons ah as) - then show ?case - proof clarsimp - fix a as' - assume ih [rule_format]: - "\a as'. a \ carrier G \ set as' \ carrier G \ wfactors G as a \ - wfactors G as' a \ essentially_equal G as as'" - and acarr: "a \ carrier G" and ahcarr: "ah \ carrier G" - and ascarr: "set as \ carrier G" and as'carr: "set as' \ carrier G" - and afs: "wfactors G (ah # as) a" - and afs': "wfactors G as' a" - then have ahdvda: "ah divides a" - by (intro wfactors_dividesI[of "ah#as" "a"]) simp_all + then have ahdvda: "ah divides a" + using wfactors_dividesI by auto then obtain a' where a'carr: "a' \ carrier G" and a: "a = ah \ a'" by blast + have carr_ah: "ah \ carrier G" "set as \ carrier G" + using Cons.prems by fastforce+ + have "ah \ foldr (\) as \ \ a" + by (rule wfactorsE[OF \wfactors G (ah # as) a\]) auto + then have "foldr (\) as \ \ a'" + by (metis Cons.prems(4) a a'carr assoc_l_cancel insert_subset list.set(2) monoid.multlist_closed monoid_axioms) + then have a'fs: "wfactors G as a'" - apply (rule wfactorsE[OF afs], rule wfactorsI, simp) - by (metis a a'carr ahcarr ascarr assoc_l_cancel factorsI factors_def factors_mult_single list.set_intros(1) list.set_intros(2) multlist_closed) - from afs have ahirr: "irreducible G ah" - by (elim wfactorsE) simp - with ascarr have ahprime: "prime G ah" - by (intro irreducible_prime ahcarr) - - note carr [simp] = acarr ahcarr ascarr as'carr a'carr - + by (meson Cons.prems(1) set_subset_Cons subset_iff wfactorsE wfactorsI) + then have ahirr: "irreducible G ah" + by (meson Cons.prems(1) list.set_intros(1) wfactorsE) + with Cons have ahprime: "prime G ah" + by (simp add: irreducible_prime) note ahdvda - also from afs' have "a divides (foldr (\) as' \)" - by (elim wfactorsE associatedE, simp) + also have "a divides (foldr (\) as' \)" + by (meson Cons.prems(2) associatedE wfactorsE) finally have "ah divides (foldr (\) as' \)" - by simp + using Cons.prems(4) by auto with ahprime have "\i carrier G" - unfolding set_conv_nth by force - note carr = carr asicarr - - from ahdvd obtain x where "x \ carrier G" and asi: "as'!i = ah \ x" + then obtain x where "x \ carrier G" and asi: "as'!i = ah \ x" by blast - with carr irrasi[simplified asi] have asiah: "as'!i \ ah" - by (metis ahprime associatedI2 irreducible_prodE primeE) + have irrasi: "irreducible G (as'!i)" + using nth_mem[OF len] wfactorsE + by (metis Cons.prems(2)) + have asicarr[simp]: "as'!i \ carrier G" + using len \set as' \ carrier G\ nth_mem by blast + have asiah: "as'!i \ ah" + by (metis \ah \ carrier G\ \x \ carrier G\ asi irrasi ahprime associatedI2 irreducible_prodE primeE) note setparts = set_take_subset[of i as'] set_drop_subset[of "Suc i" as'] - note partscarr [simp] = setparts[THEN subset_trans[OF _ as'carr]] - note carr = carr partscarr - have "\aa_1. aa_1 \ carrier G \ wfactors G (take i as') aa_1" - by (meson afs' in_set_takeD partscarr(1) wfactorsE wfactors_prod_exists) - then obtain aa_1 where aa1carr: "aa_1 \ carrier G" and aa1fs: "wfactors G (take i as') aa_1" + using Cons + by (metis setparts(1) subset_trans in_set_takeD wfactorsE wfactors_prod_exists) + then obtain aa_1 where aa1carr [simp]: "aa_1 \ carrier G" and aa1fs: "wfactors G (take i as') aa_1" by auto - - have "\aa_2. aa_2 \ carrier G \ wfactors G (drop (Suc i) as') aa_2" - by (meson afs' in_set_dropD partscarr(2) wfactors_def wfactors_prod_exists) - then obtain aa_2 where aa2carr: "aa_2 \ carrier G" + obtain aa_2 where aa2carr [simp]: "aa_2 \ carrier G" and aa2fs: "wfactors G (drop (Suc i) as') aa_2" - by auto - - note carr = carr aa1carr[simp] aa2carr[simp] - - from aa1fs aa2fs - have v1: "wfactors G (take i as' @ drop (Suc i) as') (aa_1 \ aa_2)" - by (intro wfactors_mult, simp+) - then have v1': "wfactors G (as'!i # take i as' @ drop (Suc i) as') (as'!i \ (aa_1 \ aa_2))" - using irrasi wfactors_mult_single by auto - from aa2carr carr aa1fs aa2fs have "wfactors G (as'!i # drop (Suc i) as') (as'!i \ aa_2)" - by (metis irrasi wfactors_mult_single) - with len carr aa1carr aa2carr aa1fs + by (metis Cons.prems(2) Cons.prems(5) subset_code(1) in_set_dropD wfactors_def wfactors_prod_exists) + + have set_drop: "set (drop (Suc i) as') \ carrier G" + using Cons.prems(5) setparts(2) by blast + moreover have set_take: "set (take i as') \ carrier G" + using Cons.prems(5) setparts by auto + moreover have v1: "wfactors G (take i as' @ drop (Suc i) as') (aa_1 \ aa_2)" + using aa1fs aa2fs \set as' \ carrier G\ by (force simp add: dest: in_set_takeD in_set_dropD) + ultimately have v1': "wfactors G (as'!i # take i as' @ drop (Suc i) as') (as'!i \ (aa_1 \ aa_2))" + using irrasi wfactors_mult_single + by (simp add: irrasi v1 wfactors_mult_single) + have "wfactors G (as'!i # drop (Suc i) as') (as'!i \ aa_2)" + by (simp add: aa2fs irrasi set_drop wfactors_mult_single) + with len aa1carr aa2carr aa1fs have v2: "wfactors G (take i as' @ as'!i # drop (Suc i) as') (aa_1 \ (as'!i \ aa_2))" - using wfactors_mult by auto + using wfactors_mult by (simp add: set_take set_drop) from len have as': "as' = (take i as' @ as'!i # drop (Suc i) as')" by (simp add: Cons_nth_drop_Suc) - with carr have eer: "essentially_equal G (take i as' @ as'!i # drop (Suc i) as') as'" - by simp - with v2 afs' carr aa1carr aa2carr nth_mem[OF len] have "aa_1 \ (as'!i \ aa_2) \ a" - by (metis as' ee_wfactorsD m_closed) + have eer: "essentially_equal G (take i as' @ as'!i # drop (Suc i) as') as'" + using Cons.prems(5) as' by auto + with v2 aa1carr aa2carr nth_mem[OF len] have "aa_1 \ (as'!i \ aa_2) \ a" + using Cons.prems as' comm_monoid_cancel.ee_wfactorsD is_comm_monoid_cancel by fastforce then have t1: "as'!i \ (aa_1 \ aa_2) \ a" by (metis aa1carr aa2carr asicarr m_lcomm) - from carr asiah have "ah \ (aa_1 \ aa_2) \ as'!i \ (aa_1 \ aa_2)" - by (metis associated_sym m_closed mult_cong_l) + from asiah have "ah \ (aa_1 \ aa_2) \ as'!i \ (aa_1 \ aa_2)" + by (simp add: \ah \ carrier G\ associated_sym mult_cong_l) also note t1 - finally have "ah \ (aa_1 \ aa_2) \ a" by simp - - with carr aa1carr aa2carr a'carr nth_mem[OF len] have a': "aa_1 \ aa_2 \ a'" - by (simp add: a, fast intro: assoc_l_cancel[of ah _ a']) - + finally have "ah \ (aa_1 \ aa_2) \ a" + using Cons.prems(3) carr_ah aa1carr aa2carr by blast + with aa1carr aa2carr a'carr nth_mem[OF len] have a': "aa_1 \ aa_2 \ a'" + using a assoc_l_cancel carr_ah(1) by blast note v1 also note a' finally have "wfactors G (take i as' @ drop (Suc i) as') a'" - by simp - - from a'fs this carr have "essentially_equal G as (take i as' @ drop (Suc i) as')" - by (intro ih[of a']) simp - then have ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')" - by (elim essentially_equalE) (fastforce intro: essentially_equalI) - - from carr have ee2: "essentially_equal G (ah # take i as' @ drop (Suc i) as') + by (simp add: a'carr set_drop set_take) + from a'fs this have "essentially_equal G as (take i as' @ drop (Suc i) as')" + using Cons.hyps a'carr carr_ah(2) set_drop set_take by auto + with carr_ah have ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')" + by (auto simp: essentially_equal_def) + have ee2: "essentially_equal G (ah # take i as' @ drop (Suc i) as') (as' ! i # take i as' @ drop (Suc i) as')" proof (intro essentially_equalI) show "ah # take i as' @ drop (Suc i) as' <~~> ah # take i as' @ drop (Suc i) as'" by simp next show "ah # take i as' @ drop (Suc i) as' [\] as' ! i # take i as' @ drop (Suc i) as'" - by (simp add: list_all2_append) (simp add: asiah[symmetric]) + by (simp add: asiah associated_sym set_drop set_take) qed note ee1 also note ee2 also have "essentially_equal G (as' ! i # take i as' @ drop (Suc i) as') (take i as' @ as' ! i # drop (Suc i) as')" - by (metis as' as'carr listassoc_refl essentially_equalI perm_append_Cons) + by (metis Cons.prems(5) as' essentially_equalI listassoc_refl perm_append_Cons) finally have "essentially_equal G (ah # as) (take i as' @ as' ! i # drop (Suc i) as')" - by simp - then show "essentially_equal G (ah # as) as'" - by (subst as') - qed + using Cons.prems(4) set_drop set_take by auto + then show ?case + using as' by auto qed -lemma (in primeness_condition_monoid) wfactors_unique: - assumes "wfactors G as a" "wfactors G as' a" - and "a \ carrier G" "set as \ carrier G" "set as' \ carrier G" - shows "essentially_equal G as as'" - by (rule wfactors_unique__hlp_induct[rule_format, of a]) (simp add: assms) - subsubsection \Application to factorial monoids\ @@ -2841,7 +2774,6 @@ by blast note [simp] = acarr bcarr ccarr ascarr cscarr - assume b: "b = a \ c" from afs cfs have "wfactors G (as@cs) (a \ c)" by (intro wfactors_mult) simp_all @@ -2918,9 +2850,7 @@ apply unfold_locales apply (rule wfUNIVI) apply (rule measure_induct[of "factorcount G"]) - apply simp - apply (metis properfactor_fcount) - done + using properfactor_fcount by auto sublocale factorial_monoid \ primeness_condition_monoid by standard (rule irreducible_prime) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Embedded_Algebras.thy --- a/src/HOL/Algebra/Embedded_Algebras.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Embedded_Algebras.thy Sat Aug 11 16:02:55 2018 +0200 @@ -187,7 +187,7 @@ corollary Span_is_add_subgroup: "set Us \ carrier R \ subgroup (Span K Us) (add_monoid R)" - using line_extension_is_subgroup add.normal_invE(1)[OF add.one_is_normal] by (induct Us) (auto) + using line_extension_is_subgroup normal_imp_subgroup[OF add.one_is_normal] by (induct Us) (auto) lemma line_extension_smult_closed: assumes "\k v. \ k \ K; v \ E \ \ k \ v \ E" and "E \ carrier R" "a \ carrier R" @@ -246,7 +246,7 @@ lemma line_extension_of_combine_set: assumes "u \ carrier R" - shows "line_extension K u { combine Ks Us | Ks. set Ks \ K } = + shows "line_extension K u { combine Ks Us | Ks. set Ks \ K } = { combine Ks (u # Us) | Ks. set Ks \ K }" (is "?line_extension = ?combinations") proof @@ -292,7 +292,7 @@ lemma line_extension_of_combine_set_length_version: assumes "u \ carrier R" - shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \ set Ks \ K } = + shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \ set Ks \ K } = { combine Ks (u # Us) | Ks. length Ks = length (u # Us) \ set Ks \ K }" (is "?line_extension = ?combinations") proof @@ -329,16 +329,16 @@ assumes "set Us \ carrier R" and "a \ carrier R" shows "a \ Span K Us \ (\k \ K - { \ }. \Ks. set Ks \ K \ combine (k # Ks) (a # Us) = \)" (is "?in_Span \ ?exists_combine") -proof +proof assume "?in_Span" then obtain Ks where Ks: "set Ks \ K" "a = combine Ks Us" using Span_eq_combine_set[OF assms(1)] by auto hence "((\ \) \ a) \ a = combine ((\ \) # Ks) (a # Us)" by auto moreover have "((\ \) \ a) \ a = \" - using assms(2) l_minus l_neg by auto + using assms(2) l_minus l_neg by auto moreover have "\ \ \ \" - using subfieldE(6)[OF K] l_neg by force + using subfieldE(6)[OF K] l_neg by force ultimately show "?exists_combine" using subring_props(3,5) Ks(1) by (force simp del: combine.simps) next @@ -391,14 +391,14 @@ proof (induct Ks Us rule: combine.induct) case (1 k Ks u Us) hence "k \ K" and "u \ set (u # Us)" by auto - hence "k \ u \ E" + hence "k \ u \ E" using 1(4) unfolding set_mult_def by auto moreover have "K <#> set Us \ E" using 1(4) unfolding set_mult_def by auto hence "combine Ks Us \ E" using 1 by auto ultimately show ?case - using add.subgroupE(4)[OF assms(2)] by auto + using add.subgroupE(4)[OF assms(2)] by auto next case "2_1" thus ?case using subgroup.one_closed[OF assms(2)] by auto @@ -436,7 +436,7 @@ hence "combine [ k ] (u # Us) \ Span K (u # Us)" using Span_eq_combine_set[OF Cons(2)] by (auto simp del: combine.simps) moreover have "k \ carrier R" and "u \ carrier R" - using Cons(2) k subring_props(1) by (blast, auto) + using Cons(2) k subring_props(1) by (blast, auto) ultimately show "k \ u \ Span K (u # Us)" by (auto simp del: Span.simps) qed @@ -455,7 +455,7 @@ corollary Span_same_set: assumes "set Us \ carrier R" shows "set Us = set Vs \ Span K Us = Span K Vs" - using Span_eq_generate assms by auto + using Span_eq_generate assms by auto corollary Span_incl: "set Us \ carrier R \ K <#> (set Us) \ Span K Us" using Span_eq_generate generate.incl[of _ _ "add_monoid R"] by auto @@ -583,7 +583,7 @@ moreover have "Span K Us \ Span K (u # Us)" using mono_Span independent_in_carrier[OF assms] by auto ultimately show ?thesis - using independent_backwards(1)[OF assms] by auto + using independent_backwards(1)[OF assms] by auto qed corollary independent_replacement: @@ -624,7 +624,7 @@ from assms show "Span K Us \ Span K Vs = { \ }" proof (induct Us rule: list.induct) case Nil thus ?case - using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp + using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp next case (Cons u Us) hence IH: "Span K Us \ Span K Vs = {\}" by auto @@ -653,7 +653,7 @@ hence "k \ u = (\ u') \ v'" using in_carrier(1) k(2) u'(2) v'(2) add.m_comm r_neg1 by auto hence "k \ u \ Span K (Us @ Vs)" - using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) + using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) Span_append_eq_set_add[OF in_carrier(2-3)] unfolding set_add_def' by blast hence "u \ Span K (Us @ Vs)" using Cons(2) Span_m_inv_simprule[OF _ _ in_carrier(1), of "Us @ Vs" k] @@ -678,7 +678,7 @@ hence in_carrier: "u \ carrier R" "set Us \ carrier R" "set Vs \ carrier R" "set (u # Us) \ carrier R" using Cons(2-3)[THEN independent_in_carrier] by auto - hence "Span K Us \ Span K (u # Us)" + hence "Span K Us \ Span K (u # Us)" using mono_Span by auto hence "Span K Us \ Span K Vs = { \ }" using Cons(4) Span_subgroup_props(2)[OF in_carrier(2)] by auto @@ -733,7 +733,7 @@ hence "combine Ks' Us = \" using combine_in_carrier[OF _ Us, of Ks'] Ks' u Cons(3) subring_props(1) unfolding Ks by auto hence "set (take (length Us) Ks') \ { \ }" - using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp + using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp thus ?thesis using k_zero unfolding Ks by auto qed @@ -878,10 +878,10 @@ case (Cons u Us) then obtain Vs' Vs'' where Vs: "Vs = Vs' @ (u # Vs'')" by (metis list.set_intros(1) split_list) - + have in_carrier: "u \ carrier R" "set Us \ carrier R" - using independent_in_carrier[OF Cons(2)] by auto - + using independent_in_carrier[OF Cons(2)] by auto + have "distinct Vs" using Cons(3-4) independent_distinct[OF Cons(2)] by (metis card_distinct distinct_card) @@ -905,7 +905,7 @@ shows "\Vs'. set Vs' \ set Vs \ length Vs' = length Us' \ independent K (Vs' @ Us)" using assms proof (induct "length Us'" arbitrary: Us' Us) - case 0 thus ?case by auto + case 0 thus ?case by auto next case (Suc n) then obtain u Us'' where Us'': "Us' = Us'' @ [u]" @@ -1074,9 +1074,9 @@ using space_subgroup_props(1)[OF assms(1)] li_Cons[OF _ v(2) step(4)] by auto then obtain Vs where "length (Vs @ (v # Us)) = n" "independent K (Vs @ (v # Us))" "Span K (Vs @ (v # Us)) = E" - using step(3)[of "v # Us"] step(1-2,4-6) v by auto + using step(3)[of "v # Us"] step(1-2,4-6) v by auto thus ?case - by (metis append.assoc append_Cons append_Nil) + by (metis append.assoc append_Cons append_Nil) qed } note aux_lemma = this have "length Us \ n" @@ -1119,7 +1119,7 @@ hence in_carrier: "set Us \ carrier R" "set (Vs @ Bs) \ carrier R" using independent_in_carrier[OF Us(2)] independent_in_carrier[OF Vs(2)] by auto hence "Span K Us \ (Span K (Vs @ Bs)) \ Span K Bs" - using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto + using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto hence "Span K Us \ (Span K (Vs @ Bs)) \ { \ }" using independent_split(3)[OF Us(2)] by blast hence "Span K Us \ (Span K (Vs @ Bs)) = { \ }" @@ -1147,7 +1147,7 @@ ultimately show "v \ (Span K Us) <+>\<^bsub>R\<^esub> F" using u1' unfolding set_add_def' by auto qed - ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" + ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" using Span_append_eq_set_add[OF in_carrier] Vs(3) by auto thus ?thesis using dim by simp @@ -1169,7 +1169,7 @@ by (metis One_nat_def length_0_conv length_Suc_conv) have in_carrier: "set (map (\u'. u' \ u) Us) \ carrier R" using Us(1) u(1) by (induct Us) (auto) - + have li: "independent K (map (\u'. u' \ u) Us)" proof (rule trivial_combine_imp_independent[OF assms(1) in_carrier]) fix Ks assume Ks: "set Ks \ K" and "combine Ks (map (\u'. u' \ u) Us) = \" @@ -1244,7 +1244,7 @@ ultimately have "dimension (n * Suc m) K (Span F [ v ] <+>\<^bsub>R\<^esub> Span F Vs')" using dimension_direct_sum_space[OF assms(1) _ _ inter] by auto thus "dimension (n * Suc m) K E" - using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto + using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto qed @@ -1271,14 +1271,14 @@ hence "combine Ks Us = (combine (take (length Us) Ks) Us) \ \" using combine_append[OF _ _ assms(2), of "take (length Us) Ks" "drop (length Us) Ks" "[]"] len by auto also have " ... = combine (take (length Us) Ks) Us" - using combine_in_carrier[OF set_t assms(2)] by auto + using combine_in_carrier[OF set_t assms(2)] by auto finally show "combine Ks Us = combine (take (length Us) Ks) Us" . qed *) (* lemma combine_normalize: - assumes "set Ks \ K" "set Us \ carrier R" "a = combine Ks Us" + assumes "set Ks \ K" "set Us \ carrier R" "a = combine Ks Us" shows "\Ks'. set Ks' \ K \ length Ks' = length Us \ a = combine Ks' Us" proof (cases "length Ks \ length Us") assume "\ length Ks \ length Us" @@ -1291,12 +1291,12 @@ next assume len: "length Ks \ length Us" have Ks: "set Ks \ carrier R" and set_r: "set (replicate (length Us - length Ks) \) \ carrier R" - using assms subring_props(1) zero_closed by (metis dual_order.trans, auto) + using assms subring_props(1) zero_closed by (metis dual_order.trans, auto) moreover have set_t: "set (take (length Ks) Us) \ carrier R" and set_d: "set (drop (length Ks) Us) \ carrier R" using assms(2) len dual_order.trans by (metis set_take_subset, metis set_drop_subset) - ultimately + ultimately have "combine (Ks @ (replicate (length Us - length Ks) \)) Us = (combine Ks (take (length Ks) Us)) \ (combine (replicate (length Us - length Ks) \) (drop (length Ks) Us))" diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Generated_Groups.thy --- a/src/HOL/Algebra/Generated_Groups.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Generated_Groups.thy Sat Aug 11 16:02:55 2018 +0200 @@ -466,7 +466,7 @@ shows "derived G H \ G" unfolding derived_def proof (rule normal_generateI) show "(\h1 \ H. \h2 \ H. { h1 \ h2 \ inv h1 \ inv h2 }) \ carrier G" - using subgroup.subset assms normal_invE(1) by blast + using subgroup.subset assms normal_imp_subgroup by blast next show "\h g. \ h \ derived_set G H; g \ carrier G \ \ g \ h \ inv g \ derived_set G H" proof - @@ -474,7 +474,7 @@ then obtain h1 h2 where h1: "h1 \ H" "h1 \ carrier G" and h2: "h2 \ H" "h2 \ carrier G" and h: "h = h1 \ h2 \ inv h1 \ inv h2" - using subgroup.subset assms normal_invE(1) by blast + using subgroup.subset assms normal_imp_subgroup by blast hence "g \ h \ inv g = g \ h1 \ (inv g \ g) \ h2 \ (inv g \ g) \ inv h1 \ (inv g \ g) \ inv h2 \ inv g" by (simp add: g m_assoc) @@ -486,8 +486,8 @@ have "g \ h \ inv g = (g \ h1 \ inv g) \ (g \ h2 \ inv g) \ inv (g \ h1 \ inv g) \ inv (g \ h2 \ inv g)" by (simp add: g h1 h2 inv_mult_group m_assoc) - moreover have "g \ h1 \ inv g \ H" by (simp add: assms normal_invE(2) g h1) - moreover have "g \ h2 \ inv g \ H" by (simp add: assms normal_invE(2) g h2) + moreover have "g \ h1 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h1) + moreover have "g \ h2 \ inv g \ H" by (simp add: assms normal.inv_op_closed2 g h2) ultimately show "g \ h \ inv g \ derived_set G H" by blast qed qed diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Group.thy Sat Aug 11 16:02:55 2018 +0200 @@ -763,13 +763,13 @@ and "subgroup I K" shows "subgroup (H \ I) (G \\ K)" proof (intro group.group_incl_imp_subgroup[OF DirProd_group[OF assms(1)assms(3)]]) - have "H \ carrier G" "I \ carrier K" using subgroup.subset assms apply blast+. + have "H \ carrier G" "I \ carrier K" using subgroup.subset assms by blast+ thus "(H \ I) \ carrier (G \\ K)" unfolding DirProd_def by auto have "Group.group ((G\carrier := H\) \\ (K\carrier := I\))" using DirProd_group[OF subgroup.subgroup_is_group[OF assms(2)assms(1)] subgroup.subgroup_is_group[OF assms(4)assms(3)]]. moreover have "((G\carrier := H\) \\ (K\carrier := I\)) = ((G \\ K)\carrier := H \ I\)" - unfolding DirProd_def using assms apply simp. + unfolding DirProd_def using assms by simp ultimately show "Group.group ((G \\ K)\carrier := H \ I\)" by simp qed @@ -1054,7 +1054,7 @@ lemma (in group_hom) img_is_subgroup: "subgroup (h ` (carrier G)) H" apply (rule subgroupI) apply (auto simp add: image_subsetI) - apply (metis (no_types, hide_lams) G.inv_closed hom_inv image_iff) + apply (metis G.inv_closed hom_inv image_iff) by (metis G.monoid_axioms hom_mult image_eqI monoid.m_closed) lemma (in group_hom) subgroup_img_is_subgroup: @@ -1157,9 +1157,8 @@ show "monoid (G\carrier := H\)" using submonoid.submonoid_is_monoid assms comm_monoid_axioms comm_monoid_def by blast show "\x y. x \ carrier (G\carrier := H\) \ y \ carrier (G\carrier := H\) - \ x \\<^bsub>G\carrier := H\\<^esub> y = y \\<^bsub>G\carrier := H\\<^esub> x" apply simp - using assms comm_monoid_axioms_def submonoid.mem_carrier - by (metis m_comm) + \ x \\<^bsub>G\carrier := H\\<^esub> y = y \\<^bsub>G\carrier := H\\<^esub> x" + by simp (meson assms m_comm submonoid.mem_carrier) qed locale comm_group = comm_monoid + group diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/Ideal.thy Sat Aug 11 16:02:55 2018 +0200 @@ -128,10 +128,9 @@ proof - interpret additive_subgroup I R by fact interpret cring R by fact - show ?thesis apply intro_locales - apply (intro ideal_axioms.intro) - apply (erule (1) I_l_closed) - apply (erule (1) I_r_closed) + show ?thesis + apply intro_locales + apply (simp add: I_l_closed I_r_closed ideal_axioms_def) by (simp add: I_notcarr I_prime primeideal_axioms.intro) qed diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Algebra/RingHom.thy --- a/src/HOL/Algebra/RingHom.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Algebra/RingHom.thy Sat Aug 11 16:02:55 2018 +0200 @@ -20,11 +20,10 @@ by standard (rule homh) sublocale ring_hom_ring \ abelian_group?: abelian_group_hom R S -apply (intro abelian_group_homI R.is_abelian_group S.is_abelian_group) -apply (intro group_hom.intro group_hom_axioms.intro R.a_group S.a_group) -apply (insert homh, unfold hom_def ring_hom_def) -apply simp -done +proof + show "h \ hom (add_monoid R) (add_monoid S)" + using homh by (simp add: hom_def ring_hom_def) +qed lemma (in ring_hom_ring) is_ring_hom_ring: "ring_hom_ring R S h" @@ -33,8 +32,7 @@ lemma ring_hom_ringI: fixes R (structure) and S (structure) assumes "ring R" "ring S" - assumes (* morphism: "h \ carrier R \ carrier S" *) - hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" + assumes hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" and compatible_mult: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_add: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" @@ -42,13 +40,12 @@ proof - interpret ring R by fact interpret ring S by fact - show ?thesis apply unfold_locales -apply (unfold ring_hom_def, safe) - apply (simp add: hom_closed Pi_def) - apply (erule (1) compatible_mult) - apply (erule (1) compatible_add) -apply (rule compatible_one) -done + show ?thesis + proof + show "h \ ring_hom R S" + unfolding ring_hom_def + by (auto simp: compatible_mult compatible_add compatible_one hom_closed) + qed qed lemma ring_hom_ringI2: @@ -58,11 +55,11 @@ proof - interpret R: ring R by fact interpret S: ring S by fact - show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro) - apply (rule R.is_ring) - apply (rule S.is_ring) - apply (rule h) - done + show ?thesis + proof + show "h \ ring_hom R S" + using h . + qed qed lemma ring_hom_ringI3: @@ -75,13 +72,11 @@ interpret abelian_group_hom R S h by fact interpret R: ring R by fact interpret S: ring S by fact - show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro, rule R.is_ring, rule S.is_ring) - apply (insert group_hom.homh[OF a_group_hom]) - apply (unfold hom_def ring_hom_def, simp) - apply safe - apply (erule (1) compatible_mult) - apply (rule compatible_one) - done + show ?thesis + proof + show "h \ ring_hom R S" + unfolding ring_hom_def by (auto simp: compatible_one compatible_mult) + qed qed lemma ring_hom_cringI: @@ -91,21 +86,22 @@ interpret ring_hom_ring R S h by fact interpret R: cring R by fact interpret S: cring S by fact - show ?thesis by (intro ring_hom_cring.intro ring_hom_cring_axioms.intro) - (rule R.is_cring, rule S.is_cring, rule homh) + show ?thesis + proof + show "h \ ring_hom R S" + by (simp add: homh) + qed qed subsection \The Kernel of a Ring Homomorphism\ \ \the kernel of a ring homomorphism is an ideal\ -lemma (in ring_hom_ring) kernel_is_ideal: - shows "ideal (a_kernel R S h) R" -apply (rule idealI) - apply (rule R.is_ring) - apply (rule additive_subgroup.a_subgroup[OF additive_subgroup_a_kernel]) - apply (unfold a_kernel_def', simp+) -done +lemma (in ring_hom_ring) kernel_is_ideal: "ideal (a_kernel R S h) R" + apply (rule idealI [OF R.is_ring]) + apply (rule additive_subgroup.a_subgroup[OF additive_subgroup_a_kernel]) + apply (auto simp: a_kernel_def') + done text \Elements of the kernel are mapped to zero\ lemma (in abelian_group_hom) kernel_zero [simp]: @@ -174,29 +170,10 @@ corollary (in ring_hom_ring) rcos_eq_homeq: assumes acarr: "a \ carrier R" shows "(a_kernel R S h) +> a = {x \ carrier R. h x = h a}" - apply rule defer 1 - apply clarsimp defer 1 -proof +proof - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) - - fix x - assume xrcos: "x \ a_kernel R S h +> a" - from acarr and this - have xcarr: "x \ carrier R" - by (rule a_elemrcos_carrier) - - from xrcos - have "h x = h a" by (rule rcos_imp_homeq[OF acarr]) - from xcarr and this - show "x \ {x \ carrier R. h x = h a}" by fast -next - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) - - fix x - assume xcarr: "x \ carrier R" - and hx: "h x = h a" - from acarr xcarr hx - show "x \ a_kernel R S h +> a" by (rule homeq_imp_rcos) + show ?thesis + using assms by (auto simp: intro: homeq_imp_rcos rcos_imp_homeq a_elemrcos_carrier) qed lemma (in ring_hom_ring) nat_pow_hom: diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1341,6 +1341,29 @@ lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" using has_contour_integral_trivial contour_integral_unique by blast +lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" + by (auto simp: linepath_def) + +lemma bounded_linear_linepath: + assumes "bounded_linear f" + shows "f (linepath a b x) = linepath (f a) (f b) x" +proof - + interpret f: bounded_linear f by fact + show ?thesis by (simp add: linepath_def f.add f.scale) +qed + +lemma bounded_linear_linepath': + assumes "bounded_linear f" + shows "f \ linepath a b = linepath (f a) (f b)" + using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) + +lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" + by (simp add: linepath_def) + +lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" + by (simp add: linepath_def fun_eq_iff) + + subsection\Relation to subpath construction\ @@ -1501,6 +1524,62 @@ "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) +lemma contour_integral_cong: + assumes "g = g'" "\x. x \ path_image g \ f x = f' x" + shows "contour_integral g f = contour_integral g' f'" + unfolding contour_integral_integral using assms + by (intro integral_cong) (auto simp: path_image_def) + + +text \Contour integral along a segment on the real axis\ + +lemma has_contour_integral_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f has_contour_integral I) (linepath a b) \ + ((\x. f (of_real x)) has_integral I) {Re a..Re b}" +proof - + from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" + by (simp_all add: complex_eq_iff) + from assms have "a \ b" by auto + have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ + ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" + by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) + (insert assms, simp_all add: field_simps scaleR_conv_of_real) + also have "(\x. f (a + b * of_real x - a * of_real x)) = + (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" + using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) + also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ + ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms + by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) + also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def + by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) + finally show ?thesis by simp +qed + +lemma contour_integrable_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f contour_integrable_on linepath a b) \ + (\x. f (of_real x)) integrable_on {Re a..Re b}" + using has_contour_integral_linepath_Reals_iff[OF assms, of f] + by (auto simp: contour_integrable_on_def integrable_on_def) + +lemma contour_integral_linepath_Reals_eq: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" +proof (cases "f contour_integrable_on linepath a b") + case True + thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] + using has_contour_integral_integral has_contour_integral_unique by blast +next + case False + thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + + text\Cauchy's theorem where there's a primitive\ @@ -4875,6 +4954,10 @@ apply (rule derivative_eq_intros | simp)+ done +corollary differentiable_part_circlepath: + "part_circlepath c r a b differentiable at x within A" + using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast + corollary vector_derivative_part_circlepath: "vector_derivative (part_circlepath z r s t) (at x) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" @@ -4923,6 +5006,17 @@ by (fastforce simp add: path_image_def part_circlepath_def) qed +proposition path_image_part_circlepath': + "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" +proof - + have "path_image (part_circlepath z r s t) = + (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" + by (simp add: image_image path_image_def part_circlepath_def) + also have "linepath s t ` {0..1} = closed_segment s t" + by (rule linepath_image_01) + finally show ?thesis by (simp add: cis_conv_exp) +qed + corollary path_image_part_circlepath_subset: "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) @@ -4937,6 +5031,106 @@ by (simp add: dist_norm norm_minus_commute) qed +corollary path_image_part_circlepath_subset': + assumes "r \ 0" + shows "path_image (part_circlepath z r s t) \ sphere z r" +proof (cases "s \ t") + case True + thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp +next + case False + thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms + by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all +qed + +lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" + by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) + +lemma contour_integral_bound_part_circlepath: + assumes "f contour_integrable_on part_circlepath c r a b" + assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" + shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" +proof - + let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * + exp (\ * linepath a b x))" + have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" + proof (rule integral_norm_bound_integral, goal_cases) + case 1 + with assms(1) show ?case + by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) + next + case (3 x) + with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult + by (intro mult_mono) (auto simp: path_image_def) + qed auto + also have "?I = contour_integral (part_circlepath c r a b) f" + by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) + finally show ?thesis by simp +qed + +lemma has_contour_integral_part_circlepath_iff: + assumes "a < b" + shows "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" +proof - + have "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) + (at x within {0..1})) has_integral I) {0..1}" + unfolding has_contour_integral_def .. + also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * + cis (linepath a b x)) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_part_circlepath01) + (simp_all add: cis_conv_exp) + also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * + r * \ * exp (\ * linepath (of_real a) (of_real b) x) * + vector_derivative (linepath (of_real a) (of_real b)) + (at x within {0..1})) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_linepath_within) + (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) + also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) + (linepath (of_real a) (of_real b))" + by (simp add: has_contour_integral_def) + also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms + by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) + finally show ?thesis . +qed + +lemma contour_integrable_part_circlepath_iff: + assumes "a < b" + shows "f contour_integrable_on (part_circlepath c r a b) \ + (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (auto simp: contour_integrable_on_def integrable_on_def + has_contour_integral_part_circlepath_iff) + +lemma contour_integral_part_circlepath_eq: + assumes "a < b" + shows "contour_integral (part_circlepath c r a b) f = + integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" +proof (cases "f contour_integrable_on part_circlepath c r a b") + case True + hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with True show ?thesis + using has_contour_integral_part_circlepath_iff[OF assms] + contour_integral_unique has_integral_integrable_integral by blast +next + case False + hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with False show ?thesis + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + +lemma contour_integral_part_circlepath_reverse: + "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" + by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all + +lemma contour_integral_part_circlepath_reverse': + "b < a \ contour_integral (part_circlepath c r a b) f = + -contour_integral (part_circlepath c r b a) f" + by (rule contour_integral_part_circlepath_reverse) + + proposition finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" proof (cases "w = 0") case True then show ?thesis by auto diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Sat Aug 11 16:02:55 2018 +0200 @@ -41,6 +41,24 @@ lemma linear_cnj: "linear cnj" using bounded_linear.linear[OF bounded_linear_cnj] . +lemma vector_derivative_cnj_within: + assumes "at x within A \ bot" and "f differentiable at x within A" + shows "vector_derivative (\z. cnj (f z)) (at x within A) = + cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D") +proof - + let ?D = "vector_derivative f (at x within A)" + from assms have "(f has_vector_derivative ?D) (at x within A)" + by (subst (asm) vector_derivative_works) + hence "((\x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)" + by (rule has_vector_derivative_cnj) + thus ?thesis using assms by (auto dest: vector_derivative_within) +qed + +lemma vector_derivative_cnj: + assumes "f differentiable at x" + shows "vector_derivative (\z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))" + using assms by (intro vector_derivative_cnj_within) auto + lemma lambda_zero: "(\h::'a::mult_zero. 0) = ( * ) 0" by auto @@ -286,6 +304,35 @@ "f holomorphic_on s \ g holomorphic_on t \ f ` s \ t \ (g o f) holomorphic_on s" by (metis holomorphic_on_compose holomorphic_on_subset) +lemma holomorphic_on_balls_imp_entire: + assumes "\bdd_above A" "\r. r \ A \ f holomorphic_on ball c r" + shows "f holomorphic_on B" +proof (rule holomorphic_on_subset) + show "f holomorphic_on UNIV" unfolding holomorphic_on_def + proof + fix z :: complex + from \\bdd_above A\ obtain r where r: "r \ A" "r > norm (z - c)" + by (meson bdd_aboveI not_le) + with assms(2) have "f holomorphic_on ball c r" by blast + moreover from r have "z \ ball c r" by (auto simp: dist_norm norm_minus_commute) + ultimately show "f field_differentiable at z" + by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"]) + qed +qed auto + +lemma holomorphic_on_balls_imp_entire': + assumes "\r. r > 0 \ f holomorphic_on ball c r" + shows "f holomorphic_on B" +proof (rule holomorphic_on_balls_imp_entire) + { + fix M :: real + have "\x. x > max M 0" by (intro gt_ex) + hence "\x>0. x > M" by auto + } + thus "\bdd_above {(0::real)<..}" unfolding bdd_above_def + by (auto simp: not_le) +qed (insert assms, auto) + lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \ (\z. -(f z)) holomorphic_on s" by (metis field_differentiable_minus holomorphic_on_def) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Sat Aug 11 16:02:55 2018 +0200 @@ -176,6 +176,16 @@ lemma holomorphic_on_cos: "cos holomorphic_on S" by (simp add: field_differentiable_within_cos holomorphic_on_def) +lemma holomorphic_on_sin' [holomorphic_intros]: + assumes "f holomorphic_on A" + shows "(\x. sin (f x)) holomorphic_on A" + using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def) + +lemma holomorphic_on_cos' [holomorphic_intros]: + assumes "f holomorphic_on A" + shows "(\x. cos (f x)) holomorphic_on A" + using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) + subsection\Get a nice real/imaginary separation in Euler's formula\ lemma Euler: "exp(z) = of_real(exp(Re z)) * @@ -1308,6 +1318,11 @@ lemma holomorphic_on_Ln [holomorphic_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ Ln holomorphic_on S" by (simp add: field_differentiable_within_Ln holomorphic_on_def) +lemma holomorphic_on_Ln' [holomorphic_intros]: + "(\z. z \ A \ f z \ \\<^sub>\\<^sub>0) \ f holomorphic_on A \ (\z. Ln (f z)) holomorphic_on A" + using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \\<^sub>\\<^sub>0"] + by (auto simp: o_def) + lemma tendsto_Ln [tendsto_intros]: fixes L F assumes "(f \ L) F" "L \ \\<^sub>\\<^sub>0" diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Sat Aug 11 16:02:55 2018 +0200 @@ -4612,6 +4612,93 @@ unfolding absolutely_integrable_restrict_UNIV . qed +lemma uniform_limit_set_lebesgue_integral_at_top: + fixes f :: "'a \ real \ 'b::{banach, second_countable_topology}" + and g :: "real \ real" + assumes bound: "\x y. x \ A \ y \ a \ norm (f x y) \ g y" + assumes integrable: "set_integrable M {a..} g" + assumes measurable: "\x. x \ A \ set_borel_measurable M {a..} (f x)" + assumes "sets borel \ sets M" + shows "uniform_limit A (\b x. LINT y:{a..b}|M. f x y) (\x. LINT y:{a..}|M. f x y) at_top" +proof (cases "A = {}") + case False + then obtain x where x: "x \ A" by auto + have g_nonneg: "g y \ 0" if "y \ a" for y + proof - + have "0 \ norm (f x y)" by simp + also have "\ \ g y" using bound[OF x that] by simp + finally show ?thesis . + qed + + have integrable': "set_integrable M {a..} (\y. f x y)" if "x \ A" for x + unfolding set_integrable_def + proof (rule Bochner_Integration.integrable_bound) + show "integrable M (\x. indicator {a..} x * g x)" + using integrable by (simp add: set_integrable_def) + show "(\y. indicat_real {a..} y *\<^sub>R f x y) \ borel_measurable M" using measurable[OF that] + by (simp add: set_borel_measurable_def) + show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \ norm (indicat_real {a..} y * g y)" + using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) + qed + + show ?thesis + proof (rule uniform_limitI) + fix e :: real assume e: "e > 0" + have sets [intro]: "A \ sets M" if "A \ sets borel" for A + using that assms by blast + + have "((\b. LINT y:{a..b}|M. g y) \ (LINT y:{a..}|M. g y)) at_top" + by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto + with e obtain b0 :: real where b0: "\b\b0. \(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" + by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) + define b where "b = max a b0" + have "a \ b" by (simp add: b_def) + from b0 have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" + by (auto simp: b_def) + also have "{a..} = {a..b} \ {b<..}" by (auto simp: b_def) + also have "\(LINT y:\|M. g y) - (LINT y:{a..b}|M. g y)\ = \(LINT y:{b<..}|M. g y)\" + using \a \ b\ by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) + also have "(LINT y:{b<..}|M. g y) \ 0" + using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def + by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) + hence "\(LINT y:{b<..}|M. g y)\ = (LINT y:{b<..}|M. g y)" by simp + finally have less: "(LINT y:{b<..}|M. g y) < e" . + + have "eventually (\b. b \ b0) at_top" by (rule eventually_ge_at_top) + moreover have "eventually (\b. b \ a) at_top" by (rule eventually_ge_at_top) + ultimately show "eventually (\b. \x\A. + dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top" + proof eventually_elim + case (elim b) + show ?case + proof + fix x assume x: "x \ A" + have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) = + norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))" + by (simp add: dist_norm norm_minus_commute) + also have "{a..} = {a..b} \ {b<..}" using elim by auto + also have "(LINT y:\|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)" + using elim x + by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) + also have "norm \ \ (LINT y:{b<..}|M. norm (f x y))" using elim x + by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto + also have "\ \ (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg + by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] + set_integrable_subset[OF integrable]) auto + also have "(LINT y:{b<..}|M. g y) \ 0" + using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def + by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) + hence "(LINT y:{b<..}|M. g y) = \(LINT y:{b<..}|M. g y)\" by simp + also have "\ = \(LINT y:{a..b} \ {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\" + using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) + also have "{a..b} \ {b<..} = {a..}" using elim by auto + also have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" + using b0 elim by blast + finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" . + qed + qed + qed +qed auto diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/FPS_Convergence.thy --- a/src/HOL/Analysis/FPS_Convergence.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/FPS_Convergence.thy Sat Aug 11 16:02:55 2018 +0200 @@ -193,6 +193,24 @@ by (subst analytic_on_open) auto qed +lemma continuous_eval_fps [continuous_intros]: + fixes z :: "'a::{real_normed_field,banach}" + assumes "norm z < fps_conv_radius F" + shows "continuous (at z within A) (eval_fps F)" +proof - + from ereal_dense2[OF assms] obtain K :: real where K: "norm z < K" "K < fps_conv_radius F" + by auto + have "0 \ norm z" by simp + also have "norm z < K" by fact + finally have "K > 0" . + from K and \K > 0\ have "summable (\n. fps_nth F n * of_real K ^ n)" + by (intro summable_fps) auto + from this have "isCont (eval_fps F) z" unfolding eval_fps_def + by (rule isCont_powser) (use K in auto) + thus "continuous (at z within A) (eval_fps F)" + by (simp add: continuous_at_imp_continuous_within) +qed + subsection%unimportant \Lower bounds on radius of convergence\ @@ -831,6 +849,20 @@ ultimately show ?thesis using r(1) by (auto simp: has_fps_expansion_def) qed +lemma has_fps_expansion_imp_continuous: + fixes F :: "'a::{real_normed_field,banach} fps" + assumes "f has_fps_expansion F" + shows "continuous (at 0 within A) f" +proof - + from assms have "isCont (eval_fps F) 0" + by (intro continuous_eval_fps) (auto simp: has_fps_expansion_def zero_ereal_def) + also have "?this \ isCont f 0" using assms + by (intro isCont_cong) (auto simp: has_fps_expansion_def) + finally have "isCont f 0" . + thus "continuous (at 0 within A) f" + by (simp add: continuous_at_imp_continuous_within) +qed + lemma has_fps_expansion_const [simp, intro, fps_expansion_intros]: "(\_. c) has_fps_expansion fps_const c" by (auto simp: has_fps_expansion_def) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Gamma_Function.thy --- a/src/HOL/Analysis/Gamma_Function.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Gamma_Function.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1457,6 +1457,15 @@ lemma holomorphic_rGamma [holomorphic_intros]: "rGamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) +lemma holomorphic_rGamma' [holomorphic_intros]: + assumes "f holomorphic_on A" + shows "(\x. rGamma (f x)) holomorphic_on A" +proof - + have "rGamma \ f holomorphic_on A" using assms + by (intro holomorphic_on_compose assms holomorphic_rGamma) + thus ?thesis by (simp only: o_def) +qed + lemma analytic_rGamma: "rGamma analytic_on A" unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_rGamma) @@ -1467,6 +1476,15 @@ lemma holomorphic_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) +lemma holomorphic_Gamma' [holomorphic_intros]: + assumes "f holomorphic_on A" and "\x. x \ A \ f x \ \\<^sub>\\<^sub>0" + shows "(\x. Gamma (f x)) holomorphic_on A" +proof - + have "Gamma \ f holomorphic_on A" using assms + by (intro holomorphic_on_compose assms holomorphic_Gamma) auto + thus ?thesis by (simp only: o_def) +qed + lemma analytic_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_Gamma) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Sat Aug 11 16:02:55 2018 +0200 @@ -687,6 +687,17 @@ apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) done +lemma integrable_on_cnj_iff: + "(\x. cnj (f x)) integrable_on A \ f integrable_on A" + using integrable_linear[OF _ bounded_linear_cnj, of f A] + integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] + by (auto simp: o_def) + +lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" + by (cases "f integrable_on A") + (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] + o_def integrable_on_cnj_iff not_integrable_integral) + lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" @@ -3440,6 +3451,64 @@ lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] +lemma integrable_on_affinity: + assumes "m \ 0" "f integrable_on (cbox a b)" + shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" +proof - + from assms obtain I where "(f has_integral I) (cbox a b)" + by (auto simp: integrable_on_def) + from has_integral_affinity[OF this assms(1), of c] show ?thesis + by (auto simp: integrable_on_def) +qed + +lemma has_integral_cmul_iff: + assumes "c \ 0" + shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" + using assms has_integral_cmul[of f I A c] + has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) + +lemma has_integral_affinity': + fixes a :: "'a::euclidean_space" + assumes "(f has_integral i) (cbox a b)" and "m > 0" + shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) + (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" +proof (cases "cbox a b = {}") + case True + hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" + using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) + with True and assms show ?thesis by simp +next + case False + have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) + ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" + using assms by (intro has_integral_affinity) auto + also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = + ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" + by (simp add: image_image algebra_simps) + also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False + by (subst image_smult_cbox) simp_all + also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" + by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) + finally show ?thesis using \m > 0\ by (simp add: field_simps) +qed + +lemma has_integral_affinity_iff: + fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" + assumes "m > 0" + shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) + (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ + (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") +proof + assume ?lhs + from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ + show ?rhs by (simp add: field_simps vector_add_divide_simps) +next + assume ?rhs + from has_integral_affinity'[OF this, of m c] and \m > 0\ + show ?lhs by simp +qed + + subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Path_Connected.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1177,6 +1177,21 @@ unfolding pathfinish_def linepath_def by auto +lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" + by (simp add: linepath_def algebra_simps) + +lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" + by (simp add: linepath_def) + +lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" + by (simp add: linepath_def) + +lemma linepath_0': "linepath a b 0 = a" + by (simp add: linepath_def) + +lemma linepath_1': "linepath a b 1 = b" + by (simp add: linepath_def) + lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) @@ -1200,6 +1215,9 @@ lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) +lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" + by (simp add: linepath_def) + lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Analysis/Set_Integral.thy --- a/src/HOL/Analysis/Set_Integral.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Analysis/Set_Integral.thy Sat Aug 11 16:02:55 2018 +0200 @@ -54,6 +54,15 @@ by (auto simp add: indicator_def) *) +lemma set_integrable_cong: + assumes "M = M'" "A = A'" "\x. x \ A \ f x = f' x" + shows "set_integrable M A f = set_integrable M' A' f'" +proof - + have "(\x. indicator A x *\<^sub>R f x) = (\x. indicator A' x *\<^sub>R f' x)" + using assms by (auto simp: indicator_def) + thus ?thesis by (simp add: set_integrable_def assms) +qed + lemma set_borel_measurable_sets: fixes f :: "_ \ _::real_normed_vector" assumes "set_borel_measurable M X f" "B \ sets borel" "X \ sets M" @@ -925,4 +934,127 @@ then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n]) qed (auto simp add: assms Limsup_bounded) +lemma tendsto_set_lebesgue_integral_at_right: + fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" + assumes "a < b" and sets: "\a'. a' \ {a<..b} \ {a'..b} \ sets M" + and "set_integrable M {a<..b} f" + shows "((\a'. set_lebesgue_integral M {a'..b} f) \ + set_lebesgue_integral M {a<..b} f) (at_right a)" +proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases) + case (1 S) + have eq: "(\n. {S n..b}) = {a<..b}" + proof safe + fix x n assume "x \ {S n..b}" + with 1(1,2)[of n] show "x \ {a<..b}" by auto + next + fix x assume "x \ {a<..b}" + with order_tendstoD[OF \S \ a\, of x] show "x \ (\n. {S n..b})" + by (force simp: eventually_at_top_linorder dest: less_imp_le) + qed + have "(\n. set_lebesgue_integral M {S n..b} f) \ set_lebesgue_integral M (\n. {S n..b}) f" + by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) + with eq show ?case by simp +qed + + +text \ + The next lemmas relate convergence of integrals over an interval to + improper integrals. +\ +lemma tendsto_set_lebesgue_integral_at_left: + fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" + assumes "a < b" and sets: "\b'. b' \ {a.. {a..b'} \ sets M" + and "set_integrable M {a..b'. set_lebesgue_integral M {a..b'} f) \ + set_lebesgue_integral M {a..n. {a..S n}) = {a.. {a..S n}" + with 1(1,2)[of n] show "x \ {a.. {a..S \ b\, of x] show "x \ (\n. {a..S n})" + by (force simp: eventually_at_top_linorder dest: less_imp_le) + qed + have "(\n. set_lebesgue_integral M {a..S n} f) \ set_lebesgue_integral M (\n. {a..S n}) f" + by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) + with eq show ?case by simp +qed + +lemma tendsto_set_lebesgue_integral_at_top: + fixes f :: "real \ 'a::{banach, second_countable_topology}" + assumes sets: "\b. b \ a \ {a..b} \ sets M" + and int: "set_integrable M {a..} f" + shows "((\b. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {a..} f) at_top" +proof (rule tendsto_at_topI_sequentially) + fix X :: "nat \ real" assume "filterlim X at_top sequentially" + show "(\n. set_lebesgue_integral M {a..X n} f) \ set_lebesgue_integral M {a..} f" + unfolding set_lebesgue_integral_def + proof (rule integral_dominated_convergence) + show "integrable M (\x. indicat_real {a..} x *\<^sub>R norm (f x))" + using integrable_norm[OF int[unfolded set_integrable_def]] by simp + show "AE x in M. (\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" + proof + fix x + from \filterlim X at_top sequentially\ + have "eventually (\n. x \ X n) sequentially" + unfolding filterlim_at_top_ge[where c=x] by auto + then show "(\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" + by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono) + qed + fix n show "AE x in M. norm (indicator {a..X n} x *\<^sub>R f x) \ + indicator {a..} x *\<^sub>R norm (f x)" + by (auto split: split_indicator) + next + from int show "(\x. indicat_real {a..} x *\<^sub>R f x) \ borel_measurable M" + by (simp add: set_integrable_def) + next + fix n :: nat + from sets have "{a..X n} \ sets M" by (cases "X n \ a") auto + with int have "set_integrable M {a..X n} f" + by (rule set_integrable_subset) auto + thus "(\x. indicat_real {a..X n} x *\<^sub>R f x) \ borel_measurable M" + by (simp add: set_integrable_def) + qed +qed + +lemma tendsto_set_lebesgue_integral_at_bot: + fixes f :: "real \ 'a::{banach, second_countable_topology}" + assumes sets: "\a. a \ b \ {a..b} \ sets M" + and int: "set_integrable M {..b} f" + shows "((\a. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {..b} f) at_bot" +proof (rule tendsto_at_botI_sequentially) + fix X :: "nat \ real" assume "filterlim X at_bot sequentially" + show "(\n. set_lebesgue_integral M {X n..b} f) \ set_lebesgue_integral M {..b} f" + unfolding set_lebesgue_integral_def + proof (rule integral_dominated_convergence) + show "integrable M (\x. indicat_real {..b} x *\<^sub>R norm (f x))" + using integrable_norm[OF int[unfolded set_integrable_def]] by simp + show "AE x in M. (\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" + proof + fix x + from \filterlim X at_bot sequentially\ + have "eventually (\n. x \ X n) sequentially" + unfolding filterlim_at_bot_le[where c=x] by auto + then show "(\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" + by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono) + qed + fix n show "AE x in M. norm (indicator {X n..b} x *\<^sub>R f x) \ + indicator {..b} x *\<^sub>R norm (f x)" + by (auto split: split_indicator) + next + from int show "(\x. indicat_real {..b} x *\<^sub>R f x) \ borel_measurable M" + by (simp add: set_integrable_def) + next + fix n :: nat + from sets have "{X n..b} \ sets M" by (cases "X n \ b") auto + with int have "set_integrable M {X n..b} f" + by (rule set_integrable_subset) auto + thus "(\x. indicat_real {X n..b} x *\<^sub>R f x) \ borel_measurable M" + by (simp add: set_integrable_def) + qed +qed + end diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Archimedean_Field.thy Sat Aug 11 16:02:55 2018 +0200 @@ -707,6 +707,9 @@ lemma frac_of_int [simp]: "frac (of_int z) = 0" by (simp add: frac_def) +lemma frac_frac [simp]: "frac (frac x) = frac x" + by (simp add: frac_def) + lemma floor_add: "\x + y\ = (if frac x + frac y < 1 then \x\ + \y\ else (\x\ + \y\) + 1)" proof - have "x + y < 1 + (of_int \x\ + of_int \y\) \ \x + y\ = \x\ + \y\" @@ -743,6 +746,14 @@ apply (meson frac_lt_1 less_iff_diff_less_0 not_le not_less_iff_gr_or_eq) done +lemma frac_in_Ints_iff [simp]: "frac x \ \ \ x \ \" +proof safe + assume "frac x \ \" + hence "of_int \x\ + frac x \ \" by auto + also have "of_int \x\ + frac x = x" by (simp add: frac_def) + finally show "x \ \" . +qed (auto simp: frac_def) + subsection \Rounding to the nearest integer\ diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Complex.thy Sat Aug 11 16:02:55 2018 +0200 @@ -623,6 +623,27 @@ lemma sums_cnj: "((\x. cnj(f x)) sums cnj l) \ (f sums l)" by (simp add: sums_def lim_cnj cnj_sum [symmetric] del: cnj_sum) +lemma differentiable_cnj_iff: + "(\z. cnj (f z)) differentiable at x within A \ f differentiable at x within A" +proof + assume "(\z. cnj (f z)) differentiable at x within A" + then obtain D where "((\z. cnj (f z)) has_derivative D) (at x within A)" + by (auto simp: differentiable_def) + from has_derivative_cnj[OF this] show "f differentiable at x within A" + by (auto simp: differentiable_def) +next + assume "f differentiable at x within A" + then obtain D where "(f has_derivative D) (at x within A)" + by (auto simp: differentiable_def) + from has_derivative_cnj[OF this] show "(\z. cnj (f z)) differentiable at x within A" + by (auto simp: differentiable_def) +qed + +lemma has_vector_derivative_cnj [derivative_intros]: + assumes "(f has_vector_derivative f') (at z within A)" + shows "((\z. cnj (f z)) has_vector_derivative cnj f') (at z within A)" + using assms by (auto simp: has_vector_derivative_complex_iff intro: derivative_intros) + subsection \Basic Lemmas\ @@ -778,9 +799,15 @@ lemma sgn_cis [simp]: "sgn (cis a) = cis a" by (simp add: sgn_div_norm) +lemma cis_2pi [simp]: "cis (2 * pi) = 1" + by (simp add: cis.ctr complex_eq_iff) + lemma cis_neq_zero [simp]: "cis a \ 0" by (metis norm_cis norm_zero zero_neq_one) +lemma cis_cnj: "cnj (cis t) = cis (-t)" + by (simp add: complex_eq_iff) + lemma cis_mult: "cis a * cis b = cis (a + b)" by (simp add: complex_eq_iff cos_add sin_add) @@ -802,6 +829,15 @@ lemma cis_pi [simp]: "cis pi = -1" by (simp add: complex_eq_iff) +lemma cis_pi_half[simp]: "cis (pi / 2) = \" + by (simp add: cis.ctr complex_eq_iff) + +lemma cis_minus_pi_half[simp]: "cis (-(pi / 2)) = -\" + by (simp add: cis.ctr complex_eq_iff) + +lemma cis_multiple_2pi[simp]: "n \ \ \ cis (2 * pi * n) = 1" + by (auto elim!: Ints_cases simp: cis.ctr one_complex.ctr) + subsubsection \$r(\cos \theta + i \sin \theta)$\ @@ -847,6 +883,11 @@ subsubsection \Complex exponential\ +lemma exp_Reals_eq: + assumes "z \ \" + shows "exp z = of_real (exp (Re z))" + using assms by (auto elim!: Reals_cases simp: exp_of_real) + lemma cis_conv_exp: "cis b = exp (\ * b)" proof - have "(\ * complex_of_real b) ^ n /\<^sub>R fact n = @@ -901,6 +942,10 @@ lemma exp_two_pi_i' [simp]: "exp (\ * (of_real pi * 2)) = 1" by (metis exp_two_pi_i mult.commute) +lemma continuous_on_cis [continuous_intros]: + "continuous_on A f \ continuous_on A (\x. cis (f x))" + by (auto simp: cis_conv_exp intro!: continuous_intros) + subsubsection \Complex argument\ diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/GCD.thy --- a/src/HOL/GCD.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/GCD.thy Sat Aug 11 16:02:55 2018 +0200 @@ -110,7 +110,7 @@ assumes "a \ A" shows "a \<^bold>* F A = F A" using assms by (induct A rule: infinite_finite_induct) - (auto simp add: left_commute [of a]) + (auto simp: left_commute [of a]) lemma union: "F (A \ B) = F A \<^bold>* F B" @@ -239,7 +239,7 @@ lemma is_unit_gcd_iff [simp]: "is_unit (gcd a b) \ gcd a b = 1" - by (cases "a = 0 \ b = 0") (auto simp add: unit_factor_gcd dest: is_unit_unit_factor) + by (cases "a = 0 \ b = 0") (auto simp: unit_factor_gcd dest: is_unit_unit_factor) sublocale gcd: abel_semigroup gcd proof @@ -569,20 +569,17 @@ lemma normalize_lcm_right: "lcm a (normalize b) = lcm a b" by (fact lcm.normalize_right_idem) -lemma gcd_mult_unit1: "is_unit a \ gcd (b * a) c = gcd b c" - apply (rule gcdI) - apply simp_all - apply (rule dvd_trans) - apply (rule gcd_dvd1) - apply (simp add: unit_simps) - done +lemma gcd_mult_unit1: + assumes "is_unit a" shows "gcd (b * a) c = gcd b c" +proof - + have "gcd (b * a) c dvd b" + using assms local.dvd_mult_unit_iff by blast + then show ?thesis + by (rule gcdI) simp_all +qed lemma gcd_mult_unit2: "is_unit a \ gcd b (c * a) = gcd b c" - apply (subst gcd.commute) - apply (subst gcd_mult_unit1) - apply assumption - apply (rule gcd.commute) - done + using gcd.commute gcd_mult_unit1 by auto lemma gcd_div_unit1: "is_unit a \ gcd (b div a) c = gcd b c" by (erule is_unitE [of _ b]) (simp add: gcd_mult_unit1) @@ -652,13 +649,13 @@ "a dvd d \ b dvd d \ normalize d = d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" by rule (auto intro: lcmI simp: lcm_least lcm_eq_0_iff) -lemma lcm_proj1_if_dvd: "b dvd a \ lcm a b = normalize a" - apply (cases "a = 0") - apply simp - apply (rule sym) - apply (rule lcmI) - apply simp_all - done +lemma lcm_proj1_if_dvd: + assumes "b dvd a" shows "lcm a b = normalize a" +proof (cases "a = 0") + case False + then show ?thesis + using assms gcd_proj2_if_dvd lcm_mult_gcd local.lcm_gcd by auto +qed auto lemma lcm_proj2_if_dvd: "a dvd b \ lcm a b = normalize b" using lcm_proj1_if_dvd [of a b] by (simp add: ac_simps) @@ -841,14 +838,12 @@ by (blast intro: Lcm_least dvd_Lcm) lemma Lcm_Un: "Lcm (A \ B) = lcm (Lcm A) (Lcm B)" - apply (rule lcmI) - apply (blast intro: Lcm_subset) - apply (blast intro: Lcm_subset) - apply (intro Lcm_least ballI, elim UnE) - apply (rule dvd_trans, erule dvd_Lcm, assumption) - apply (rule dvd_trans, erule dvd_Lcm, assumption) - apply simp - done +proof - + have "\d. \Lcm A dvd d; Lcm B dvd d\ \ Lcm (A \ B) dvd d" + by (meson UnE local.Lcm_least local.dvd_Lcm local.dvd_trans) + then show ?thesis + by (meson Lcm_subset local.lcm_unique local.normalize_Lcm sup.cobounded1 sup.cobounded2) +qed lemma Gcd_0_iff [simp]: "Gcd A = 0 \ A \ {0}" (is "?P \ ?Q") @@ -963,7 +958,7 @@ next case False with assms show ?thesis - by (induct A rule: finite_ne_induct) (auto simp add: lcm_eq_0_iff) + by (induct A rule: finite_ne_induct) (auto simp: lcm_eq_0_iff) qed lemma Gcd_image_normalize [simp]: "Gcd (normalize ` A) = Gcd A" @@ -996,25 +991,25 @@ lemma dvd_Gcd_iff: "x dvd Gcd A \ (\y\A. x dvd y)" by (blast dest: dvd_GcdD intro: Gcd_greatest) -lemma Gcd_mult: "Gcd (( * ) c ` A) = normalize c * Gcd A" +lemma Gcd_mult: "Gcd (( *) c ` A) = normalize c * Gcd A" proof (cases "c = 0") case True then show ?thesis by auto next case [simp]: False - have "Gcd (( * ) c ` A) div c dvd Gcd A" + have "Gcd (( *) c ` A) div c dvd Gcd A" by (intro Gcd_greatest, subst div_dvd_iff_mult) (auto intro!: Gcd_greatest Gcd_dvd simp: mult.commute[of _ c]) - then have "Gcd (( * ) c ` A) dvd c * Gcd A" + then have "Gcd (( *) c ` A) dvd c * Gcd A" by (subst (asm) div_dvd_iff_mult) (auto intro: Gcd_greatest simp: mult_ac) also have "c * Gcd A = (normalize c * Gcd A) * unit_factor c" by (subst unit_factor_mult_normalize [symmetric]) (simp only: mult_ac) - also have "Gcd (( * ) c ` A) dvd \ \ Gcd (( * ) c ` A) dvd normalize c * Gcd A" + also have "Gcd (( *) c ` A) dvd \ \ Gcd (( *) c ` A) dvd normalize c * Gcd A" by (simp add: dvd_mult_unit_iff) - finally have "Gcd (( * ) c ` A) dvd normalize c * Gcd A" . - moreover have "normalize c * Gcd A dvd Gcd (( * ) c ` A)" + finally have "Gcd (( *) c ` A) dvd normalize c * Gcd A" . + moreover have "normalize c * Gcd A dvd Gcd (( *) c ` A)" by (intro Gcd_greatest) (auto intro: mult_dvd_mono Gcd_dvd) - ultimately have "normalize (Gcd (( * ) c ` A)) = normalize (normalize c * Gcd A)" + ultimately have "normalize (Gcd (( *) c ` A)) = normalize (normalize c * Gcd A)" by (rule associatedI) then show ?thesis by (simp add: normalize_mult) @@ -1035,10 +1030,10 @@ lemma Lcm_mult: assumes "A \ {}" - shows "Lcm (( * ) c ` A) = normalize c * Lcm A" + shows "Lcm (( *) c ` A) = normalize c * Lcm A" proof (cases "c = 0") case True - with assms have "( * ) c ` A = {0}" + with assms have "( *) c ` A = {0}" by auto with True show ?thesis by auto next @@ -1047,23 +1042,23 @@ by blast have "c dvd c * x" by simp - also from x have "c * x dvd Lcm (( * ) c ` A)" + also from x have "c * x dvd Lcm (( *) c ` A)" by (intro dvd_Lcm) auto - finally have dvd: "c dvd Lcm (( * ) c ` A)" . - - have "Lcm A dvd Lcm (( * ) c ` A) div c" + finally have dvd: "c dvd Lcm (( *) c ` A)" . + + have "Lcm A dvd Lcm (( *) c ` A) div c" by (intro Lcm_least dvd_mult_imp_div) (auto intro!: Lcm_least dvd_Lcm simp: mult.commute[of _ c]) - then have "c * Lcm A dvd Lcm (( * ) c ` A)" + then have "c * Lcm A dvd Lcm (( *) c ` A)" by (subst (asm) dvd_div_iff_mult) (auto intro!: Lcm_least simp: mult_ac dvd) also have "c * Lcm A = (normalize c * Lcm A) * unit_factor c" by (subst unit_factor_mult_normalize [symmetric]) (simp only: mult_ac) - also have "\ dvd Lcm (( * ) c ` A) \ normalize c * Lcm A dvd Lcm (( * ) c ` A)" + also have "\ dvd Lcm (( *) c ` A) \ normalize c * Lcm A dvd Lcm (( *) c ` A)" by (simp add: mult_unit_dvd_iff) - finally have "normalize c * Lcm A dvd Lcm (( * ) c ` A)" . - moreover have "Lcm (( * ) c ` A) dvd normalize c * Lcm A" + finally have "normalize c * Lcm A dvd Lcm (( *) c ` A)" . + moreover have "Lcm (( *) c ` A) dvd normalize c * Lcm A" by (intro Lcm_least) (auto intro: mult_dvd_mono dvd_Lcm) - ultimately have "normalize (normalize c * Lcm A) = normalize (Lcm (( * ) c ` A))" + ultimately have "normalize (normalize c * Lcm A) = normalize (Lcm (( *) c ` A))" by (rule associatedI) then show ?thesis by (simp add: normalize_mult) @@ -1240,7 +1235,7 @@ lemma Lcm_fin_0_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 0 \ 0 \ A" if "finite A" - using that by (induct A) (auto simp add: lcm_eq_0_iff) + using that by (induct A) (auto simp: lcm_eq_0_iff) lemma Lcm_fin_1_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 1 \ (\a\A. is_unit a) \ finite A" @@ -1452,7 +1447,7 @@ from this [symmetric] have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by (simp_all add: mult.assoc mult.left_commute [of "gcd a b"]) then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" - by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)] dvd_mult_div_cancel [OF dvdg(2)] dvd_def) + by (auto simp: dvd_mult_div_cancel [OF dvdg(1)] dvd_mult_div_cancel [OF dvdg(2)] dvd_def) have "?g \ 0" using assms by simp moreover from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . @@ -1480,11 +1475,12 @@ lemma gcd_coprime_exists: assumes "gcd a b \ 0" shows "\a' b'. a = a' * gcd a b \ b = b' * gcd a b \ coprime a' b'" - apply (rule_tac x = "a div gcd a b" in exI) - apply (rule_tac x = "b div gcd a b" in exI) - using assms - apply (auto intro: div_gcd_coprime) - done +proof - + have "coprime (a div gcd a b) (b div gcd a b)" + using assms div_gcd_coprime by auto + then show ?thesis + by force +qed lemma pow_divides_pow_iff [simp]: "a ^ n dvd b ^ n \ a dvd b" if "n > 0" @@ -1628,7 +1624,7 @@ by simp also note gcd_mult_distrib also have "unit_factor (gcd a b ^ n) = 1" - using False by (auto simp add: unit_factor_power unit_factor_gcd) + using False by (auto simp: unit_factor_power unit_factor_gcd) also have "(gcd a b) ^ n * (a div gcd a b) ^ n = a ^ n" by (simp add: ac_simps div_power dvd_power_same) also have "(gcd a b) ^ n * (b div gcd a b) ^ n = b ^ n" @@ -1809,16 +1805,16 @@ for i j :: int by (simp only: lcm_int_def) -lemma gcd_nat_induct: +lemma gcd_nat_induct [case_names base step]: fixes m n :: nat assumes "\m. P m 0" and "\m n. 0 < n \ P n (m mod n) \ P m n" shows "P m n" - apply (rule gcd_nat.induct) - apply (case_tac "y = 0") - using assms - apply simp_all - done +proof (induction m n rule: gcd_nat.induct) + case (1 x y) + then show ?case + using assms neq0_conv by blast +qed lemma gcd_neg1_int [simp]: "gcd (- x) y = gcd x y" for x y :: int @@ -1856,7 +1852,7 @@ and "x \ 0 \ y \ 0 \ P (lcm (- x) y)" and "x \ 0 \ y \ 0 \ P (lcm (- x) (- y))" shows "P (lcm x y)" - using assms by (auto simp add: lcm_neg1_int lcm_neg2_int) arith + using assms by (auto simp: lcm_neg1_int lcm_neg2_int) arith lemma lcm_ge_0_int [simp]: "lcm x y \ 0" for x y :: int @@ -1907,7 +1903,7 @@ lemma gcd_idem_int: "gcd x x = \x\" for x :: int - by (auto simp add: gcd_int_def) + by (auto simp: gcd_int_def) declare gcd_nat.simps [simp del] @@ -1921,13 +1917,10 @@ fix m n :: nat show "gcd m n dvd m" and "gcd m n dvd n" proof (induct m n rule: gcd_nat_induct) - fix m n :: nat - assume "gcd n (m mod n) dvd m mod n" - and "gcd n (m mod n) dvd n" + case (step m n) then have "gcd n (m mod n) dvd m" - by (rule dvd_mod_imp_dvd) - moreover assume "0 < n" - ultimately show "gcd m n dvd m" + by (metis dvd_mod_imp_dvd) + with step show "gcd m n dvd m" by (simp add: gcd_non_0_nat) qed (simp_all add: gcd_0_nat gcd_non_0_nat) next @@ -1979,25 +1972,16 @@ lemma gcd_unique_nat: "d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: nat - apply auto - apply (rule dvd_antisym) - apply (erule (1) gcd_greatest) - apply auto - done + using gcd_unique by fastforce lemma gcd_unique_int: "d \ 0 \ d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: int - apply (cases "d = 0") - apply simp - apply (rule iffI) - apply (rule zdvd_antisym_nonneg) - apply (auto intro: gcd_greatest) - done + using zdvd_antisym_nonneg by auto interpretation gcd_nat: semilattice_neutr_order gcd "0::nat" Rings.dvd "\m n. m dvd n \ m \ n" - by standard (auto simp add: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans) + by standard (auto simp: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans) lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \ gcd x y = \x\" for x y :: int @@ -2013,11 +1997,11 @@ lemma gcd_mult_distrib_nat: "k * gcd m n = gcd (k * m) (k * n)" for k m n :: nat \ \@{cite \page 27\ davenport92}\ - apply (induct m n rule: gcd_nat_induct) - apply simp - apply (cases "k = 0") - apply (simp_all add: gcd_non_0_nat) - done +proof (induct m n rule: gcd_nat_induct) + case (step m n) + then show ?case + by (metis gcd_mult_distrib' gcd_red_nat) +qed auto lemma gcd_mult_distrib_int: "\k\ * gcd m n = gcd (k * m) (k * n)" for k m n :: int @@ -2033,34 +2017,49 @@ lemma gcd_diff2_nat: "n \ m \ gcd (n - m) n = gcd m n" for m n :: nat - apply (subst gcd.commute) - apply (subst gcd_diff1_nat [symmetric]) - apply auto - apply (subst gcd.commute) - apply (subst gcd_diff1_nat) - apply assumption - apply (rule gcd.commute) - done - -lemma gcd_non_0_int: "y > 0 \ gcd x y = gcd y (x mod y)" - for x y :: int - apply (frule_tac b = y and a = x in pos_mod_sign) - apply (simp del: Euclidean_Division.pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib) - apply (auto simp add: gcd_non_0_nat nat_mod_distrib [symmetric] zmod_zminus1_eq_if) - apply (frule_tac a = x in pos_mod_bound) - apply (subst (1 2) gcd.commute) - apply (simp del: Euclidean_Division.pos_mod_bound add: nat_diff_distrib gcd_diff2_nat nat_le_eq_zle) - done + by (metis gcd.commute gcd_add2 gcd_diff1_nat le_add_diff_inverse2) + +lemma gcd_non_0_int: + fixes x y :: int + assumes "y > 0" shows "gcd x y = gcd y (x mod y)" +proof (cases "x mod y = 0") + case False + then have neg: "x mod y = y - (- x) mod y" + by (simp add: zmod_zminus1_eq_if) + have xy: "0 \ x mod y" + by (simp add: assms) + show ?thesis + proof (cases "x < 0") + case True + have "nat (- x mod y) \ nat y" + by (simp add: assms dual_order.order_iff_strict) + moreover have "gcd (nat (- x)) (nat y) = gcd (nat (- x mod y)) (nat y)" + using True assms gcd_non_0_nat nat_mod_distrib by auto + ultimately have "gcd (nat (- x)) (nat y) = gcd (nat y) (nat (x mod y))" + using assms + by (simp add: neg nat_diff_distrib') (metis gcd.commute gcd_diff2_nat) + with assms \0 \ x mod y\ show ?thesis + by (simp add: True dual_order.order_iff_strict gcd_int_def) + next + case False + with assms xy have "gcd (nat x) (nat y) = gcd (nat y) (nat x mod nat y)" + using gcd_red_nat by blast + with False assms show ?thesis + by (simp add: gcd_int_def nat_mod_distrib) + qed +qed (use assms in auto) lemma gcd_red_int: "gcd x y = gcd y (x mod y)" for x y :: int - apply (cases "y = 0") - apply force - apply (cases "y > 0") - apply (subst gcd_non_0_int, auto) - apply (insert gcd_non_0_int [of "- y" "- x"]) - apply auto - done +proof (cases y "0::int" rule: linorder_cases) + case less + with gcd_non_0_int [of "- y" "- x"] show ?thesis + by auto +next + case greater + with gcd_non_0_int [of y x] show ?thesis + by auto +qed auto (* TODO: differences, and all variations of addition rules as simplification rules for nat and int *) @@ -2092,34 +2091,34 @@ qed lemma Max_divisors_self_nat [simp]: "n \ 0 \ Max {d::nat. d dvd n} = n" - apply (rule antisym) - apply (fastforce intro: Max_le_iff[THEN iffD2] simp: dvd_imp_le) - apply simp - done - -lemma Max_divisors_self_int [simp]: "n \ 0 \ Max {d::int. d dvd n} = \n\" - apply (rule antisym) - apply (rule Max_le_iff [THEN iffD2]) - apply (auto intro: abs_le_D1 dvd_imp_le_int) - done - -lemma gcd_is_Max_divisors_nat: "m > 0 \ n > 0 \ gcd m n = Max {d. d dvd m \ d dvd n}" - for m n :: nat - apply (rule Max_eqI[THEN sym]) - apply (metis finite_Collect_conjI finite_divisors_nat) - apply simp - apply (metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff gcd_pos_nat) - apply simp - done - -lemma gcd_is_Max_divisors_int: "m \ 0 \ n \ 0 \ gcd m n = Max {d. d dvd m \ d dvd n}" - for m n :: int - apply (rule Max_eqI[THEN sym]) - apply (metis finite_Collect_conjI finite_divisors_int) - apply simp - apply (metis gcd_greatest_iff gcd_pos_int zdvd_imp_le) - apply simp - done + by (fastforce intro: antisym Max_le_iff[THEN iffD2] simp: dvd_imp_le) + +lemma Max_divisors_self_int [simp]: + assumes "n \ 0" shows "Max {d::int. d dvd n} = \n\" +proof (rule antisym) + show "Max {d. d dvd n} \ \n\" + using assms by (auto intro: abs_le_D1 dvd_imp_le_int intro!: Max_le_iff [THEN iffD2]) +qed (simp add: assms) + +lemma gcd_is_Max_divisors_nat: + fixes m n :: nat + assumes "n > 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" +proof (rule Max_eqI[THEN sym], simp_all) + show "finite {d. d dvd m \ d dvd n}" + by (simp add: \n > 0\) + show "\y. y dvd m \ y dvd n \ y \ gcd m n" + by (simp add: \n > 0\ dvd_imp_le) +qed + +lemma gcd_is_Max_divisors_int: + fixes m n :: int + assumes "n \ 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" +proof (rule Max_eqI[THEN sym], simp_all) + show "finite {d. d dvd m \ d dvd n}" + by (simp add: \n \ 0\) + show "\y. y dvd m \ y dvd n \ y \ gcd m n" + by (simp add: \n \ 0\ zdvd_imp_le) +qed lemma gcd_code_int [code]: "gcd k l = \if l = 0 then k else gcd l (\k\ mod \l\)\" for k l :: int @@ -2178,25 +2177,22 @@ declare bezw.simps [simp del] -lemma bezw_aux: "fst (bezw x y) * int x + snd (bezw x y) * int y = int (gcd x y)" + +lemma bezw_aux: "int (gcd x y) = fst (bezw x y) * int x + snd (bezw x y) * int y" proof (induct x y rule: gcd_nat_induct) - fix m :: nat - show "fst (bezw m 0) * int m + snd (bezw m 0) * int 0 = int (gcd m 0)" - by auto -next - fix m n :: nat - assume ngt0: "n > 0" - and ih: "fst (bezw n (m mod n)) * int n + snd (bezw n (m mod n)) * int (m mod n) = - int (gcd n (m mod n))" - then show "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)" - apply (simp add: bezw_non_0 gcd_non_0_nat) - apply (erule subst) - apply (simp add: field_simps) - apply (subst div_mult_mod_eq [of m n, symmetric]) - (* applying simp here undoes the last substitution! what is procedure cancel_div_mod? *) - apply (simp only: NO_MATCH_def field_simps of_nat_add of_nat_mult) - done -qed + case (step m n) + then have "fst (bezw m n) * int m + snd (bezw m n) * int n - int (gcd m n) + = int m * snd (bezw n (m mod n)) - + (int (m mod n) * snd (bezw n (m mod n)) + int n * (int (m div n) * snd (bezw n (m mod n))))" + by (simp add: bezw_non_0 gcd_non_0_nat field_simps) + also have "\ = int m * snd (bezw n (m mod n)) - (int (m mod n) + int (n * (m div n))) * snd (bezw n (m mod n))" + by (simp add: distrib_right) + also have "\ = 0" + by (metis cancel_comm_monoid_add_class.diff_cancel mod_mult_div_eq of_nat_add) + finally show ?case + by simp +qed auto + lemma bezout_int: "\u v. u * x + v * y = gcd x y" for x y :: int @@ -2204,13 +2200,9 @@ have aux: "x \ 0 \ y \ 0 \ \u v. u * x + v * y = gcd x y" for x y :: int apply (rule_tac x = "fst (bezw (nat x) (nat y))" in exI) apply (rule_tac x = "snd (bezw (nat x) (nat y))" in exI) - apply (unfold gcd_int_def) - apply simp - apply (subst bezw_aux [symmetric]) - apply auto - done + by (simp add: bezw_aux gcd_int_def) consider "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" - by atomize_elim auto + using linear by blast then show ?thesis proof cases case 1 @@ -2218,48 +2210,29 @@ next case 2 then show ?thesis - apply - - apply (insert aux [of x "-y"]) - apply auto - apply (rule_tac x = u in exI) - apply (rule_tac x = "-v" in exI) - apply (subst gcd_neg2_int [symmetric]) - apply auto - done + using aux [of x "-y"] + by (metis gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 3 then show ?thesis - apply - - apply (insert aux [of "-x" y]) - apply auto - apply (rule_tac x = "-u" in exI) - apply (rule_tac x = v in exI) - apply (subst gcd_neg1_int [symmetric]) - apply auto - done + using aux [of "-x" y] + by (metis gcd.commute gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 4 then show ?thesis - apply - - apply (insert aux [of "-x" "-y"]) - apply auto - apply (rule_tac x = "-u" in exI) - apply (rule_tac x = "-v" in exI) - apply (subst gcd_neg1_int [symmetric]) - apply (subst gcd_neg2_int [symmetric]) - apply auto - done + using aux [of "-x" "-y"] + by (metis diff_0 diff_ge_0_iff_ge gcd_neg1_int gcd_neg2_int mult.commute mult_minus_right) qed qed text \Versions of Bezout for \nat\, by Amine Chaieb.\ -lemma ind_euclid: +lemma Euclid_induct [case_names swap zero add]: fixes P :: "nat \ nat \ bool" - assumes c: " \a b. P a b \ P b a" - and z: "\a. P a 0" - and add: "\a b. P a b \ P a (a + b)" + assumes c: "\a b. P a b \ P b a" + and z: "\a. P a 0" + and add: "\a b. P a b \ P a (a + b)" shows "P a b" proof (induct "a + b" arbitrary: a b rule: less_induct) case less @@ -2302,53 +2275,32 @@ qed lemma bezout_lemma_nat: - assumes ex: "\(d::nat) x y. d dvd a \ d dvd b \ - (a * x = b * y + d \ b * x = a * y + d)" - shows "\d x y. d dvd a \ d dvd a + b \ - (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" - using ex - apply clarsimp - apply (rule_tac x="d" in exI) - apply simp - apply (case_tac "a * x = b * y + d") - apply simp_all - apply (rule_tac x="x + y" in exI) - apply (rule_tac x="y" in exI) - apply algebra - apply (rule_tac x="x" in exI) - apply (rule_tac x="x + y" in exI) - apply algebra - done - -lemma bezout_add_nat: "\(d::nat) x y. d dvd a \ d dvd b \ - (a * x = b * y + d \ b * x = a * y + d)" - apply (induct a b rule: ind_euclid) - apply blast - apply clarify - apply (rule_tac x="a" in exI) - apply simp - apply clarsimp - apply (rule_tac x="d" in exI) - apply (case_tac "a * x = b * y + d") - apply simp_all - apply (rule_tac x="x+y" in exI) - apply (rule_tac x="y" in exI) - apply algebra - apply (rule_tac x="x" in exI) - apply (rule_tac x="x+y" in exI) - apply algebra - done - -lemma bezout1_nat: "\(d::nat) x y. d dvd a \ d dvd b \ - (a * x - b * y = d \ b * x - a * y = d)" - using bezout_add_nat[of a b] - apply clarsimp - apply (rule_tac x="d" in exI) - apply simp - apply (rule_tac x="x" in exI) - apply (rule_tac x="y" in exI) + fixes d::nat + shows "\d dvd a; d dvd b; a * x = b * y + d \ b * x = a * y + d\ + \ \x y. d dvd a \ d dvd a + b \ (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" apply auto - done + apply (metis add_mult_distrib2 left_add_mult_distrib) + apply (rule_tac x=x in exI) + by (metis add_mult_distrib2 mult.commute add.assoc) + +lemma bezout_add_nat: + "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" +proof (induct a b rule: Euclid_induct) + case (swap a b) + then show ?case + by blast +next + case (zero a) + then show ?case + by fastforce +next + case (add a b) + then show ?case + by (meson bezout_lemma_nat) +qed + +lemma bezout1_nat: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x - b * y = d \ b * x - a * y = d)" + using bezout_add_nat[of a b] by (metis add_diff_cancel_left') lemma bezout_add_strong_nat: fixes a b :: nat @@ -2356,7 +2308,7 @@ shows "\d x y. d dvd a \ d dvd b \ a * x = b * y + d" proof - consider d x y where "d dvd a" "d dvd b" "a * x = b * y + d" - | d x y where "d dvd a" "d dvd b" "b * x = a * y + d" + | d x y where "d dvd a" "d dvd b" "b * x = a * y + d" using bezout_add_nat [of a b] by blast then show ?thesis proof cases @@ -2377,13 +2329,7 @@ proof cases case 1 with a H show ?thesis - apply simp - apply (rule exI[where x = b]) - apply simp - apply (rule exI[where x = b]) - apply (rule exI[where x = "a - 1"]) - apply (simp add: diff_mult_distrib2) - done + by (metis Suc_pred add.commute mult.commute mult_Suc_right neq0_conv) next case 2 show ?thesis @@ -2410,13 +2356,7 @@ then have "a * ((b - 1) * y) = b * (x * (b - 1) - d) + d" by (simp only: diff_mult_distrib2 ac_simps) with H(1,2) show ?thesis - apply - - apply (rule exI [where x = d]) - apply simp - apply (rule exI [where x = "(b - 1) * y"]) - apply (rule exI [where x = "x * (b - 1) - d"]) - apply simp - done + by blast qed qed qed @@ -2451,17 +2391,11 @@ lemma prod_gcd_lcm_nat: "m * n = gcd m n * lcm m n" for m n :: nat - unfolding lcm_nat_def - by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod]) + by simp lemma prod_gcd_lcm_int: "\m\ * \n\ = gcd m n * lcm m n" for m n :: int - unfolding lcm_int_def gcd_int_def - apply (subst of_nat_mult [symmetric]) - apply (subst prod_gcd_lcm_nat [symmetric]) - apply (subst nat_abs_mult_distrib [symmetric]) - apply (simp add: abs_mult) - done + by simp lemma lcm_pos_nat: "m > 0 \ n > 0 \ lcm m n > 0" for m n :: nat @@ -2473,7 +2407,7 @@ lemma dvd_pos_nat: "n > 0 \ m dvd n \ m > 0" (* FIXME move *) for m n :: nat - by (cases m) auto + by auto lemma lcm_unique_nat: "a dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" @@ -2487,17 +2421,11 @@ lemma lcm_proj2_if_dvd_nat [simp]: "x dvd y \ lcm x y = y" for x y :: nat - apply (rule sym) - apply (subst lcm_unique_nat [symmetric]) - apply auto - done + by (simp add: lcm_proj2_if_dvd) lemma lcm_proj2_if_dvd_int [simp]: "x dvd y \ lcm x y = \y\" for x y :: int - apply (rule sym) - apply (subst lcm_unique_int [symmetric]) - apply auto - done + by (simp add: lcm_proj2_if_dvd) lemma lcm_proj1_if_dvd_nat [simp]: "x dvd y \ lcm y x = y" for x y :: nat @@ -2551,7 +2479,7 @@ by (simp add: Lcm_nat_def del: One_nat_def) lemma Lcm_nat_insert: "Lcm (insert n M) = lcm n (Lcm M)" for n :: nat - by (cases "finite M") (auto simp add: Lcm_nat_def simp del: One_nat_def) + by (cases "finite M") (auto simp: Lcm_nat_def simp del: One_nat_def) lemma Lcm_nat_infinite: "infinite M \ Lcm M = 0" for M :: "nat set" by (simp add: Lcm_nat_def) @@ -2595,9 +2523,9 @@ fix N :: "nat set" fix n :: nat show "Gcd N dvd n" if "n \ N" - using that by (induct N rule: infinite_finite_induct) (auto simp add: Gcd_nat_def) + using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Gcd N" if "\m. m \ N \ n dvd m" - using that by (induct N rule: infinite_finite_induct) (auto simp add: Gcd_nat_def) + using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Lcm N" if "n \ N" using that by (induct N rule: infinite_finite_induct) auto show "Lcm N dvd n" if "\m. m \ N \ m dvd n" @@ -2629,52 +2557,51 @@ from fin show "Gcd M \ Max (\m\M. {d. d dvd m})" by (auto intro: Max_ge Gcd_dvd) from fin show "Max (\m\M. {d. d dvd m}) \ Gcd M" - apply (rule Max.boundedI) - apply auto - apply (meson Gcd_dvd Gcd_greatest \0 < m\ \m \ M\ dvd_imp_le dvd_pos_nat) - done + proof (rule Max.boundedI, simp_all) + show "(\m\M. {d. d dvd m}) \ {}" + by auto + show "\a. \x\M. a dvd x \ a \ Gcd M" + by (meson Gcd_dvd Gcd_greatest \0 < m\ \m \ M\ dvd_imp_le dvd_pos_nat) + qed qed lemma Gcd_remove0_nat: "finite M \ Gcd M = Gcd (M - {0})" for M :: "nat set" - apply (induct pred: finite) - apply simp - apply (case_tac "x = 0") - apply simp - apply (subgoal_tac "insert x F - {0} = insert x (F - {0})") - apply simp - apply blast - done +proof (induct pred: finite) + case (insert x M) + then show ?case + by (simp add: insert_Diff_if) +qed auto lemma Lcm_in_lcm_closed_set_nat: - "finite M \ M \ {} \ \m n. m \ M \ n \ M \ lcm m n \ M \ Lcm M \ M" - for M :: "nat set" - apply (induct rule: finite_linorder_min_induct) - apply simp - apply simp - apply (subgoal_tac "\m n. m \ A \ n \ A \ lcm m n \ A") - apply simp - apply(case_tac "A = {}") - apply simp - apply simp - apply (metis lcm_pos_nat lcm_unique_nat linorder_neq_iff nat_dvd_not_less not_less0) - done + fixes M :: "nat set" + assumes "finite M" "M \ {}" "\m n. \m \ M; n \ M\ \ lcm m n \ M" + shows "Lcm M \ M" + using assms +proof (induction M rule: finite_linorder_min_induct) + case (insert x M) + then have "\m n. m \ M \ n \ M \ lcm m n \ M" + by (metis dvd_lcm1 gr0I insert_iff lcm_pos_nat nat_dvd_not_less) + with insert show ?case + by simp (metis Lcm_nat_empty One_nat_def dvd_1_left dvd_lcm2) +qed auto lemma Lcm_eq_Max_nat: - "finite M \ M \ {} \ 0 \ M \ \m n. m \ M \ n \ M \ lcm m n \ M \ Lcm M = Max M" - for M :: "nat set" - apply (rule antisym) - apply (rule Max_ge) - apply assumption - apply (erule (2) Lcm_in_lcm_closed_set_nat) - apply (auto simp add: not_le Lcm_0_iff dvd_imp_le leD le_neq_trans) - done + fixes M :: "nat set" + assumes M: "finite M" "M \ {}" "0 \ M" and lcm: "\m n. \m \ M; n \ M\ \ lcm m n \ M" + shows "Lcm M = Max M" +proof (rule antisym) + show "Lcm M \ Max M" + by (simp add: Lcm_in_lcm_closed_set_nat \finite M\ \M \ {}\ lcm) + show "Max M \ Lcm M" + by (meson Lcm_0_iff Max_in M dvd_Lcm dvd_imp_le le_0_eq not_le) +qed lemma mult_inj_if_coprime_nat: - "inj_on f A \ inj_on g B \ \a\A. \b\B. coprime (f a) (g b) \ + "inj_on f A \ inj_on g B \ (\a b. \a\A; b\B\ \ coprime (f a) (g b)) \ inj_on (\(a, b). f a * g b) (A \ B)" for f :: "'a \ nat" and g :: "'b \ nat" - by (auto simp add: inj_on_def coprime_crossproduct_nat simp del: One_nat_def) + by (auto simp: inj_on_def coprime_crossproduct_nat simp del: One_nat_def) subsubsection \Setwise GCD and LCM for integers\ diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Int.thy --- a/src/HOL/Int.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Int.thy Sat Aug 11 16:02:55 2018 +0200 @@ -892,6 +892,9 @@ apply (rule of_int_minus [symmetric]) done +lemma minus_in_Ints_iff: "-x \ \ \ x \ \" + using Ints_minus[of x] Ints_minus[of "-x"] by auto + lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" apply (auto simp add: Ints_def) apply (rule range_eqI) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Library/FuncSet.thy Sat Aug 11 16:02:55 2018 +0200 @@ -163,8 +163,6 @@ lemma compose_assoc: assumes "f \ A \ B" - and "g \ B \ C" - and "h \ C \ D" shows "compose A h (compose A g f) = compose A (compose B h g) f" using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Library/datatype_records.ML --- a/src/HOL/Library/datatype_records.ML Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Library/datatype_records.ML Sat Aug 11 16:02:55 2018 +0200 @@ -7,10 +7,10 @@ val mk_update_defs: string -> local_theory -> local_theory - val bnf_record: binding -> ctr_options -> (binding option * (typ * sort)) list -> + val record: binding -> ctr_options -> (binding option * (typ * sort)) list -> (binding * typ) list -> local_theory -> local_theory - val bnf_record_cmd: binding -> ctr_options_cmd -> + val record_cmd: binding -> ctr_options_cmd -> (binding option * (string * string option)) list -> (binding * string) list -> local_theory -> local_theory @@ -35,21 +35,31 @@ val extend = I ) +fun mk_eq_dummy (lhs, rhs) = + Const (@{const_name HOL.eq}, dummyT --> dummyT --> @{typ bool}) $ lhs $ rhs + +val dummify = map_types (K dummyT) +fun repeat_split_tac ctxt thm = REPEAT_ALL_NEW (CHANGED o Splitter.split_tac ctxt [thm]) + fun mk_update_defs typ_name lthy = let val short_name = Long_Name.base_name typ_name + val {ctrs, casex, selss, split, sel_thmss, injects, ...} = + the (Ctr_Sugar.ctr_sugar_of lthy typ_name) + val ctr = case ctrs of [ctr] => ctr | _ => error "Datatype_Records.mk_update_defs: expected only single constructor" + val sels = case selss of [sels] => sels | _ => error "Datatype_Records.mk_update_defs: expected selectors" + val sels_dummy = map dummify sels + val ctr_dummy = dummify ctr + val casex_dummy = dummify casex + val len = length sels - val {ctrs, casex, selss, ...} = the (Ctr_Sugar.ctr_sugar_of lthy typ_name) - val ctr = case ctrs of [ctr] => ctr | _ => error "BNF_Record.mk_update_defs: expected only single constructor" - val sels = case selss of [sels] => sels | _ => error "BNF_Record.mk_update_defs: expected selectors" - val ctr_dummy = Const (fst (dest_Const ctr), dummyT) - val casex_dummy = Const (fst (dest_Const casex), dummyT) - - val len = length sels + val simp_thms = flat sel_thmss @ injects fun mk_name sel = Binding.name ("update_" ^ Long_Name.base_name (fst (dest_Const sel))) + val thms_binding = (@{binding record_simps}, @{attributes [simp]}) + fun mk_t idx = let val body = @@ -59,22 +69,143 @@ Abs ("f", dummyT, casex_dummy $ body) end - fun define name t = - Local_Theory.define ((name, NoSyn), ((Binding.empty, @{attributes [datatype_record_update, code]}), t)) #> snd + fun simp_only_tac ctxt = + REPEAT_ALL_NEW (resolve_tac ctxt @{thms impI allI}) THEN' + asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp_thms) + + fun prove ctxt defs ts n = + let + val t = nth ts n + + val sel_dummy = nth sels_dummy n + val t_dummy = dummify t + fun tac {context = ctxt, ...} = + Goal.conjunction_tac 1 THEN + Local_Defs.unfold_tac ctxt defs THEN + PARALLEL_ALLGOALS (repeat_split_tac ctxt split THEN' simp_only_tac ctxt) + + val sel_upd_same_thm = + let + val ([f, x], ctxt') = Variable.add_fixes ["f", "x"] ctxt + val f = Free (f, dummyT) + val x = Free (x, dummyT) + + val lhs = sel_dummy $ (t_dummy $ f $ x) + val rhs = f $ (sel_dummy $ x) + val prop = Syntax.check_term ctxt' (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) + in + [Goal.prove_future ctxt' [] [] prop tac] + |> Variable.export ctxt' ctxt + end + + val sel_upd_diff_thms = + let + val ([f, x], ctxt') = Variable.add_fixes ["f", "x"] ctxt + val f = Free (f, dummyT) + val x = Free (x, dummyT) + + fun lhs sel = sel $ (t_dummy $ f $ x) + fun rhs sel = sel $ x + fun eq sel = (lhs sel, rhs sel) + fun is_n i = i = n + val props = + sels_dummy ~~ (0 upto len - 1) + |> filter_out (is_n o snd) + |> map (HOLogic.mk_Trueprop o mk_eq_dummy o eq o fst) + |> Syntax.check_terms ctxt' + in + if length props > 0 then + Goal.prove_common ctxt' (SOME ~1) [] [] props tac + |> Variable.export ctxt' ctxt + else + [] + end + + val upd_comp_thm = + let + val ([f, g, x], ctxt') = Variable.add_fixes ["f", "g", "x"] ctxt + val f = Free (f, dummyT) + val g = Free (g, dummyT) + val x = Free (x, dummyT) - val lthy' = - Local_Theory.map_background_naming (Name_Space.qualified_path false (Binding.name short_name)) lthy + val lhs = t_dummy $ f $ (t_dummy $ g $ x) + val rhs = t_dummy $ Abs ("a", dummyT, f $ (g $ Bound 0)) $ x + val prop = Syntax.check_term ctxt' (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) + in + [Goal.prove_future ctxt' [] [] prop tac] + |> Variable.export ctxt' ctxt + end + + val upd_comm_thms = + let + fun prop i ctxt = + let + val ([f, g, x], ctxt') = Variable.variant_fixes ["f", "g", "x"] ctxt + val self = t_dummy $ Free (f, dummyT) + val other = dummify (nth ts i) $ Free (g, dummyT) + val lhs = other $ (self $ Free (x, dummyT)) + val rhs = self $ (other $ Free (x, dummyT)) + in + (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs)), ctxt') + end + val (props, ctxt') = fold_map prop (0 upto n - 1) ctxt + val props = Syntax.check_terms ctxt' props + in + if length props > 0 then + Goal.prove_common ctxt' (SOME ~1) [] [] props tac + |> Variable.export ctxt' ctxt + else + [] + end + + val upd_sel_thm = + let + val ([x], ctxt') = Variable.add_fixes ["x"] ctxt + + val lhs = t_dummy $ Abs("_", dummyT, (sel_dummy $ Free(x, dummyT))) $ Free (x, dummyT) + val rhs = Free (x, dummyT) + val prop = Syntax.check_term ctxt (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) + in + [Goal.prove_future ctxt [] [] prop tac] + |> Variable.export ctxt' ctxt + end + in + sel_upd_same_thm @ sel_upd_diff_thms @ upd_comp_thm @ upd_comm_thms @ upd_sel_thm + end + + fun define name t = + Local_Theory.define ((name, NoSyn), ((Binding.empty, @{attributes [datatype_record_update, code]}),t)) + #> apfst (apsnd snd) + + val (updates, (lthy'', lthy')) = + lthy + |> Local_Theory.open_target + |> snd + |> Local_Theory.map_background_naming (Name_Space.qualified_path false (Binding.name short_name)) + |> @{fold_map 2} define (map mk_name sels) (Syntax.check_terms lthy (map mk_t (0 upto len - 1))) + ||> `Local_Theory.close_target + + val phi = Proof_Context.export_morphism lthy' lthy'' + + val (update_ts, update_defs) = + split_list updates + |>> map (Morphism.term phi) + ||> map (Morphism.thm phi) + + val thms = flat (map (prove lthy'' update_defs update_ts) (0 upto len-1)) fun insert sel = Symtab.insert op = (fst (dest_Const sel), Local_Theory.full_name lthy' (mk_name sel)) in - lthy' - |> @{fold 2} define (map mk_name sels) (Syntax.check_terms lthy (map mk_t (0 upto len - 1))) + lthy'' + |> Local_Theory.map_background_naming (Name_Space.mandatory_path short_name) + |> Local_Theory.note (thms_binding, thms) + |> snd + |> Local_Theory.restore_background_naming lthy |> Local_Theory.background_theory (Data.map (fold insert sels)) - |> Local_Theory.restore_background_naming lthy end -fun bnf_record binding opts tyargs args lthy = +fun record binding opts tyargs args lthy = let val constructor = (((Binding.empty, Binding.map_name (fn c => "make_" ^ c) binding), args), NoSyn) @@ -93,8 +224,8 @@ lthy' end -fun bnf_record_cmd binding opts tyargs args lthy = - bnf_record binding (opts lthy) +fun record_cmd binding opts tyargs args lthy = + record binding (opts lthy) (map (apsnd (apfst (Syntax.parse_typ lthy) o apsnd (Typedecl.read_constraint lthy))) tyargs) (map (apsnd (Syntax.parse_typ lthy)) args) lthy @@ -172,7 +303,7 @@ @{command_keyword datatype_record} "Defines a record based on the BNF/datatype machinery" (parser >> (fn (((ctr_options, tyargs), binding), args) => - bnf_record_cmd binding ctr_options tyargs args)) + record_cmd binding ctr_options tyargs args)) val setup = (Sign.parse_translation diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Limits.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1316,6 +1316,16 @@ and filterlim_compose[OF filterlim_uminus_at_top_at_bot, of "\x. - f x" F] by auto +lemma tendsto_at_botI_sequentially: + fixes f :: "real \ 'b::first_countable_topology" + assumes *: "\X. filterlim X at_bot sequentially \ (\n. f (X n)) \ y" + shows "(f \ y) at_bot" + unfolding filterlim_at_bot_mirror +proof (rule tendsto_at_topI_sequentially) + fix X :: "nat \ real" assume "filterlim X at_top sequentially" + thus "(\n. f (-X n)) \ y" by (intro *) (auto simp: filterlim_uminus_at_top) +qed + lemma filterlim_at_infinity_imp_filterlim_at_top: assumes "filterlim (f :: 'a \ real) at_infinity F" assumes "eventually (\x. f x > 0) F" diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/List.thy --- a/src/HOL/List.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/List.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1,5 +1,5 @@ (* Title: HOL/List.thy - Author: Tobias Nipkow + Author: Tobias Nipkow; proofs tidied by LCP *) section \The datatype of finite lists\ @@ -166,7 +166,7 @@ primrec upt :: "nat \ nat \ nat list" ("(1[_.. j then [i.. 'a list \ 'a list" where "insert x xs = (if x \ set xs then xs else x # xs)" @@ -834,11 +834,9 @@ lemma Suc_length_conv: "(Suc n = length xs) = (\y ys. xs = y # ys \ length ys = n)" -apply (induct xs, simp, simp) -apply blast -done - -lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False" + by (induct xs; simp; blast) + +lemma impossible_Cons: "length xs \ length ys ==> xs = x # ys = False" by (induct xs) auto lemma list_induct2 [consumes 1, case_names Nil Cons]: @@ -846,10 +844,8 @@ (\x xs y ys. length xs = length ys \ P xs ys \ P (x#xs) (y#ys)) \ P xs ys" proof (induct xs arbitrary: ys) - case Nil then show ?case by simp -next case (Cons x xs ys) then show ?case by (cases ys) simp_all -qed +qed simp lemma list_induct3 [consumes 2, case_names Nil Cons]: "length xs = length ys \ length ys = length zs \ P [] [] [] \ @@ -960,19 +956,15 @@ lemma append_eq_append_conv [simp]: "length xs = length ys \ length us = length vs ==> (xs@us = ys@vs) = (xs=ys \ us=vs)" -apply (induct xs arbitrary: ys) - apply (case_tac ys, simp, force) -apply (case_tac ys, force, simp) -done + by (induct xs arbitrary: ys; case_tac ys; force) lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) = (\us. xs = zs @ us \ us @ ys = ts \ xs @ us = zs \ ys = us @ ts)" -apply (induct xs arbitrary: ys zs ts) - apply fastforce -apply(case_tac zs) - apply simp -apply fastforce -done +proof (induct xs arbitrary: ys zs ts) + case (Cons x xs) + then show ?case + by (case_tac zs) auto +qed fastforce lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)" by simp @@ -1152,15 +1144,14 @@ qed lemma map_inj_on: - "[| map f xs = map f ys; inj_on f (set xs Un set ys) |] - ==> xs = ys" -apply(frule map_eq_imp_length_eq) -apply(rotate_tac -1) -apply(induct rule:list_induct2) - apply simp -apply(simp) -apply (blast intro:sym) -done + assumes map: "map f xs = map f ys" and inj: "inj_on f (set xs Un set ys)" + shows "xs = ys" + using map_eq_imp_length_eq [OF map] assms +proof (induct rule: list_induct2) + case (Cons x xs y ys) + then show ?case + by (auto intro: sym) +qed auto lemma inj_on_map_eq_map: "inj_on f (set xs Un set ys) \ (map f xs = map f ys) = (xs = ys)" @@ -1177,21 +1168,13 @@ by (iprover dest: map_injective injD intro: inj_onI) lemma inj_mapD: "inj (map f) ==> inj f" - apply (unfold inj_def) - apply clarify - apply (erule_tac x = "[x]" in allE) - apply (erule_tac x = "[y]" in allE) - apply auto - done + by (metis (no_types, hide_lams) injI list.inject list.simps(9) the_inv_f_f) lemma inj_map[iff]: "inj (map f) = inj f" by (blast dest: inj_mapD intro: inj_mapI) lemma inj_on_mapI: "inj_on f (\(set ` A)) \ inj_on (map f) A" -apply(rule inj_onI) -apply(erule map_inj_on) -apply(blast intro:inj_onI dest:inj_onD) -done + by (blast intro:inj_onI dest:inj_onD map_inj_on) lemma map_idI: "(\x. x \ set xs \ f x = x) \ map f xs = xs" by (induct xs, auto) @@ -1248,9 +1231,11 @@ by (cases xs) auto lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)" -apply (induct xs arbitrary: ys, force) -apply (case_tac ys, simp, force) -done +proof (induct xs arbitrary: ys) + case (Cons a xs) + then show ?case + by (case_tac ys) auto +qed force lemma inj_on_rev[iff]: "inj_on rev A" by(simp add:inj_on_def) @@ -1481,11 +1466,12 @@ by (induct xs) simp_all lemma filter_id_conv: "(filter P xs = xs) = (\x\set xs. P x)" -apply (induct xs) - apply auto -apply(cut_tac P=P and xs=xs in length_filter_le) -apply simp -done +proof (induct xs) + case (Cons x xs) + then show ?case + using length_filter_le + by (simp add: impossible_Cons) +qed auto lemma filter_map: "filter P (map f xs) = map f (filter (P \ f) xs)" by (induct xs) simp_all @@ -1503,9 +1489,7 @@ case Nil thus ?case by simp next case (Cons x xs) thus ?case - apply (auto split:if_split_asm) - using length_filter_le[of P xs] apply arith - done + using Suc_le_eq by fastforce qed lemma length_filter_conv_card: @@ -1573,17 +1557,17 @@ lemma filter_eq_ConsD: "filter P ys = x#xs \ \us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs" -by(rule Cons_eq_filterD) simp + by(rule Cons_eq_filterD) simp lemma filter_eq_Cons_iff: "(filter P ys = x#xs) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" -by(auto dest:filter_eq_ConsD) + by(auto dest:filter_eq_ConsD) lemma Cons_eq_filter_iff: "(x#xs = filter P ys) = (\us vs. ys = us @ x # vs \ (\u\set us. \ P u) \ P x \ xs = filter P vs)" -by(auto dest:Cons_eq_filterD) + by(auto dest:Cons_eq_filterD) lemma inj_on_filter_key_eq: assumes "inj_on f (insert y (set xs))" @@ -1592,24 +1576,22 @@ lemma filter_cong[fundef_cong]: "xs = ys \ (\x. x \ set ys \ P x = Q x) \ filter P xs = filter Q ys" -apply simp -apply(erule thin_rl) -by (induct ys) simp_all + by (induct ys arbitrary: xs) auto subsubsection \List partitioning\ primrec partition :: "('a \ bool) \'a list \ 'a list \ 'a list" where -"partition P [] = ([], [])" | -"partition P (x # xs) = + "partition P [] = ([], [])" | + "partition P (x # xs) = (let (yes, no) = partition P xs in if P x then (x # yes, no) else (yes, x # no))" lemma partition_filter1: "fst (partition P xs) = filter P xs" -by (induct xs) (auto simp add: Let_def split_def) + by (induct xs) (auto simp add: Let_def split_def) lemma partition_filter2: "snd (partition P xs) = filter (Not \ P) xs" -by (induct xs) (auto simp add: Let_def split_def) + by (induct xs) (auto simp add: Let_def split_def) lemma partition_P: assumes "partition P xs = (yes, no)" @@ -1631,8 +1613,8 @@ lemma partition_filter_conv[simp]: "partition f xs = (filter f xs,filter (Not \ f) xs)" -unfolding partition_filter2[symmetric] -unfolding partition_filter1[symmetric] by simp + unfolding partition_filter2[symmetric] + unfolding partition_filter1[symmetric] by simp declare partition.simps[simp del] @@ -1640,28 +1622,28 @@ subsubsection \@{const concat}\ lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys" -by (induct xs) auto + by (induct xs) auto lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\xs \ set xss. xs = [])" -by (induct xss) auto + by (induct xss) auto lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\xs \ set xss. xs = [])" -by (induct xss) auto + by (induct xss) auto lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)" -by (induct xs) auto + by (induct xs) auto lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs" -by (induct xs) auto + by (induct xs) auto lemma map_concat: "map f (concat xs) = concat (map (map f) xs)" -by (induct xs) auto + by (induct xs) auto lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)" -by (induct xs) auto + by (induct xs) auto lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))" -by (induct xs) auto + by (induct xs) auto lemma concat_eq_concat_iff: "\(x, y) \ set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)" proof (induct xs arbitrary: ys) @@ -1670,66 +1652,78 @@ qed (auto) lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \(x, y) \ set (zip xs ys). length x = length y ==> xs = ys" -by (simp add: concat_eq_concat_iff) + by (simp add: concat_eq_concat_iff) subsubsection \@{const nth}\ lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x" -by auto + by auto lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n" -by auto + by auto declare nth.simps [simp del] lemma nth_Cons_pos[simp]: "0 < n \ (x#xs) ! n = xs ! (n - 1)" -by(auto simp: Nat.gr0_conv_Suc) + by(auto simp: Nat.gr0_conv_Suc) lemma nth_append: "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))" -apply (induct xs arbitrary: n, simp) -apply (case_tac n, auto) -done +proof (induct xs arbitrary: n) + case (Cons x xs) + then show ?case + using less_Suc_eq_0_disj by auto +qed simp lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x" -by (induct xs) auto + by (induct xs) auto lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n" -by (induct xs) auto + by (induct xs) auto lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)" -apply (induct xs arbitrary: n, simp) -apply (case_tac n, auto) -done +proof (induct xs arbitrary: n) + case (Cons x xs) + then show ?case + using less_Suc_eq_0_disj by auto +qed simp lemma nth_tl: "n < length (tl xs) \ tl xs ! n = xs ! Suc n" -by (induction xs) auto + by (induction xs) auto lemma hd_conv_nth: "xs \ [] \ hd xs = xs!0" -by(cases xs) simp_all + by(cases xs) simp_all lemma list_eq_iff_nth_eq: "(xs = ys) = (length xs = length ys \ (\i ?R" + by force + show "?R \ ?L" + using less_Suc_eq_0_disj by auto + qed + with Cons show ?case + by simp +qed simp lemma in_set_conv_nth: "(x \ set xs) = (\i < length xs. xs!i = x)" -by(auto simp:set_conv_nth) + by(auto simp:set_conv_nth) lemma nth_equal_first_eq: assumes "x \ set xs" @@ -1761,18 +1755,18 @@ qed lemma list_ball_nth: "\n < length xs; \x \ set xs. P x\ \ P(xs!n)" -by (auto simp add: set_conv_nth) + by (auto simp add: set_conv_nth) lemma nth_mem [simp]: "n < length xs \ xs!n \ set xs" -by (auto simp add: set_conv_nth) + by (auto simp add: set_conv_nth) lemma all_nth_imp_all_set: "\\i < length xs. P(xs!i); x \ set xs\ \ P x" -by (auto simp add: set_conv_nth) + by (auto simp add: set_conv_nth) lemma all_set_conv_all_nth: "(\x \ set xs. P x) = (\i. i < length xs \ P (xs ! i))" -by (auto simp add: set_conv_nth) + by (auto simp add: set_conv_nth) lemma rev_nth: "n < size xs \ rev xs ! n = xs ! (length xs - Suc n)" @@ -1816,149 +1810,141 @@ subsubsection \@{const list_update}\ lemma length_list_update [simp]: "length(xs[i:=x]) = length xs" -by (induct xs arbitrary: i) (auto split: nat.split) + by (induct xs arbitrary: i) (auto split: nat.split) lemma nth_list_update: -"i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)" -by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) + "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)" + by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x" -by (simp add: nth_list_update) + by (simp add: nth_list_update) lemma nth_list_update_neq [simp]: "i \ j ==> xs[i:=x]!j = xs!j" -by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) + by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split) lemma list_update_id[simp]: "xs[i := xs!i] = xs" -by (induct xs arbitrary: i) (simp_all split:nat.splits) + by (induct xs arbitrary: i) (simp_all split:nat.splits) lemma list_update_beyond[simp]: "length xs \ i \ xs[i:=x] = xs" -apply (induct xs arbitrary: i) - apply simp -apply (case_tac i) -apply simp_all -done +proof (induct xs arbitrary: i) + case (Cons x xs i) + then show ?case + by (metis leD length_list_update list_eq_iff_nth_eq nth_list_update_neq) +qed simp lemma list_update_nonempty[simp]: "xs[k:=x] = [] \ xs=[]" -by (simp only: length_0_conv[symmetric] length_list_update) + by (simp only: length_0_conv[symmetric] length_list_update) lemma list_update_same_conv: "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)" -by (induct xs arbitrary: i) (auto split: nat.split) + by (induct xs arbitrary: i) (auto split: nat.split) lemma list_update_append1: "i < size xs \ (xs @ ys)[i:=x] = xs[i:=x] @ ys" -by (induct xs arbitrary: i)(auto split:nat.split) + by (induct xs arbitrary: i)(auto split:nat.split) lemma list_update_append: "(xs @ ys) [n:= x] = (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))" -by (induct xs arbitrary: n) (auto split:nat.splits) + by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_length [simp]: "(xs @ x # ys)[length xs := y] = (xs @ y # ys)" -by (induct xs, auto) + by (induct xs, auto) lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]" -by(induct xs arbitrary: k)(auto split:nat.splits) + by(induct xs arbitrary: k)(auto split:nat.splits) lemma rev_update: "k < length xs \ rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]" -by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits) + by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits) lemma update_zip: "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])" -by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split) - -lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)" -by (induct xs arbitrary: i) (auto split: nat.split) + by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split) + +lemma set_update_subset_insert: "set(xs[i:=x]) \ insert x (set xs)" + by (induct xs arbitrary: i) (auto split: nat.split) lemma set_update_subsetI: "\set xs \ A; x \ A\ \ set(xs[i := x]) \ A" -by (blast dest!: set_update_subset_insert [THEN subsetD]) + by (blast dest!: set_update_subset_insert [THEN subsetD]) lemma set_update_memI: "n < length xs \ x \ set (xs[n := x])" -by (induct xs arbitrary: n) (auto split:nat.splits) + by (induct xs arbitrary: n) (auto split:nat.splits) lemma list_update_overwrite[simp]: "xs [i := x, i := y] = xs [i := y]" -apply (induct xs arbitrary: i) apply simp -apply (case_tac i, simp_all) -done + by (induct xs arbitrary: i) (simp_all split: nat.split) lemma list_update_swap: "i \ i' \ xs [i := x, i' := x'] = xs [i' := x', i := x]" -apply (induct xs arbitrary: i i') - apply simp -apply (case_tac i, case_tac i') - apply auto -apply (case_tac i') -apply auto -done + by (induct xs arbitrary: i i') (simp_all split: nat.split) lemma list_update_code [code]: "[][i := y] = []" "(x # xs)[0 := y] = y # xs" "(x # xs)[Suc i := y] = x # xs[i := y]" -by simp_all + by simp_all subsubsection \@{const last} and @{const butlast}\ lemma last_snoc [simp]: "last (xs @ [x]) = x" -by (induct xs) auto + by (induct xs) auto lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs" -by (induct xs) auto + by (induct xs) auto lemma last_ConsL: "xs = [] \ last(x#xs) = x" -by simp + by simp lemma last_ConsR: "xs \ [] \ last(x#xs) = last xs" -by simp + by simp lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)" -by (induct xs) (auto) + by (induct xs) (auto) lemma last_appendL[simp]: "ys = [] \ last(xs @ ys) = last xs" -by(simp add:last_append) + by(simp add:last_append) lemma last_appendR[simp]: "ys \ [] \ last(xs @ ys) = last ys" -by(simp add:last_append) + by(simp add:last_append) lemma last_tl: "xs = [] \ tl xs \ [] \last (tl xs) = last xs" -by (induct xs) simp_all + by (induct xs) simp_all lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)" -by (induct xs) simp_all + by (induct xs) simp_all lemma hd_rev: "xs \ [] \ hd(rev xs) = last xs" -by(rule rev_exhaust[of xs]) simp_all + by(rule rev_exhaust[of xs]) simp_all lemma last_rev: "xs \ [] \ last(rev xs) = hd xs" -by(cases xs) simp_all + by(cases xs) simp_all lemma last_in_set[simp]: "as \ [] \ last as \ set as" -by (induct as) auto + by (induct as) auto lemma length_butlast [simp]: "length (butlast xs) = length xs - 1" -by (induct xs rule: rev_induct) auto + by (induct xs rule: rev_induct) auto lemma butlast_append: "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)" -by (induct xs arbitrary: ys) auto + by (induct xs arbitrary: ys) auto lemma append_butlast_last_id [simp]: "xs \ [] \ butlast xs @ [last xs] = xs" -by (induct xs) auto + by (induct xs) auto lemma in_set_butlastD: "x \ set (butlast xs) \ x \ set xs" -by (induct xs) (auto split: if_split_asm) + by (induct xs) (auto split: if_split_asm) lemma in_set_butlast_appendI: "x \ set (butlast xs) \ x \ set (butlast ys) \ x \ set (butlast (xs @ ys))" -by (auto dest: in_set_butlastD simp add: butlast_append) + by (auto dest: in_set_butlastD simp add: butlast_append) lemma last_drop[simp]: "n < length xs \ last (drop n xs) = last xs" -by (induct xs arbitrary: n)(auto split:nat.split) + by (induct xs arbitrary: n)(auto split:nat.split) lemma nth_butlast: assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n" @@ -1970,164 +1956,173 @@ qed simp lemma last_conv_nth: "xs\[] \ last xs = xs!(length xs - 1)" -by(induct xs)(auto simp:neq_Nil_conv) + by(induct xs)(auto simp:neq_Nil_conv) lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs" -by (induction xs rule: induct_list012) simp_all + by (induction xs rule: induct_list012) simp_all lemma last_list_update: "xs \ [] \ last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)" -by (auto simp: last_conv_nth) + by (auto simp: last_conv_nth) lemma butlast_list_update: "butlast(xs[k:=x]) = (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])" -by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits) + by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits) lemma last_map: "xs \ [] \ last (map f xs) = f (last xs)" -by (cases xs rule: rev_cases) simp_all + by (cases xs rule: rev_cases) simp_all lemma map_butlast: "map f (butlast xs) = butlast (map f xs)" -by (induct xs) simp_all + by (induct xs) simp_all lemma snoc_eq_iff_butlast: "xs @ [x] = ys \ (ys \ [] \ butlast ys = xs \ last ys = x)" -by fastforce + by fastforce corollary longest_common_suffix: "\ss xs' ys'. xs = xs' @ ss \ ys = ys' @ ss \ (xs' = [] \ ys' = [] \ last xs' \ last ys')" -using longest_common_prefix[of "rev xs" "rev ys"] -unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) + using longest_common_prefix[of "rev xs" "rev ys"] + unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) subsubsection \@{const take} and @{const drop}\ lemma take_0: "take 0 xs = []" -by (induct xs) auto + by (induct xs) auto lemma drop_0: "drop 0 xs = xs" -by (induct xs) auto + by (induct xs) auto lemma take0[simp]: "take 0 = (\xs. [])" -by(rule ext) (rule take_0) + by(rule ext) (rule take_0) lemma drop0[simp]: "drop 0 = (\x. x)" -by(rule ext) (rule drop_0) + by(rule ext) (rule drop_0) lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs" -by simp + by simp lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs" -by simp + by simp declare take_Cons [simp del] and drop_Cons [simp del] lemma take_Suc: "xs \ [] \ take (Suc n) xs = hd xs # take n (tl xs)" -by(clarsimp simp add:neq_Nil_conv) + by(clarsimp simp add:neq_Nil_conv) lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)" -by(cases xs, simp_all) + by(cases xs, simp_all) lemma hd_take[simp]: "j > 0 \ hd (take j xs) = hd xs" -by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc) + by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc) lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)" -by (induct xs arbitrary: n) simp_all + by (induct xs arbitrary: n) simp_all lemma drop_tl: "drop n (tl xs) = tl(drop n xs)" -by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split) + by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split) lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)" -by (cases n, simp, cases xs, auto) + by (cases n, simp, cases xs, auto) lemma tl_drop: "tl (drop n xs) = drop n (tl xs)" -by (simp only: drop_tl) + by (simp only: drop_tl) lemma nth_via_drop: "drop n xs = y#ys \ xs!n = y" -by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits) + by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits) lemma take_Suc_conv_app_nth: "i < length xs \ take (Suc i) xs = take i xs @ [xs!i]" -apply (induct xs arbitrary: i, simp) -apply (case_tac i, auto) -done +proof (induct xs arbitrary: i) + case (Cons x xs) then show ?case + by (case_tac i, auto) +qed simp lemma Cons_nth_drop_Suc: "i < length xs \ (xs!i) # (drop (Suc i) xs) = drop i xs" -apply (induct xs arbitrary: i, simp) -apply (case_tac i, auto) -done +proof (induct xs arbitrary: i) + case (Cons x xs) then show ?case + by (case_tac i, auto) +qed simp lemma length_take [simp]: "length (take n xs) = min (length xs) n" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) + by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma length_drop [simp]: "length (drop n xs) = (length xs - n)" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) - -lemma take_all [simp]: "length xs <= n ==> take n xs = xs" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) - -lemma drop_all [simp]: "length xs <= n ==> drop n xs = []" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) + by (induct n arbitrary: xs) (auto, case_tac xs, auto) + +lemma take_all [simp]: "length xs \ n ==> take n xs = xs" + by (induct n arbitrary: xs) (auto, case_tac xs, auto) + +lemma drop_all [simp]: "length xs \ n ==> drop n xs = []" + by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_append [simp]: "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) + by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma drop_append [simp]: "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys" -by (induct n arbitrary: xs) (auto, case_tac xs, auto) + by (induct n arbitrary: xs) (auto, case_tac xs, auto) lemma take_take [simp]: "take n (take m xs) = take (min n m) xs" -apply (induct m arbitrary: xs n, auto) - apply (case_tac xs, auto) -apply (case_tac n, auto) -done +proof (induct m arbitrary: xs n) + case (Suc m) then show ?case + by (case_tac xs; case_tac n; simp) +qed auto lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs" -apply (induct m arbitrary: xs, auto) - apply (case_tac xs, auto) -done +proof (induct m arbitrary: xs) + case (Suc m) then show ?case + by (case_tac xs; simp) +qed auto lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)" -apply (induct m arbitrary: xs n, auto) - apply (case_tac xs, auto) -done +proof (induct m arbitrary: xs n) + case (Suc m) then show ?case + by (case_tac xs; case_tac n; simp) +qed auto lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)" -by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split) + by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split) lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs" -apply (induct n arbitrary: xs, auto) -apply (case_tac xs, auto) -done +proof (induct n arbitrary: xs) + case (Suc n) then show ?case + by (case_tac xs; simp) +qed auto lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \ xs = [])" -by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) - -lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)" -by (induct xs arbitrary: n) (auto simp: drop_Cons split:nat.split) + by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) + +lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs \ n)" + by (induct xs arbitrary: n) (auto simp: drop_Cons split:nat.split) lemma take_map: "take n (map f xs) = map f (take n xs)" -apply (induct n arbitrary: xs, auto) - apply (case_tac xs, auto) -done +proof (induct n arbitrary: xs) + case (Suc n) then show ?case + by (case_tac xs; simp) +qed auto lemma drop_map: "drop n (map f xs) = map f (drop n xs)" -apply (induct n arbitrary: xs, auto) - apply (case_tac xs, auto) -done +proof (induct n arbitrary: xs) + case (Suc n) then show ?case + by (case_tac xs; simp) +qed auto lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)" -apply (induct xs arbitrary: i, auto) - apply (case_tac i, auto) -done +proof (induct xs arbitrary: i) + case (Cons x xs) then show ?case + by (case_tac i, auto) +qed simp lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)" -apply (induct xs arbitrary: i, auto) - apply (case_tac i, auto) -done +proof (induct xs arbitrary: i) + case (Cons x xs) then show ?case + by (case_tac i, auto) +qed simp lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)" by (cases "length xs < n") (auto simp: rev_take) @@ -2136,87 +2131,87 @@ by (cases "length xs < n") (auto simp: rev_drop) lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i" -apply (induct xs arbitrary: i n, auto) - apply (case_tac n, blast) -apply (case_tac i, auto) -done +proof (induct xs arbitrary: i n) + case (Cons x xs) then show ?case + by (case_tac n; case_tac i; simp) +qed auto lemma nth_drop [simp]: - "n <= length xs ==> (drop n xs)!i = xs!(n + i)" -apply (induct n arbitrary: xs i, auto) - apply (case_tac xs, auto) -done + "n \ length xs ==> (drop n xs)!i = xs!(n + i)" +proof (induct n arbitrary: xs) + case (Suc n) then show ?case + by (case_tac xs; simp) +qed auto lemma butlast_take: - "n <= length xs ==> butlast (take n xs) = take (n - 1) xs" -by (simp add: butlast_conv_take min.absorb1 min.absorb2) + "n \ length xs ==> butlast (take n xs) = take (n - 1) xs" + by (simp add: butlast_conv_take min.absorb1 min.absorb2) lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)" -by (simp add: butlast_conv_take drop_take ac_simps) + by (simp add: butlast_conv_take drop_take ac_simps) lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs" -by (simp add: butlast_conv_take min.absorb1) + by (simp add: butlast_conv_take min.absorb1) lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" -by (simp add: butlast_conv_take drop_take ac_simps) + by (simp add: butlast_conv_take drop_take ac_simps) lemma hd_drop_conv_nth: "n < length xs \ hd(drop n xs) = xs!n" -by(simp add: hd_conv_nth) + by(simp add: hd_conv_nth) lemma set_take_subset_set_take: - "m <= n \ set(take m xs) <= set(take n xs)" -apply (induct xs arbitrary: m n) - apply simp -apply (case_tac n) -apply (auto simp: take_Cons) -done + "m \ n \ set(take m xs) \ set(take n xs)" +proof (induct xs arbitrary: m n) + case (Cons x xs m n) then show ?case + by (cases n) (auto simp: take_Cons) +qed simp lemma set_take_subset: "set(take n xs) \ set xs" -by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split) + by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split) lemma set_drop_subset: "set(drop n xs) \ set xs" -by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split) + by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split) lemma set_drop_subset_set_drop: - "m >= n \ set(drop m xs) <= set(drop n xs)" -apply(induct xs arbitrary: m n) - apply(auto simp:drop_Cons split:nat.split) -by (metis set_drop_subset subset_iff) + "m \ n \ set(drop m xs) \ set(drop n xs)" +proof (induct xs arbitrary: m n) + case (Cons x xs m n) + then show ?case + by (clarsimp simp: drop_Cons split: nat.split) (metis set_drop_subset subset_iff) +qed simp lemma in_set_takeD: "x \ set(take n xs) \ x \ set xs" -using set_take_subset by fast + using set_take_subset by fast lemma in_set_dropD: "x \ set(drop n xs) \ x \ set xs" -using set_drop_subset by fast + using set_drop_subset by fast lemma append_eq_conv_conj: "(xs @ ys = zs) = (xs = take (length xs) zs \ ys = drop (length xs) zs)" -apply (induct xs arbitrary: zs, simp, clarsimp) - apply (case_tac zs, auto) -done +proof (induct xs arbitrary: zs) + case (Cons x xs zs) then show ?case + by (cases zs, auto) +qed auto lemma take_add: "take (i+j) xs = take i xs @ take j (drop i xs)" -apply (induct xs arbitrary: i, auto) - apply (case_tac i, simp_all) -done +proof (induct xs arbitrary: i) + case (Cons x xs i) then show ?case + by (cases i, auto) +qed auto lemma append_eq_append_conv_if: "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) = (if size xs\<^sub>1 \ size ys\<^sub>1 then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \ xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2 else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \ drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)" -apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1) - apply simp -apply(case_tac ys\<^sub>1) -apply simp_all -done +proof (induct xs\<^sub>1 arbitrary: ys\<^sub>1) + case (Cons a xs\<^sub>1 ys\<^sub>1) then show ?case + by (cases ys\<^sub>1, auto) +qed auto lemma take_hd_drop: "n < length xs \ take n xs @ [hd (drop n xs)] = take (Suc n) xs" -apply(induct xs arbitrary: n) - apply simp -apply(simp add:drop_Cons split:nat.split) -done + by (induct xs arbitrary: n) (simp_all add:drop_Cons split:nat.split) lemma id_take_nth_drop: "i < length xs \ xs = take i xs @ xs!i # drop (Suc i) xs" @@ -2225,15 +2220,15 @@ hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto moreover from si have "take (Suc i) xs = take i xs @ [xs!i]" - apply (rule_tac take_Suc_conv_app_nth) by arith + using take_Suc_conv_app_nth by blast ultimately show ?thesis by auto qed lemma take_update_cancel[simp]: "n \ m \ take n (xs[m := y]) = take n xs" -by(simp add: list_eq_iff_nth_eq) + by(simp add: list_eq_iff_nth_eq) lemma drop_update_cancel[simp]: "n < m \ drop m (xs[n := x]) = drop m xs" -by(simp add: list_eq_iff_nth_eq) + by(simp add: list_eq_iff_nth_eq) lemma upd_conv_take_nth_drop: "i < length xs \ xs[i:=a] = take i xs @ a # drop (Suc i) xs" @@ -2247,105 +2242,107 @@ qed lemma take_update_swap: "take m (xs[n := x]) = (take m xs)[n := x]" -apply(cases "n \ length xs") - apply simp -apply(simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc - split: nat.split) -done - -lemma drop_update_swap: "m \ n \ drop m (xs[n := x]) = (drop m xs)[n-m := x]" -apply(cases "n \ length xs") - apply simp -apply(simp add: upd_conv_take_nth_drop drop_take) -done +proof (cases "n \ length xs") + case False + then show ?thesis + by (simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc split: nat.split) +qed auto + +lemma drop_update_swap: + assumes "m \ n" shows "drop m (xs[n := x]) = (drop m xs)[n-m := x]" +proof (cases "n \ length xs") + case False + with assms show ?thesis + by (simp add: upd_conv_take_nth_drop drop_take) +qed auto lemma nth_image: "l \ size xs \ nth xs ` {0..@{const takeWhile} and @{const dropWhile}\ lemma length_takeWhile_le: "length (takeWhile P xs) \ length xs" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ takeWhile P (xs @ ys) = takeWhile P xs" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_append2 [simp]: "(\x. x \ set xs \ P x) \ takeWhile P (xs @ ys) = xs @ takeWhile P ys" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_tail: "\ P x \ takeWhile P (xs @ (x#l)) = takeWhile P xs" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_nth: "j < length (takeWhile P xs) \ takeWhile P xs ! j = xs ! j" -apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto + by (metis nth_append takeWhile_dropWhile_id) lemma dropWhile_nth: "j < length (dropWhile P xs) \ dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))" -apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto + by (metis add.commute nth_append_length_plus takeWhile_dropWhile_id) lemma length_dropWhile_le: "length (dropWhile P xs) \ length xs" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_append1 [simp]: "\x \ set xs; \P(x)\ \ dropWhile P (xs @ ys) = (dropWhile P xs)@ys" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_append2 [simp]: "(\x. x \ set xs \ P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_append3: "\ P y \dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_last: "x \ set xs \ \ P x \ last (dropWhile P xs) = last xs" -by (auto simp add: dropWhile_append3 in_set_conv_decomp) + by (auto simp add: dropWhile_append3 in_set_conv_decomp) lemma set_dropWhileD: "x \ set (dropWhile P xs) \ x \ set xs" -by (induct xs) (auto split: if_split_asm) + by (induct xs) (auto split: if_split_asm) lemma set_takeWhileD: "x \ set (takeWhile P xs) \ x \ set xs \ P x" -by (induct xs) (auto split: if_split_asm) + by (induct xs) (auto split: if_split_asm) lemma takeWhile_eq_all_conv[simp]: "(takeWhile P xs = xs) = (\x \ set xs. P x)" -by(induct xs, auto) + by(induct xs, auto) lemma dropWhile_eq_Nil_conv[simp]: "(dropWhile P xs = []) = (\x \ set xs. P x)" -by(induct xs, auto) + by(induct xs, auto) lemma dropWhile_eq_Cons_conv: "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \ \ P y)" -by(induct xs, auto) + by(induct xs, auto) lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)" -by (induct xs) (auto dest: set_takeWhileD) + by (induct xs) (auto dest: set_takeWhileD) lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \ f) xs)" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \ f) xs)" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs" -by (induct xs) auto + by (induct xs) auto lemma hd_dropWhile: "dropWhile P xs \ [] \ \ P (hd (dropWhile P xs))" -by (induct xs) auto + by (induct xs) auto lemma takeWhile_eq_filter: assumes "\ x. x \ set (dropWhile P xs) \ \ P x" @@ -2386,12 +2383,12 @@ thus "\ P (xs ! n')" using Cons by auto qed ultimately show ?thesis by simp - qed + qed qed lemma nth_length_takeWhile: "length (takeWhile P xs) < length xs \ \ P (xs ! length (takeWhile P xs))" -by (induct xs) auto + by (induct xs) auto lemma length_takeWhile_less_P_nth: assumes all: "\ i. i < j \ P (xs ! i)" and "j \ length xs" @@ -2404,47 +2401,46 @@ lemma takeWhile_neq_rev: "\distinct xs; x \ set xs\ \ takeWhile (\y. y \ x) (rev xs) = rev (tl (dropWhile (\y. y \ x) xs))" -by(induct xs) (auto simp: takeWhile_tail[where l="[]"]) + by(induct xs) (auto simp: takeWhile_tail[where l="[]"]) lemma dropWhile_neq_rev: "\distinct xs; x \ set xs\ \ dropWhile (\y. y \ x) (rev xs) = x # rev (takeWhile (\y. y \ x) xs)" -apply(induct xs) - apply simp -apply auto -apply(subst dropWhile_append2) -apply auto -done +proof (induct xs) + case (Cons a xs) + then show ?case + by(auto, subst dropWhile_append2, auto) +qed simp lemma takeWhile_not_last: "distinct xs \ takeWhile (\y. y \ last xs) xs = butlast xs" -by(induction xs rule: induct_list012) auto + by(induction xs rule: induct_list012) auto lemma takeWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ takeWhile P l = takeWhile Q k" -by (induct k arbitrary: l) (simp_all) + by (induct k arbitrary: l) (simp_all) lemma dropWhile_cong [fundef_cong]: "\l = k; \x. x \ set l \ P x = Q x\ \ dropWhile P l = dropWhile Q k" -by (induct k arbitrary: l, simp_all) + by (induct k arbitrary: l, simp_all) lemma takeWhile_idem [simp]: "takeWhile P (takeWhile P xs) = takeWhile P xs" -by (induct xs) auto + by (induct xs) auto lemma dropWhile_idem [simp]: "dropWhile P (dropWhile P xs) = dropWhile P xs" -by (induct xs) auto + by (induct xs) auto subsubsection \@{const zip}\ lemma zip_Nil [simp]: "zip [] ys = []" -by (induct ys) auto + by (induct ys) auto lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys" -by simp + by simp declare zip_Cons [simp del] @@ -2452,15 +2448,15 @@ "zip [] ys = []" "zip xs [] = []" "zip (x # xs) (y # ys) = (x, y) # zip xs ys" -by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+ + by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+ lemma zip_Cons1: "zip (x#xs) ys = (case ys of [] \ [] | y#ys \ (x,y)#zip xs ys)" -by(auto split:list.split) + by(auto split:list.split) lemma length_zip [simp]: "length (zip xs ys) = min (length xs) (length ys)" -by (induct xs ys rule:list_induct2') auto + by (induct xs ys rule:list_induct2') auto lemma zip_obtain_same_length: assumes "\zs ws n. length zs = length ws \ n = min (length xs) (length ys) @@ -2482,21 +2478,21 @@ lemma zip_append1: "zip (xs @ ys) zs = zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)" -by (induct xs zs rule:list_induct2') auto + by (induct xs zs rule:list_induct2') auto lemma zip_append2: "zip xs (ys @ zs) = zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs" -by (induct xs ys rule:list_induct2') auto + by (induct xs ys rule:list_induct2') auto lemma zip_append [simp]: "[| length xs = length us |] ==> zip (xs@ys) (us@vs) = zip xs us @ zip ys vs" -by (simp add: zip_append1) + by (simp add: zip_append1) lemma zip_rev: "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)" -by (induct rule:list_induct2, simp_all) + by (induct rule:list_induct2, simp_all) lemma zip_map_map: "zip (map f xs) (map g ys) = map (\ (x, y). (f x, g y)) (zip xs ys)" @@ -2511,66 +2507,66 @@ lemma zip_map1: "zip (map f xs) ys = map (\(x, y). (f x, y)) (zip xs ys)" -using zip_map_map[of f xs "\x. x" ys] by simp + using zip_map_map[of f xs "\x. x" ys] by simp lemma zip_map2: "zip xs (map f ys) = map (\(x, y). (x, f y)) (zip xs ys)" -using zip_map_map[of "\x. x" xs f ys] by simp + using zip_map_map[of "\x. x" xs f ys] by simp lemma map_zip_map: "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)" -by (auto simp: zip_map1) + by (auto simp: zip_map1) lemma map_zip_map2: "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)" -by (auto simp: zip_map2) + by (auto simp: zip_map2) text\Courtesy of Andreas Lochbihler:\ lemma zip_same_conv_map: "zip xs xs = map (\x. (x, x)) xs" -by(induct xs) auto + by(induct xs) auto lemma nth_zip [simp]: "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)" -apply (induct ys arbitrary: i xs, simp) -apply (case_tac xs) - apply (simp_all add: nth.simps split: nat.split) -done +proof (induct ys arbitrary: i xs) + case (Cons y ys) + then show ?case + by (cases xs) (simp_all add: nth.simps split: nat.split) +qed auto lemma set_zip: "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}" -by(simp add: set_conv_nth cong: rev_conj_cong) + by(simp add: set_conv_nth cong: rev_conj_cong) lemma zip_same: "((a,b) \ set (zip xs xs)) = (a \ set xs \ a = b)" -by(induct xs) auto - -lemma zip_update: - "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]" -by(rule sym, simp add: update_zip) + by(induct xs) auto + +lemma zip_update: "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]" + by (simp add: update_zip) lemma zip_replicate [simp]: "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)" -apply (induct i arbitrary: j, auto) -apply (case_tac j, auto) -done +proof (induct i arbitrary: j) + case (Suc i) + then show ?case + by (cases j, auto) +qed auto lemma zip_replicate1: "zip (replicate n x) ys = map (Pair x) (take n ys)" -by(induction ys arbitrary: n)(case_tac [2] n, simp_all) - -lemma take_zip: - "take n (zip xs ys) = zip (take n xs) (take n ys)" -apply (induct n arbitrary: xs ys) - apply simp -apply (case_tac xs, simp) -apply (case_tac ys, simp_all) -done - -lemma drop_zip: - "drop n (zip xs ys) = zip (drop n xs) (drop n ys)" -apply (induct n arbitrary: xs ys) - apply simp -apply (case_tac xs, simp) -apply (case_tac ys, simp_all) -done + by(induction ys arbitrary: n)(case_tac [2] n, simp_all) + +lemma take_zip: "take n (zip xs ys) = zip (take n xs) (take n ys)" +proof (induct n arbitrary: xs ys) + case (Suc n) + then show ?case + by (case_tac xs; case_tac ys; simp) +qed simp + +lemma drop_zip: "drop n (zip xs ys) = zip (drop n xs) (drop n ys)" +proof (induct n arbitrary: xs ys) + case (Suc n) + then show ?case + by (case_tac xs; case_tac ys; simp) +qed simp lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \ fst) (zip xs ys)" proof (induct xs arbitrary: ys) @@ -2583,26 +2579,26 @@ qed simp lemma set_zip_leftD: "(x,y)\ set (zip xs ys) \ x \ set xs" -by (induct xs ys rule:list_induct2') auto + by (induct xs ys rule:list_induct2') auto lemma set_zip_rightD: "(x,y)\ set (zip xs ys) \ y \ set ys" -by (induct xs ys rule:list_induct2') auto + by (induct xs ys rule:list_induct2') auto lemma in_set_zipE: "(x,y) \ set(zip xs ys) \ (\ x \ set xs; y \ set ys \ \ R) \ R" -by(blast dest: set_zip_leftD set_zip_rightD) + by(blast dest: set_zip_leftD set_zip_rightD) lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs" -by (induct zs) simp_all + by (induct zs) simp_all lemma zip_eq_conv: "length xs = length ys \ zip xs ys = zs \ map fst zs = xs \ map snd zs = ys" -by (auto simp add: zip_map_fst_snd) + by (auto simp add: zip_map_fst_snd) lemma in_set_zip: "p \ set (zip xs ys) \ (\n. xs ! n = fst p \ ys ! n = snd p \ n < length xs \ n < length ys)" -by (cases p) (auto simp add: set_zip) + by (cases p) (auto simp add: set_zip) lemma in_set_impl_in_set_zip1: assumes "length xs = length ys" @@ -2636,25 +2632,25 @@ lemma list_all2_lengthD [intro?]: "list_all2 P xs ys ==> length xs = length ys" -by (simp add: list_all2_iff) + by (simp add: list_all2_iff) lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])" -by (simp add: list_all2_iff) + by (simp add: list_all2_iff) lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])" -by (simp add: list_all2_iff) + by (simp add: list_all2_iff) lemma list_all2_Cons [iff, code]: "list_all2 P (x # xs) (y # ys) = (P x y \ list_all2 P xs ys)" -by (auto simp add: list_all2_iff) + by (auto simp add: list_all2_iff) lemma list_all2_Cons1: "list_all2 P (x # xs) ys = (\z zs. ys = z # zs \ P x z \ list_all2 P xs zs)" -by (cases ys) auto + by (cases ys) auto lemma list_all2_Cons2: "list_all2 P xs (y # ys) = (\z zs. xs = z # zs \ P z y \ list_all2 P zs ys)" -by (cases xs) auto + by (cases xs) auto lemma list_all2_induct [consumes 1, case_names Nil Cons, induct set: list_all2]: @@ -2663,59 +2659,69 @@ assumes Cons: "\x xs y ys. \P x y; list_all2 P xs ys; R xs ys\ \ R (x # xs) (y # ys)" shows "R xs ys" -using P -by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons) + using P + by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons) lemma list_all2_rev [iff]: "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys" -by (simp add: list_all2_iff zip_rev cong: conj_cong) + by (simp add: list_all2_iff zip_rev cong: conj_cong) lemma list_all2_rev1: "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)" -by (subst list_all2_rev [symmetric]) simp + by (subst list_all2_rev [symmetric]) simp lemma list_all2_append1: "list_all2 P (xs @ ys) zs = (\us vs. zs = us @ vs \ length us = length xs \ length vs = length ys \ - list_all2 P xs us \ list_all2 P ys vs)" -apply (simp add: list_all2_iff zip_append1) -apply (rule iffI) - apply (rule_tac x = "take (length xs) zs" in exI) - apply (rule_tac x = "drop (length xs) zs" in exI) - apply (force split: nat_diff_split simp add: min_def, clarify) -apply (simp add: ball_Un) -done + list_all2 P xs us \ list_all2 P ys vs)" (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (rule_tac x = "take (length xs) zs" in exI) + apply (rule_tac x = "drop (length xs) zs" in exI) + apply (force split: nat_diff_split simp add: list_all2_iff zip_append1) + done +next + assume ?rhs + then show ?lhs + by (auto simp add: list_all2_iff) +qed lemma list_all2_append2: "list_all2 P xs (ys @ zs) = (\us vs. xs = us @ vs \ length us = length ys \ length vs = length zs \ - list_all2 P us ys \ list_all2 P vs zs)" -apply (simp add: list_all2_iff zip_append2) -apply (rule iffI) - apply (rule_tac x = "take (length ys) xs" in exI) - apply (rule_tac x = "drop (length ys) xs" in exI) - apply (force split: nat_diff_split simp add: min_def, clarify) -apply (simp add: ball_Un) -done + list_all2 P us ys \ list_all2 P vs zs)" (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (rule_tac x = "take (length ys) xs" in exI) + apply (rule_tac x = "drop (length ys) xs" in exI) + apply (force split: nat_diff_split simp add: list_all2_iff zip_append2) + done +next + assume ?rhs + then show ?lhs + by (auto simp add: list_all2_iff) +qed lemma list_all2_append: "length xs = length ys \ list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \ list_all2 P us vs)" -by (induct rule:list_induct2, simp_all) + by (induct rule:list_induct2, simp_all) lemma list_all2_appendI [intro?, trans]: "\ list_all2 P a b; list_all2 P c d \ \ list_all2 P (a@c) (b@d)" -by (simp add: list_all2_append list_all2_lengthD) + by (simp add: list_all2_append list_all2_lengthD) lemma list_all2_conv_all_nth: "list_all2 P xs ys = (length xs = length ys \ (\i < length xs. P (xs!i) (ys!i)))" -by (force simp add: list_all2_iff set_zip) + by (force simp add: list_all2_iff set_zip) lemma list_all2_trans: assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c" shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs" - (is "!!bs cs. PROP ?Q as bs cs") + (is "!!bs cs. PROP ?Q as bs cs") proof (induct as) fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs" show "!!cs. PROP ?Q (x # xs) bs cs" @@ -2728,100 +2734,98 @@ lemma list_all2_all_nthI [intro?]: "length a = length b \ (\n. n < length a \ P (a!n) (b!n)) \ list_all2 P a b" -by (simp add: list_all2_conv_all_nth) + by (simp add: list_all2_conv_all_nth) lemma list_all2I: "\x \ set (zip a b). case_prod P x \ length a = length b \ list_all2 P a b" -by (simp add: list_all2_iff) + by (simp add: list_all2_iff) lemma list_all2_nthD: "\ list_all2 P xs ys; p < size xs \ \ P (xs!p) (ys!p)" -by (simp add: list_all2_conv_all_nth) + by (simp add: list_all2_conv_all_nth) lemma list_all2_nthD2: "\list_all2 P xs ys; p < size ys\ \ P (xs!p) (ys!p)" -by (frule list_all2_lengthD) (auto intro: list_all2_nthD) + by (frule list_all2_lengthD) (auto intro: list_all2_nthD) lemma list_all2_map1: "list_all2 P (map f as) bs = list_all2 (\x y. P (f x) y) as bs" -by (simp add: list_all2_conv_all_nth) + by (simp add: list_all2_conv_all_nth) lemma list_all2_map2: "list_all2 P as (map f bs) = list_all2 (\x y. P x (f y)) as bs" -by (auto simp add: list_all2_conv_all_nth) + by (auto simp add: list_all2_conv_all_nth) lemma list_all2_refl [intro?]: "(\x. P x x) \ list_all2 P xs xs" -by (simp add: list_all2_conv_all_nth) + by (simp add: list_all2_conv_all_nth) lemma list_all2_update_cong: "\ list_all2 P xs ys; P x y \ \ list_all2 P (xs[i:=x]) (ys[i:=y])" -by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update) + by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update) lemma list_all2_takeI [simp,intro?]: "list_all2 P xs ys \ list_all2 P (take n xs) (take n ys)" -apply (induct xs arbitrary: n ys) - apply simp -apply (clarsimp simp add: list_all2_Cons1) -apply (case_tac n) -apply auto -done +proof (induct xs arbitrary: n ys) + case (Cons x xs) + then show ?case + by (cases n) (auto simp: list_all2_Cons1) +qed auto lemma list_all2_dropI [simp,intro?]: - "list_all2 P as bs \ list_all2 P (drop n as) (drop n bs)" -apply (induct as arbitrary: n bs, simp) -apply (clarsimp simp add: list_all2_Cons1) -apply (case_tac n, simp, simp) -done + "list_all2 P xs ys \ list_all2 P (drop n xs) (drop n ys)" +proof (induct xs arbitrary: n ys) + case (Cons x xs) + then show ?case + by (cases n) (auto simp: list_all2_Cons1) +qed auto lemma list_all2_mono [intro?]: "list_all2 P xs ys \ (\xs ys. P xs ys \ Q xs ys) \ list_all2 Q xs ys" -apply (induct xs arbitrary: ys, simp) -apply (case_tac ys, auto) -done + by (rule list.rel_mono_strong) lemma list_all2_eq: "xs = ys \ list_all2 (=) xs ys" -by (induct xs ys rule: list_induct2') auto + by (induct xs ys rule: list_induct2') auto lemma list_eq_iff_zip_eq: "xs = ys \ length xs = length ys \ (\(x,y) \ set (zip xs ys). x = y)" -by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong) + by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong) lemma list_all2_same: "list_all2 P xs xs \ (\x\set xs. P x x)" -by(auto simp add: list_all2_conv_all_nth set_conv_nth) + by(auto simp add: list_all2_conv_all_nth set_conv_nth) lemma zip_assoc: "zip xs (zip ys zs) = map (\((x, y), z). (x, y, z)) (zip (zip xs ys) zs)" -by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all + by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_commute: "zip xs ys = map (\(x, y). (y, x)) (zip ys xs)" -by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all + by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_left_commute: "zip xs (zip ys zs) = map (\(y, (x, z)). (x, y, z)) (zip ys (zip xs zs))" -by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all + by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all lemma zip_replicate2: "zip xs (replicate n y) = map (\x. (x, y)) (take n xs)" -by(subst zip_commute)(simp add: zip_replicate1) + by(subst zip_commute)(simp add: zip_replicate1) subsubsection \@{const List.product} and @{const product_lists}\ lemma product_concat_map: "List.product xs ys = concat (map (\x. map (\y. (x,y)) ys) xs)" -by(induction xs) (simp)+ + by(induction xs) (simp)+ lemma set_product[simp]: "set (List.product xs ys) = set xs \ set ys" -by (induct xs) auto + by (induct xs) auto lemma length_product [simp]: "length (List.product xs ys) = length xs * length ys" -by (induct xs) simp_all + by (induct xs) simp_all lemma product_nth: assumes "n < length xs * length ys" shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))" -using assms proof (induct xs arbitrary: n) + using assms proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs n) @@ -2832,7 +2836,7 @@ lemma in_set_product_lists_length: "xs \ set (product_lists xss) \ length xs = length xss" -by (induct xss arbitrary: xs) auto + by (induct xss arbitrary: xs) auto lemma product_lists_set: "set (product_lists xss) = {xs. list_all2 (\x ys. x \ set ys) xs xss}" (is "?L = Collect ?R") @@ -2851,25 +2855,25 @@ lemma fold_simps [code]: \ \eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\ "fold f [] s = s" "fold f (x # xs) s = fold f xs (f x s)" -by simp_all + by simp_all lemma fold_remove1_split: "\ \x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x; x \ set xs \ \ fold f xs = fold f (remove1 x xs) \ f x" -by (induct xs) (auto simp add: comp_assoc) + by (induct xs) (auto simp add: comp_assoc) lemma fold_cong [fundef_cong]: "a = b \ xs = ys \ (\x. x \ set xs \ f x = g x) \ fold f xs a = fold g ys b" -by (induct ys arbitrary: a b xs) simp_all + by (induct ys arbitrary: a b xs) simp_all lemma fold_id: "(\x. x \ set xs \ f x = id) \ fold f xs = id" -by (induct xs) simp_all + by (induct xs) simp_all lemma fold_commute: "(\x. x \ set xs \ h \ g x = f x \ h) \ h \ fold g xs = fold f xs \ h" -by (induct xs) (simp_all add: fun_eq_iff) + by (induct xs) (simp_all add: fun_eq_iff) lemma fold_commute_apply: assumes "\x. x \ set xs \ h \ g x = f x \ h" @@ -2882,41 +2886,41 @@ lemma fold_invariant: "\ \x. x \ set xs \ Q x; P s; \x s. Q x \ P s \ P (f x s) \ \ P (fold f xs s)" -by (induct xs arbitrary: s) simp_all + by (induct xs arbitrary: s) simp_all lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \ fold f xs" -by (induct xs) simp_all + by (induct xs) simp_all lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \ f) xs" -by (induct xs) simp_all + by (induct xs) simp_all lemma fold_filter: "fold f (filter P xs) = fold (\x. if P x then f x else id) xs" -by (induct xs) simp_all + by (induct xs) simp_all lemma fold_rev: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ fold f (rev xs) = fold f xs" -by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) + by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) lemma fold_Cons_rev: "fold Cons xs = append (rev xs)" -by (induct xs) simp_all + by (induct xs) simp_all lemma rev_conv_fold [code]: "rev xs = fold Cons xs []" -by (simp add: fold_Cons_rev) + by (simp add: fold_Cons_rev) lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))" -by (induct xss) simp_all + by (induct xss) simp_all text \@{const Finite_Set.fold} and @{const fold}\ lemma (in comp_fun_commute) fold_set_fold_remdups: "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" -by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb) + by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb) lemma (in comp_fun_idem) fold_set_fold: "Finite_Set.fold f y (set xs) = fold f xs y" -by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm) + by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm) lemma union_set_fold [code]: "set xs \ A = fold Set.insert xs A" proof - @@ -2927,7 +2931,7 @@ lemma union_coset_filter [code]: "List.coset xs \ A = List.coset (List.filter (\x. x \ A) xs)" -by auto + by auto lemma minus_set_fold [code]: "A - set xs = fold Set.remove xs A" proof - @@ -2939,15 +2943,15 @@ lemma minus_coset_filter [code]: "A - List.coset xs = set (List.filter (\x. x \ A) xs)" -by auto + by auto lemma inter_set_filter [code]: "A \ set xs = set (List.filter (\x. x \ A) xs)" -by auto + by auto lemma inter_coset_fold [code]: "A \ List.coset xs = fold Set.remove xs A" -by (simp add: Diff_eq [symmetric] minus_set_fold) + by (simp add: Diff_eq [symmetric] minus_set_fold) lemma (in semilattice_set) set_eq_fold [code]: "F (set (x # xs)) = fold f xs x" @@ -2995,82 +2999,82 @@ text \Correspondence\ lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)" -by (induct xs) simp_all + by (induct xs) simp_all lemma foldl_conv_fold: "foldl f s xs = fold (\x s. f s x) xs s" -by (induct xs arbitrary: s) simp_all + by (induct xs arbitrary: s) simp_all lemma foldr_conv_foldl: \ \The ``Third Duality Theorem'' in Bird \& Wadler:\ "foldr f xs a = foldl (\x y. f y x) a (rev xs)" -by (simp add: foldr_conv_fold foldl_conv_fold) + by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldl_conv_foldr: "foldl f a xs = foldr (\x y. f y x) (rev xs) a" -by (simp add: foldr_conv_fold foldl_conv_fold) + by (simp add: foldr_conv_fold foldl_conv_fold) lemma foldr_fold: "(\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y) \ foldr f xs = fold f xs" -unfolding foldr_conv_fold by (rule fold_rev) + unfolding foldr_conv_fold by (rule fold_rev) lemma foldr_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f x a = g x a) \ foldr f l a = foldr g k b" -by (auto simp add: foldr_conv_fold intro!: fold_cong) + by (auto simp add: foldr_conv_fold intro!: fold_cong) lemma foldl_cong [fundef_cong]: "a = b \ l = k \ (\a x. x \ set l \ f a x = g a x) \ foldl f a l = foldl g b k" -by (auto simp add: foldl_conv_fold intro!: fold_cong) + by (auto simp add: foldl_conv_fold intro!: fold_cong) lemma foldr_append [simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)" -by (simp add: foldr_conv_fold) + by (simp add: foldr_conv_fold) lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys" -by (simp add: foldl_conv_fold) + by (simp add: foldl_conv_fold) lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \ f) xs a" -by (simp add: foldr_conv_fold fold_map rev_map) + by (simp add: foldr_conv_fold fold_map rev_map) lemma foldr_filter: "foldr f (filter P xs) = foldr (\x. if P x then f x else id) xs" -by (simp add: foldr_conv_fold rev_filter fold_filter) + by (simp add: foldr_conv_fold rev_filter fold_filter) lemma foldl_map [code_unfold]: "foldl g a (map f xs) = foldl (\a x. g a (f x)) a xs" -by (simp add: foldl_conv_fold fold_map comp_def) + by (simp add: foldl_conv_fold fold_map comp_def) lemma concat_conv_foldr [code]: "concat xss = foldr append xss []" -by (simp add: fold_append_concat_rev foldr_conv_fold) + by (simp add: fold_append_concat_rev foldr_conv_fold) subsubsection \@{const upt}\ lemma upt_rec[code]: "[i.. \simp does not terminate!\ -by (induct j) auto + \ \simp does not terminate!\ + by (induct j) auto lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n -lemma upt_conv_Nil [simp]: "j <= i ==> [i.. j <= i)" -by(induct j)simp_all +lemma upt_conv_Nil [simp]: "j \ i ==> [i.. j \ i)" + by(induct j)simp_all lemma upt_eq_Cons_conv: - "([i.. i = x \ [i+1.. [i..<(Suc j)] = [i.. \Only needed if \upt_Suc\ is deleted from the simpset.\ -by simp + "([i.. i = x \ [i+1.. j ==> [i..<(Suc j)] = [i.. \Only needed if \upt_Suc\ is deleted from the simpset.\ + by simp lemma upt_conv_Cons: "i < j ==> [i.. \no precondition\ "m # n # ns = [m.. n # ns = [Suc m.. [i.. \LOOPS as a simprule, since \j <= j\.\ -by (induct k) auto + \ \LOOPS as a simprule, since \j \ j\.\ + by (induct k) auto lemma length_upt [simp]: "length [i.. [i.. hd[i.. last[i.. take m [i.. n ==> take m [i..i. i + n) [0.. (map f [m..n. n - Suc 0) [Suc m..i. f (Suc i)) [0 ..< n]" -by (induct n arbitrary: f) auto - + by (induct n arbitrary: f) auto lemma nth_take_lemma: "k \ length xs \ k \ length ys \ (\i. i < k \ xs!i = ys!i) \ take k xs = take k ys" -apply (atomize, induct k arbitrary: xs ys) -apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify) -txt \Both lists must be non-empty\ -apply (case_tac xs, simp) -apply (case_tac ys, clarify) - apply (simp (no_asm_use)) -apply clarify -txt \prenexing's needed, not miniscoping\ -apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps) -apply blast -done +proof (induct k arbitrary: xs ys) + case (Suc k) + then show ?case + apply (simp add: less_Suc_eq_0_disj) + by (simp add: Suc.prems(3) take_Suc_conv_app_nth) +qed simp lemma nth_equalityI: "[| length xs = length ys; \i < length xs. xs!i = ys!i |] ==> xs = ys" -by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all + by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all lemma map_nth: "map (\i. xs ! i) [0.. (\x y. \P x y; Q y x\ \ x = y); list_all2 P xs ys; list_all2 Q ys xs \ \ xs = ys" -apply (simp add: list_all2_conv_all_nth) -apply (rule nth_equalityI, blast, simp) -done + by (simp add: list_all2_conv_all_nth nth_equalityI) lemma take_equalityI: "(\i. take i xs = take i ys) ==> xs = ys" \ \The famous take-lemma.\ -apply (drule_tac x = "max (length xs) (length ys)" in spec) -apply (simp add: le_max_iff_disj) -done - + by (metis length_take min.commute order_refl take_all) lemma take_Cons': "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)" @@ -3241,7 +3234,7 @@ qed lemma nth_upto: "i + int k \ j \ [i..j] ! k = i + int k" -apply(induction i j arbitrary: k rule: upto.induct) + apply(induction i j arbitrary: k rule: upto.induct) apply(subst upto_rec1) apply(auto simp add: nth_Cons') done @@ -3307,17 +3300,16 @@ lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])" by (induct x, auto) -lemma length_remdups_leq[iff]: "length(remdups xs) <= length xs" +lemma length_remdups_leq[iff]: "length(remdups xs) \ length xs" by (induct xs) auto lemma length_remdups_eq[iff]: "(length (remdups xs) = length xs) = (remdups xs = xs)" -apply(induct xs) - apply auto -apply(subgoal_tac "length (remdups xs) <= length xs") - apply arith -apply(rule length_remdups_leq) -done +proof (induct xs) + case (Cons a xs) + then show ?case + by simp (metis Suc_n_not_le_n impossible_Cons length_remdups_leq) +qed auto lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)" by (induct xs) auto @@ -3343,31 +3335,38 @@ done lemma distinct_take[simp]: "distinct xs \ distinct (take i xs)" -apply(induct xs arbitrary: i) - apply simp -apply (case_tac i) - apply simp_all -apply(blast dest:in_set_takeD) -done +proof (induct xs arbitrary: i) + case (Cons a xs) + then show ?case + by (metis Cons.prems append_take_drop_id distinct_append) +qed auto lemma distinct_drop[simp]: "distinct xs \ distinct (drop i xs)" -apply(induct xs arbitrary: i) - apply simp -apply (case_tac i) - apply simp_all -done +proof (induct xs arbitrary: i) + case (Cons a xs) + then show ?case + by (metis Cons.prems append_take_drop_id distinct_append) +qed auto lemma distinct_list_update: -assumes d: "distinct xs" and a: "a \ set xs - {xs!i}" -shows "distinct (xs[i:=a])" + assumes d: "distinct xs" and a: "a \ set xs - {xs!i}" + shows "distinct (xs[i:=a])" proof (cases "i < length xs") case True - with a have "a \ set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}" - apply (drule_tac id_take_nth_drop) by simp - with d True show ?thesis - apply (simp add: upd_conv_take_nth_drop) - apply (drule subst [OF id_take_nth_drop]) apply assumption - apply simp apply (cases "a = xs!i") apply simp by blast + with a have anot: "a \ set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}" + by simp (metis in_set_dropD in_set_takeD) + show ?thesis + proof (cases "a = xs!i") + case True + with d show ?thesis + by auto + next + case False + then show ?thesis + using d anot \i < length xs\ + apply (simp add: upd_conv_take_nth_drop) + by (metis disjoint_insert(1) distinct_append id_take_nth_drop set_simps(2)) + qed next case False with d show ?thesis by auto qed @@ -3382,22 +3381,14 @@ text \It is best to avoid this indexed version of distinct, but sometimes it is useful.\ -lemma distinct_conv_nth: -"distinct xs = (\i < size xs. \j < size xs. i \ j \ xs!i \ xs!j)" -apply (induct xs, simp, simp) -apply (rule iffI, clarsimp) - apply (case_tac i) -apply (case_tac j, simp) -apply (simp add: set_conv_nth) - apply (case_tac j) -apply (clarsimp simp add: set_conv_nth, simp) -apply (rule conjI) - apply (clarsimp simp add: set_conv_nth) - apply (erule_tac x = 0 in allE, simp) - apply (erule_tac x = "Suc i" in allE, simp, clarsimp) -apply (erule_tac x = "Suc i" in allE, simp) -apply (erule_tac x = "Suc j" in allE, simp) -done +lemma distinct_conv_nth: "distinct xs = (\i < size xs. \j < size xs. i \ j \ xs!i \ xs!j)" +proof (induct xs) + case (Cons x xs) + show ?case + apply (auto simp add: Cons nth_Cons split: nat.split_asm) + apply (metis Suc_less_eq2 in_set_conv_nth less_not_refl zero_less_Suc)+ + done +qed auto lemma nth_eq_iff_index_eq: "\ distinct xs; i < length xs; j < length xs \ \ (xs!i = xs!j) = (i = j)" @@ -3422,7 +3413,7 @@ lemma distinct_swap[simp]: "\ i < size xs; j < size xs \ \ distinct(xs[i := xs!j, j := xs!i]) = distinct xs" -apply (simp add: distinct_conv_nth nth_list_update) + apply (simp add: distinct_conv_nth nth_list_update) apply safe apply metis+ done @@ -3436,8 +3427,6 @@ lemma card_distinct: "card (set xs) = size xs ==> distinct xs" proof (induct xs) - case Nil thus ?case by simp -next case (Cons x xs) show ?case proof (cases "x \ set xs") @@ -3450,17 +3439,20 @@ ultimately have False by simp thus ?thesis .. qed -qed +qed simp lemma distinct_length_filter: "distinct xs \ length (filter P xs) = card ({x. P x} Int set xs)" by (induct xs) (auto) lemma not_distinct_decomp: "\ distinct ws \ \xs ys zs y. ws = xs@[y]@ys@[y]@zs" -apply (induct n == "length ws" arbitrary:ws) apply simp -apply(case_tac ws) apply simp -apply (simp split:if_split_asm) -apply (metis Cons_eq_appendI eq_Nil_appendI split_list) -done +proof (induct n == "length ws" arbitrary:ws) + case (Suc n ws) + then show ?case + using length_Suc_conv [of ws n] + apply (auto simp: eq_commute) + apply (metis append_Nil in_set_conv_decomp_first) + by (metis append_Cons) +qed simp lemma not_distinct_conv_prefix: defines "dec as xs y ys \ y \ set xs \ distinct xs \ as = xs @ y # ys" @@ -3692,7 +3684,6 @@ then have "Suc 0 \ f ` {0 ..< length (x1 # x2 # xs)}" by auto then show False using f_img \2 \ length ys\ by auto qed - obtain ys' where "ys = x1 # x2 # ys'" using \2 \ length ys\ f_nth[of 0] f_nth[of 1] apply (cases ys) @@ -3717,10 +3708,7 @@ using Suc0_le_f_Suc f_mono by (auto simp: f'_def mono_iff_le_Suc le_diff_iff) next have "f' ` {0 ..< length (x2 # xs)} = (\x. f x - 1) ` {0 ..< length (x1 # x2 #xs)}" - apply safe - apply (rename_tac [!] n, case_tac [!] n) - apply (auto simp: f'_def \f 0 = 0\ \f (Suc 0) = Suc 0\ intro: rev_image_eqI) - done + by (auto simp: f'_def \f 0 = 0\ \f (Suc 0) = Suc 0\ image_def Bex_def less_Suc_eq_0_disj) also have "\ = (\x. x - 1) ` f ` {0 ..< length (x1 # x2 #xs)}" by (auto simp: image_comp) also have "\ = (\x. x - 1) ` {0 ..< length ys}" @@ -3887,10 +3875,12 @@ lemma sum_count_set: "set xs \ X \ finite X \ sum (count_list xs) X = length xs" -apply(induction xs arbitrary: X) - apply simp -apply (simp add: sum.If_cases) -by (metis Diff_eq sum.remove) +proof (induction xs arbitrary: X) + case (Cons x xs) + then show ?case + apply (auto simp: sum.If_cases sum.remove) + by (metis (no_types) Cons.IH Cons.prems(2) diff_eq sum.remove) +qed simp subsubsection \@{const List.extract}\ @@ -3933,23 +3923,13 @@ lemma in_set_remove1[simp]: "a \ b \ a \ set(remove1 b xs) = (a \ set xs)" -apply (induct xs) - apply auto -done + by (induct xs) auto lemma set_remove1_subset: "set(remove1 x xs) \ set xs" -apply(induct xs) - apply simp -apply simp -apply blast -done + by (induct xs) auto lemma set_remove1_eq [simp]: "distinct xs \ set(remove1 x xs) = set xs - {x}" -apply(induct xs) - apply simp -apply simp -apply blast -done + by (induct xs) auto lemma length_remove1: "length(remove1 x xs) = (if x \ set xs then length xs - 1 else length xs)" @@ -4069,9 +4049,7 @@ text\Courtesy of Matthias Daum:\ lemma append_replicate_commute: "replicate n x @ replicate k x = replicate k x @ replicate n x" -apply (simp add: replicate_add [symmetric]) -apply (simp add: add.commute) -done + by (metis add.commute replicate_add) text\Courtesy of Andreas Lochbihler:\ lemma filter_replicate: @@ -4092,23 +4070,24 @@ text\Courtesy of Matthias Daum (2 lemmas):\ lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x" -apply (case_tac "k \ i") - apply (simp add: min_def) -apply (drule not_le_imp_less) -apply (simp add: min_def) -apply (subgoal_tac "replicate k x = replicate i x @ replicate (k - i) x") - apply simp -apply (simp add: replicate_add [symmetric]) -done +proof (cases "k \ i") + case True + then show ?thesis + by (simp add: min_def) +next + case False + then have "replicate k x = replicate i x @ replicate (k - i) x" + by (simp add: replicate_add [symmetric]) + then show ?thesis + by (simp add: min_def) +qed lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x" -apply (induct k arbitrary: i) - apply simp -apply clarsimp -apply (case_tac i) - apply simp -apply clarsimp -done +proof (induct k arbitrary: i) + case (Suc k) + then show ?case + by (simp add: drop_Cons') +qed simp lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}" by (induct n) auto @@ -4150,11 +4129,11 @@ lemma replicate_eq_replicate[simp]: "(replicate m x = replicate n y) \ (m=n \ (m\0 \ x=y))" -apply(induct m arbitrary: n) - apply simp -apply(induct_tac n) -apply auto -done +proof (induct m arbitrary: n) + case (Suc m n) + then show ?case + by (induct n) auto +qed simp lemma replicate_length_filter: "replicate (length (filter (\y. x = y) xs)) x = filter (\y. x = y) xs" @@ -4241,10 +4220,7 @@ lemma enumerate_simps [simp, code]: "enumerate n [] = []" "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs" -apply (auto simp add: enumerate_eq_zip not_le) -apply (cases "n < n + length xs") - apply (auto simp add: upt_conv_Cons) -done + by (simp_all add: enumerate_eq_zip upt_rec) lemma length_enumerate [simp]: "length (enumerate n xs) = length xs" @@ -4290,12 +4266,7 @@ lemma enumerate_append_eq: "enumerate n (xs @ ys) = enumerate n xs @ enumerate (n + length xs) ys" -unfolding enumerate_eq_zip -apply auto - apply (subst zip_append [symmetric]) apply simp - apply (subst upt_add_eq_append [symmetric]) - apply (simp_all add: ac_simps) -done + by (simp add: enumerate_eq_zip add.assoc zip_append2) lemma enumerate_map_upt: "enumerate n (map f [n..k. (k, f k)) [n.. rotate1 xs = xs" +lemma rotate1_length01[simp]: "length xs \ 1 \ rotate1 xs = xs" by(cases xs) simp_all -lemma rotate_length01[simp]: "length xs <= 1 \ rotate n xs = xs" -apply(induct n) - apply simp -apply (simp add:rotate_def) -done +lemma rotate_length01[simp]: "length xs \ 1 \ rotate n xs = xs" + by (induct n) (simp_all add:rotate_def) lemma rotate1_hd_tl: "xs \ [] \ rotate1 xs = tl xs @ [hd xs]" by (cases xs) simp_all lemma rotate_drop_take: "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs" -apply(induct n) - apply simp -apply(simp add:rotate_def) -apply(cases "xs = []") - apply (simp) -apply(case_tac "n mod length xs = 0") - apply(simp add:mod_Suc) - apply(simp add: rotate1_hd_tl drop_Suc take_Suc) -apply(simp add:mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric] - take_hd_drop linorder_not_le) -done +proof (induct n) + case (Suc n) + show ?case + proof (cases "xs = []") + case False + then show ?thesis + proof (cases "n mod length xs = 0") + case True + then show ?thesis + apply (simp add: mod_Suc) + by (simp add: False Suc.hyps drop_Suc rotate1_hd_tl take_Suc) + next + case False + with \xs \ []\ Suc + show ?thesis + by (simp add: rotate_def mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric] + take_hd_drop linorder_not_le) + qed + qed simp +qed simp lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs" by(simp add:rotate_drop_take) @@ -4387,11 +4364,14 @@ by(simp add:rotate_drop_take rev_drop rev_take) qed force -lemma hd_rotate_conv_nth: "xs \ [] \ hd(rotate n xs) = xs!(n mod length xs)" -apply(simp add:rotate_drop_take hd_append hd_drop_conv_nth hd_conv_nth) -apply(subgoal_tac "length xs \ 0") - prefer 2 apply simp -using mod_less_divisor[of "length xs" n] by arith +lemma hd_rotate_conv_nth: + assumes "xs \ []" shows "hd(rotate n xs) = xs!(n mod length xs)" +proof - + have "n mod length xs < length xs" + using assms by simp + then show ?thesis + by (metis drop_eq_Nil hd_append2 hd_drop_conv_nth leD rotate_drop_take) +qed lemma rotate_append: "rotate (length l) (l @ q) = q @ l" by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap) @@ -4410,14 +4390,13 @@ by(simp add: nths_def length_filter_conv_card cong:conj_cong) lemma nths_shift_lemma_Suc: - "map fst (filter (%p. P(Suc(snd p))) (zip xs is)) = - map fst (filter (%p. P(snd p)) (zip xs (map Suc is)))" -apply(induct xs arbitrary: "is") - apply simp -apply (case_tac "is") - apply simp -apply simp -done + "map fst (filter (\p. P(Suc(snd p))) (zip xs is)) = + map fst (filter (\p. P(snd p)) (zip xs (map Suc is)))" +proof (induct xs arbitrary: "is") + case (Cons x xs "is") + show ?case + by (cases "is") (auto simp add: Cons.hyps) +qed simp lemma nths_shift_lemma: "map fst (filter (\p. snd p \ A) (zip xs [i.. A}" -apply (unfold nths_def) -apply (induct l' rule: rev_induct, simp) -apply (simp add: upt_add_eq_append[of 0] nths_shift_lemma) -apply (simp add: add.commute) -done + unfolding nths_def +proof (induct l' rule: rev_induct) + case (snoc x xs) + then show ?case + by (simp add: upt_add_eq_append[of 0] nths_shift_lemma add.commute) +qed auto lemma nths_Cons: "nths (x # l) A = (if 0 \ A then [x] else []) @ nths l {j. Suc j \ A}" -apply (induct l rule: rev_induct) - apply (simp add: nths_def) -apply (simp del: append_Cons add: append_Cons[symmetric] nths_append) -done +proof (induct l rule: rev_induct) + case (snoc x xs) + then show ?case + by (simp flip: append_Cons add: nths_append) +qed (auto simp: nths_def) lemma nths_map: "nths (map f xs) I = map f (nths xs I)" by(induction xs arbitrary: I) (simp_all add: nths_Cons) lemma set_nths: "set(nths xs I) = {xs!i|i. i i \ I}" -apply(induct xs arbitrary: I) -apply(auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc) -done + by (induct xs arbitrary: I) (auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc) lemma set_nths_subset: "set(nths xs I) \ set xs" by(auto simp add:set_nths) @@ -4794,8 +4773,7 @@ show ?thesis unfolding transpose.simps \i = Suc j\ nth_Cons_Suc "3.hyps"[OF j_less] apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric]) - apply (rule list.exhaust) - by auto + by (simp add: nth_tl) qed qed simp_all @@ -4919,11 +4897,7 @@ qed lemma infinite_UNIV_listI: "\ finite(UNIV::'a list set)" -apply (rule notI) -apply (drule finite_maxlen) -apply clarsimp -apply (erule_tac x = "replicate n undefined" in allE) -by simp + by (metis UNIV_I finite_maxlen length_replicate less_irrefl) subsection \Sorting\ @@ -4938,10 +4912,11 @@ by(simp) lemma sorted_wrt2: "transp P \ sorted_wrt P (x # y # zs) = (P x y \ sorted_wrt P (y # zs))" -apply(induction zs arbitrary: x y) -apply(auto dest: transpD) -apply (meson transpD) -done +proof (induction zs arbitrary: x y) + case (Cons z zs) + then show ?case + by simp (meson transpD)+ +qed auto lemmas sorted_wrt2_simps = sorted_wrt1 sorted_wrt2 @@ -4971,9 +4946,7 @@ lemma sorted_wrt_iff_nth_less: "sorted_wrt P xs = (\i j. i < j \ j < length xs \ P (xs ! i) (xs ! j))" -apply(induction xs) -apply(auto simp add: in_set_conv_nth Ball_def nth_Cons split: nat.split) -done + by (induction xs) (auto simp add: in_set_conv_nth Ball_def nth_Cons split: nat.split) lemma sorted_wrt_nth_less: "\ sorted_wrt P xs; i < j; j < length xs \ \ P (xs ! i) (xs ! j)" @@ -4983,10 +4956,11 @@ by(induction n) (auto simp: sorted_wrt_append) lemma sorted_wrt_upto[simp]: "sorted_wrt (<) [i..j]" -apply(induction i j rule: upto.induct) -apply(subst upto.simps) -apply(simp) -done +proof(induct i j rule:upto.induct) + case (1 i j) + from this show ?case + unfolding upto.simps[of i j] by auto +qed text \Each element is greater or equal to its index:\ @@ -5315,12 +5289,13 @@ qed lemma finite_sorted_distinct_unique: -shows "finite A \ \!xs. set xs = A \ sorted xs \ distinct xs" -apply(drule finite_distinct_list) -apply clarify -apply(rule_tac a="sort xs" in ex1I) -apply (auto simp: sorted_distinct_set_unique) -done + assumes "finite A" shows "\!xs. set xs = A \ sorted xs \ distinct xs" +proof - + obtain xs where "distinct xs" "A = set xs" + using finite_distinct_list [OF assms] by metis + then show ?thesis + by (rule_tac a="sort xs" in ex1I) (auto simp: sorted_distinct_set_unique) +qed lemma insort_insert_key_triv: "f x \ f ` set xs \ insort_insert_key f x xs = xs" @@ -5719,7 +5694,7 @@ lemmas in_listsI [intro!] = in_listspI [to_set] -lemma lists_eq_set: "lists A = {xs. set xs <= A}" +lemma lists_eq_set: "lists A = {xs. set xs \ A}" by auto lemma lists_empty [simp]: "lists {} = {[]}" @@ -5743,12 +5718,12 @@ | insert: "ListMem x xs \ ListMem x (y # xs)" lemma ListMem_iff: "(ListMem x xs) = (x \ set xs)" -apply (rule iffI) - apply (induct set: ListMem) - apply auto -apply (induct xs) - apply (auto intro: ListMem.intros) -done +proof + show "ListMem x xs \ x \ set xs" + by (induct set: ListMem) auto + show "x \ set xs \ ListMem x xs" + by (induct xs) (auto intro: ListMem.intros) +qed subsubsection \Lists as Cartesian products\ @@ -5791,20 +5766,23 @@ "lenlex r = inv_image (less_than <*lex*> lex r) (\xs. (length xs, xs))" \ \Compares lists by their length and then lexicographically\ -lemma wf_lexn: "wf r ==> wf (lexn r n)" -apply (induct n, simp, simp) -apply(rule wf_subset) - prefer 2 apply (rule Int_lower1) -apply(rule wf_map_prod_image) - prefer 2 apply (rule inj_onI, auto) -done +lemma wf_lexn: assumes "wf r" shows "wf (lexn r n)" +proof (induct n) + case (Suc n) + have inj: "inj (\(x, xs). x # xs)" + using assms by (auto simp: inj_on_def) + have wf: "wf (map_prod (\(x, xs). x # xs) (\(x, xs). x # xs) ` (r <*lex*> lexn r n))" + by (simp add: Suc.hyps assms wf_lex_prod wf_map_prod_image [OF _ inj]) + then show ?case + by (rule wf_subset) auto +qed auto lemma lexn_length: "(xs, ys) \ lexn r n \ length xs = n \ length ys = n" by (induct n arbitrary: xs ys) auto lemma wf_lex [intro!]: "wf r ==> wf (lex r)" - apply (unfold lex_def) + unfolding lex_def apply (rule wf_UN) apply (simp add: wf_lexn) apply (metis DomainE Int_emptyI RangeE lexn_length) @@ -5907,14 +5885,12 @@ by (simp add:lex_conv) lemma Cons_in_lex [simp]: - "((x # xs, y # ys) \ lex r) = + "((x # xs, y # ys) \ lex r) = ((x, y) \ r \ length xs = length ys \ x = y \ (xs, ys) \ lex r)" -apply (simp add: lex_conv) -apply (rule iffI) - prefer 2 apply (blast intro: Cons_eq_appendI, clarify) -apply (case_tac xys, simp, simp) -apply blast - done + apply (simp add: lex_conv) + apply (rule iffI) + prefer 2 apply (blast intro: Cons_eq_appendI, clarify) + by (metis hd_append append_Nil list.sel(1) list.sel(3) tl_append2) lemma lex_append_rightI: "(xs, ys) \ lex r \ length vs = length us \ (xs @ us, ys @ vs) \ lex r" @@ -5962,12 +5938,13 @@ by (unfold lexord_def, induct_tac x, auto) lemma lexord_cons_cons[simp]: - "((a # x, b # y) \ lexord r) = ((a,b)\ r \ (a = b \ (x,y)\ lexord r))" - apply (unfold lexord_def, safe, simp_all) - apply (case_tac u, simp, simp) - apply (case_tac u, simp, clarsimp, blast, blast, clarsimp) - apply (erule_tac x="b # u" in allE) - by force + "((a # x, b # y) \ lexord r) = ((a,b)\ r \ (a = b \ (x,y)\ lexord r))" + unfolding lexord_def + apply (safe, simp_all) + apply (metis hd_append list.sel(1)) + apply (metis hd_append list.sel(1) list.sel(3) tl_append2) + apply blast + by (meson Cons_eq_appendI) lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons @@ -5991,24 +5968,17 @@ (\i. i < min(length x)(length y) \ take i x = take i y \ (x!i,y!i) \ r))" apply (unfold lexord_def Let_def, clarsimp) apply (rule_tac f = "(% a b. a \ b)" in arg_cong2) + apply (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le) apply auto - apply (rule_tac x="hd (drop (length x) y)" in exI) - apply (rule_tac x="tl (drop (length x) y)" in exI) - apply (erule subst, simp add: min_def) apply (rule_tac x ="length u" in exI, simp) - apply (rule_tac x ="take i x" in exI) - apply (rule_tac x ="x ! i" in exI) - apply (rule_tac x ="y ! i" in exI, safe) - apply (rule_tac x="drop (Suc i) x" in exI) - apply (drule sym, simp add: Cons_nth_drop_Suc) - apply (rule_tac x="drop (Suc i) y" in exI) - by (simp add: Cons_nth_drop_Suc) + by (metis id_take_nth_drop) \ \lexord is extension of partial ordering List.lex\ lemma lexord_lex: "(x,y) \ lex r = ((x,y) \ lexord r \ length x = length y)" - apply (rule_tac x = y in spec) - apply (induct_tac x, clarsimp) - by (clarify, case_tac x, simp, force) +proof (induction x arbitrary: y) + case (Cons a x y) then show ?case + by (cases y) (force+) +qed auto lemma lexord_irreflexive: "\x. (x,x) \ r \ (xs,xs) \ lexord r" by (induct xs) auto @@ -6051,12 +6021,15 @@ lemma lexord_transI: "trans r \ trans (lexord r)" by (rule transI, drule lexord_trans, blast) -lemma lexord_linear: "(\a b. (a,b)\ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" - apply (rule_tac x = y in spec) - apply (induct_tac x, rule allI) - apply (case_tac x, simp, simp) - apply (rule allI, case_tac x, simp, simp) - by blast +lemma lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" +proof (induction x arbitrary: y) + case Nil + then show ?case + by (metis lexord_Nil_left list.exhaust) +next + case (Cons a x y) then show ?case + by (cases y) (force+) +qed lemma lexord_irrefl: "irrefl R \ irrefl (lexord R)" @@ -6222,27 +6195,17 @@ lemma lexordp_eq_trans: assumes "lexordp_eq xs ys" and "lexordp_eq ys zs" shows "lexordp_eq xs zs" -using assms -apply(induct arbitrary: zs) -apply(case_tac [2-3] zs) -apply auto -done + using assms + by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_trans: assumes "lexordp xs ys" "lexordp ys zs" shows "lexordp xs zs" -using assms -apply(induct arbitrary: zs) -apply(case_tac [2-3] zs) -apply auto -done + using assms + by (induct arbitrary: zs) (case_tac zs; auto)+ lemma lexordp_linear: "lexordp xs ys \ xs = ys \ lexordp ys xs" -proof(induct xs arbitrary: ys) - case Nil thus ?case by(cases ys) simp_all -next - case Cons thus ?case by(cases ys) auto -qed + by(induct xs arbitrary: ys; case_tac ys; fastforce) lemma lexordp_conv_lexordp_eq: "lexordp xs ys \ lexordp_eq xs ys \ \ lexordp_eq ys xs" (is "?lhs \ ?rhs") @@ -6260,13 +6223,11 @@ by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym) lemma lexordp_eq_linear: "lexordp_eq xs ys \ lexordp_eq ys xs" -apply(induct xs arbitrary: ys) -apply(case_tac [!] ys) -apply auto -done + by (induct xs arbitrary: ys) (case_tac ys; auto)+ lemma lexordp_linorder: "class.linorder lexordp_eq lexordp" -by unfold_locales(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl lexordp_eq_antisym intro: lexordp_eq_trans del: disjCI intro: lexordp_eq_linear) + by unfold_locales + (auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl lexordp_eq_antisym intro: lexordp_eq_trans del: disjCI intro: lexordp_eq_linear) end @@ -6294,7 +6255,7 @@ lemma measures_less: "f x < f y ==> (x, y) \ measures (f#fs)" by simp -lemma measures_lesseq: "f x <= f y ==> (x, y) \ measures fs ==> (x, y) \ measures (f#fs)" +lemma measures_lesseq: "f x \ f y ==> (x, y) \ measures fs ==> (x, y) \ measures (f#fs)" by auto @@ -6404,17 +6365,20 @@ apply (induct arbitrary: xs set: Wellfounded.acc) apply (erule thin_rl) apply (erule acc_induct) -apply (rule accI) + apply (rule accI) apply (blast) done lemma lists_accD: "xs \ lists (Wellfounded.acc r) \ xs \ Wellfounded.acc (listrel1 r)" -apply (induct set: lists) - apply (rule accI) - apply simp -apply (rule accI) -apply (fast dest: acc_downward) -done +proof (induct set: lists) + case Nil + then show ?case + by (meson acc.intros not_listrel1_Nil) +next + case (Cons a l) + then show ?case + by blast +qed lemma lists_accI: "xs \ Wellfounded.acc (listrel1 r) \ xs \ lists (Wellfounded.acc r)" apply (induct set: Wellfounded.acc) @@ -6461,10 +6425,7 @@ lemma listrel_mono: "r \ s \ listrel r \ listrel s" -apply clarify -apply (erule listrel.induct) -apply (blast intro: listrel.intros)+ -done + by (meson listrel_iff_nth subrelI subset_eq) lemma listrel_subset: "r \ A \ A \ listrel r \ lists A \ lists A" apply clarify @@ -6479,10 +6440,7 @@ done lemma listrel_sym: "sym r \ sym (listrel r)" -apply (auto simp add: sym_def) -apply (erule listrel.induct) -apply (blast intro: listrel.intros)+ -done + by (simp add: listrel_iff_nth sym_def) lemma listrel_trans: "trans r \ trans (listrel r)" apply (simp add: trans_def) @@ -7138,7 +7096,7 @@ lemma subset_code [code]: "set xs \ B \ (\x\set xs. x \ B)" "A \ List.coset ys \ (\y\set ys. y \ A)" - "List.coset [] \ set [] \ False" + "List.coset [] \ set [] \ False" by auto text \A frequent case -- avoid intermediate sets\ @@ -7367,10 +7325,7 @@ "(rel_set A ===> rel_set (list_all2 A) ===> rel_set (list_all2 A)) set_Cons set_Cons" unfolding rel_fun_def rel_set_def set_Cons_def - apply safe - apply (simp add: list_all2_Cons1, fast) - apply (simp add: list_all2_Cons2, fast) - done + by (fastforce simp add: list_all2_Cons1 list_all2_Cons2) lemma listset_transfer [transfer_rule]: "(list_all2 (rel_set A) ===> rel_set (list_all2 A)) listset listset" diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Number_Theory/Cong.thy Sat Aug 11 16:02:55 2018 +0200 @@ -395,11 +395,8 @@ lemma cong_dvd_modulus_nat: "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" for x y :: nat - apply (auto simp add: cong_iff_lin_nat dvd_def) - apply (rule_tac x= "k1 * k" in exI) - apply (rule_tac x= "k2 * k" in exI) - apply (simp add: field_simps) - done + unfolding cong_iff_lin_nat dvd_def + by (metis mult.commute mult.left_commute) lemma cong_to_1_nat: fixes a :: nat @@ -433,41 +430,36 @@ for x y :: nat by (auto simp add: cong_altdef_nat le_imp_diff_is_add elim!: dvdE) + lemma cong_solve_nat: fixes a :: nat - assumes "a \ 0" shows "\x. [a * x = gcd a n] (mod n)" -proof (cases "n = 0") +proof (cases "a = 0 \ n = 0") case True - then show ?thesis by force + then show ?thesis + by (force simp add: cong_0_iff cong_sym) next case False then show ?thesis - using bezout_nat [of a n, OF \a \ 0\] + using bezout_nat [of a n] by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) qed -lemma cong_solve_int: "a \ 0 \ \x. [a * x = gcd a n] (mod n)" - for a :: int - apply (cases "n = 0") - apply (cases "a \ 0") - apply auto - apply (rule_tac x = "-1" in exI) - apply auto - apply (insert bezout_int [of a n], auto) - apply (metis cong_iff_lin mult.commute) - done +lemma cong_solve_int: + fixes a :: int + shows "\x. [a * x = gcd a n] (mod n)" + by (metis bezout_int cong_iff_lin mult.commute) lemma cong_solve_dvd_nat: fixes a :: nat - assumes a: "a \ 0" and b: "gcd a n dvd d" + assumes "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - - from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)" + from cong_solve_nat [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast - also from b have "(d div gcd a n) * gcd a n = d" + also from assms have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" by auto @@ -476,10 +468,11 @@ qed lemma cong_solve_dvd_int: - assumes a: "(a::int) \ 0" and b: "gcd a n dvd d" + fixes a::int + assumes b: "gcd a n dvd d" shows "\x. [a * x = d] (mod n)" proof - - from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)" + from cong_solve_int [of a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" using cong_scalar_left by blast @@ -493,12 +486,11 @@ lemma cong_solve_coprime_nat: "\x. [a * x = Suc 0] (mod n)" if "coprime a n" - using that cong_solve_nat [of a n] by (cases "a = 0") simp_all + using that cong_solve_nat [of a n] by auto lemma cong_solve_coprime_int: "\x. [a * x = 1] (mod n)" if "coprime a n" for a n x :: int - using that cong_solve_int [of a n] by (cases "a = 0") - (auto simp add: zabs_def split: if_splits) + using that cong_solve_int [of a n] by (auto simp add: zabs_def split: if_splits) lemma coprime_iff_invertible_nat: "coprime a m \ (\x. [a * x = Suc 0] (mod m))" (is "?P \ ?Q") @@ -529,27 +521,29 @@ qed lemma coprime_iff_invertible'_nat: - "m > 0 \ coprime a m \ (\x. 0 \ x \ x < m \ [a * x = Suc 0] (mod m))" - apply (subst coprime_iff_invertible_nat) - apply auto - apply (auto simp add: cong_def) - apply (metis mod_less_divisor mod_mult_right_eq) - done + assumes "m > 0" + shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = Suc 0] (mod m))" +proof - + have "\b. \0 < m; [a * b = Suc 0] (mod m)\ \ \b' 0 \ coprime a m \ (\x. 0 \ x \ x < m \ [a * x = 1] (mod m))" - for m :: int - apply (subst coprime_iff_invertible_int) - apply (auto simp add: cong_def) - apply (metis mod_mult_right_eq pos_mod_conj) - done + fixes m :: int + assumes "m > 0" + shows "coprime a m \ (\x. 0 \ x \ x < m \ [a * x = 1] (mod m))" +proof - + have "\b. \0 < m; [a * b = 1] (mod m)\ \ \b' [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: nat - apply (cases "y \ x") - apply (simp add: cong_altdef_nat) - apply (meson cong_altdef_nat cong_sym lcm_least_iff nat_le_linear) - done + by (meson cong_altdef_nat' lcm_least) lemma cong_cong_lcm_int: "[x = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: int @@ -636,10 +630,7 @@ lemma cong_modulus_mult_nat: "[x = y] (mod m * n) \ [x = y] (mod m)" for x y :: nat - apply (cases "y \ x") - apply (auto simp add: cong_altdef_nat elim: dvd_mult_left) - apply (metis cong_def mod_mult_cong_right) - done + by (metis cong_def mod_mult_cong_right) lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \ x < m \ y < m \ x = y" for x y :: nat @@ -651,50 +642,28 @@ and nz: "m1 \ 0" "m2 \ 0" shows "\!x. x < m1 * m2 \ [x = u1] (mod m1) \ [x = u2] (mod m2)" proof - - from binary_chinese_remainder_nat [OF a] obtain y - where "[y = u1] (mod m1)" and "[y = u2] (mod m2)" - by blast + obtain y where y1: "[y = u1] (mod m1)" and y2: "[y = u2] (mod m2)" + using binary_chinese_remainder_nat [OF a] by blast let ?x = "y mod (m1 * m2)" from nz have less: "?x < m1 * m2" by auto have 1: "[?x = u1] (mod m1)" - apply (rule cong_trans) - prefer 2 - apply (rule \[y = u1] (mod m1)\) - apply (rule cong_modulus_mult_nat [of _ _ _ m2]) - apply simp - done + using y1 mod_mult_cong_right by blast have 2: "[?x = u2] (mod m2)" - apply (rule cong_trans) - prefer 2 - apply (rule \[y = u2] (mod m2)\) - apply (subst mult.commute) - apply (rule cong_modulus_mult_nat [of _ _ _ m1]) - apply simp - done - have "\z. z < m1 * m2 \ [z = u1] (mod m1) \ [z = u2] (mod m2) \ z = ?x" - proof clarify - fix z - assume "z < m1 * m2" - assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)" + using y2 mod_mult_cong_left by blast + have "z = ?x" if "z < m1 * m2" "[z = u1] (mod m1)" "[z = u2] (mod m2)" for z + proof - have "[?x = z] (mod m1)" - apply (rule cong_trans) - apply (rule \[?x = u1] (mod m1)\) - apply (rule cong_sym) - apply (rule \[z = u1] (mod m1)\) - done + by (metis "1" cong_def that(2)) moreover have "[?x = z] (mod m2)" - apply (rule cong_trans) - apply (rule \[?x = u2] (mod m2)\) - apply (rule cong_sym) - apply (rule \[z = u2] (mod m2)\) - done + by (metis "2" cong_def that(3)) ultimately have "[?x = z] (mod m1 * m2)" using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) with \z < m1 * m2\ \?x < m1 * m2\ show "z = ?x" by (auto simp add: cong_def) qed - with less 1 2 show ?thesis by auto + with less 1 2 show ?thesis + by blast qed lemma chinese_remainder_nat: @@ -720,7 +689,7 @@ ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod prod m (A - {i}))" by blast qed - then obtain b where b: "\i \ A. [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j))" + then obtain b where b: "\i. i \ A \ [b i = 1] (mod m i) \ [b i = 0] (mod (\j \ A - {i}. m j))" by blast let ?x = "\i\A. (u i) * (b i)" show ?thesis @@ -735,29 +704,32 @@ by auto also have "[u i * b i + (\j \ A - {i}. u j * b j) = u i * 1 + (\j \ A - {i}. u j * 0)] (mod m i)" - apply (rule cong_add) - apply (rule cong_scalar_left) - using b a apply blast - apply (rule cong_sum) - apply (rule cong_scalar_left) - using b apply (auto simp add: mod_eq_0_iff_dvd cong_def) - apply (rule dvd_trans [of _ "prod m (A - {x})" "b x" for x]) - using a fin apply auto - done + proof (intro cong_add cong_scalar_left cong_sum) + show "[b i = 1] (mod m i)" + using a b by blast + show "[b x = 0] (mod m i)" if "x \ A - {i}" for x + proof - + have "x \ A" "x \ i" + using that by auto + then show ?thesis + using a b [OF \x \ A\] cong_dvd_modulus_nat fin by blast + qed + qed finally show ?thesis by simp qed qed qed -lemma coprime_cong_prod_nat: - "[x = y] (mod (\i\A. m i))" - if "\i\A. (\j\A. i \ j \ coprime (m i) (m j))" - and "\i\A. [x = y] (mod m i)" for x y :: nat - using that apply (induct A rule: infinite_finite_induct) - apply auto - apply (metis coprime_cong_mult_nat prod_coprime_right) - done +lemma coprime_cong_prod_nat: "[x = y] (mod (\i\A. m i))" + if "\i j. \i \ A; j \ A; i \ j\ \ coprime (m i) (m j)" + and "\i. i \ A \ [x = y] (mod m i)" for x y :: nat + using that +proof (induct A rule: infinite_finite_induct) + case (insert x A) + then show ?case + by simp (metis coprime_cong_mult_nat prod_coprime_right) +qed auto lemma chinese_remainder_unique_nat: fixes A :: "'a set" @@ -795,94 +767,4 @@ by blast qed - -subsection \Aliasses\ - -lemma cong_altdef_int: - "[a = b] (mod m) \ m dvd (a - b)" - for a b :: int - by (fact cong_iff_dvd_diff) - -lemma cong_iff_lin_int: "[a = b] (mod m) \ (\k. b = a + m * k)" - for a b :: int - by (fact cong_iff_lin) - -lemma cong_minus_int: "[a = b] (mod - m) \ [a = b] (mod m)" - for a b :: int - by (fact cong_modulus_minus_iff) - -lemma cong_add_lcancel_int: "[a + x = a + y] (mod n) \ [x = y] (mod n)" - for a x y :: int - by (fact cong_add_lcancel) - -lemma cong_add_rcancel_int: "[x + a = y + a] (mod n) \ [x = y] (mod n)" - for a x y :: int - by (fact cong_add_rcancel) - -lemma cong_add_lcancel_0_int: - "[a + x = a] (mod n) \ [x = 0] (mod n)" - for a x :: int - by (fact cong_add_lcancel_0) - -lemma cong_add_rcancel_0_int: - "[x + a = a] (mod n) \ [x = 0] (mod n)" - for a x :: int - by (fact cong_add_rcancel_0) - -lemma cong_dvd_modulus_int: "[x = y] (mod m) \ n dvd m \ [x = y] (mod n)" - for x y :: int - by (fact cong_dvd_modulus) - -lemma cong_abs_int: - "[x = y] (mod \m\) \ [x = y] (mod m)" - for x y :: int - by (fact cong_abs) - -lemma cong_square_int: - "prime p \ 0 < a \ [a * a = 1] (mod p) \ [a = 1] (mod p) \ [a = - 1] (mod p)" - for a :: int - by (fact cong_square) - -lemma cong_mult_rcancel_int: - "[a * k = b * k] (mod m) \ [a = b] (mod m)" - if "coprime k m" for a k m :: int - using that by (fact cong_mult_rcancel) - -lemma cong_mult_lcancel_int: - "[k * a = k * b] (mod m) = [a = b] (mod m)" - if "coprime k m" for a k m :: int - using that by (fact cong_mult_lcancel) - -lemma coprime_cong_mult_int: - "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" - for a b :: int - by (fact coprime_cong_mult) - -lemma cong_gcd_eq_nat: "[a = b] (mod m) \ gcd a m = gcd b m" - for a b :: nat - by (fact cong_gcd_eq) - -lemma cong_gcd_eq_int: "[a = b] (mod m) \ gcd a m = gcd b m" - for a b :: int - by (fact cong_gcd_eq) - -lemma cong_imp_coprime_nat: "[a = b] (mod m) \ coprime a m \ coprime b m" - for a b :: nat - by (fact cong_imp_coprime) - -lemma cong_imp_coprime_int: "[a = b] (mod m) \ coprime a m \ coprime b m" - for a b :: int - by (fact cong_imp_coprime) - -lemma cong_cong_prod_coprime_int: - "[x = y] (mod (\i\A. m i))" if - "(\i\A. [x = y] (mod m i))" - "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" - for x y :: int - using that by (fact cong_cong_prod_coprime) - -lemma cong_modulus_mult_int: "[x = y] (mod m * n) \ [x = y] (mod m)" - for x y :: int - by (fact cong_modulus_mult) - end diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Number_Theory/Euler_Criterion.thy --- a/src/HOL/Number_Theory/Euler_Criterion.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Number_Theory/Euler_Criterion.thy Sat Aug 11 16:02:55 2018 +0200 @@ -48,8 +48,8 @@ have "[nat \b\ ^ (p - 1) = 1] (mod p)" using p_prime proof (rule fermat_theorem) from b p_a_relprime show "\ p dvd nat \b\" - by (auto simp add: cong_altdef_int power2_eq_square) - (metis cong_altdef_int cong_dvd_iff dvd_mult2) + by (auto simp add: cong_iff_dvd_diff power2_eq_square) + (metis cong_iff_dvd_diff cong_dvd_iff dvd_mult2) qed then have "nat \b\ ^ (p - 1) mod p = 1 mod p" by (simp add: cong_def) @@ -90,13 +90,13 @@ cong_scalar_right [of "x * y'" 1 "int p" a] by (auto simp add: cong_def ac_simps) moreover have "y \ {0 .. int p - 1}" unfolding y_def using p_ge_2 by auto - hence "y \ S1" using calculation cong_altdef_int p_a_relprime S1_def by auto + hence "y \ S1" using calculation cong_iff_dvd_diff p_a_relprime S1_def cong_dvd_iff by fastforce ultimately have "P x y" unfolding P_def by blast moreover { fix y1 y2 assume "P x y1" "P x y2" moreover hence "[y1 = y2] (mod p)" unfolding P_def - using co_xp cong_mult_lcancel_int[of x p y1 y2] cong_sym cong_trans by blast + using co_xp cong_mult_lcancel[of x p y1 y2] cong_sym cong_trans by blast ultimately have "y1 = y2" unfolding P_def S1_def using cong_less_imp_eq_int by auto } ultimately show ?thesis by blast @@ -200,7 +200,8 @@ moreover have "(0::int) ^ ((p - 1) div 2) = 0" using zero_power [of "(p - 1) div 2"] assms(2) by simp ultimately have "[a ^ ((p - 1) div 2) = 0] (mod p)" - using True assms(1) cong_altdef_int prime_dvd_power_int_iff by auto + using True assms(1) prime_dvd_power_int_iff + by (simp add: cong_iff_dvd_diff) then show ?thesis unfolding Legendre_def using True cong_sym by auto next diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Number_Theory/Gauss.thy --- a/src/HOL/Number_Theory/Gauss.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Number_Theory/Gauss.thy Sat Aug 11 16:02:55 2018 +0200 @@ -18,7 +18,7 @@ lemma cong_prime_prod_zero_int: "[a * b = 0] (mod p) \ prime p \ [a = 0] (mod p) \ [b = 0] (mod p)" for a :: int - by (auto simp add: cong_altdef_int prime_dvd_mult_iff) + by (simp add: cong_0_iff prime_dvd_mult_iff) locale GAUSS = @@ -114,11 +114,11 @@ for x y proof - from p_a_relprime have "\ p dvd a" - by (simp add: cong_altdef_int) + by (simp add: cong_0_iff) with p_prime prime_imp_coprime [of _ "nat \a\"] have "coprime a (int p)" by (simp_all add: ac_simps) - with a cong_mult_rcancel_int [of a "int p" x y] have "[x = y] (mod p)" + with a cong_mult_rcancel [of a "int p" x y] have "[x = y] (mod p)" by simp with cong_less_imp_eq_int [of x y p] p_minus_one_l order_le_less_trans [of x "(int p - 1) div 2" p] @@ -127,12 +127,8 @@ by (metis b c cong_less_imp_eq_int d e zero_less_imp_eq_int of_nat_0_le_iff) qed show ?thesis - apply (insert p_ge_2 p_a_relprime p_minus_one_l) - apply (auto simp add: B_def) - apply (rule ResSet_image) - apply (auto simp add: A_res) - apply (auto simp add: A_def *) - done + using p_ge_2 p_a_relprime p_minus_one_l + by (metis "*" A_def A_res B_def GAUSS.ResSet_image GAUSS_axioms greaterThanAtMost_iff odd_p odd_pos of_nat_0_less_iff) qed lemma SR_B_inj: "inj_on (\x. x mod p) B" @@ -149,11 +145,11 @@ from a have a': "[x * a = y * a](mod p)" using cong_def by blast from p_a_relprime have "\p dvd a" - by (simp add: cong_altdef_int) + by (simp add: cong_0_iff) with p_prime prime_imp_coprime [of _ "nat \a\"] have "coprime a (int p)" by (simp_all add: ac_simps) - with a' cong_mult_rcancel_int [of a "int p" x y] + with a' cong_mult_rcancel [of a "int p" x y] have "[x = y] (mod p)" by simp with cong_less_imp_eq_int [of x y p] p_minus_one_l order_le_less_trans [of x "(int p - 1) div 2" p] @@ -224,7 +220,7 @@ "coprime x p" if "x \ A" proof - from A_ncong_p [OF that] have "\ int p dvd x" - by (simp add: cong_altdef_int) + by (simp add: cong_0_iff) with p_prime show ?thesis by (metis (no_types) coprime_commute prime_imp_coprime prime_nat_int_transfer) qed @@ -370,7 +366,7 @@ then have "[prod id A * (-1)^(card E) = prod id A * a^(card A)](mod p)" by (rule cong_trans) (simp add: aux cong del: prod.strong_cong) with A_prod_relprime have "[(- 1) ^ card E = a ^ card A](mod p)" - by (metis cong_mult_lcancel_int) + by (metis cong_mult_lcancel) then show ?thesis by (simp add: A_card_eq cong_sym) qed @@ -390,7 +386,8 @@ moreover have "(-1::int) ^ (card E) = 1 \ (-1::int) ^ (card E) = -1" using neg_one_even_power neg_one_odd_power by blast moreover have "[1 \ - 1] (mod int p)" - using cong_altdef_int nonzero_mod_p[of 2] p_odd_int by fastforce + using cong_iff_dvd_diff [where 'a=int] nonzero_mod_p[of 2] p_odd_int + by fastforce ultimately show ?thesis by (auto simp add: cong_sym) qed diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Number_Theory/Pocklington.thy --- a/src/HOL/Number_Theory/Pocklington.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Number_Theory/Pocklington.thy Sat Aug 11 16:02:55 2018 +0200 @@ -121,7 +121,7 @@ obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" "\y. ?P y \ y = x" by blast from ma nb x have "coprime x a" "coprime x b" - using cong_imp_coprime_nat cong_sym by blast+ + using cong_imp_coprime cong_sym by blast+ then have "coprime x (a*b)" by simp with x show ?thesis @@ -209,7 +209,7 @@ by arith+ from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp - from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime_nat an1 + from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime an1 have an: "coprime a n" "coprime (a ^ (n - 1)) n" using \n \ 2\ by simp_all have False if H0: "\m. 0 < m \ m < n - 1 \ [a ^ m = 1] (mod n)" (is "\m. ?P m") @@ -716,7 +716,7 @@ by simp qed then have b1: "b \ 1" by arith - from cong_imp_coprime_nat[OF Cong.cong_diff_nat[OF cong_sym [OF b(1)] cong_refl [of 1] b1]] + from cong_imp_coprime[OF Cong.cong_diff_nat[OF cong_sym [OF b(1)] cong_refl [of 1] b1]] ath b1 b nqr have "coprime (a ^ ((n - 1) div p) - 1) n" by simp @@ -858,7 +858,7 @@ have "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)" by (simp add: cong_diff_nat) then show ?thesis - by (metis cong_imp_coprime_nat eq1 p') + by (metis cong_imp_coprime eq1 p') qed with pocklington[OF n qrn[symmetric] nq2 an1] show ?thesis by blast diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Number_Theory/Quadratic_Reciprocity.thy --- a/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Sat Aug 11 16:02:55 2018 +0200 @@ -93,7 +93,7 @@ lemma Gpq: "GAUSS p q" using p_prime pq_neq p_ge_2 q_prime - by (auto simp: GAUSS_def cong_altdef_int dest: primes_dvd_imp_eq) + by (auto simp: GAUSS_def cong_iff_dvd_diff dest: primes_dvd_imp_eq) lemma Gqp: "GAUSS q p" by (simp add: QRqp QR.Gpq) @@ -304,7 +304,7 @@ by (simp add: f_2_def) lemma f_2_lemma_2: "[f_2 x = int p - x] (mod p)" - by (simp add: f_2_def cong_altdef_int) + by (simp add: f_2_def cong_iff_dvd_diff) lemma f_2_lemma_3: "f_2 x \ S \ x \ f_2 ` S" using f_2_lemma_1[of x] image_eqI[of x f_2 "f_2 x" S] by presburger diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Real_Vector_Spaces.thy Sat Aug 11 16:02:55 2018 +0200 @@ -1012,6 +1012,9 @@ lemma dist_triangle3: "dist x y \ dist a x + dist a y" using dist_triangle2 [of x y a] by (simp add: dist_commute) +lemma abs_dist_diff_le: "\dist a b - dist b c\ \ dist a c" + using dist_triangle3[of b c a] dist_triangle2[of a b c] by simp + lemma dist_pos_lt: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Series.thy --- a/src/HOL/Series.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Series.thy Sat Aug 11 16:02:55 2018 +0200 @@ -703,6 +703,27 @@ qed qed +lemma summable_Cauchy': + fixes f :: "nat \ 'a :: banach" + assumes "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" + assumes "filterlim g (nhds 0) sequentially" + shows "summable f" +proof (subst summable_Cauchy, intro allI impI, goal_cases) + case (1 e) + from order_tendstoD(2)[OF assms(2) this] and assms(1) + have "eventually (\m. \n. norm (sum f {m.. m") auto + qed + qed + thus ?case by (auto simp: eventually_at_top_linorder) +qed + context fixes f :: "nat \ 'a::banach" begin diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/Topological_Spaces.thy Sat Aug 11 16:02:55 2018 +0200 @@ -2131,6 +2131,9 @@ lemma isCont_def: "isCont f a \ f \a\ f a" by (rule continuous_at) +lemma isContD: "isCont f x \ f \x\ f x" + by (simp add: isCont_def) + lemma isCont_cong: assumes "eventually (\x. f x = g x) (nhds x)" shows "isCont f x \ isCont g x" diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/HOL/ex/Datatype_Record_Examples.thy --- a/src/HOL/ex/Datatype_Record_Examples.thy Tue Aug 07 11:39:40 2018 +0200 +++ b/src/HOL/ex/Datatype_Record_Examples.thy Sat Aug 11 16:02:55 2018 +0200 @@ -45,4 +45,23 @@ lemma "b_set \ field_1 = True, field_2 = False \ = {False}" by simp +text \More tests\ + +datatype_record ('a, 'b) test1 = + field_t11 :: 'a + field_t12 :: 'b + field_t13 :: nat + field_t14 :: int + +thm test1.record_simps + +definition ID where "ID x = x" +lemma ID_cong[cong]: "ID x = ID x" by (rule refl) + +lemma "update_field_t11 f (update_field_t12 g (update_field_t11 h x)) = ID (update_field_t12 g (update_field_t11 (\x. f (h x)) x))" + apply (simp only: test1.record_simps) + apply (subst ID_def) + apply (rule refl) + done + end \ No newline at end of file diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/ROOT --- a/src/Pure/ROOT Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/ROOT Sat Aug 11 16:02:55 2018 +0200 @@ -4,7 +4,7 @@ description {* The Pure logical framework *} - options [threads = 1] + options [threads = 1, export_theory] theories Pure (global) ML_Bootstrap (global) diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/Thy/export_theory.ML Sat Aug 11 16:02:55 2018 +0200 @@ -13,6 +13,28 @@ structure Export_Theory: EXPORT_THEORY = struct +(* names for bound variables *) + +local + fun declare_names (Abs (_, _, b)) = declare_names b + | declare_names (t $ u) = declare_names t #> declare_names u + | declare_names (Const (c, _)) = Name.declare (Long_Name.base_name c) + | declare_names (Free (x, _)) = Name.declare x + | declare_names _ = I; + + fun variant_abs bs (Abs (x, T, t)) = + let + val names = fold Name.declare bs (declare_names t Name.context); + val x' = #1 (Name.variant x names); + val t' = variant_abs (x' :: bs) t; + in Abs (x', T, t') end + | variant_abs bs (t $ u) = variant_abs bs t $ variant_abs bs u + | variant_abs _ t = t; +in + val named_bounds = variant_abs []; +end; + + (* general setup *) fun setup_presentation f = @@ -59,9 +81,12 @@ else (case export name decl of NONE => I - | SOME body => cons (name, XML.Elem (entity_markup space name, body)))) - |> sort_by #1 |> map #2 + | SOME body => + cons (#serial (Name_Space.the_entry space name), + XML.Elem (entity_markup space name, body)))) + |> sort (int_ord o apply2 #1) |> map #2 end; + in if null elems then () else export_body thy export_name elems end; @@ -86,7 +111,7 @@ val encode_const = let open XML.Encode Term_XML.Encode - in triple (list string) typ (option term) end; + in triple (list string) typ (option (term o named_bounds)) end; fun export_const c (T, abbrev) = let @@ -102,31 +127,36 @@ (* axioms and facts *) - val standard_prop_of = - Thm.transfer thy #> - Thm.check_hyps (Context.Theory thy) #> - Drule.sort_constraint_intr_shyps #> - Thm.full_prop_of; - - val encode_props = - let open XML.Encode Term_XML.Encode - in triple (list (pair string sort)) (list (pair string typ)) (list term) end; + fun standard_prop_of raw_thm = + let + val thm = raw_thm + |> Thm.transfer thy + |> Thm.check_hyps (Context.Theory thy) + |> Thm.strip_shyps; + val prop = thm + |> Thm.full_prop_of + |> Term_Subst.zero_var_indexes; + in (Thm.extra_shyps thm, prop) end; - fun export_props props = + fun encode_prop (Ss, prop) = let - val props' = map Logic.unvarify_global props; - val typargs = rev (fold Term.add_tfrees props' []); - val args = rev (fold Term.add_frees props' []); - in encode_props (typargs, args, props') end; + val prop' = Logic.unvarify_global (named_bounds prop); + val typargs = rev (Term.add_tfrees prop' []); + val sorts = Name.invent (Name.make_context (map #1 typargs)) Name.aT (length Ss) ~~ Ss; + val args = rev (Term.add_frees prop' []); + in + (sorts @ typargs, args, prop') |> + let open XML.Encode Term_XML.Encode + in triple (list (pair string sort)) (list (pair string typ)) term end + end; - val export_axiom = export_props o single; - val export_fact = export_props o Term_Subst.zero_var_indexes_list o map standard_prop_of; + val encode_fact = XML.Encode.list encode_prop o map standard_prop_of; val _ = - export_entities "axioms" (K (SOME o export_axiom)) Theory.axiom_space + export_entities "axioms" (fn _ => fn t => SOME (encode_prop ([], t))) Theory.axiom_space (Theory.axioms_of thy); val _ = - export_entities "facts" (K (SOME o export_fact)) (Facts.space_of o Global_Theory.facts_of) + export_entities "facts" (K (SOME o encode_fact)) (Facts.space_of o Global_Theory.facts_of) (Facts.dest_static true [] (Global_Theory.facts_of thy)); @@ -134,7 +164,7 @@ val encode_class = let open XML.Encode Term_XML.Encode - in pair (list (pair string typ)) (list term) end; + in pair (list (pair string typ)) (list (term o named_bounds)) end; fun export_class name = (case try (Axclass.get_info thy) name of diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/Thy/export_theory.scala --- a/src/Pure/Thy/export_theory.scala Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/Thy/export_theory.scala Sat Aug 11 16:02:55 2018 +0200 @@ -66,8 +66,8 @@ sealed case class Theory(name: String, parents: List[String], types: List[Type], consts: List[Const], - axioms: List[Axiom], - facts: List[Fact], + axioms: List[Fact_Single], + facts: List[Fact_Multi], classes: List[Class], typedefs: List[Typedef], classrel: List[Classrel], @@ -75,6 +75,14 @@ { override def toString: String = name + lazy val entities: Set[Long] = + Set.empty[Long] ++ + types.iterator.map(_.entity.serial) ++ + consts.iterator.map(_.entity.serial) ++ + axioms.iterator.map(_.entity.serial) ++ + facts.iterator.map(_.entity.serial) ++ + classes.iterator.map(_.entity.serial) + def cache(cache: Term.Cache): Theory = Theory(cache.string(name), parents.map(cache.string(_)), @@ -122,18 +130,41 @@ if (cache.isDefined) theory.cache(cache.get) else theory } + def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory = + { + val session_name = Thy_Header.PURE + val theory_name = Thy_Header.PURE + + using(store.open_database(session_name))(db => + { + db.transaction { + read_theory(Export.Provider.database(db, session_name, theory_name), + session_name, theory_name, cache = cache) + } + }) + } + /* entities */ - sealed case class Entity(name: String, serial: Long, pos: Position.T) + object Kind extends Enumeration { - override def toString: String = name + val TYPE = Value("type") + val CONST = Value("const") + val AXIOM = Value("axiom") + val FACT = Value("fact") + val CLASS = Value("class") + } + + sealed case class Entity(kind: Kind.Value, name: String, serial: Long, pos: Position.T) + { + override def toString: String = kind.toString + " " + quote(name) def cache(cache: Term.Cache): Entity = - Entity(cache.string(name), serial, cache.position(pos)) + Entity(kind, cache.string(name), serial, cache.position(pos)) } - def decode_entity(tree: XML.Tree): (Entity, XML.Body) = + def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) = { def err(): Nothing = throw new XML.XML_Body(List(tree)) @@ -142,7 +173,7 @@ val name = Markup.Name.unapply(props) getOrElse err() val serial = Markup.Serial.unapply(props) getOrElse err() val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) }) - (Entity(name, serial, pos), body) + (Entity(kind, name, serial, pos), body) case _ => err() } } @@ -161,7 +192,7 @@ def read_types(provider: Export.Provider): List[Type] = provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) => { - val (entity, body) = decode_entity(tree) + val (entity, body) = decode_entity(Kind.TYPE, tree) val (args, abbrev) = { import XML.Decode._ @@ -186,7 +217,7 @@ def read_consts(provider: Export.Provider): List[Const] = provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) => { - val (entity, body) = decode_entity(tree) + val (entity, body) = decode_entity(Kind.CONST, tree) val (args, typ, abbrev) = { import XML.Decode._ @@ -196,56 +227,65 @@ }) - /* axioms and facts */ + /* facts */ - def decode_props(body: XML.Body): - (List[(String, Term.Sort)], List[(String, Term.Typ)], List[Term.Term]) = + sealed case class Prop( + typargs: List[(String, Term.Sort)], + args: List[(String, Term.Typ)], + term: Term.Term) { - import XML.Decode._ - import Term_XML.Decode._ - triple(list(pair(string, sort)), list(pair(string, typ)), list(term))(body) + def cache(cache: Term.Cache): Prop = + Prop( + typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), + args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), + cache.term(term)) } - sealed case class Axiom( - entity: Entity, - typargs: List[(String, Term.Sort)], - args: List[(String, Term.Typ)], - prop: Term.Term) + def decode_prop(body: XML.Body): Prop = { - def cache(cache: Term.Cache): Axiom = - Axiom(entity.cache(cache), - typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), - args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), - cache.term(prop)) + val (typargs, args, t) = + { + import XML.Decode._ + import Term_XML.Decode._ + triple(list(pair(string, sort)), list(pair(string, typ)), term)(body) + } + Prop(typargs, args, t) } - def read_axioms(provider: Export.Provider): List[Axiom] = + sealed case class Fact_Single(entity: Entity, prop: Prop) + { + def cache(cache: Term.Cache): Fact_Single = + Fact_Single(entity.cache(cache), prop.cache(cache)) + } + + sealed case class Fact_Multi(entity: Entity, props: List[Prop]) + { + def cache(cache: Term.Cache): Fact_Multi = + Fact_Multi(entity.cache(cache), props.map(_.cache(cache))) + + def split: List[Fact_Single] = + props match { + case List(prop) => List(Fact_Single(entity, prop)) + case _ => + for ((prop, i) <- props.zipWithIndex) + yield Fact_Single(entity.copy(name = entity.name + "(" + (i + 1) + ")"), prop) + } + } + + def read_axioms(provider: Export.Provider): List[Fact_Single] = provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) => { - val (entity, body) = decode_entity(tree) - val (typargs, args, List(prop)) = decode_props(body) - Axiom(entity, typargs, args, prop) + val (entity, body) = decode_entity(Kind.AXIOM, tree) + val prop = decode_prop(body) + Fact_Single(entity, prop) }) - sealed case class Fact( - entity: Entity, - typargs: List[(String, Term.Sort)], - args: List[(String, Term.Typ)], - props: List[Term.Term]) - { - def cache(cache: Term.Cache): Fact = - Fact(entity.cache(cache), - typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), - args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), - props.map(cache.term(_))) - } - - def read_facts(provider: Export.Provider): List[Fact] = + def read_facts(provider: Export.Provider): List[Fact_Multi] = provider.uncompressed_yxml(export_prefix + "facts").map((tree: XML.Tree) => { - val (entity, body) = decode_entity(tree) - val (typargs, args, props) = decode_props(body) - Fact(entity, typargs, args, props) + val (entity, body) = decode_entity(Kind.FACT, tree) + val props = XML.Decode.list(decode_prop)(body) + Fact_Multi(entity, props) }) @@ -263,7 +303,7 @@ def read_classes(provider: Export.Provider): List[Class] = provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) => { - val (entity, body) = decode_entity(tree) + val (entity, body) = decode_entity(Kind.CLASS, tree) val (params, axioms) = { import XML.Decode._ diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/build-jars --- a/src/Pure/build-jars Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/build-jars Sat Aug 11 16:02:55 2018 +0200 @@ -159,6 +159,7 @@ Tools/update_then.scala Tools/update_theorems.scala library.scala + pure_thy.scala term.scala term_xml.scala ../Tools/Graphview/graph_file.scala diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/drule.ML --- a/src/Pure/drule.ML Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/drule.ML Sat Aug 11 16:02:55 2018 +0200 @@ -98,8 +98,6 @@ val is_sort_constraint: term -> bool val sort_constraintI: thm val sort_constraint_eq: thm - val sort_constraint_intr: indexname * sort -> thm -> thm - val sort_constraint_intr_shyps: thm -> thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm val rename_bvars: (string * string) list -> thm -> thm @@ -649,26 +647,6 @@ (Thm.unvarify_global (Context.the_global_context ()) sort_constraintI))) (implies_intr_list [A, C] (Thm.assume A))); -val sort_constraint_eq' = Thm.symmetric sort_constraint_eq; - -fun sort_constraint_intr tvar thm = - Thm.equal_elim - (Thm.instantiate - ([((("'a", 0), []), Thm.global_ctyp_of (Thm.theory_of_thm thm) (TVar tvar))], - [((("A", 0), propT), Thm.cprop_of thm)]) - sort_constraint_eq') thm; - -fun sort_constraint_intr_shyps raw_thm = - let val thm = Thm.strip_shyps raw_thm in - (case Thm.extra_shyps thm of - [] => thm - | shyps => - let - val names = Name.make_context (map #1 (Thm.fold_terms Term.add_tvar_names thm [])); - val constraints = map (rpair 0) (Name.invent names Name.aT (length shyps)) ~~ shyps; - in Thm.strip_shyps (fold_rev sort_constraint_intr constraints thm) end) - end; - end; diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/library.scala --- a/src/Pure/library.scala Tue Aug 07 11:39:40 2018 +0200 +++ b/src/Pure/library.scala Sat Aug 11 16:02:55 2018 +0200 @@ -259,6 +259,15 @@ result.toList } + def replicate[A](n: Int, a: A): List[A] = + if (n < 0) throw new IllegalArgumentException + else if (n == 0) Nil + else { + val res = new mutable.ListBuffer[A] + (1 to n).foreach(_ => res += a) + res.toList + } + /* proper values */ diff -r 0c62e3b4f4c0 -r 682ff0e84387 src/Pure/pure_thy.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/pure_thy.scala Sat Aug 11 16:02:55 2018 +0200 @@ -0,0 +1,20 @@ +/* Title: Pure/pure_thy.scala + Author: Makarius + +Pure theory content. +*/ + +package isabelle + + +object Pure_Thy +{ + val DUMMY: String = "dummy" + val FUN: String = "fun" + val PROP: String = "prop" + val ITSELF: String = "itself" + + val ALL: String = "Pure.all" + val IMP: String = "Pure.imp" + val EQ: String = "Pure.eq" +}