merged
authorwenzelm
Sat, 28 Mar 2020 21:54:31 +0100
changeset 71615 74c874b5aed0
parent 71609 beef2e221c26 (diff)
parent 71614 e6dead7d5334 (current diff)
child 71616 a9de39608b1a
child 71617 01166f13c2c0
merged
--- a/src/HOL/Equiv_Relations.thy	Sat Mar 28 19:58:19 2020 +0100
+++ b/src/HOL/Equiv_Relations.thy	Sat Mar 28 21:54:31 2020 +0100
@@ -34,21 +34,21 @@
   unfolding refl_on_def by blast
 
 lemma equiv_comp_eq: "equiv A r \<Longrightarrow> r\<inverse> 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 \<open>Second half.\<close>
 
-lemma comp_equivI: "r\<inverse> O r = r \<Longrightarrow> Domain r = A \<Longrightarrow> equiv A r"
-  apply (unfold equiv_def refl_on_def sym_def trans_def)
-  apply (erule equalityE)
-  apply (subgoal_tac "\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r")
-   apply fast
-  apply fast
-  done
+lemma comp_equivI:
+  assumes "r\<inverse> O r = r" "Domain r = A"
+  shows "equiv A r"
+proof -
+  have *: "\<And>x y. (x, y) \<in> r \<Longrightarrow> (y, x) \<in> r"
+    using assms by blast
+  show ?thesis
+    unfolding equiv_def refl_on_def sym_def trans_def
+    using assms by (auto intro: *)
+qed
 
 
 subsection \<open>Equivalence classes\<close>
@@ -58,10 +58,7 @@
   unfolding equiv_def trans_def sym_def by blast
 
 theorem equiv_class_eq: "equiv A r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> 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 \<Longrightarrow> a \<in> A \<Longrightarrow> a \<in> r``{a}"
   unfolding equiv_def refl_on_def by blast
@@ -91,7 +88,7 @@
 definition quotient :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set set"  (infixl "'/'/" 90)
   where "A//r = (\<Union>x \<in> A. {r``{x}})"  \<comment> \<open>set of equiv classes\<close>
 
-lemma quotientI: "x \<in> A ==> r``{x} \<in> A//r"
+lemma quotientI: "x \<in> A \<Longrightarrow> r``{x} \<in> A//r"
   unfolding quotient_def by blast
 
 lemma quotientE: "X \<in> A//r \<Longrightarrow> (\<And>x. X = r``{x} \<Longrightarrow> x \<in> A \<Longrightarrow> P) \<Longrightarrow> P"
@@ -101,32 +98,31 @@
   unfolding equiv_def refl_on_def quotient_def by blast
 
 lemma quotient_disj: "equiv A r \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> X = Y \<or> X \<inter> 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 \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> 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 \<in> A//r" "Y \<in> A//r" and xy: "x \<in> X" "y \<in> Y" "(x, y) \<in> r"
+  shows "X = Y"
+proof -
+  obtain a b where "a \<in> A" and a: "X = r `` {a}" and "b \<in> A" and b: "Y = r `` {b}"
+    using assms by (auto elim!: quotientE)
+  then have "(a,b) \<in> r"
+      using xy \<open>equiv A r\<close> unfolding equiv_def sym_def trans_def by blast
+  then show ?thesis
+    unfolding a b by (rule equiv_class_eq [OF \<open>equiv A r\<close>])
+qed
 
 lemma quotient_eq_iff:
-  "equiv A r \<Longrightarrow> X \<in> A//r \<Longrightarrow> Y \<in> A//r \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> X = Y \<longleftrightarrow> (x, y) \<in> 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 \<in> A//r" "Y \<in> A//r" and xy: "x \<in> X" "y \<in> Y" 
+  shows "X = Y \<longleftrightarrow> (x, y) \<in> r"
+proof
+  assume L: "X = Y" 
+  with assms show "(x, y) \<in> r" 
+    unfolding equiv_def sym_def trans_def by (blast elim!: quotientE)
+next
+  assume \<section>: "(x, y) \<in> r" show "X = Y"
+    by (rule quotient_eqI) (use \<section> assms in \<open>blast+\<close>)
+qed
 
 lemma eq_equiv_class_iff2: "equiv A r \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> {x}//r = {y}//r \<longleftrightarrow> (x, y) \<in> r"
   by (simp add: quotient_def eq_equiv_class_iff)
@@ -189,22 +185,22 @@
   \<comment> \<open>lemma required to prove \<open>UN_equiv_class\<close>\<close>
   by auto
 
-lemma UN_equiv_class: "equiv A r \<Longrightarrow> f respects r \<Longrightarrow> a \<in> A \<Longrightarrow> (\<Union>x \<in> r``{a}. f x) = f a"
+lemma UN_equiv_class:
+  assumes "equiv A r" "f respects r" "a \<in> A"
+  shows "(\<Union>x \<in> r``{a}. f x) = f a"
   \<comment> \<open>Conversion rule\<close>
-  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 \<section>: "\<forall>x\<in>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 \<section>])
+qed
 
 lemma UN_equiv_class_type:
-  "equiv A r \<Longrightarrow> f respects r \<Longrightarrow> X \<in> A//r \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> (\<Union>x \<in> X. f x) \<in> 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 \<in> A//r" and AB: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B"
+  shows "(\<Union>x \<in> X. f x) \<in> B"
+  using assms unfolding quotient_def
+  by (auto simp: UN_equiv_class [OF r])
 
 text \<open>
   Sufficient conditions for injectiveness.  Could weaken premises!
@@ -213,19 +209,23 @@
 \<close>
 
 lemma UN_equiv_class_inject:
-  "equiv A r \<Longrightarrow> f respects r \<Longrightarrow>
-    (\<Union>x \<in> X. f x) = (\<Union>y \<in> Y. f y) \<Longrightarrow> X \<in> A//r ==> Y \<in> A//r
-    \<Longrightarrow> (\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> (x, y) \<in> r)
-    \<Longrightarrow> 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: "(\<Union>x \<in> X. f x) = (\<Union>y \<in> Y. f y)" 
+    and X: "X \<in> A//r" and Y: "Y \<in> A//r" 
+    and fr: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f x = f y \<Longrightarrow> (x, y) \<in> r"
+  shows "X = Y"
+proof -
+  obtain a b where "a \<in> A" and a: "X = r `` {a}" and "b \<in> A" and b: "Y = r `` {b}"
+    using assms by (auto elim!: quotientE)
+  then have "\<Union> (f ` r `` {a}) = f a" "\<Union> (f ` r `` {b}) = f b"
+    by (iprover intro: UN_equiv_class [OF \<open>equiv A r\<close>] assms)+
+  then have "f a = f b"
+    using eq unfolding a b by (iprover intro: trans sym)
+  then have "(a,b) \<in> r"
+    using fr \<open>a \<in> A\<close> \<open>b \<in> A\<close> by blast
+  then show ?thesis
+    unfolding a b by (rule equiv_class_eq [OF \<open>equiv A r\<close>])
+qed
 
 
 subsection \<open>Defining binary operations upon equivalence classes\<close>
@@ -253,15 +253,20 @@
   unfolding congruent_def congruent2_def equiv_def refl_on_def by blast
 
 lemma congruent2_implies_congruent_UN:
-  "equiv A1 r1 \<Longrightarrow> equiv A2 r2 \<Longrightarrow> congruent2 r1 r2 f \<Longrightarrow> a \<in> A2 \<Longrightarrow>
-    congruent r1 (\<lambda>x1. \<Union>x2 \<in> 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 \<in> A2" 
+  shows "congruent r1 (\<lambda>x1. \<Union>x2 \<in> r2``{a}. f x1 x2)"
+  unfolding congruent_def
+proof clarify
+  fix c d
+  assume cd: "(c,d) \<in> r1"
+  then have "c \<in> A1" "d \<in> A1"
+    using \<open>equiv A1 r1\<close> by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2])
+  with assms show "\<Union> (f c ` r2 `` {a}) = \<Union> (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 \<Longrightarrow> equiv A2 r2 \<Longrightarrow> congruent2 r1 r2 f \<Longrightarrow> a1 \<in> A1 \<Longrightarrow> a2 \<in> A2 \<Longrightarrow>
@@ -273,11 +278,10 @@
     \<Longrightarrow> X1 \<in> A1//r1 \<Longrightarrow> X2 \<in> A2//r2
     \<Longrightarrow> (\<And>x1 x2. x1 \<in> A1 \<Longrightarrow> x2 \<in> A2 \<Longrightarrow> f x1 x2 \<in> B)
     \<Longrightarrow> (\<Union>x1 \<in> X1. \<Union>x2 \<in> X2. f x1 x2) \<in> 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:
   "(\<Union>(x1, x2) \<in> X. \<Union>(y1, y2) \<in> Y. A x1 x2 y1 y2) =
@@ -293,60 +297,63 @@
     \<Longrightarrow> congruent2 r1 r2 f"
   \<comment> \<open>Suggested by John Harrison -- the two subproofs may be\<close>
   \<comment> \<open>\<^emph>\<open>much\<close> simpler than the direct proof.\<close>
-  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: "\<And>y z. y \<in> A \<Longrightarrow> z \<in> A \<Longrightarrow> f y z = f z y"
     and congt: "\<And>y z w. w \<in> A \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> 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 "\<And>y z w. \<lbrakk>w \<in> A; (y, z) \<in> r\<rbrakk> \<Longrightarrow> f y w = f z w"
+    by (iprover intro: commute [THEN trans] sym congt elim: eqv)
+  show "\<And>y z w. \<lbrakk>w \<in> A; (y, z) \<in> r\<rbrakk> \<Longrightarrow> f w y = f w z"
+    by (iprover intro: congt elim: eqv)
+qed
 
 
 subsection \<open>Quotients and finiteness\<close>
 
 text \<open>Suggested by Florian Kammüller\<close>
 
-lemma finite_quotient: "finite A \<Longrightarrow> r \<subseteq> A \<times> A \<Longrightarrow> finite (A//r)"
-  \<comment> \<open>recall @{thm equiv_type}\<close>
-  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 \<subseteq> A \<times> A"
+  shows "finite (A//r)"
+    \<comment> \<open>recall @{thm equiv_type}\<close>
+proof -
+  have "A//r \<subseteq> 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 \<Longrightarrow> r \<subseteq> A \<times> A \<Longrightarrow> X \<in> A//r \<Longrightarrow> 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 \<Longrightarrow> equiv A r \<Longrightarrow> \<forall>X \<in> A//r. k dvd card X \<Longrightarrow> k dvd card A"
-  apply (rule Union_quotient [THEN subst [where P="\<lambda>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" "\<And>X. X \<in> A//r \<Longrightarrow> k dvd card X"
+  shows "k dvd card A"
+proof (rule Union_quotient [THEN subst])
+  show "k dvd card (\<Union> (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 \<Longrightarrow> inj_on (\<lambda>x. {x} // r) A \<Longrightarrow> 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 (\<lambda>x. {x} // r) A"
+  shows "card (A//r) = card A"
+proof -
+  have "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> r `` {j} \<noteq> 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 \<open>Projection\<close>
--- a/src/HOL/HOL.thy	Sat Mar 28 19:58:19 2020 +0100
+++ b/src/HOL/HOL.thy	Sat Mar 28 21:54:31 2020 +0100
@@ -249,11 +249,7 @@
         |   |
         c = d   *)
 lemma box_equals: "\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d"
-  apply (rule trans)
-   apply (rule trans)
-    apply (rule sym)
-    apply assumption+
-  done
+  by (iprover intro: sym trans)
 
 text \<open>For calculational reasoning:\<close>
 
@@ -268,25 +264,17 @@
 
 text \<open>Similar to \<open>AP_THM\<close> in Gordon's HOL.\<close>
 lemma fun_cong: "(f :: 'a \<Rightarrow> 'b) = g \<Longrightarrow> f x = g x"
-  apply (erule subst)
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 text \<open>Similar to \<open>AP_TERM\<close> in Gordon's HOL and FOL's \<open>subst_context\<close>.\<close>
 lemma arg_cong: "x = y \<Longrightarrow> f x = f y"
-  apply (erule subst)
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 lemma arg_cong2: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> f a c = f b d"
-  apply (erule ssubst)+
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 lemma cong: "\<lbrakk>f = g; (x::'a) = y\<rbrakk> \<Longrightarrow> f x = g y"
-  apply (erule subst)+
-  apply (rule refl)
-  done
+  by (iprover intro: refl elim: subst)
 
 ML \<open>fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\<close>
 
@@ -324,20 +312,15 @@
 subsubsection \<open>Universal quantifier (1)\<close>
 
 lemma spec: "\<forall>x::'a. P x \<Longrightarrow> 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: "\<forall>x. P x"
-    and minor: "P x \<Longrightarrow> R"
+  assumes major: "\<forall>x. P x" and minor: "P x \<Longrightarrow> R"
   shows R
   by (iprover intro: minor major [THEN spec])
 
 lemma all_dupE:
-  assumes major: "\<forall>x. P x"
-    and minor: "\<lbrakk>P x; \<forall>x. P x\<rbrakk> \<Longrightarrow> R"
+  assumes major: "\<forall>x. P x" and minor: "\<lbrakk>P x; \<forall>x. P x\<rbrakk> \<Longrightarrow> R"
   shows R
   by (iprover intro: minor major major [THEN spec])
 
@@ -350,9 +333,7 @@
 \<close>
 
 lemma FalseE: "False \<Longrightarrow> P"
-  apply (unfold False_def)
-  apply (erule spec)
-  done
+  unfolding False_def by (erule spec)
 
 lemma False_neq_True: "False = True \<Longrightarrow> P"
   by (erule eqTrueE [THEN FalseE])
@@ -363,29 +344,17 @@
 lemma notI:
   assumes "P \<Longrightarrow> False"
   shows "\<not> P"
-  apply (unfold not_def)
-  apply (iprover intro: impI assms)
-  done
+  unfolding not_def by (iprover intro: impI assms)
 
 lemma False_not_True: "False \<noteq> True"
-  apply (rule notI)
-  apply (erule False_neq_True)
-  done
+  by (iprover intro: notI elim: False_neq_True)
 
 lemma True_not_False: "True \<noteq> 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: "\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R"
-  apply (unfold not_def)
-  apply (erule mp [THEN FalseE])
-  apply assumption
-  done
-
-lemma notI2: "(P \<Longrightarrow> \<not> Pa) \<Longrightarrow> (P \<Longrightarrow> Pa) \<Longrightarrow> \<not> P"
-  by (erule notE [THEN notI]) (erule meta_mp)
+  unfolding not_def
+  by (iprover intro: mp [THEN FalseE])
 
 
 subsubsection \<open>Implication\<close>
@@ -397,7 +366,7 @@
 
 text \<open>Reduces \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, allowing substitution in \<open>P\<close>.\<close>
 lemma rev_mp: "\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
-  by (iprover intro: mp)
+  by (rule mp)
 
 lemma contrapos_nn:
   assumes major: "\<not> Q"
@@ -510,10 +479,10 @@
   assumes major: "P \<and> Q"
     and minor: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> Q"
@@ -533,14 +502,18 @@
 subsubsection \<open>Classical logic\<close>
 
 lemma classical:
-  assumes prem: "\<not> P \<Longrightarrow> P"
+  assumes "\<not> P \<Longrightarrow> 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: "\<not> R \<Longrightarrow> \<not> P"
   shows R
-  apply (rule ccontr)
-  apply (erule notE [OF premnot premp])
-  done
+  by (iprover intro: ccontr notE [OF premnot premp])
+
 
 text \<open>Double negation law.\<close>
 lemma notnotD: "\<not>\<not> P \<Longrightarrow> 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: "\<exists>!x. P x"
-    and minor: "\<And>x. \<lbrakk>P x; \<forall>y. P y \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R"
+  assumes major: "\<exists>!x. P x" and minor: "\<And>x. \<lbrakk>P x; \<forall>y. P y \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> 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 "\<And>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x) \<Longrightarrow> R"
+    by (iprover intro: minor elim: conjE)
+qed
 
 lemma ex1_implies_ex: "\<exists>!x. P x \<Longrightarrow> \<exists>x. P x"
-  apply (erule ex1E)
-  apply (rule exI)
-  apply assumption
-  done
+  by (iprover intro: exI elim: ex1E)
 
 
 subsubsection \<open>Classical intro rules for disjunction and existential quantifiers\<close>
@@ -613,22 +578,18 @@
   Note that \<open>\<not> P\<close> is the second case, not the first.
 \<close>
 lemma case_split [case_names True False]:
-  assumes prem1: "P \<Longrightarrow> Q"
-    and prem2: "\<not> P \<Longrightarrow> Q"
+  assumes "P \<Longrightarrow> Q" "\<not> P \<Longrightarrow> 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 \<open>Classical implies (\<open>\<longrightarrow>\<close>) elimination.\<close>
 lemma impCE:
   assumes major: "P \<longrightarrow> Q"
     and minor: "\<not> P \<Longrightarrow> R" "Q \<Longrightarrow> 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 \<open>
   This version of \<open>\<longrightarrow>\<close> elimination works on \<open>Q\<close> before \<open>P\<close>.  It works best for
@@ -639,9 +600,8 @@
   assumes major: "P \<longrightarrow> Q"
     and minor: "Q \<Longrightarrow> R" "\<not> P \<Longrightarrow> 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 \<open>Classical \<open>\<longleftrightarrow>\<close> elimination.\<close>
 lemma iffCE:
@@ -874,9 +834,7 @@
 ML \<open>val HOL_cs = claset_of \<^context>\<close>
 
 lemma contrapos_np: "\<not> Q \<Longrightarrow> (\<not> P \<Longrightarrow> Q) \<Longrightarrow> 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 \<open>Better than \<open>ex1E\<close> for classical reasoner: needs no quantifier duplication!\<close>
 lemma alt_ex1E [elim!]:
   assumes major: "\<exists>!x. P x"
-    and prem: "\<And>x. \<lbrakk>P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R"
+    and minor: "\<And>x. \<lbrakk>P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R"
   shows R
-  apply (rule ex1E [OF major])
-  apply (rule prem)
-   apply assumption
-  apply (rule allI)+
-  apply (tactic \<open>eresolve_tac \<^context> [Classical.dup_elim \<^context> @{thm allE}] 1\<close>)
-  apply iprover
-  done
+proof (rule ex1E [OF major minor])
+  show "\<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'" if "P x" and \<section>: "\<forall>y. P y \<longrightarrow> y = x" for x
+    using \<open>P x\<close> \<section> \<section> by fast
+qed assumption
+
 
 ML \<open>
   structure Blast = Blast
@@ -1101,12 +1057,12 @@
   unfolding If_def by blast
 
 lemma if_split: "P (if Q then x else y) = ((Q \<longrightarrow> P x) \<and> (\<not> Q \<longrightarrow> 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 "\<not> Q"
+    using that by (simplesubst if_not_P) blast+
+qed
 
 lemma if_split_asm: "P (if Q then x else y) = (\<not> ((Q \<and> \<not> P x) \<or> (\<not> Q \<and> \<not> P y)))"
   by (simplesubst if_split) blast
@@ -1150,20 +1106,15 @@
 lemma simp_impliesI:
   assumes PQ: "(PROP P \<Longrightarrow> 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 \<Longrightarrow> 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 \<equiv> PROP P'"
@@ -1658,22 +1609,32 @@
 lemma ex1_eq [iff]: "\<exists>!x. x = t" "\<exists>!x. t = x"
   by blast+
 
-lemma choice_eq: "(\<forall>x. \<exists>!y. P x y) = (\<exists>!f. \<forall>x. P x (f x))"
-  apply (rule iffI)
-   apply (rule_tac a = "\<lambda>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 = "\<lambda>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: "(\<forall>x. \<exists>!y. P x y) = (\<exists>!f. \<forall>x. P x (f x))" (is "?lhs = ?rhs")
+proof (intro iffI allI)
+  assume L: ?lhs
+  then have \<section>: "\<forall>x. P x (THE y. P x y)"
+    by (best intro: theI')
+  show ?rhs
+    by (rule ex1I) (use L \<section> in \<open>fast+\<close>)
+next
+  fix x
+  assume R: ?rhs
+  then obtain f where f: "\<forall>x. P x (f x)" and f1: "\<And>y. (\<forall>x. P x (y x)) \<Longrightarrow> y = f"
+    by (blast elim: ex1E)
+  show "\<exists>!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 "\<lambda>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