# HG changeset patch # User paulson # Date 1585416428 0 # Node ID beef2e221c2603a666552d218521e966f2a6307a # Parent d97f504c8145bf8ec96e77c7f9204fe135128e42# Parent 856c68ab6f13e063342f025352fd4deaf5ce65d2 merged diff -r d97f504c8145 -r beef2e221c26 src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy Sat Mar 28 14:01:45 2020 +0100 +++ b/src/HOL/Equiv_Relations.thy Sat Mar 28 17:27:08 2020 +0000 @@ -34,21 +34,21 @@ unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" - apply (unfold equiv_def) - apply clarify - apply (rule equalityI) - apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+ - done + unfolding equiv_def + by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI) text \Second half.\ -lemma comp_equivI: "r\ O r = r \ Domain r = A \ equiv A r" - apply (unfold equiv_def refl_on_def sym_def trans_def) - apply (erule equalityE) - apply (subgoal_tac "\x y. (x, y) \ r \ (y, x) \ r") - apply fast - apply fast - done +lemma comp_equivI: + assumes "r\ O r = r" "Domain r = A" + shows "equiv A r" +proof - + have *: "\x y. (x, y) \ r \ (y, x) \ r" + using assms by blast + show ?thesis + unfolding equiv_def refl_on_def sym_def trans_def + using assms by (auto intro: *) +qed subsection \Equivalence classes\ @@ -58,10 +58,7 @@ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" - apply (assumption | rule equalityI equiv_class_subset)+ - apply (unfold equiv_def sym_def) - apply blast - done + by (intro equalityI equiv_class_subset; force simp add: equiv_def sym_def) lemma equiv_class_self: "equiv A r \ a \ A \ a \ r``{a}" unfolding equiv_def refl_on_def by blast @@ -91,7 +88,7 @@ definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where "A//r = (\x \ A. {r``{x}})" \ \set of equiv classes\ -lemma quotientI: "x \ A ==> r``{x} \ A//r" +lemma quotientI: "x \ A \ r``{x} \ A//r" unfolding quotient_def by blast lemma quotientE: "X \ A//r \ (\x. X = r``{x} \ x \ A \ P) \ P" @@ -101,32 +98,31 @@ unfolding equiv_def refl_on_def quotient_def by blast lemma quotient_disj: "equiv A r \ X \ A//r \ Y \ A//r \ X = Y \ X \ Y = {}" - apply (unfold quotient_def) - apply clarify - apply (rule equiv_class_eq) - apply assumption - apply (unfold equiv_def trans_def sym_def) - apply blast - done + unfolding quotient_def equiv_def trans_def sym_def by blast lemma quotient_eqI: - "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ (x, y) \ r \ X = Y" - apply (clarify elim!: quotientE) - apply (rule equiv_class_eq) - apply assumption - apply (unfold equiv_def sym_def trans_def) - apply blast - done + assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" "(x, y) \ r" + shows "X = Y" +proof - + obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" + using assms by (auto elim!: quotientE) + then have "(a,b) \ r" + using xy \equiv A r\ unfolding equiv_def sym_def trans_def by blast + then show ?thesis + unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) +qed lemma quotient_eq_iff: - "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ X = Y \ (x, y) \ r" - apply (rule iffI) - prefer 2 - apply (blast del: equalityI intro: quotient_eqI) - apply (clarify elim!: quotientE) - apply (unfold equiv_def sym_def trans_def) - apply blast - done + assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" + shows "X = Y \ (x, y) \ r" +proof + assume L: "X = Y" + with assms show "(x, y) \ r" + unfolding equiv_def sym_def trans_def by (blast elim!: quotientE) +next + assume \
: "(x, y) \ r" show "X = Y" + by (rule quotient_eqI) (use \
assms in \blast+\) +qed lemma eq_equiv_class_iff2: "equiv A r \ x \ A \ y \ A \ {x}//r = {y}//r \ (x, y) \ r" by (simp add: quotient_def eq_equiv_class_iff) @@ -189,22 +185,22 @@ \ \lemma required to prove \UN_equiv_class\\ by auto -lemma UN_equiv_class: "equiv A r \ f respects r \ a \ A \ (\x \ r``{a}. f x) = f a" +lemma UN_equiv_class: + assumes "equiv A r" "f respects r" "a \ A" + shows "(\x \ r``{a}. f x) = f a" \ \Conversion rule\ - apply (rule equiv_class_self [THEN UN_constant_eq]) - apply assumption - apply assumption - apply (unfold equiv_def congruent_def sym_def) - apply (blast del: equalityI) - done +proof - + have \
: "\x\r `` {a}. f x = f a" + using assms unfolding equiv_def congruent_def sym_def by blast + show ?thesis + by (iprover intro: assms UN_constant_eq [OF equiv_class_self \
]) +qed lemma UN_equiv_class_type: - "equiv A r \ f respects r \ X \ A//r \ (\x. x \ A \ f x \ B) \ (\x \ X. f x) \ B" - apply (unfold quotient_def) - apply clarify - apply (subst UN_equiv_class) - apply auto - done + assumes r: "equiv A r" "f respects r" and X: "X \ A//r" and AB: "\x. x \ A \ f x \ B" + shows "(\x \ X. f x) \ B" + using assms unfolding quotient_def + by (auto simp: UN_equiv_class [OF r]) text \ Sufficient conditions for injectiveness. Could weaken premises! @@ -213,19 +209,23 @@ \ lemma UN_equiv_class_inject: - "equiv A r \ f respects r \ - (\x \ X. f x) = (\y \ Y. f y) \ X \ A//r ==> Y \ A//r - \ (\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r) - \ X = Y" - apply (unfold quotient_def) - apply clarify - apply (rule equiv_class_eq) - apply assumption - apply (subgoal_tac "f x = f xa") - apply blast - apply (erule box_equals) - apply (assumption | rule UN_equiv_class)+ - done + assumes "equiv A r" "f respects r" + and eq: "(\x \ X. f x) = (\y \ Y. f y)" + and X: "X \ A//r" and Y: "Y \ A//r" + and fr: "\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r" + shows "X = Y" +proof - + obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" + using assms by (auto elim!: quotientE) + then have "\ (f ` r `` {a}) = f a" "\ (f ` r `` {b}) = f b" + by (iprover intro: UN_equiv_class [OF \equiv A r\] assms)+ + then have "f a = f b" + using eq unfolding a b by (iprover intro: trans sym) + then have "(a,b) \ r" + using fr \a \ A\ \b \ A\ by blast + then show ?thesis + unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) +qed subsection \Defining binary operations upon equivalence classes\ @@ -253,15 +253,20 @@ unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: - "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a \ A2 \ - congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" - apply (unfold congruent_def) - apply clarify - apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+) - apply (simp add: UN_equiv_class congruent2_implies_congruent) - apply (unfold congruent2_def equiv_def refl_on_def) - apply (blast del: equalityI) - done + assumes "equiv A1 r1" "equiv A2 r2" "congruent2 r1 r2 f" "a \ A2" + shows "congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" + unfolding congruent_def +proof clarify + fix c d + assume cd: "(c,d) \ r1" + then have "c \ A1" "d \ A1" + using \equiv A1 r1\ by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2]) + with assms show "\ (f c ` r2 `` {a}) = \ (f d ` r2 `` {a})" + proof (simp add: UN_equiv_class congruent2_implies_congruent) + show "f c a = f d a" + using assms cd unfolding congruent2_def equiv_def refl_on_def by blast + qed +qed lemma UN_equiv_class2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a1 \ A1 \ a2 \ A2 \ @@ -273,11 +278,10 @@ \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" - apply (unfold quotient_def) - apply clarify - apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN - congruent2_implies_congruent quotientI) - done + unfolding quotient_def + by (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN + congruent2_implies_congruent quotientI) + lemma UN_UN_split_split_eq: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = @@ -293,60 +297,63 @@ \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ - apply (unfold congruent2_def equiv_def refl_on_def) - apply clarify - apply (blast intro: trans) - done + unfolding congruent2_def equiv_def refl_on_def + by (blast intro: trans) lemma congruent2_commuteI: assumes equivA: "equiv A r" and commute: "\y z. y \ A \ z \ A \ f y z = f z y" and congt: "\y z w. w \ A \ (y,z) \ r \ f w y = f w z" shows "f respects2 r" - apply (rule congruent2I [OF equivA equivA]) - apply (rule commute [THEN trans]) - apply (rule_tac [3] commute [THEN trans, symmetric]) - apply (rule_tac [5] sym) - apply (rule congt | assumption | - erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+ - done +proof (rule congruent2I [OF equivA equivA]) + note eqv = equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2] + show "\y z w. \w \ A; (y, z) \ r\ \ f y w = f z w" + by (iprover intro: commute [THEN trans] sym congt elim: eqv) + show "\y z w. \w \ A; (y, z) \ r\ \ f w y = f w z" + by (iprover intro: congt elim: eqv) +qed subsection \Quotients and finiteness\ text \Suggested by Florian Kammüller\ -lemma finite_quotient: "finite A \ r \ A \ A \ finite (A//r)" - \ \recall @{thm equiv_type}\ - apply (rule finite_subset) - apply (erule_tac [2] finite_Pow_iff [THEN iffD2]) - apply (unfold quotient_def) - apply blast - done +lemma finite_quotient: + assumes "finite A" "r \ A \ A" + shows "finite (A//r)" + \ \recall @{thm equiv_type}\ +proof - + have "A//r \ Pow A" + using assms unfolding quotient_def by blast + moreover have "finite (Pow A)" + using assms by simp + ultimately show ?thesis + by (iprover intro: finite_subset) +qed lemma finite_equiv_class: "finite A \ r \ A \ A \ X \ A//r \ finite X" - apply (unfold quotient_def) - apply (rule finite_subset) - prefer 2 apply assumption - apply blast - done + unfolding quotient_def + by (erule rev_finite_subset) blast -lemma equiv_imp_dvd_card: "finite A \ equiv A r \ \X \ A//r. k dvd card X \ k dvd card A" - apply (rule Union_quotient [THEN subst [where P="\A. k dvd card A"]]) - apply assumption - apply (rule dvd_partition) - prefer 3 apply (blast dest: quotient_disj) - apply (simp_all add: Union_quotient equiv_type) - done +lemma equiv_imp_dvd_card: + assumes "finite A" "equiv A r" "\X. X \ A//r \ k dvd card X" + shows "k dvd card A" +proof (rule Union_quotient [THEN subst]) + show "k dvd card (\ (A // r))" + apply (rule dvd_partition) + using assms + by (auto simp: Union_quotient dest: quotient_disj) +qed (use assms in blast) -lemma card_quotient_disjoint: "finite A \ inj_on (\x. {x} // r) A \ card (A//r) = card A" - apply (simp add:quotient_def) - apply (subst card_UN_disjoint) - apply assumption - apply simp - apply (fastforce simp add:inj_on_def) - apply simp - done +lemma card_quotient_disjoint: + assumes "finite A" "inj_on (\x. {x} // r) A" + shows "card (A//r) = card A" +proof - + have "\i\A. \j\A. i \ j \ r `` {j} \ r `` {i}" + using assms by (fastforce simp add: quotient_def inj_on_def) + with assms show ?thesis + by (simp add: quotient_def card_UN_disjoint) +qed subsection \Projection\ diff -r d97f504c8145 -r beef2e221c26 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Sat Mar 28 14:01:45 2020 +0100 +++ b/src/HOL/HOL.thy Sat Mar 28 17:27:08 2020 +0000 @@ -249,11 +249,7 @@ | | c = d *) lemma box_equals: "\a = b; a = c; b = d\ \ c = d" - apply (rule trans) - apply (rule trans) - apply (rule sym) - apply assumption+ - done + by (iprover intro: sym trans) text \For calculational reasoning:\ @@ -268,25 +264,17 @@ text \Similar to \AP_THM\ in Gordon's HOL.\ lemma fun_cong: "(f :: 'a \ 'b) = g \ f x = g x" - apply (erule subst) - apply (rule refl) - done + by (iprover intro: refl elim: subst) text \Similar to \AP_TERM\ in Gordon's HOL and FOL's \subst_context\.\ lemma arg_cong: "x = y \ f x = f y" - apply (erule subst) - apply (rule refl) - done + by (iprover intro: refl elim: subst) lemma arg_cong2: "\a = b; c = d\ \ f a c = f b d" - apply (erule ssubst)+ - apply (rule refl) - done + by (iprover intro: refl elim: subst) lemma cong: "\f = g; (x::'a) = y\ \ f x = g y" - apply (erule subst)+ - apply (rule refl) - done + by (iprover intro: refl elim: subst) ML \fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\ @@ -324,20 +312,15 @@ subsubsection \Universal quantifier (1)\ lemma spec: "\x::'a. P x \ P x" - apply (unfold All_def) - apply (rule eqTrueE) - apply (erule fun_cong) - done + unfolding All_def by (iprover intro: eqTrueE fun_cong) lemma allE: - assumes major: "\x. P x" - and minor: "P x \ R" + assumes major: "\x. P x" and minor: "P x \ R" shows R by (iprover intro: minor major [THEN spec]) lemma all_dupE: - assumes major: "\x. P x" - and minor: "\P x; \x. P x\ \ R" + assumes major: "\x. P x" and minor: "\P x; \x. P x\ \ R" shows R by (iprover intro: minor major major [THEN spec]) @@ -350,9 +333,7 @@ \ lemma FalseE: "False \ P" - apply (unfold False_def) - apply (erule spec) - done + unfolding False_def by (erule spec) lemma False_neq_True: "False = True \ P" by (erule eqTrueE [THEN FalseE]) @@ -363,29 +344,17 @@ lemma notI: assumes "P \ False" shows "\ P" - apply (unfold not_def) - apply (iprover intro: impI assms) - done + unfolding not_def by (iprover intro: impI assms) lemma False_not_True: "False \ True" - apply (rule notI) - apply (erule False_neq_True) - done + by (iprover intro: notI elim: False_neq_True) lemma True_not_False: "True \ False" - apply (rule notI) - apply (drule sym) - apply (erule False_neq_True) - done + by (iprover intro: notI dest: sym elim: False_neq_True) lemma notE: "\\ P; P\ \ R" - apply (unfold not_def) - apply (erule mp [THEN FalseE]) - apply assumption - done - -lemma notI2: "(P \ \ Pa) \ (P \ Pa) \ \ P" - by (erule notE [THEN notI]) (erule meta_mp) + unfolding not_def + by (iprover intro: mp [THEN FalseE]) subsubsection \Implication\ @@ -397,7 +366,7 @@ text \Reduces \Q\ to \P \ Q\, allowing substitution in \P\.\ lemma rev_mp: "\P; P \ Q\ \ Q" - by (iprover intro: mp) + by (rule mp) lemma contrapos_nn: assumes major: "\ Q" @@ -510,10 +479,10 @@ assumes major: "P \ Q" and minor: "\P; Q\ \ R" shows R - apply (rule minor) - apply (rule major [THEN conjunct1]) - apply (rule major [THEN conjunct2]) - done +proof (rule minor) + show P by (rule major [THEN conjunct1]) + show Q by (rule major [THEN conjunct2]) +qed lemma context_conjI: assumes P "P \ Q" @@ -533,14 +502,18 @@ subsubsection \Classical logic\ lemma classical: - assumes prem: "\ P \ P" + assumes "\ P \ P" shows P - apply (rule True_or_False [THEN disjE, THEN eqTrueE]) - apply assumption - apply (rule notI [THEN prem, THEN eqTrueI]) - apply (erule subst) - apply assumption - done +proof (rule True_or_False [THEN disjE]) + show P if "P = True" + using that by (iprover intro: eqTrueE) + show P if "P = False" + proof (intro notI assms) + assume P + with that show False + by (iprover elim: subst) + qed +qed lemmas ccontr = FalseE [THEN classical] @@ -550,16 +523,12 @@ assumes premp: P and premnot: "\ R \ \ P" shows R - apply (rule ccontr) - apply (erule notE [OF premnot premp]) - done + by (iprover intro: ccontr notE [OF premnot premp]) + text \Double negation law.\ lemma notnotD: "\\ P \ P" - apply (rule classical) - apply (erule notE) - apply assumption - done + by (iprover intro: ccontr notE ) lemma contrapos_pp: assumes p1: Q @@ -583,19 +552,15 @@ by (iprover intro: ex_prem [THEN exE] ex1I eq) lemma ex1E: - assumes major: "\!x. P x" - and minor: "\x. \P x; \y. P y \ y = x\ \ R" + assumes major: "\!x. P x" and minor: "\x. \P x; \y. P y \ y = x\ \ R" shows R - apply (rule major [unfolded Ex1_def, THEN exE]) - apply (erule conjE) - apply (iprover intro: minor) - done +proof (rule major [unfolded Ex1_def, THEN exE]) + show "\x. P x \ (\y. P y \ y = x) \ R" + by (iprover intro: minor elim: conjE) +qed lemma ex1_implies_ex: "\!x. P x \ \x. P x" - apply (erule ex1E) - apply (rule exI) - apply assumption - done + by (iprover intro: exI elim: ex1E) subsubsection \Classical intro rules for disjunction and existential quantifiers\ @@ -613,22 +578,18 @@ Note that \\ P\ is the second case, not the first. \ lemma case_split [case_names True False]: - assumes prem1: "P \ Q" - and prem2: "\ P \ Q" + assumes "P \ Q" "\ P \ Q" shows Q - apply (rule excluded_middle [THEN disjE]) - apply (erule prem2) - apply (erule prem1) - done + using excluded_middle [of P] + by (iprover intro: assms elim: disjE) text \Classical implies (\\\) elimination.\ lemma impCE: assumes major: "P \ Q" and minor: "\ P \ R" "Q \ R" shows R - apply (rule excluded_middle [of P, THEN disjE]) - apply (iprover intro: minor major [THEN mp])+ - done + using excluded_middle [of P] + by (iprover intro: minor major [THEN mp] elim: disjE)+ text \ This version of \\\ elimination works on \Q\ before \P\. It works best for @@ -639,9 +600,8 @@ assumes major: "P \ Q" and minor: "Q \ R" "\ P \ R" shows R - apply (rule excluded_middle [of P, THEN disjE]) - apply (iprover intro: minor major [THEN mp])+ - done + using assms by (elim impCE) + text \Classical \\\ elimination.\ lemma iffCE: @@ -874,9 +834,7 @@ ML \val HOL_cs = claset_of \<^context>\ lemma contrapos_np: "\ Q \ (\ P \ Q) \ P" - apply (erule swap) - apply (erule (1) meta_mp) - done + by (erule swap) declare ex_ex1I [rule del, intro! 2] and ex1I [intro] @@ -889,15 +847,13 @@ text \Better than \ex1E\ for classical reasoner: needs no quantifier duplication!\ lemma alt_ex1E [elim!]: assumes major: "\!x. P x" - and prem: "\x. \P x; \y y'. P y \ P y' \ y = y'\ \ R" + and minor: "\x. \P x; \y y'. P y \ P y' \ y = y'\ \ R" shows R - apply (rule ex1E [OF major]) - apply (rule prem) - apply assumption - apply (rule allI)+ - apply (tactic \eresolve_tac \<^context> [Classical.dup_elim \<^context> @{thm allE}] 1\) - apply iprover - done +proof (rule ex1E [OF major minor]) + show "\y y'. P y \ P y' \ y = y'" if "P x" and \
: "\y. P y \ y = x" for x + using \P x\ \
\
by fast +qed assumption + ML \ structure Blast = Blast @@ -1101,12 +1057,12 @@ unfolding If_def by blast lemma if_split: "P (if Q then x else y) = ((Q \ P x) \ (\ Q \ P y))" - apply (rule case_split [of Q]) - apply (simplesubst if_P) - prefer 3 - apply (simplesubst if_not_P) - apply blast+ - done +proof (rule case_split [of Q]) + show ?thesis if Q + using that by (simplesubst if_P) blast+ + show ?thesis if "\ Q" + using that by (simplesubst if_not_P) blast+ +qed lemma if_split_asm: "P (if Q then x else y) = (\ ((Q \ \ P x) \ (\ Q \ \ P y)))" by (simplesubst if_split) blast @@ -1150,20 +1106,15 @@ lemma simp_impliesI: assumes PQ: "(PROP P \ PROP Q)" shows "PROP P =simp=> PROP Q" - apply (unfold simp_implies_def) - apply (rule PQ) - apply assumption - done + unfolding simp_implies_def + by (iprover intro: PQ) lemma simp_impliesE: assumes PQ: "PROP P =simp=> PROP Q" and P: "PROP P" and QR: "PROP Q \ PROP R" shows "PROP R" - apply (rule QR) - apply (rule PQ [unfolded simp_implies_def]) - apply (rule P) - done + by (iprover intro: QR P PQ [unfolded simp_implies_def]) lemma simp_implies_cong: assumes PP' :"PROP P \ PROP P'" @@ -1658,22 +1609,32 @@ lemma ex1_eq [iff]: "\!x. x = t" "\!x. t = x" by blast+ -lemma choice_eq: "(\x. \!y. P x y) = (\!f. \x. P x (f x))" - apply (rule iffI) - apply (rule_tac a = "\x. THE y. P x y" in ex1I) - apply (fast dest!: theI') - apply (fast intro: the1_equality [symmetric]) - apply (erule ex1E) - apply (rule allI) - apply (rule ex1I) - apply (erule spec) - apply (erule_tac x = "\z. if z = x then y else f z" in allE) - apply (erule impE) - apply (rule allI) - apply (case_tac "xa = x") - apply (drule_tac [3] x = x in fun_cong) - apply simp_all - done +lemma choice_eq: "(\x. \!y. P x y) = (\!f. \x. P x (f x))" (is "?lhs = ?rhs") +proof (intro iffI allI) + assume L: ?lhs + then have \
: "\x. P x (THE y. P x y)" + by (best intro: theI') + show ?rhs + by (rule ex1I) (use L \
in \fast+\) +next + fix x + assume R: ?rhs + then obtain f where f: "\x. P x (f x)" and f1: "\y. (\x. P x (y x)) \ y = f" + by (blast elim: ex1E) + show "\!y. P x y" + proof (rule ex1I) + show "P x (f x)" + using f by blast + show "y = f x" if "P x y" for y + proof - + have "P z (if z = x then y else f z)" for z + using f that by (auto split: if_split) + with f1 [of "\z. if z = x then y else f z"] f + show ?thesis + by (auto simp add: split: if_split_asm dest: fun_cong) + qed + qed +qed lemmas eq_sym_conv = eq_commute