# HG changeset patch # User haftmann # Date 1546166096 0 # Node ID 27dae626822b664bc4d8379c5789bf154b80c5a0 # Parent 4aed40ecfb436b025633f1e1b4b688ed1918272c prefer naming convention from datatype package for strong congruence rules diff -r 4aed40ecfb43 -r 27dae626822b NEWS --- a/NEWS Sun Dec 30 10:34:56 2018 +0000 +++ b/NEWS Sun Dec 30 10:34:56 2018 +0000 @@ -84,7 +84,8 @@ INCOMPATIBILITY. * Strong congruence rules (with =simp=> in the premises) for constant f -are now uniformly called f_cong_strong. +are now uniformly called f_cong_simp, in accordance with congruence +rules produced for mappers by the datatype package. INCOMPATIBILITY. * Sledgehammer: The URL for SystemOnTPTP, which is used by remote provers, has been updated. diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Bochner_Integration.thy Sun Dec 30 10:34:56 2018 +0000 @@ -524,7 +524,7 @@ assumes "M = N" "\x. x \ space N \ f x = g x" "x = y" shows "has_bochner_integral M f x \ has_bochner_integral N g y" unfolding has_bochner_integral.simps assms(1,3) - using assms(2) by (simp cong: measurable_cong_strong nn_integral_cong_strong) + using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp) lemma%unimportant has_bochner_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Caratheodory.thy --- a/src/HOL/Analysis/Caratheodory.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Caratheodory.thy Sun Dec 30 10:34:56 2018 +0000 @@ -737,7 +737,7 @@ using f C Ai unfolding bij_betw_def by auto then have "\range f = A i" using f C Ai unfolding bij_betw_def - by (auto simp add: f_def cong del: SUP_cong_strong) + by (auto simp add: f_def cong del: SUP_cong_simp) moreover { have "(\j. \_r (f j)) = (\j. if j \ {..< card C} then \_r (f j) else 0)" using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def) diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/L2_Norm.thy --- a/src/HOL/Analysis/L2_Norm.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/L2_Norm.thy Sun Dec 30 10:34:56 2018 +0000 @@ -14,7 +14,7 @@ "\A = B; \x. x \ B \ f x = g x\ \ L2_set f A = L2_set g B" unfolding L2_set_def by simp -lemma L2_set_cong_strong: +lemma L2_set_cong_simp: "\A = B; \x. x \ B =simp=> f x = g x\ \ L2_set f A = L2_set g B" unfolding L2_set_def simp_implies_def by simp diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Lebesgue_Measure.thy Sun Dec 30 10:34:56 2018 +0000 @@ -407,7 +407,7 @@ lemma measurable_lebesgue_cong: assumes "\x. x \ S \ f x = g x" shows "f \ measurable (lebesgue_on S) M \ g \ measurable (lebesgue_on S) M" - by (metis (mono_tags, lifting) IntD1 assms measurable_cong_strong space_restrict_space) + by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space) text%unimportant \Measurability of continuous functions\ diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Measure_Space.thy Sun Dec 30 10:34:56 2018 +0000 @@ -1143,7 +1143,7 @@ "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto -lemma AE_cong_strong: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" +lemma AE_cong_simp: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" by (auto simp: simp_implies_def) lemma AE_all_countable: @@ -2589,7 +2589,7 @@ using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" - using A by (auto cong del: SUP_cong_strong) + using A by (auto cong del: SUP_cong_simp) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Sun Dec 30 10:34:56 2018 +0000 @@ -2058,7 +2058,7 @@ by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) qed then show ?thesis - unfolding nn_integral_def_finite by (simp cong del: SUP_cong_strong) + unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp) qed lemma nn_integral_count_space_indicator: diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Analysis/Sigma_Algebra.thy --- a/src/HOL/Analysis/Sigma_Algebra.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Analysis/Sigma_Algebra.thy Sun Dec 30 10:34:56 2018 +0000 @@ -436,10 +436,10 @@ by (auto simp add: binary_def) lemma Un_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: range_binary_eq cong del: SUP_cong_strong) + by (simp add: range_binary_eq cong del: SUP_cong_simp) lemma Int_range_binary: "a \ b = (\i::nat. binary a b i)" - by (simp add: range_binary_eq cong del: INF_cong_strong) + by (simp add: range_binary_eq cong del: INF_cong_simp) lemma sigma_algebra_iff2: "sigma_algebra \ M \ @@ -943,7 +943,7 @@ done lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" - by (simp add: range_binaryset_eq cong del: SUP_cong_strong) + by (simp add: range_binaryset_eq cong del: SUP_cong_simp) subsubsection \Closed CDI\ @@ -1802,7 +1802,7 @@ unfolding measurable_def using assms by (simp cong: vimage_inter_cong Pi_cong add: simp_implies_def) -lemma measurable_cong_strong: +lemma measurable_cong_simp: "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ f \ measurable M M' \ g \ measurable N N'" by (metis measurable_cong) diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Complete_Lattices.thy Sun Dec 30 10:34:56 2018 +0000 @@ -61,7 +61,7 @@ lemma INF_cong: "A = B \ (\x. x \ B \ C x = D x) \ \(C ` A) = \(D ` B)" by (simp add: image_def) -lemma INF_cong_strong [cong]: +lemma INF_cong_simp [cong]: "A = B \ (\x. x \ B =simp=> C x = D x) \ \(C ` A) = \(D ` B)" unfolding simp_implies_def by (fact INF_cong) @@ -82,9 +82,9 @@ lemma SUP_cong: "A = B \ (\x. x \ B \ C x = D x) \ \(C ` A) = \(D ` B)" by (fact Inf.INF_cong) -lemma SUP_cong_strong [cong]: +lemma SUP_cong_simp [cong]: "A = B \ (\x. x \ B =simp=> C x = D x) \ \(C ` A) = \(D ` B)" -by (fact Inf.INF_cong_strong) +by (fact Inf.INF_cong_simp) end @@ -179,19 +179,19 @@ by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) lemma INF_insert [simp]: "(\x\insert a A. f x) = f a \ \(f ` A)" - by (simp cong del: INF_cong_strong) + by (simp cong del: INF_cong_simp) lemma Sup_insert [simp]: "\insert a A = a \ \A" by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) lemma SUP_insert [simp]: "(\x\insert a A. f x) = f a \ \(f ` A)" - by (simp cong del: SUP_cong_strong) + by (simp cong del: SUP_cong_simp) lemma INF_empty [simp]: "(\x\{}. f x) = \" - by (simp cong del: INF_cong_strong) + by (simp cong del: INF_cong_simp) lemma SUP_empty [simp]: "(\x\{}. f x) = \" - by (simp cong del: SUP_cong_strong) + by (simp cong del: SUP_cong_simp) lemma Inf_UNIV [simp]: "\UNIV = \" by (auto intro!: antisym Inf_lower) @@ -437,11 +437,11 @@ lemma INF_eq_iff: "I \ {} \ (\i. i \ I \ f i \ c) \ \(f ` I) = c \ (\i\I. f i = c)" using INF_eq_const [of I f c] INF_lower [of _ I f] - by (auto intro: antisym cong del: INF_cong_strong) + by (auto intro: antisym cong del: INF_cong_simp) lemma SUP_eq_iff: "I \ {} \ (\i. i \ I \ c \ f i) \ \(f ` I) = c \ (\i\I. f i = c)" using SUP_eq_const [of I f c] SUP_upper [of _ I f] - by (auto intro: antisym cong del: SUP_cong_strong) + by (auto intro: antisym cong del: SUP_cong_simp) end diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Data_Structures/Braun_Tree.thy --- a/src/HOL/Data_Structures/Braun_Tree.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Data_Structures/Braun_Tree.thy Sun Dec 30 10:34:56 2018 +0000 @@ -252,7 +252,7 @@ by (simp add: insert_ident Icc_eq_insert_lb_nat nat_in_image braun_indices1) thus ?case using Node.IH even_odd_decomp[OF _ _ 3] by(simp add: card_image inj_on_def card_braun_indices Let_def 1 2 inj_image_eq_iff image_comp - cong: image_cong_strong) + cong: image_cong_simp) qed lemma braun_iff_braun_indices: "braun t \ braun_indices t = {1..size t}" diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Hoare_Parallel/RG_Examples.thy --- a/src/HOL/Hoare_Parallel/RG_Examples.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Sun Dec 30 10:34:56 2018 +0000 @@ -321,7 +321,7 @@ \\iX!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i)\]" apply (rule Parallel) -apply (auto cong del: INF_cong_strong SUP_cong_strong) +apply (auto cong del: INF_cong_simp SUP_cong_simp) apply force apply (rule While) apply force diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Library/Complete_Partial_Order2.thy --- a/src/HOL/Library/Complete_Partial_Order2.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Library/Complete_Partial_Order2.thy Sun Dec 30 10:34:56 2018 +0000 @@ -188,7 +188,7 @@ proof(cases "Y = {}") case True then show ?thesis - by (simp add: image_constant_conv cong del: SUP_cong_strong) + by (simp add: image_constant_conv cong del: SUP_cong_simp) next case False have chain1: "\f. f \ Z \ Complete_Partial_Order.chain (\) (f ` Y)" @@ -518,7 +518,7 @@ context ccpo begin lemma cont_const [simp, cont_intro]: "cont luba orda Sup (\) (\x. c)" -by (rule contI) (simp add: image_constant_conv cong del: SUP_cong_strong) +by (rule contI) (simp add: image_constant_conv cong del: SUP_cong_simp) lemma mcont_const [cont_intro, simp]: "mcont luba orda Sup (\) (\x. c)" @@ -695,7 +695,7 @@ hence "\Y \ bound" using chain by(auto intro: ccpo_Sup_least) moreover have "Y \ {x. \ x \ bound} = {}" using True by auto ultimately show ?thesis using True Y - by (auto simp add: image_constant_conv cong del: c.SUP_cong_strong) + by (auto simp add: image_constant_conv cong del: c.SUP_cong_simp) next case False let ?Y = "Y \ {x. \ x \ bound}" diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Probability/Giry_Monad.thy --- a/src/HOL/Probability/Giry_Monad.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Probability/Giry_Monad.thy Sun Dec 30 10:34:56 2018 +0000 @@ -1763,7 +1763,7 @@ qed qed -lemma bind_cong_strong: "M = N \ (\x. x\space M =simp=> f x = g x) \ bind M f = bind N g" +lemma bind_cong_simp: "M = N \ (\x. x\space M =simp=> f x = g x) \ bind M f = bind N g" by (auto simp: simp_implies_def intro!: bind_cong) lemma sets_bind_measurable: diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Set.thy --- a/src/HOL/Set.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Set.thy Sun Dec 30 10:34:56 2018 +0000 @@ -426,7 +426,7 @@ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Ball_def) -lemma ball_cong_strong [cong]: +lemma ball_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Ball_def) @@ -436,7 +436,7 @@ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Bex_def cong: conj_cong) -lemma bex_cong_strong [cong]: +lemma bex_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Bex_def cong: conj_cong) @@ -939,7 +939,7 @@ lemma image_cong: "\ M = N; \x. x \ N \ f x = g x \ \ f ` M = g ` N" by (simp add: image_def) -lemma image_cong_strong: "\ M = N; \x. x \ N =simp=> f x = g x\ \ f ` M = g ` N" +lemma image_cong_simp: "\ M = N; \x. x \ N =simp=> f x = g x\ \ f ` M = g ` N" by (simp add: image_def simp_implies_def) lemma image_Int_subset: "f ` (A \ B) \ f ` A \ f ` B" diff -r 4aed40ecfb43 -r 27dae626822b src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Sun Dec 30 10:34:56 2018 +0000 +++ b/src/HOL/Topological_Spaces.thy Sun Dec 30 10:34:56 2018 +0000 @@ -1882,7 +1882,7 @@ unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_at_filter) -lemma continuous_on_cong_strong: +lemma continuous_on_cong_simp: "s = t \ (\x. x \ t =simp=> f x = g x) \ continuous_on s f \ continuous_on t g" unfolding simp_implies_def by (rule continuous_on_cong)