more symbols;
authorwenzelm
Sun, 26 Nov 2017 21:08:32 +0100
changeset 67091 1393c2340eec
parent 67090 0ec94bb9cec4
child 67092 d7b3876d3ab1
more symbols;
src/HOL/Algebra/AbelCoset.thy
src/HOL/Algebra/Bij.thy
src/HOL/Algebra/Complete_Lattice.thy
src/HOL/Algebra/Congruence.thy
src/HOL/Algebra/Coset.thy
src/HOL/Algebra/FiniteProduct.thy
src/HOL/Algebra/Group.thy
src/HOL/Algebra/IntRing.thy
src/HOL/Algebra/Lattice.thy
src/HOL/Algebra/Order.thy
src/HOL/Algebra/Ring.thy
src/HOL/Algebra/UnivPoly.thy
src/HOL/BNF_Cardinal_Order_Relation.thy
src/HOL/BNF_Composition.thy
src/HOL/BNF_Def.thy
src/HOL/BNF_Fixpoint_Base.thy
src/HOL/BNF_Greatest_Fixpoint.thy
src/HOL/BNF_Least_Fixpoint.thy
src/HOL/BNF_Wellorder_Constructions.thy
src/HOL/BNF_Wellorder_Embedding.thy
src/HOL/Basic_BNFs.thy
src/HOL/Computational_Algebra/Polynomial.thy
src/HOL/Computational_Algebra/Primes.thy
src/HOL/Conditionally_Complete_Lattices.thy
src/HOL/Divides.thy
src/HOL/Enum.thy
src/HOL/Fields.thy
src/HOL/Groebner_Basis.thy
src/HOL/HOL.thy
src/HOL/Library/BNF_Corec.thy
src/HOL/Library/Cardinality.thy
src/HOL/Library/DAList.thy
src/HOL/Library/Diagonal_Subsequence.thy
src/HOL/Library/Extended.thy
src/HOL/Library/Extended_Nat.thy
src/HOL/Library/Extended_Real.thy
src/HOL/Library/FSet.thy
src/HOL/Library/Old_Datatype.thy
src/HOL/Library/Old_Recdef.thy
src/HOL/Library/Option_ord.thy
src/HOL/Library/Predicate_Compile_Alternative_Defs.thy
src/HOL/Library/Product_Order.thy
src/HOL/Library/RBT_Set.thy
src/HOL/Library/Ramsey.thy
src/HOL/Library/Stream.thy
src/HOL/Library/Sublist.thy
src/HOL/Library/Sum_of_Squares/sum_of_squares.ML
src/HOL/Library/While_Combinator.thy
src/HOL/Library/positivstellensatz.ML
src/HOL/Limited_Sequence.thy
src/HOL/Limits.thy
src/HOL/List.thy
src/HOL/MacLaurin.thy
src/HOL/Map.thy
src/HOL/Meson.thy
src/HOL/Nat.thy
src/HOL/Nitpick.thy
src/HOL/Nonstandard_Analysis/HTranscendental.thy
src/HOL/Nonstandard_Analysis/HyperNat.thy
src/HOL/Nonstandard_Analysis/NSCA.thy
src/HOL/Nonstandard_Analysis/NatStar.thy
src/HOL/Number_Theory/Euler_Criterion.thy
src/HOL/Number_Theory/Gauss.thy
src/HOL/Number_Theory/Pocklington.thy
src/HOL/Number_Theory/Residues.thy
src/HOL/Numeral_Simprocs.thy
src/HOL/Option.thy
src/HOL/Orderings.thy
src/HOL/Partial_Function.thy
src/HOL/Predicate.thy
src/HOL/Presburger.thy
src/HOL/Quickcheck_Exhaustive.thy
src/HOL/Quickcheck_Narrowing.thy
src/HOL/Quotient.thy
src/HOL/Random_Pred.thy
src/HOL/Random_Sequence.thy
src/HOL/Record.thy
src/HOL/SMT.thy
src/HOL/Set.thy
src/HOL/Set_Interval.thy
src/HOL/String.thy
src/HOL/Tools/ATP/atp_problem_generate.ML
src/HOL/Tools/ATP/atp_proof_reconstruct.ML
src/HOL/Tools/Argo/argo_real.ML
src/HOL/Tools/Argo/argo_tactic.ML
src/HOL/Tools/BNF/bnf_comp.ML
src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML
src/HOL/Tools/BNF/bnf_fp_util.ML
src/HOL/Tools/BNF/bnf_gfp.ML
src/HOL/Tools/BNF/bnf_gfp_tactics.ML
src/HOL/Tools/BNF/bnf_lfp.ML
src/HOL/Tools/BNF/bnf_lift.ML
src/HOL/Tools/Meson/meson.ML
src/HOL/Tools/Meson/meson_clausify.ML
src/HOL/Tools/Metis/metis_reconstruct.ML
src/HOL/Tools/Old_Datatype/old_datatype.ML
src/HOL/Tools/Quotient/quotient_tacs.ML
src/HOL/Tools/SMT/conj_disj_perm.ML
src/HOL/Tools/SMT/z3_replay.ML
src/HOL/Tools/SMT/z3_replay_methods.ML
src/HOL/Tools/cnf.ML
src/HOL/Tools/inductive.ML
src/HOL/Tools/rewrite_hol_proof.ML
src/HOL/Tools/sat.ML
src/HOL/Transcendental.thy
--- a/src/HOL/Algebra/AbelCoset.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/AbelCoset.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -290,7 +290,7 @@
 text\<open>Alternative characterization of normal subgroups\<close>
 lemma (in abelian_group) a_normal_inv_iff:
      "(N \<lhd> \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>) = 
-      (subgroup N \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> & (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<oplus> h \<oplus> (\<ominus> x) \<in> N))"
+      (subgroup N \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> \<and> (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<oplus> h \<oplus> (\<ominus> x) \<in> N))"
       (is "_ = ?rhs")
 by (rule group.normal_inv_iff [OF a_group,
     folded a_inv_def, simplified monoid_record_simps])
--- a/src/HOL/Algebra/Bij.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Bij.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -67,7 +67,7 @@
  shows "\<lbrakk>h \<in> Bij S;  g \<in> S \<rightarrow> S \<rightarrow> S;  x \<in> S;  y \<in> S\<rbrakk>
         \<Longrightarrow> 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 "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' & y = h y'", clarify)
+apply (subgoal_tac "\<exists>x'\<in>S. \<exists>y'\<in>S. x = h x' \<and> y = h y'", clarify)
  apply (simp add: eq [symmetric] inv_f_f funcset_mem [THEN funcset_mem], blast)
 done
 
--- a/src/HOL/Algebra/Complete_Lattice.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Complete_Lattice.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -14,9 +14,9 @@
 
 locale weak_complete_lattice = weak_partial_order +
   assumes sup_exists:
-    "[| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
+    "[| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
     and inf_exists:
-    "[| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
+    "[| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
 
 sublocale weak_complete_lattice \<subseteq> weak_lattice
 proof
@@ -32,9 +32,9 @@
 
 lemma (in weak_partial_order) weak_complete_latticeI:
   assumes sup_exists:
-    "!!A. [| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
+    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
     and inf_exists:
-    "!!A. [| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
+    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
   shows "weak_complete_lattice L"
   by standard (auto intro: sup_exists inf_exists)
 
@@ -111,9 +111,9 @@
 qed
 
 theorem (in weak_partial_order) weak_complete_lattice_criterion1:
-  assumes top_exists: "EX g. greatest L g (carrier L)"
+  assumes top_exists: "\<exists>g. greatest L g (carrier L)"
     and inf_exists:
-      "!!A. [| A \<subseteq> carrier L; A ~= {} |] ==> EX i. greatest L i (Lower L A)"
+      "\<And>A. [| A \<subseteq> carrier L; A \<noteq> {} |] ==> \<exists>i. greatest L i (Lower L A)"
   shows "weak_complete_lattice L"
 proof (rule weak_complete_latticeI)
   from top_exists obtain top where top: "greatest L top (carrier L)" ..
@@ -121,7 +121,7 @@
   assume L: "A \<subseteq> carrier L"
   let ?B = "Upper L A"
   from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
-  then have B_non_empty: "?B ~= {}" by fast
+  then have B_non_empty: "?B \<noteq> {}" by fast
   have B_L: "?B \<subseteq> carrier L" by simp
   from inf_exists [OF B_L B_non_empty]
   obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
@@ -139,11 +139,11 @@
  apply (rule L)
 apply (rule greatest_closed [OF b_inf_B])
 done
-  then show "EX s. least L s (Upper L A)" ..
+  then show "\<exists>s. least L s (Upper L A)" ..
 next
   fix A
   assume L: "A \<subseteq> carrier L"
-  show "EX i. greatest L i (Lower L A)"
+  show "\<exists>i. greatest L i (Lower L A)"
   proof (cases "A = {}")
     case True then show ?thesis
       by (simp add: top_exists)
@@ -547,9 +547,9 @@
 
 locale complete_lattice = partial_order +
   assumes sup_exists:
-    "[| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
+    "[| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
     and inf_exists:
-    "[| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
+    "[| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
 
 sublocale complete_lattice \<subseteq> lattice
 proof
@@ -578,16 +578,16 @@
 
 lemma (in partial_order) complete_latticeI:
   assumes sup_exists:
-    "!!A. [| A \<subseteq> carrier L |] ==> EX s. least L s (Upper L A)"
+    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>s. least L s (Upper L A)"
     and inf_exists:
-    "!!A. [| A \<subseteq> carrier L |] ==> EX i. greatest L i (Lower L A)"
+    "!!A. [| A \<subseteq> carrier L |] ==> \<exists>i. greatest L i (Lower L A)"
   shows "complete_lattice L"
   by standard (auto intro: sup_exists inf_exists)
 
 theorem (in partial_order) complete_lattice_criterion1:
-  assumes top_exists: "EX g. greatest L g (carrier L)"
+  assumes top_exists: "\<exists>g. greatest L g (carrier L)"
     and inf_exists:
-      "!!A. [| A \<subseteq> carrier L; A ~= {} |] ==> EX i. greatest L i (Lower L A)"
+      "!!A. [| A \<subseteq> carrier L; A \<noteq> {} |] ==> \<exists>i. greatest L i (Lower L A)"
   shows "complete_lattice L"
 proof (rule complete_latticeI)
   from top_exists obtain top where top: "greatest L top (carrier L)" ..
@@ -595,7 +595,7 @@
   assume L: "A \<subseteq> carrier L"
   let ?B = "Upper L A"
   from L top have "top \<in> ?B" by (fast intro!: Upper_memI intro: greatest_le)
-  then have B_non_empty: "?B ~= {}" by fast
+  then have B_non_empty: "?B \<noteq> {}" by fast
   have B_L: "?B \<subseteq> carrier L" by simp
   from inf_exists [OF B_L B_non_empty]
   obtain b where b_inf_B: "greatest L b (Lower L ?B)" ..
@@ -613,11 +613,11 @@
  apply (rule L)
 apply (rule greatest_closed [OF b_inf_B])
 done
-  then show "EX s. least L s (Upper L A)" ..
+  then show "\<exists>s. least L s (Upper L A)" ..
 next
   fix A
   assume L: "A \<subseteq> carrier L"
-  show "EX i. greatest L i (Lower L A)"
+  show "\<exists>i. greatest L i (Lower L A)"
   proof (cases "A = {}")
     case True then show ?thesis
       by (simp add: top_exists)
@@ -1181,7 +1181,7 @@
   assume "B \<subseteq> carrier ?L"
   then have "least ?L (\<Union> B) (Upper ?L B)"
     by (fastforce intro!: least_UpperI simp: Upper_def)
-  then show "EX s. least ?L s (Upper ?L B)" ..
+  then show "\<exists>s. least ?L s (Upper ?L B)" ..
 next
   fix B
   assume "B \<subseteq> carrier ?L"
@@ -1189,7 +1189,7 @@
     txt \<open>@{term "\<Inter> B"} is not the infimum of @{term B}:
       @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! \<close>
     by (fastforce intro!: greatest_LowerI simp: Lower_def)
-  then show "EX i. greatest ?L i (Lower ?L B)" ..
+  then show "\<exists>i. greatest ?L i (Lower ?L B)" ..
 qed
 
 text \<open>Another example, that of the lattice of subgroups of a group,
--- a/src/HOL/Algebra/Congruence.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Congruence.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -52,15 +52,15 @@
 
 abbreviation
   not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".\<noteq>\<index>" 50)
-  where "x .\<noteq>\<^bsub>S\<^esub> y == ~(x .=\<^bsub>S\<^esub> y)"
+  where "x .\<noteq>\<^bsub>S\<^esub> y \<equiv> \<not>(x .=\<^bsub>S\<^esub> y)"
 
 abbreviation
   not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
-  where "x .\<notin>\<^bsub>S\<^esub> A == ~(x .\<in>\<^bsub>S\<^esub> A)"
+  where "x .\<notin>\<^bsub>S\<^esub> A \<equiv> \<not>(x .\<in>\<^bsub>S\<^esub> A)"
 
 abbreviation
   set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
-  where "A {.\<noteq>}\<^bsub>S\<^esub> B == ~(A {.=}\<^bsub>S\<^esub> B)"
+  where "A {.\<noteq>}\<^bsub>S\<^esub> B \<equiv> \<not>(A {.=}\<^bsub>S\<^esub> B)"
 
 locale equivalence =
   fixes S (structure)
--- a/src/HOL/Algebra/Coset.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Coset.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -410,7 +410,7 @@
 text\<open>Alternative characterization of normal subgroups\<close>
 lemma (in group) normal_inv_iff:
      "(N \<lhd> G) = 
-      (subgroup N G & (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<otimes> h \<otimes> (inv x) \<in> N))"
+      (subgroup N G \<and> (\<forall>x \<in> carrier G. \<forall>h \<in> N. x \<otimes> h \<otimes> (inv x) \<in> N))"
       (is "_ = ?rhs")
 proof
   assume N: "N \<lhd> G"
@@ -476,7 +476,7 @@
   assume "\<exists>h\<in>H. y = x \<otimes> h"
     and x: "x \<in> carrier G"
     and sb: "subgroup H G"
-  then obtain h' where h': "h' \<in> H & x \<otimes> h' = y" by blast
+  then obtain h' where h': "h' \<in> H \<and> x \<otimes> h' = y" by blast
   show "\<exists>h\<in>H. x = y \<otimes> h"
   proof
     show "x = y \<otimes> inv h'" using h' x sb
@@ -593,7 +593,7 @@
 
 definition
   r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set"  ("rcong\<index> _")
-  where "rcong\<^bsub>G\<^esub> H = {(x,y). x \<in> carrier G & y \<in> carrier G & inv\<^bsub>G\<^esub> x \<otimes>\<^bsub>G\<^esub> y \<in> H}"
+  where "rcong\<^bsub>G\<^esub> H = {(x,y). x \<in> carrier G \<and> y \<in> carrier G \<and> inv\<^bsub>G\<^esub> x \<otimes>\<^bsub>G\<^esub> y \<in> H}"
 
 
 lemma (in subgroup) equiv_rcong:
@@ -784,7 +784,7 @@
 lemma (in group) inj_on_f:
     "\<lbrakk>H \<subseteq> carrier G;  a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> inv a) (H #> a)"
 apply (rule inj_onI)
-apply (subgoal_tac "x \<in> carrier G & y \<in> carrier G")
+apply (subgoal_tac "x \<in> carrier G \<and> y \<in> carrier G")
  prefer 2 apply (blast intro: r_coset_subset_G [THEN subsetD])
 apply (simp add: subsetD)
 done
@@ -898,7 +898,7 @@
 definition
   kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow>  ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
     \<comment>\<open>the kernel of a homomorphism\<close>
-  where "kernel G H h = {x. x \<in> carrier G & h x = \<one>\<^bsub>H\<^esub>}"
+  where "kernel G H h = {x. x \<in> carrier G \<and> h x = \<one>\<^bsub>H\<^esub>}"
 
 lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G"
 apply (rule subgroup.intro) 
--- a/src/HOL/Algebra/FiniteProduct.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/FiniteProduct.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -48,7 +48,7 @@
 
 lemma finite_imp_foldSetD:
   "[| finite A; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D |] ==>
-   EX x. (A, x) \<in> foldSetD D f e"
+   \<exists>x. (A, x) \<in> foldSetD D f e"
 proof (induct set: finite)
   case empty then show ?case by auto
 next
@@ -89,7 +89,7 @@
   by (induct set: foldSetD) auto
 
 lemma (in LCD) finite_imp_foldSetD:
-  "[| finite A; A \<subseteq> B; e \<in> D |] ==> EX x. (A, x) \<in> foldSetD D f e"
+  "[| finite A; A \<subseteq> B; e \<in> D |] ==> \<exists>x. (A, x) \<in> foldSetD D f e"
 proof (induct set: finite)
   case empty then show ?case by auto
 next
@@ -102,8 +102,8 @@
 qed
 
 lemma (in LCD) foldSetD_determ_aux:
-  "e \<in> D ==> \<forall>A x. A \<subseteq> B & card A < n --> (A, x) \<in> foldSetD D f e -->
-    (\<forall>y. (A, y) \<in> foldSetD D f e --> y = x)"
+  "e \<in> D \<Longrightarrow> \<forall>A x. A \<subseteq> B \<and> card A < n \<longrightarrow> (A, x) \<in> foldSetD D f e \<longrightarrow>
+    (\<forall>y. (A, y) \<in> foldSetD D f e \<longrightarrow> y = x)"
   apply (induct n)
    apply (auto simp add: less_Suc_eq) (* slow *)
   apply (erule foldSetD.cases)
@@ -120,7 +120,7 @@
     prefer 2 apply (blast elim!: equalityE)
    apply blast
   txt \<open>case @{prop "xa \<notin> xb"}.\<close>
-  apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb \<in> Aa & xa \<in> Ab")
+  apply (subgoal_tac "Aa - {xb} = Ab - {xa} \<and> xb \<in> Aa \<and> xa \<in> Ab")
    prefer 2 apply (blast elim!: equalityE)
   apply clarify
   apply (subgoal_tac "Aa = insert xb Ab - {xa}")
@@ -166,9 +166,9 @@
   by (unfold foldD_def) blast
 
 lemma (in LCD) foldD_insert_aux:
-  "[| x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
+  "[| x \<notin> A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
     ((insert x A, v) \<in> foldSetD D f e) =
-    (EX y. (A, y) \<in> foldSetD D f e & v = f x y)"
+    (\<exists>y. (A, y) \<in> foldSetD D f e \<and> v = f x y)"
   apply auto
   apply (rule_tac A1 = A in finite_imp_foldSetD [THEN exE])
      apply (fastforce dest: foldSetD_imp_finite)
@@ -291,7 +291,7 @@
   finprod :: "[('b, 'm) monoid_scheme, 'a => 'b, 'a set] => 'b"
   where "finprod G f A =
    (if finite A
-    then foldD (carrier G) (mult G o f) \<one>\<^bsub>G\<^esub> A
+    then foldD (carrier G) (mult G \<circ> f) \<one>\<^bsub>G\<^esub> A
     else \<one>\<^bsub>G\<^esub>)"
 
 syntax
@@ -360,7 +360,7 @@
   by fast
 
 lemma funcset_Un_left [iff]:
-  "(f \<in> A Un B \<rightarrow> C) = (f \<in> A \<rightarrow> C & f \<in> B \<rightarrow> C)"
+  "(f \<in> A Un B \<rightarrow> C) = (f \<in> A \<rightarrow> C \<and> f \<in> B \<rightarrow> C)"
   by fast
 
 lemma finprod_Un_Int:
--- a/src/HOL/Algebra/Group.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Group.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -22,12 +22,12 @@
 
 definition
   m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\<index> _" [81] 80)
-  where "inv\<^bsub>G\<^esub> x = (THE y. y \<in> carrier G & x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> & y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)"
+  where "inv\<^bsub>G\<^esub> x = (THE y. y \<in> carrier G \<and> x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> \<and> y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)"
 
 definition
   Units :: "_ => 'a set"
   \<comment>\<open>The set of invertible elements\<close>
-  where "Units G = {y. y \<in> carrier G & (\<exists>x \<in> carrier G. x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> & y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)}"
+  where "Units G = {y. y \<in> carrier G \<and> (\<exists>x \<in> carrier G. x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> \<and> y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)}"
 
 consts
   pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a"  (infixr "'(^')\<index>" 75)
@@ -277,7 +277,7 @@
     with x xG show "x \<otimes> \<one> = x" by simp
   qed
   have inv_ex:
-    "!!x. x \<in> carrier G ==> \<exists>y \<in> carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>"
+    "\<And>x. x \<in> carrier G \<Longrightarrow> \<exists>y \<in> carrier G. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>"
   proof -
     fix x
     assume x: "x \<in> carrier G"
@@ -287,10 +287,10 @@
       by (simp add: m_assoc [symmetric] l_inv r_one)
     with x y have r_inv: "x \<otimes> y = \<one>"
       by simp
-    from x y show "\<exists>y \<in> carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>"
+    from x y show "\<exists>y \<in> carrier G. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>"
       by (fast intro: l_inv r_inv)
   qed
-  then have carrier_subset_Units: "carrier G <= Units G"
+  then have carrier_subset_Units: "carrier G \<subseteq> Units G"
     by (unfold Units_def) fast
   show ?thesis
     by standard (auto simp: r_one m_assoc carrier_subset_Units)
@@ -305,9 +305,9 @@
 lemma (in group) Units_eq [simp]:
   "Units G = carrier G"
 proof
-  show "Units G <= carrier G" by fast
+  show "Units G \<subseteq> carrier G" by fast
 next
-  show "carrier G <= Units G" by (rule Units)
+  show "carrier G \<subseteq> Units G" by (rule Units)
 qed
 
 lemma (in group) inv_closed [intro, simp]:
@@ -510,13 +510,13 @@
   monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp]
 
 lemma subgroup_nonempty:
-  "~ subgroup {} G"
+  "\<not> subgroup {} G"
   by (blast dest: subgroup.one_closed)
 
 lemma (in subgroup) finite_imp_card_positive:
   "finite (carrier G) ==> 0 < card H"
 proof (rule classical)
-  assume "finite (carrier G)" and a: "~ 0 < card H"
+  assume "finite (carrier G)" and a: "\<not> 0 < card H"
   then have "finite H" by (blast intro: finite_subset [OF subset])
   with is_subgroup a have "subgroup {} G" by simp
   with subgroup_nonempty show ?thesis by contradiction
@@ -591,7 +591,7 @@
 definition
   hom :: "_ => _ => ('a => 'b) set" where
   "hom G H =
-    {h. h \<in> carrier G \<rightarrow> carrier H &
+    {h. h \<in> carrier G \<rightarrow> carrier H \<and>
       (\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y)}"
 
 lemma (in group) hom_compose:
@@ -600,9 +600,9 @@
 
 definition
   iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60)
-  where "G \<cong> H = {h. h \<in> hom G H & bij_betw h (carrier G) (carrier H)}"
+  where "G \<cong> H = {h. h \<in> hom G H \<and> bij_betw h (carrier G) (carrier H)}"
 
-lemma iso_refl: "(%x. x) \<in> G \<cong> G"
+lemma iso_refl: "(\<lambda>x. x) \<in> G \<cong> G"
 by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def)
 
 lemma (in group) iso_sym:
@@ -798,15 +798,15 @@
 done
 
 theorem (in group) subgroups_Inter:
-  assumes subgr: "(!!H. H \<in> A ==> subgroup H G)"
-    and not_empty: "A ~= {}"
+  assumes subgr: "(\<And>H. H \<in> A \<Longrightarrow> subgroup H G)"
+    and not_empty: "A \<noteq> {}"
   shows "subgroup (\<Inter>A) G"
 proof (rule subgroupI)
   from subgr [THEN subgroup.subset] and not_empty
   show "\<Inter>A \<subseteq> carrier G" by blast
 next
   from subgr [THEN subgroup.one_closed]
-  show "\<Inter>A ~= {}" by blast
+  show "\<Inter>A \<noteq> {}" by blast
 next
   fix x assume "x \<in> \<Inter>A"
   with subgr [THEN subgroup.m_inv_closed]
@@ -828,7 +828,7 @@
   then show "\<exists>G. greatest ?L G (carrier ?L)" ..
 next
   fix A
-  assume L: "A \<subseteq> carrier ?L" and non_empty: "A ~= {}"
+  assume L: "A \<subseteq> carrier ?L" and non_empty: "A \<noteq> {}"
   then have Int_subgroup: "subgroup (\<Inter>A) G"
     by (fastforce intro: subgroups_Inter)
   have "greatest ?L (\<Inter>A) (Lower ?L A)" (is "greatest _ ?Int _")
--- a/src/HOL/Algebra/IntRing.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/IntRing.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -166,7 +166,7 @@
   "x \<in> UNIV \<longleftrightarrow> True"
   "A \<subseteq> UNIV \<longleftrightarrow> True"
   "(\<forall>x \<in> UNIV. P x) \<longleftrightarrow> (\<forall>x. P x)"
-  "(EX x : UNIV. P x) \<longleftrightarrow> (EX x. P x)"
+  "(\<exists>x \<in> UNIV. P x) \<longleftrightarrow> (\<exists>x. P x)"
   "(True \<longrightarrow> Q) \<longleftrightarrow> Q"
   "(True \<Longrightarrow> PROP R) \<equiv> PROP R"
   by simp_all
@@ -247,7 +247,7 @@
   then show "(\<exists>x. a = x * p) \<or> (\<exists>x. b = x * p)"
     by (metis dvd_def mult.commute)
 next
-  assume "UNIV = {uu. EX x. uu = x * p}"
+  assume "UNIV = {uu. \<exists>x. uu = x * p}"
   then obtain x where "1 = x * p" by best
   then have "\<bar>p * x\<bar> = 1" by (simp add: mult.commute)
   then show False using prime
--- a/src/HOL/Algebra/Lattice.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Lattice.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -98,11 +98,11 @@
 
 locale weak_upper_semilattice = weak_partial_order +
   assumes sup_of_two_exists:
-    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. least L s (Upper L {x, y})"
+    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. least L s (Upper L {x, y})"
 
 locale weak_lower_semilattice = weak_partial_order +
   assumes inf_of_two_exists:
-    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
+    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. greatest L s (Lower L {x, y})"
 
 locale weak_lattice = weak_upper_semilattice + weak_lower_semilattice
 
@@ -250,7 +250,7 @@
 qed
 
 lemma (in weak_upper_semilattice) finite_sup_least:
-  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> least L (\<Squnion>A) (Upper L A)"
+  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> least L (\<Squnion>A) (Upper L A)"
 proof (induct set: finite)
   case empty
   then show ?case by simp
@@ -284,7 +284,7 @@
 qed
 
 lemma (in weak_upper_semilattice) finite_sup_closed [simp]:
-  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Squnion>A \<in> carrier L"
+  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> \<Squnion>A \<in> carrier L"
 proof (induct set: finite)
   case empty then show ?case by simp
 next
@@ -495,7 +495,7 @@
 qed
 
 lemma (in weak_lower_semilattice) finite_inf_greatest:
-  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> greatest L (\<Sqinter>A) (Lower L A)"
+  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> greatest L (\<Sqinter>A) (Lower L A)"
 proof (induct set: finite)
   case empty then show ?case by simp
 next
@@ -528,7 +528,7 @@
 qed
 
 lemma (in weak_lower_semilattice) finite_inf_closed [simp]:
-  "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Sqinter>A \<in> carrier L"
+  "[| finite A; A \<subseteq> carrier L; A \<noteq> {} |] ==> \<Sqinter>A \<in> carrier L"
 proof (induct set: finite)
   case empty then show ?case by simp
 next
@@ -611,7 +611,7 @@
 proof
   fix x y
   assume L: "x \<in> carrier L"  "y \<in> carrier L"
-  show "EX s. least L s (Upper L {x, y})"
+  show "\<exists>s. least L s (Upper L {x, y})"
   proof -
     note total L
     moreover
@@ -631,7 +631,7 @@
 next
   fix x y
   assume L: "x \<in> carrier L"  "y \<in> carrier L"
-  show "EX i. greatest L i (Lower L {x, y})"
+  show "\<exists>i. greatest L i (Lower L {x, y})"
   proof -
     note total L
     moreover
@@ -687,14 +687,14 @@
 
 locale upper_semilattice = partial_order +
   assumes sup_of_two_exists:
-    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. least L s (Upper L {x, y})"
+    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. least L s (Upper L {x, y})"
 
 sublocale upper_semilattice \<subseteq> weak?: weak_upper_semilattice
   by unfold_locales (rule sup_of_two_exists)
 
 locale lower_semilattice = partial_order +
   assumes inf_of_two_exists:
-    "[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
+    "[| x \<in> carrier L; y \<in> carrier L |] ==> \<exists>s. greatest L s (Lower L {x, y})"
 
 sublocale lower_semilattice \<subseteq> weak?: weak_lower_semilattice
   by unfold_locales (rule inf_of_two_exists)
--- a/src/HOL/Algebra/Order.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Order.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -42,7 +42,7 @@
 
 definition
   lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50)
-  where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y & x .\<noteq>\<^bsub>L\<^esub> y"
+  where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y \<and> x .\<noteq>\<^bsub>L\<^esub> y"
 
 
 subsubsection \<open>The order relation\<close>
@@ -65,7 +65,7 @@
 
 lemma weak_llessI:
   fixes R (structure)
-  assumes "x \<sqsubseteq> y" and "~(x .= y)"
+  assumes "x \<sqsubseteq> y" and "\<not>(x .= y)"
   shows "x \<sqsubset> y"
   using assms unfolding lless_def by simp
 
@@ -137,11 +137,11 @@
 
 definition
   Upper :: "[_, 'a set] => 'a set"
-  where "Upper L A = {u. (ALL x. x \<in> A \<inter> carrier L --> x \<sqsubseteq>\<^bsub>L\<^esub> u)} \<inter> carrier L"
+  where "Upper L A = {u. (\<forall>x. x \<in> A \<inter> carrier L \<longrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> u)} \<inter> carrier L"
 
 definition
   Lower :: "[_, 'a set] => 'a set"
-  where "Lower L A = {l. (ALL x. x \<in> A \<inter> carrier L --> l \<sqsubseteq>\<^bsub>L\<^esub> x)} \<inter> carrier L"
+  where "Lower L A = {l. (\<forall>x. x \<in> A \<inter> carrier L \<longrightarrow> l \<sqsubseteq>\<^bsub>L\<^esub> x)} \<inter> carrier L"
 
 lemma Upper_closed [intro!, simp]:
   "Upper L A \<subseteq> carrier L"
@@ -320,11 +320,11 @@
 
 definition
   least :: "[_, 'a, 'a set] => bool"
-  where "least L l A \<longleftrightarrow> A \<subseteq> carrier L & l \<in> A & (ALL x : A. l \<sqsubseteq>\<^bsub>L\<^esub> x)"
+  where "least L l A \<longleftrightarrow> A \<subseteq> carrier L \<and> l \<in> A \<and> (\<forall>x\<in>A. l \<sqsubseteq>\<^bsub>L\<^esub> x)"
 
 definition
   greatest :: "[_, 'a, 'a set] => bool"
-  where "greatest L g A \<longleftrightarrow> A \<subseteq> carrier L & g \<in> A & (ALL x : A. x \<sqsubseteq>\<^bsub>L\<^esub> g)"
+  where "greatest L g A \<longleftrightarrow> A \<subseteq> carrier L \<and> g \<in> A \<and> (\<forall>x\<in>A. x \<sqsubseteq>\<^bsub>L\<^esub> g)"
 
 text (in weak_partial_order) \<open>Could weaken these to @{term "l \<in> carrier L \<and> l
   .\<in> A"} and @{term "g \<in> carrier L \<and> g .\<in> A"}.\<close>
@@ -601,7 +601,7 @@
   using weak_le_antisym unfolding eq_is_equal .
 
 lemma lless_eq:
-  "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y & x \<noteq> y"
+  "x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y"
   unfolding lless_def by (simp add: eq_is_equal)
 
 lemma set_eq_is_eq: "A {.=} B \<longleftrightarrow> A = B"
--- a/src/HOL/Algebra/Ring.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/Ring.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -238,9 +238,9 @@
 locale cring = ring + comm_monoid R
 
 locale "domain" = cring +
-  assumes one_not_zero [simp]: "\<one> ~= \<zero>"
+  assumes one_not_zero [simp]: "\<one> \<noteq> \<zero>"
     and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==>
-                  a = \<zero> | b = \<zero>"
+                  a = \<zero> \<or> b = \<zero>"
 
 locale field = "domain" +
   assumes field_Units: "Units R = carrier R - {\<zero>}"
@@ -427,7 +427,7 @@
   a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus
 
 lemma (in semiring) nat_pow_zero:
-  "(n::nat) ~= 0 ==> \<zero> (^) n = \<zero>"
+  "(n::nat) \<noteq> 0 \<Longrightarrow> \<zero> (^) n = \<zero>"
   by (induct n) simp_all
 
 context semiring begin
@@ -469,7 +469,7 @@
   fixes R (structure) and S (structure)
   assumes "ring R" "cring S"
   assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"
-  shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
+  shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b \<and> c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
 proof -
   interpret ring R by fact
   interpret cring S by fact
@@ -513,27 +513,27 @@
 context "domain" begin
 
 lemma zero_not_one [simp]:
-  "\<zero> ~= \<one>"
+  "\<zero> \<noteq> \<one>"
   by (rule not_sym) simp
 
 lemma integral_iff: (* not by default a simp rule! *)
-  "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> | b = \<zero>)"
+  "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> \<or> b = \<zero>)"
 proof
   assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"
-  then show "a = \<zero> | b = \<zero>" by (simp add: integral)
+  then show "a = \<zero> \<or> b = \<zero>" by (simp add: integral)
 next
-  assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> | b = \<zero>"
+  assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> \<or> b = \<zero>"
   then show "a \<otimes> b = \<zero>" by auto
 qed
 
 lemma m_lcancel:
-  assumes prem: "a ~= \<zero>"
+  assumes prem: "a \<noteq> \<zero>"
     and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
   shows "(a \<otimes> b = a \<otimes> c) = (b = c)"
 proof
   assume eq: "a \<otimes> b = a \<otimes> c"
   with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra
-  with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff)
+  with R have "a = \<zero> \<or> (b \<ominus> c) = \<zero>" by (simp add: integral_iff)
   with prem and R have "b \<ominus> c = \<zero>" by auto 
   with R have "b = b \<ominus> (b \<ominus> c)" by algebra 
   also from R have "b \<ominus> (b \<ominus> c) = c" by algebra
@@ -543,7 +543,7 @@
 qed
 
 lemma m_rcancel:
-  assumes prem: "a ~= \<zero>"
+  assumes prem: "a \<noteq> \<zero>"
     and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
   shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)"
 proof -
@@ -609,9 +609,9 @@
 definition
   ring_hom :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme] => ('a => 'b) set"
   where "ring_hom R S =
-    {h. h \<in> carrier R \<rightarrow> carrier S &
-      (ALL x y. x \<in> carrier R & y \<in> carrier R -->
-        h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y & h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) &
+    {h. h \<in> carrier R \<rightarrow> carrier S \<and>
+      (\<forall>x y. x \<in> carrier R \<and> y \<in> carrier R \<longrightarrow>
+        h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y \<and> h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) \<and>
       h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}"
 
 lemma ring_hom_memI:
@@ -675,13 +675,13 @@
 qed
 
 lemma (in ring_hom_cring) hom_finsum [simp]:
-  "f \<in> A \<rightarrow> carrier R ==>
-  h (finsum R f A) = finsum S (h o f) A"
+  "f \<in> A \<rightarrow> carrier R \<Longrightarrow>
+  h (finsum R f A) = finsum S (h \<circ> f) A"
   by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
 
 lemma (in ring_hom_cring) hom_finprod:
-  "f \<in> A \<rightarrow> carrier R ==>
-  h (finprod R f A) = finprod S (h o f) A"
+  "f \<in> A \<rightarrow> carrier R \<Longrightarrow>
+  h (finprod R f A) = finprod S (h \<circ> f) A"
   by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
 
 declare ring_hom_cring.hom_finprod [simp]
--- a/src/HOL/Algebra/UnivPoly.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Algebra/UnivPoly.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -42,7 +42,7 @@
 lemma bound_below:
   assumes bound: "bound z m f" and nonzero: "f n \<noteq> z" shows "n \<le> m"
 proof (rule classical)
-  assume "~ ?thesis"
+  assume "\<not> ?thesis"
   then have "m < n" by arith
   with bound have "f n = z" ..
   with nonzero show ?thesis by contradiction
@@ -54,7 +54,7 @@
 
 definition
   up :: "('a, 'm) ring_scheme => (nat => 'a) set"
-  where "up R = {f. f \<in> UNIV \<rightarrow> carrier R & (EX n. bound \<zero>\<^bsub>R\<^esub> n f)}"
+  where "up R = {f. f \<in> UNIV \<rightarrow> carrier R \<and> (\<exists>n. bound \<zero>\<^bsub>R\<^esub> n f)}"
 
 definition UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring"
   where "UP R = \<lparr>
@@ -72,7 +72,7 @@
 \<close>
 
 lemma mem_upI [intro]:
-  "[| !!n. f n \<in> carrier R; EX n. bound (zero R) n f |] ==> f \<in> up R"
+  "[| \<And>n. f n \<in> carrier R; \<exists>n. bound (zero R) n f |] ==> f \<in> up R"
   by (simp add: up_def Pi_def)
 
 lemma mem_upD [dest]:
@@ -82,7 +82,7 @@
 context ring
 begin
 
-lemma bound_upD [dest]: "f \<in> up R ==> EX n. bound \<zero> n f" by (simp add: up_def)
+lemma bound_upD [dest]: "f \<in> up R \<Longrightarrow> \<exists>n. bound \<zero> n f" by (simp add: up_def)
 
 lemma up_one_closed: "(\<lambda>n. if n = 0 then \<one> else \<zero>) \<in> up R" using up_def by force
 
@@ -97,7 +97,7 @@
     by auto
 next
   assume UP: "p \<in> up R" "q \<in> up R"
-  show "EX n. bound \<zero> n (\<lambda>i. p i \<oplus> q i)"
+  show "\<exists>n. bound \<zero> n (\<lambda>i. p i \<oplus> q i)"
   proof -
     from UP obtain n where boundn: "bound \<zero> n p" by fast
     from UP obtain m where boundm: "bound \<zero> m q" by fast
@@ -117,7 +117,7 @@
   assume R: "p \<in> up R"
   then obtain n where "bound \<zero> n p" by auto
   then have "bound \<zero> n (\<lambda>i. \<ominus> p i)" by auto
-  then show "EX n. bound \<zero> n (\<lambda>i. \<ominus> p i)" by auto
+  then show "\<exists>n. bound \<zero> n (\<lambda>i. \<ominus> p i)" by auto
 qed auto
 
 lemma up_minus_closed:
@@ -135,7 +135,7 @@
     by (simp add: mem_upD  funcsetI)
 next
   assume UP: "p \<in> up R" "q \<in> up R"
-  show "EX n. bound \<zero> n (\<lambda>n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i))"
+  show "\<exists>n. bound \<zero> n (\<lambda>n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i))"
   proof -
     from UP obtain n where boundn: "bound \<zero> n p" by fast
     from UP obtain m where boundm: "bound \<zero> m q" by fast
@@ -751,33 +751,33 @@
 qed
 
 lemma deg_belowI:
-  assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"
+  assumes non_zero: "n \<noteq> 0 \<Longrightarrow> coeff P p n \<noteq> \<zero>"
     and R: "p \<in> carrier P"
-  shows "n <= deg R p"
+  shows "n \<le> deg R p"
 \<comment> \<open>Logically, this is a slightly stronger version of
    @{thm [source] deg_aboveD}\<close>
 proof (cases "n=0")
   case True then show ?thesis by simp
 next
-  case False then have "coeff P p n ~= \<zero>" by (rule non_zero)
-  then have "~ deg R p < n" by (fast dest: deg_aboveD intro: R)
+  case False then have "coeff P p n \<noteq> \<zero>" by (rule non_zero)
+  then have "\<not> deg R p < n" by (fast dest: deg_aboveD intro: R)
   then show ?thesis by arith
 qed
 
 lemma lcoeff_nonzero_deg:
-  assumes deg: "deg R p ~= 0" and R: "p \<in> carrier P"
-  shows "coeff P p (deg R p) ~= \<zero>"
+  assumes deg: "deg R p \<noteq> 0" and R: "p \<in> carrier P"
+  shows "coeff P p (deg R p) \<noteq> \<zero>"
 proof -
-  from R obtain m where "deg R p <= m" and m_coeff: "coeff P p m ~= \<zero>"
+  from R obtain m where "deg R p \<le> m" and m_coeff: "coeff P p m \<noteq> \<zero>"
   proof -
-    have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"
+    have minus: "\<And>(n::nat) m. n \<noteq> 0 \<Longrightarrow> (n - Suc 0 < m) = (n \<le> m)"
       by arith
     from deg have "deg R p - 1 < (LEAST n. bound \<zero> n (coeff P p))"
       by (unfold deg_def P_def) simp
-    then have "~ bound \<zero> (deg R p - 1) (coeff P p)" by (rule not_less_Least)
-    then have "EX m. deg R p - 1 < m & coeff P p m ~= \<zero>"
+    then have "\<not> bound \<zero> (deg R p - 1) (coeff P p)" by (rule not_less_Least)
+    then have "\<exists>m. deg R p - 1 < m \<and> coeff P p m \<noteq> \<zero>"
       by (unfold bound_def) fast
-    then have "EX m. deg R p <= m & coeff P p m ~= \<zero>" by (simp add: deg minus)
+    then have "\<exists>m. deg R p \<le> m \<and> coeff P p m \<noteq> \<zero>" by (simp add: deg minus)
     then show ?thesis by (auto intro: that)
   qed
   with deg_belowI R have "deg R p = m" by fastforce
@@ -785,24 +785,24 @@
 qed
 
 lemma lcoeff_nonzero_nonzero:
-  assumes deg: "deg R p = 0" and nonzero: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
-  shows "coeff P p 0 ~= \<zero>"
+  assumes deg: "deg R p = 0" and nonzero: "p \<noteq> \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
+  shows "coeff P p 0 \<noteq> \<zero>"
 proof -
-  have "EX m. coeff P p m ~= \<zero>"
+  have "\<exists>m. coeff P p m \<noteq> \<zero>"
   proof (rule classical)
-    assume "~ ?thesis"
+    assume "\<not> ?thesis"
     with R have "p = \<zero>\<^bsub>P\<^esub>" by (auto intro: up_eqI)
     with nonzero show ?thesis by contradiction
   qed
-  then obtain m where coeff: "coeff P p m ~= \<zero>" ..
-  from this and R have "m <= deg R p" by (rule deg_belowI)
+  then obtain m where coeff: "coeff P p m \<noteq> \<zero>" ..
+  from this and R have "m \<le> deg R p" by (rule deg_belowI)
   then have "m = 0" by (simp add: deg)
   with coeff show ?thesis by simp
 qed
 
 lemma lcoeff_nonzero:
-  assumes neq: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
-  shows "coeff P p (deg R p) ~= \<zero>"
+  assumes neq: "p \<noteq> \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
+  shows "coeff P p (deg R p) \<noteq> \<zero>"
 proof (cases "deg R p = 0")
   case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero)
 next
@@ -810,55 +810,55 @@
 qed
 
 lemma deg_eqI:
-  "[| !!m. n < m ==> coeff P p m = \<zero>;
-      !!n. n ~= 0 ==> coeff P p n ~= \<zero>; p \<in> carrier P |] ==> deg R p = n"
+  "[| \<And>m. n < m \<Longrightarrow> coeff P p m = \<zero>;
+      \<And>n. n \<noteq> 0 \<Longrightarrow> coeff P p n \<noteq> \<zero>; p \<in> carrier P |] ==> deg R p = n"
 by (fast intro: le_antisym deg_aboveI deg_belowI)
 
 text \<open>Degree and polynomial operations\<close>
 
 lemma deg_add [simp]:
   "p \<in> carrier P \<Longrightarrow> q \<in> carrier P \<Longrightarrow>
-  deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
+  deg R (p \<oplus>\<^bsub>P\<^esub> q) \<le> max (deg R p) (deg R q)"
 by(rule deg_aboveI)(simp_all add: deg_aboveD)
 
 lemma deg_monom_le:
-  "a \<in> carrier R ==> deg R (monom P a n) <= n"
+  "a \<in> carrier R \<Longrightarrow> deg R (monom P a n) \<le> n"
   by (intro deg_aboveI) simp_all
 
 lemma deg_monom [simp]:
-  "[| a ~= \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"
+  "[| a \<noteq> \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"
   by (fastforce intro: le_antisym deg_aboveI deg_belowI)
 
 lemma deg_const [simp]:
   assumes R: "a \<in> carrier R" shows "deg R (monom P a 0) = 0"
 proof (rule le_antisym)
-  show "deg R (monom P a 0) <= 0" by (rule deg_aboveI) (simp_all add: R)
+  show "deg R (monom P a 0) \<le> 0" by (rule deg_aboveI) (simp_all add: R)
 next
-  show "0 <= deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R)
+  show "0 \<le> deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R)
 qed
 
 lemma deg_zero [simp]:
   "deg R \<zero>\<^bsub>P\<^esub> = 0"
 proof (rule le_antisym)
-  show "deg R \<zero>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
+  show "deg R \<zero>\<^bsub>P\<^esub> \<le> 0" by (rule deg_aboveI) simp_all
 next
-  show "0 <= deg R \<zero>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
+  show "0 \<le> deg R \<zero>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
 qed
 
 lemma deg_one [simp]:
   "deg R \<one>\<^bsub>P\<^esub> = 0"
 proof (rule le_antisym)
-  show "deg R \<one>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
+  show "deg R \<one>\<^bsub>P\<^esub> \<le> 0" by (rule deg_aboveI) simp_all
 next
-  show "0 <= deg R \<one>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
+  show "0 \<le> deg R \<one>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
 qed
 
 lemma deg_uminus [simp]:
   assumes R: "p \<in> carrier P" shows "deg R (\<ominus>\<^bsub>P\<^esub> p) = deg R p"
 proof (rule le_antisym)
-  show "deg R (\<ominus>\<^bsub>P\<^esub> p) <= deg R p" by (simp add: deg_aboveI deg_aboveD R)
+  show "deg R (\<ominus>\<^bsub>P\<^esub> p) \<le> deg R p" by (simp add: deg_aboveI deg_aboveD R)
 next
-  show "deg R p <= deg R (\<ominus>\<^bsub>P\<^esub> p)"
+  show "deg R p \<le> deg R (\<ominus>\<^bsub>P\<^esub> p)"
     by (simp add: deg_belowI lcoeff_nonzero_deg
       inj_on_eq_iff [OF R.a_inv_inj, of _ "\<zero>", simplified] R)
 qed
@@ -868,7 +868,7 @@
 
 lemma deg_smult_ring [simp]:
   "[| a \<in> carrier R; p \<in> carrier P |] ==>
-  deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
+  deg R (a \<odot>\<^bsub>P\<^esub> p) \<le> (if a = \<zero> then 0 else deg R p)"
   by (cases "a = \<zero>") (simp add: deg_aboveI deg_aboveD)+
 
 end
@@ -880,10 +880,10 @@
   assumes R: "a \<in> carrier R" "p \<in> carrier P"
   shows "deg R (a \<odot>\<^bsub>P\<^esub> p) = (if a = \<zero> then 0 else deg R p)"
 proof (rule le_antisym)
-  show "deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
+  show "deg R (a \<odot>\<^bsub>P\<^esub> p) \<le> (if a = \<zero> then 0 else deg R p)"
     using R by (rule deg_smult_ring)
 next
-  show "(if a = \<zero> then 0 else deg R p) <= deg R (a \<odot>\<^bsub>P\<^esub> p)"
+  show "(if a = \<zero> then 0 else deg R p) \<le> deg R (a \<odot>\<^bsub>P\<^esub> p)"
   proof (cases "a = \<zero>")
   qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R)
 qed
@@ -895,7 +895,7 @@
 
 lemma deg_mult_ring:
   assumes R: "p \<in> carrier P" "q \<in> carrier P"
-  shows "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q"
+  shows "deg R (p \<otimes>\<^bsub>P\<^esub> q) \<le> deg R p + deg R q"
 proof (rule deg_aboveI)
   fix m
   assume boundm: "deg R p + deg R q < m"
@@ -919,16 +919,16 @@
 begin
 
 lemma deg_mult [simp]:
-  "[| p ~= \<zero>\<^bsub>P\<^esub>; q ~= \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==>
+  "[| p \<noteq> \<zero>\<^bsub>P\<^esub>; q \<noteq> \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==>
   deg R (p \<otimes>\<^bsub>P\<^esub> q) = deg R p + deg R q"
 proof (rule le_antisym)
   assume "p \<in> carrier P" " q \<in> carrier P"
-  then show "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_ring)
+  then show "deg R (p \<otimes>\<^bsub>P\<^esub> q) \<le> deg R p + deg R q" by (rule deg_mult_ring)
 next
   let ?s = "(\<lambda>i. coeff P p i \<otimes> coeff P q (deg R p + deg R q - i))"
-  assume R: "p \<in> carrier P" "q \<in> carrier P" and nz: "p ~= \<zero>\<^bsub>P\<^esub>" "q ~= \<zero>\<^bsub>P\<^esub>"
+  assume R: "p \<in> carrier P" "q \<in> carrier P" and nz: "p \<noteq> \<zero>\<^bsub>P\<^esub>" "q \<noteq> \<zero>\<^bsub>P\<^esub>"
   have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith
-  show "deg R p + deg R q <= deg R (p \<otimes>\<^bsub>P\<^esub> q)"
+  show "deg R p + deg R q \<le> deg R (p \<otimes>\<^bsub>P\<^esub> q)"
   proof (rule deg_belowI, simp add: R)
     have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
       = (\<Oplus>i \<in> {..< deg R p} \<union> {deg R p .. deg R p + deg R q}. ?s i)"
@@ -942,7 +942,7 @@
       by (simp cong: R.finsum_cong add: deg_aboveD R Pi_def)
     finally have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
       = coeff P p (deg R p) \<otimes> coeff P q (deg R q)" .
-    with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) ~= \<zero>"
+    with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) \<noteq> \<zero>"
       by (simp add: integral_iff lcoeff_nonzero R)
   qed (simp add: R)
 qed
@@ -969,7 +969,7 @@
   from R have RR: "!!i. (if i = k then coeff P p i else \<zero>) \<in> carrier R"
     by simp
   show "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k = coeff P p k"
-  proof (cases "k <= deg R p")
+  proof (cases "k \<le> deg R p")
     case True
     hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
           coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k} \<union> {k<..deg R p}. ?s i) k"
@@ -1017,9 +1017,9 @@
 
 lemma domainI:
   assumes cring: "cring R"
-    and one_not_zero: "one R ~= zero R"
-    and integral: "!!a b. [| mult R a b = zero R; a \<in> carrier R;
-      b \<in> carrier R |] ==> a = zero R | b = zero R"
+    and one_not_zero: "one R \<noteq> zero R"
+    and integral: "\<And>a b. [| mult R a b = zero R; a \<in> carrier R;
+      b \<in> carrier R |] ==> a = zero R \<or> b = zero R"
   shows "domain R"
   by (auto intro!: domain.intro domain_axioms.intro cring.axioms assms
     del: disjCI)
@@ -1028,7 +1028,7 @@
 begin
 
 lemma UP_one_not_zero:
-  "\<one>\<^bsub>P\<^esub> ~= \<zero>\<^bsub>P\<^esub>"
+  "\<one>\<^bsub>P\<^esub> \<noteq> \<zero>\<^bsub>P\<^esub>"
 proof
   assume "\<one>\<^bsub>P\<^esub> = \<zero>\<^bsub>P\<^esub>"
   hence "coeff P \<one>\<^bsub>P\<^esub> 0 = (coeff P \<zero>\<^bsub>P\<^esub> 0)" by simp
@@ -1037,17 +1037,17 @@
 qed
 
 lemma UP_integral:
-  "[| p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==> p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
+  "[| p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==> p = \<zero>\<^bsub>P\<^esub> \<or> q = \<zero>\<^bsub>P\<^esub>"
 proof -
   fix p q
   assume pq: "p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P" "q \<in> carrier P"
-  show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
+  show "p = \<zero>\<^bsub>P\<^esub> \<or> q = \<zero>\<^bsub>P\<^esub>"
   proof (rule classical)
-    assume c: "~ (p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>)"
+    assume c: "\<not> (p = \<zero>\<^bsub>P\<^esub> \<or> q = \<zero>\<^bsub>P\<^esub>)"
     with R have "deg R p + deg R q = deg R (p \<otimes>\<^bsub>P\<^esub> q)" by simp
     also from pq have "... = 0" by simp
     finally have "deg R p + deg R q = 0" .
-    then have f1: "deg R p = 0 & deg R q = 0" by simp
+    then have f1: "deg R p = 0 \<and> deg R q = 0" by simp
     from f1 R have "p = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P p i) i)"
       by (simp only: up_repr_le)
     also from R have "... = monom P (coeff P p 0) 0" by simp
@@ -1059,9 +1059,9 @@
     from R have "coeff P p 0 \<otimes> coeff P q 0 = coeff P (p \<otimes>\<^bsub>P\<^esub> q) 0" by simp
     also from pq have "... = \<zero>" by simp
     finally have "coeff P p 0 \<otimes> coeff P q 0 = \<zero>" .
-    with R have "coeff P p 0 = \<zero> | coeff P q 0 = \<zero>"
+    with R have "coeff P p 0 = \<zero> \<or> coeff P q 0 = \<zero>"
       by (simp add: R.integral_iff)
-    with p q show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>" by fastforce
+    with p q show "p = \<zero>\<^bsub>P\<^esub> \<or> q = \<zero>\<^bsub>P\<^esub>" by fastforce
   qed
 qed
 
@@ -1206,8 +1206,8 @@
   maybe it is not that necessary.\<close>
 
 lemma (in ring_hom_ring) hom_finsum [simp]:
-  "f \<in> A \<rightarrow> carrier R ==>
-  h (finsum R f A) = finsum S (h o f) A"
+  "f \<in> A \<rightarrow> carrier R \<Longrightarrow>
+  h (finsum R f A) = finsum S (h \<circ> f) A"
   by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
 
 context UP_pre_univ_prop
@@ -1429,9 +1429,9 @@
 
 theorem UP_universal_property:
   assumes S: "s \<in> carrier S"
-  shows "\<exists>!Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) &
-    Phi (monom P \<one> 1) = s &
-    (ALL r : carrier R. Phi (monom P r 0) = h r)"
+  shows "\<exists>!Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) \<and>
+    Phi (monom P \<one> 1) = s \<and>
+    (\<forall>r \<in> carrier R. Phi (monom P r 0) = h r)"
   using S eval_monom1
   apply (auto intro: eval_ring_hom eval_const eval_extensional)
   apply (rule extensionalityI)
@@ -1550,14 +1550,14 @@
 lemma long_div_theorem:
   assumes g_in_P [simp]: "g \<in> carrier P" and f_in_P [simp]: "f \<in> carrier P"
   and g_not_zero: "g \<noteq> \<zero>\<^bsub>P\<^esub>"
-  shows "\<exists> q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
+  shows "\<exists>q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g)"
   using f_in_P
 proof (induct "deg R f" arbitrary: "f" rule: nat_less_induct)
   case (1 f)
   note f_in_P [simp] = "1.prems"
   let ?pred = "(\<lambda> q r (k::nat).
     (q \<in> carrier P) \<and> (r \<in> carrier P)
-    \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
+    \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R g))"
   let ?lg = "lcoeff g" and ?lf = "lcoeff f"
   show ?case
   proof (cases "deg R f < deg R g")
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -387,8 +387,8 @@
 lemma card_of_image:
 "|f ` A| <=o |A|"
 proof(cases "A = {}", simp add: card_of_empty)
-  assume "A ~= {}"
-  hence "f ` A ~= {}" by auto
+  assume "A \<noteq> {}"
+  hence "f ` A \<noteq> {}" by auto
   thus "|f ` A| \<le>o |A|"
   using card_of_ordLeq2[of "f ` A" A] by auto
 qed
@@ -886,7 +886,7 @@
   moreover have "under r a \<le> Field r"
   using under_Field .
   ultimately have "under r a < Field r" by blast
-  then obtain b where 1: "b : Field r \<and> ~ (b,a) : r"
+  then obtain b where 1: "b \<in> Field r \<and> \<not> (b,a) \<in> r"
   unfolding under_def by blast
   moreover have ba: "b \<noteq> a"
   using 1 r unfolding card_order_on_def well_order_on_def
@@ -1100,7 +1100,7 @@
     using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto
     moreover have "?A' = A <+> B" using Case1 by blast
     ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp
-    hence "bij_betw (g o ?Inl) A (A <+> B)"
+    hence "bij_betw (g \<circ> ?Inl) A (A <+> B)"
     using 2 by (auto simp add: bij_betw_trans)
     thus ?thesis using card_of_ordIso ordIso_symmetric by blast
   next
@@ -1510,11 +1510,11 @@
 
 definition regularCard where
 "regularCard r \<equiv>
- ALL K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r"
+ \<forall>K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r"
 
 definition relChain where
 "relChain r As \<equiv>
- ALL i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j"
+ \<forall>i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j"
 
 lemma regularCard_UNION:
 assumes r: "Card_order r"   "regularCard r"
@@ -1523,17 +1523,17 @@
 and cardB: "|B| <o r"
 shows "EX i : Field r. B \<le> As i"
 proof-
-  let ?phi = "%b j. j : Field r \<and> b : As j"
-  have "ALL b : B. EX j. ?phi b j" using Bsub by blast
+  let ?phi = "\<lambda>b j. j \<in> Field r \<and> b \<in> As j"
+  have "\<forall>b\<in>B. \<exists>j. ?phi b j" using Bsub by blast
   then obtain f where f: "!! b. b : B \<Longrightarrow> ?phi b (f b)"
   using bchoice[of B ?phi] by blast
   let ?K = "f ` B"
-  {assume 1: "!! i. i : Field r \<Longrightarrow> ~ B \<le> As i"
+  {assume 1: "\<And>i. i \<in> Field r \<Longrightarrow> \<not> B \<le> As i"
    have 2: "cofinal ?K r"
    unfolding cofinal_def proof auto
      fix i assume i: "i : Field r"
      with 1 obtain b where b: "b : B \<and> b \<notin> As i" by blast
-     hence "i \<noteq> f b \<and> ~ (f b,i) : r"
+     hence "i \<noteq> f b \<and> \<not> (f b,i) \<in> r"
      using As f unfolding relChain_def by auto
      hence "i \<noteq> f b \<and> (i, f b) : r" using r
      unfolding card_order_on_def well_order_on_def linear_order_on_def
--- a/src/HOL/BNF_Composition.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Composition.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -18,13 +18,13 @@
 lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X"
   by simp
 
-lemma empty_natural: "(\<lambda>_. {}) o f = image g o (\<lambda>_. {})"
+lemma empty_natural: "(\<lambda>_. {}) \<circ> f = image g \<circ> (\<lambda>_. {})"
   by (rule ext) simp
 
-lemma Union_natural: "Union o image (image f) = image f o Union"
+lemma Union_natural: "Union \<circ> image (image f) = image f \<circ> Union"
   by (rule ext) (auto simp only: comp_apply)
 
-lemma in_Union_o_assoc: "x \<in> (Union o gset o gmap) A \<Longrightarrow> x \<in> (Union o (gset o gmap)) A"
+lemma in_Union_o_assoc: "x \<in> (Union \<circ> gset \<circ> gmap) A \<Longrightarrow> x \<in> (Union \<circ> (gset \<circ> gmap)) A"
   by (unfold comp_assoc)
 
 lemma comp_single_set_bd:
@@ -66,7 +66,7 @@
 lemma Union_image_empty: "A \<union> \<Union>(f ` {}) = A"
   by simp
 
-lemma image_o_collect: "collect ((\<lambda>f. image g o f) ` F) = image g o collect F"
+lemma image_o_collect: "collect ((\<lambda>f. image g \<circ> f) ` F) = image g \<circ> collect F"
   by (rule ext) (auto simp add: collect_def)
 
 lemma conj_subset_def: "A \<subseteq> {x. P x \<and> Q x} = (A \<subseteq> {x. P x} \<and> A \<subseteq> {x. Q x})"
@@ -91,19 +91,19 @@
   vimage2p f g R OO vimage2p g h S \<le> vimage2p f h T"
   unfolding vimage2p_def by auto
 
-lemma type_copy_map_cong0: "M (g x) = N (h x) \<Longrightarrow> (f o M o g) x = (f o N o h) x"
+lemma type_copy_map_cong0: "M (g x) = N (h x) \<Longrightarrow> (f \<circ> M \<circ> g) x = (f \<circ> N \<circ> h) x"
   by auto
 
-lemma type_copy_set_bd: "(\<And>y. |S y| \<le>o bd) \<Longrightarrow> |(S o Rep) x| \<le>o bd"
+lemma type_copy_set_bd: "(\<And>y. |S y| \<le>o bd) \<Longrightarrow> |(S \<circ> Rep) x| \<le>o bd"
   by auto
 
 lemma vimage2p_cong: "R = S \<Longrightarrow> vimage2p f g R = vimage2p f g S"
   by simp
 
-lemma Ball_comp_iff: "(\<lambda>x. Ball (A x) f) o g = (\<lambda>x. Ball ((A o g) x) f)"
+lemma Ball_comp_iff: "(\<lambda>x. Ball (A x) f) \<circ> g = (\<lambda>x. Ball ((A \<circ> g) x) f)"
   unfolding o_def by auto
 
-lemma conj_comp_iff: "(\<lambda>x. P x \<and> Q x) o g = (\<lambda>x. (P o g) x \<and> (Q o g) x)"
+lemma conj_comp_iff: "(\<lambda>x. P x \<and> Q x) \<circ> g = (\<lambda>x. (P \<circ> g) x \<and> (Q \<circ> g) x)"
   unfolding o_def by auto
 
 context
@@ -111,26 +111,26 @@
   assumes type_copy: "type_definition Rep Abs UNIV"
 begin
 
-lemma type_copy_map_id0: "M = id \<Longrightarrow> Abs o M o Rep = id"
+lemma type_copy_map_id0: "M = id \<Longrightarrow> Abs \<circ> M \<circ> Rep = id"
   using type_definition.Rep_inverse[OF type_copy] by auto
 
-lemma type_copy_map_comp0: "M = M1 o M2 \<Longrightarrow> f o M o g = (f o M1 o Rep) o (Abs o M2 o g)"
+lemma type_copy_map_comp0: "M = M1 \<circ> M2 \<Longrightarrow> f \<circ> M \<circ> g = (f \<circ> M1 \<circ> Rep) \<circ> (Abs \<circ> M2 \<circ> g)"
   using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
 
-lemma type_copy_set_map0: "S o M = image f o S' \<Longrightarrow> (S o Rep) o (Abs o M o g) = image f o (S' o g)"
+lemma type_copy_set_map0: "S \<circ> M = image f \<circ> S' \<Longrightarrow> (S \<circ> Rep) \<circ> (Abs \<circ> M \<circ> g) = image f \<circ> (S' \<circ> g)"
   using type_definition.Abs_inverse[OF type_copy UNIV_I] by (auto simp: o_def fun_eq_iff)
 
-lemma type_copy_wit: "x \<in> (S o Rep) (Abs y) \<Longrightarrow> x \<in> S y"
+lemma type_copy_wit: "x \<in> (S \<circ> Rep) (Abs y) \<Longrightarrow> x \<in> S y"
   using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
 
 lemma type_copy_vimage2p_Grp_Rep: "vimage2p f Rep (Grp (Collect P) h) =
-    Grp (Collect (\<lambda>x. P (f x))) (Abs o h o f)"
+    Grp (Collect (\<lambda>x. P (f x))) (Abs \<circ> h \<circ> f)"
   unfolding vimage2p_def Grp_def fun_eq_iff
   by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
    type_definition.Rep_inverse[OF type_copy] dest: sym)
 
 lemma type_copy_vimage2p_Grp_Abs:
-  "\<And>h. vimage2p g Abs (Grp (Collect P) h) = Grp (Collect (\<lambda>x. P (g x))) (Rep o h o g)"
+  "\<And>h. vimage2p g Abs (Grp (Collect P) h) = Grp (Collect (\<lambda>x. P (g x))) (Rep \<circ> h \<circ> g)"
   unfolding vimage2p_def Grp_def fun_eq_iff
   by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
    type_definition.Rep_inverse[OF type_copy] dest: sym)
--- a/src/HOL/BNF_Def.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Def.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -180,7 +180,7 @@
 lemma case_sum_o_inj: "case_sum f g \<circ> Inl = f" "case_sum f g \<circ> Inr = g"
   by auto
 
-lemma map_sum_o_inj: "map_sum f g o Inl = Inl o f" "map_sum f g o Inr = Inr o g"
+lemma map_sum_o_inj: "map_sum f g \<circ> Inl = Inl \<circ> f" "map_sum f g \<circ> Inr = Inr \<circ> g"
   by auto
 
 lemma card_order_csum_cone_cexp_def:
@@ -262,11 +262,11 @@
 lemma eq_onp_True: "eq_onp (\<lambda>_. True) = (op =)"
   unfolding eq_onp_def by simp
 
-lemma Ball_image_comp: "Ball (f ` A) g = Ball A (g o f)"
+lemma Ball_image_comp: "Ball (f ` A) g = Ball A (g \<circ> f)"
   by auto
 
 lemma rel_fun_Collect_case_prodD:
-  "rel_fun A B f g \<Longrightarrow> X \<subseteq> Collect (case_prod A) \<Longrightarrow> x \<in> X \<Longrightarrow> B ((f o fst) x) ((g o snd) x)"
+  "rel_fun A B f g \<Longrightarrow> X \<subseteq> Collect (case_prod A) \<Longrightarrow> x \<in> X \<Longrightarrow> B ((f \<circ> fst) x) ((g \<circ> snd) x)"
   unfolding rel_fun_def by auto
 
 lemma eq_onp_mono_iff: "eq_onp P \<le> eq_onp Q \<longleftrightarrow> P \<le> Q"
--- a/src/HOL/BNF_Fixpoint_Base.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Fixpoint_Base.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -46,7 +46,7 @@
   fix b
   have "b = f (g b)"
   using fg unfolding fun_eq_iff by simp
-  thus "EX a. b = f a" by blast
+  thus "\<exists>a. b = f a" by blast
 qed
 
 lemma case_sum_step:
@@ -248,27 +248,27 @@
 lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x"
   using snd_convol unfolding convol_def by simp
 
-lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f"
+lemma convol_expand_snd: "fst \<circ> f = g \<Longrightarrow> \<langle>g, snd \<circ> f\<rangle> = f"
   unfolding convol_def by auto
 
 lemma convol_expand_snd':
-  assumes "(fst o f = g)"
-  shows "h = snd o f \<longleftrightarrow> \<langle>g, h\<rangle> = f"
+  assumes "(fst \<circ> f = g)"
+  shows "h = snd \<circ> f \<longleftrightarrow> \<langle>g, h\<rangle> = f"
 proof -
-  from assms have *: "\<langle>g, snd o f\<rangle> = f" by (rule convol_expand_snd)
-  then have "h = snd o f \<longleftrightarrow> h = snd o \<langle>g, snd o f\<rangle>" by simp
-  moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol)
+  from assms have *: "\<langle>g, snd \<circ> f\<rangle> = f" by (rule convol_expand_snd)
+  then have "h = snd \<circ> f \<longleftrightarrow> h = snd \<circ> \<langle>g, snd \<circ> f\<rangle>" by simp
+  moreover have "\<dots> \<longleftrightarrow> h = snd \<circ> f" by (simp add: snd_convol)
   moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
   ultimately show ?thesis by simp
 qed
 
-lemma case_sum_expand_Inr_pointfree: "f o Inl = g \<Longrightarrow> case_sum g (f o Inr) = f"
+lemma case_sum_expand_Inr_pointfree: "f \<circ> Inl = g \<Longrightarrow> case_sum g (f \<circ> Inr) = f"
   by (auto split: sum.splits)
 
-lemma case_sum_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> case_sum g h = f"
+lemma case_sum_expand_Inr': "f \<circ> Inl = g \<Longrightarrow> h = f \<circ> Inr \<longleftrightarrow> case_sum g h = f"
   by (rule iffI) (auto simp add: fun_eq_iff split: sum.splits)
 
-lemma case_sum_expand_Inr: "f o Inl = g \<Longrightarrow> f x = case_sum g (f o Inr) x"
+lemma case_sum_expand_Inr: "f \<circ> Inl = g \<Longrightarrow> f x = case_sum g (f \<circ> Inr) x"
   by (auto split: sum.splits)
 
 lemma id_transfer: "rel_fun A A id id"
--- a/src/HOL/BNF_Greatest_Fixpoint.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Greatest_Fixpoint.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -36,7 +36,7 @@
 
 lemma equiv_proj:
   assumes e: "equiv A R" and m: "z \<in> R"
-  shows "(proj R o fst) z = (proj R o snd) z"
+  shows "(proj R \<circ> fst) z = (proj R \<circ> snd) z"
 proof -
   from m have z: "(fst z, snd z) \<in> R" by auto
   with e have "\<And>x. (fst z, x) \<in> R \<Longrightarrow> (snd z, x) \<in> R" "\<And>x. (snd z, x) \<in> R \<Longrightarrow> (fst z, x) \<in> R"
@@ -139,10 +139,10 @@
 lemma fst_diag_id: "(fst \<circ> (\<lambda>x. (x, x))) z = id z" by simp
 lemma snd_diag_id: "(snd \<circ> (\<lambda>x. (x, x))) z = id z" by simp
 
-lemma fst_diag_fst: "fst o ((\<lambda>x. (x, x)) o fst) = fst" by auto
-lemma snd_diag_fst: "snd o ((\<lambda>x. (x, x)) o fst) = fst" by auto
-lemma fst_diag_snd: "fst o ((\<lambda>x. (x, x)) o snd) = snd" by auto
-lemma snd_diag_snd: "snd o ((\<lambda>x. (x, x)) o snd) = snd" by auto
+lemma fst_diag_fst: "fst \<circ> ((\<lambda>x. (x, x)) \<circ> fst) = fst" by auto
+lemma snd_diag_fst: "snd \<circ> ((\<lambda>x. (x, x)) \<circ> fst) = fst" by auto
+lemma fst_diag_snd: "fst \<circ> ((\<lambda>x. (x, x)) \<circ> snd) = snd" by auto
+lemma snd_diag_snd: "snd \<circ> ((\<lambda>x. (x, x)) \<circ> snd) = snd" by auto
 
 definition Succ where "Succ Kl kl = {k . kl @ [k] \<in> Kl}"
 definition Shift where "Shift Kl k = {kl. k # kl \<in> Kl}"
--- a/src/HOL/BNF_Least_Fixpoint.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Least_Fixpoint.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -54,7 +54,7 @@
   unfolding bij_betw_def inj_on_def by blast
 
 lemma surj_fun_eq:
-  assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x"
+  assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 \<circ> f) x = (g2 \<circ> f) x"
   shows "g1 = g2"
 proof (rule ext)
   fix y
--- a/src/HOL/BNF_Wellorder_Constructions.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Wellorder_Constructions.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -368,7 +368,7 @@
   where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
         "embed r r' f" and "embed r' r'' f'"
   using * ** unfolding ordLeq_def by blast
-  hence "embed r r'' (f' o f)"
+  hence "embed r r'' (f' \<circ> f)"
   using comp_embed[of r r' f r'' f'] by auto
   thus "r \<le>o r''" unfolding ordLeq_def using 1 by auto
 qed
@@ -389,7 +389,7 @@
   where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
         "iso r r' f" and 3: "iso r' r'' f'"
   using * ** unfolding ordIso_def by auto
-  hence "iso r r'' (f' o f)"
+  hence "iso r r'' (f' \<circ> f)"
   using comp_iso[of r r' f r'' f'] by auto
   thus "r =o r''" unfolding ordIso_def using 1 by auto
 qed
@@ -691,10 +691,10 @@
   ultimately
   have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset)
   moreover
-  {have "embed r1 r3 (f23 o f12)"
+  {have "embed r1 r3 (f23 \<circ> f12)"
    using 1 EMB23 0 by (auto simp add: comp_embed)
    hence "\<forall>a \<in> ?A1. f23(f12 a) = f13 a"
-   using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto
+   using EMB13 0 embed_unique[of r1 r3 "f23 \<circ> f12" f13] by auto
    hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force
   }
   ultimately
--- a/src/HOL/BNF_Wellorder_Embedding.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/BNF_Wellorder_Embedding.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -128,7 +128,7 @@
 lemma comp_embed:
 assumes WELL: "Well_order r" and
         EMB: "embed r r' f" and EMB': "embed r' r'' f'"
-shows "embed r r'' (f' o f)"
+shows "embed r r'' (f' \<circ> f)"
 proof(unfold embed_def, auto)
   fix a assume *: "a \<in> Field r"
   hence "bij_betw f (under r a) (under r' (f a))"
@@ -147,7 +147,7 @@
 lemma comp_iso:
 assumes WELL: "Well_order r" and
         EMB: "iso r r' f" and EMB': "iso r' r'' f'"
-shows "iso r r'' (f' o f)"
+shows "iso r r'' (f' \<circ> f)"
 using assms unfolding iso_def
 by (auto simp add: comp_embed bij_betw_trans)
 
@@ -672,25 +672,25 @@
   have Well': "wo_rel r'" using WELL' unfolding wo_rel_def .
   (* Main proof *)
   obtain H where H_def: "H =
-  (\<lambda>h a. if False \<notin> (snd o h)`(underS r a) \<and> (fst o h)`(underS r a) \<noteq> Field r'
-                then (wo_rel.suc r' ((fst o h)`(underS r a)), True)
+  (\<lambda>h a. if False \<notin> (snd \<circ> h)`(underS r a) \<and> (fst \<circ> h)`(underS r a) \<noteq> Field r'
+                then (wo_rel.suc r' ((fst \<circ> h)`(underS r a)), True)
                 else (undefined, False))" by blast
   have Adm: "wo_rel.adm_wo r H"
   using Well
   proof(unfold wo_rel.adm_wo_def, clarify)
     fix h1::"'a \<Rightarrow> 'a' * bool" and h2::"'a \<Rightarrow> 'a' * bool" and x
     assume "\<forall>y\<in>underS r x. h1 y = h2 y"
-    hence "\<forall>y\<in>underS r x. (fst o h1) y = (fst o h2) y \<and>
-                          (snd o h1) y = (snd o h2) y" by auto
-    hence "(fst o h1)`(underS r x) = (fst o h2)`(underS r x) \<and>
-           (snd o h1)`(underS r x) = (snd o h2)`(underS r x)"
+    hence "\<forall>y\<in>underS r x. (fst \<circ> h1) y = (fst \<circ> h2) y \<and>
+                          (snd \<circ> h1) y = (snd \<circ> h2) y" by auto
+    hence "(fst \<circ> h1)`(underS r x) = (fst \<circ> h2)`(underS r x) \<and>
+           (snd \<circ> h1)`(underS r x) = (snd \<circ> h2)`(underS r x)"
       by (auto simp add: image_def)
     thus "H h1 x = H h2 x" by (simp add: H_def del: not_False_in_image_Ball)
   qed
   (* More constant definitions:  *)
   obtain h::"'a \<Rightarrow> 'a' * bool" and f::"'a \<Rightarrow> 'a'" and g::"'a \<Rightarrow> bool"
   where h_def: "h = wo_rel.worec r H" and
-        f_def: "f = fst o h" and g_def: "g = snd o h" by blast
+        f_def: "f = fst \<circ> h" and g_def: "g = snd \<circ> h" by blast
   obtain test where test_def:
   "test = (\<lambda> a. False \<notin> (g`(underS r a)) \<and> f`(underS r a) \<noteq> Field r')" by blast
   (*  *)
@@ -813,19 +813,19 @@
         EMB: "embed r r' f" and EMB': "embed r' r f'"
 shows "(\<forall>a \<in> Field r. f'(f a) = a) \<and> (\<forall>a' \<in> Field r'. f(f' a') = a')"
 proof-
-  have "embed r r (f' o f)" using assms
+  have "embed r r (f' \<circ> f)" using assms
   by(auto simp add: comp_embed)
   moreover have "embed r r id" using assms
   by (auto simp add: id_embed)
   ultimately have "\<forall>a \<in> Field r. f'(f a) = a"
-  using assms embed_unique[of r r "f' o f" id] id_def by auto
+  using assms embed_unique[of r r "f' \<circ> f" id] id_def by auto
   moreover
-  {have "embed r' r' (f o f')" using assms
+  {have "embed r' r' (f \<circ> f')" using assms
    by(auto simp add: comp_embed)
    moreover have "embed r' r' id" using assms
    by (auto simp add: id_embed)
    ultimately have "\<forall>a' \<in> Field r'. f(f' a') = a'"
-   using assms embed_unique[of r' r' "f o f'" id] id_def by auto
+   using assms embed_unique[of r' r' "f \<circ> f'" id] id_def by auto
   }
   ultimately show ?thesis by blast
 qed
@@ -836,11 +836,11 @@
 shows "bij_betw f (Field r) (Field r')"
 proof-
   let ?A = "Field r"  let ?A' = "Field r'"
-  have "embed r r (g o f) \<and> embed r' r' (f o g)"
+  have "embed r r (g \<circ> f) \<and> embed r' r' (f \<circ> g)"
   using assms by (auto simp add: comp_embed)
   hence 1: "(\<forall>a \<in> ?A. g(f a) = a) \<and> (\<forall>a' \<in> ?A'. f(g a') = a')"
-  using WELL id_embed[of r] embed_unique[of r r "g o f" id]
-        WELL' id_embed[of r'] embed_unique[of r' r' "f o g" id]
+  using WELL id_embed[of r] embed_unique[of r r "g \<circ> f" id]
+        WELL' id_embed[of r'] embed_unique[of r' r' "f \<circ> g" id]
         id_def by auto
   have 2: "(\<forall>a \<in> ?A. f a \<in> ?A') \<and> (\<forall>a' \<in> ?A'. g a' \<in> ?A)"
   using assms embed_Field[of r r' f] embed_Field[of r' r g] by blast
@@ -883,9 +883,9 @@
 lemma embedS_comp_embed:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "embedS r r' f" and EMB': "embed r' r'' f'"
-shows "embedS r r'' (f' o f)"
+shows "embedS r r'' (f' \<circ> f)"
 proof-
-  let ?g = "(f' o f)"  let ?h = "inv_into (Field r) ?g"
+  let ?g = "(f' \<circ> f)"  let ?h = "inv_into (Field r) ?g"
   have 1: "embed r r' f \<and> \<not> (bij_betw f (Field r) (Field r'))"
   using EMB by (auto simp add: embedS_def)
   hence 2: "embed r r'' ?g"
@@ -894,7 +894,7 @@
   {assume "bij_betw ?g (Field r) (Field r'')"
    hence "embed r'' r ?h" using 2 WELL
    by (auto simp add: inv_into_Field_embed_bij_betw)
-   hence "embed r' r (?h o f')" using WELL' EMB'
+   hence "embed r' r (?h \<circ> f')" using WELL' EMB'
    by (auto simp add: comp_embed)
    hence "bij_betw f (Field r) (Field r')" using WELL WELL' 1
    by (auto simp add: embed_bothWays_Field_bij_betw)
@@ -906,9 +906,9 @@
 lemma embed_comp_embedS:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "embed r r' f" and EMB': "embedS r' r'' f'"
-shows "embedS r r'' (f' o f)"
+shows "embedS r r'' (f' \<circ> f)"
 proof-
-  let ?g = "(f' o f)"  let ?h = "inv_into (Field r) ?g"
+  let ?g = "(f' \<circ> f)"  let ?h = "inv_into (Field r) ?g"
   have 1: "embed r' r'' f' \<and> \<not> (bij_betw f' (Field r') (Field r''))"
   using EMB' by (auto simp add: embedS_def)
   hence 2: "embed r r'' ?g"
@@ -917,7 +917,7 @@
   {assume "bij_betw ?g (Field r) (Field r'')"
    hence "embed r'' r ?h" using 2 WELL
    by (auto simp add: inv_into_Field_embed_bij_betw)
-   hence "embed r'' r' (f o ?h)" using WELL'' EMB
+   hence "embed r'' r' (f \<circ> ?h)" using WELL'' EMB
    by (auto simp add: comp_embed)
    hence "bij_betw f' (Field r') (Field r'')" using WELL' WELL'' 1
    by (auto simp add: embed_bothWays_Field_bij_betw)
@@ -929,28 +929,28 @@
 lemma embed_comp_iso:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "embed r r' f" and EMB': "iso r' r'' f'"
-shows "embed r r'' (f' o f)"
+shows "embed r r'' (f' \<circ> f)"
 using assms unfolding iso_def
 by (auto simp add: comp_embed)
 
 lemma iso_comp_embed:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "iso r r' f" and EMB': "embed r' r'' f'"
-shows "embed r r'' (f' o f)"
+shows "embed r r'' (f' \<circ> f)"
 using assms unfolding iso_def
 by (auto simp add: comp_embed)
 
 lemma embedS_comp_iso:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "embedS r r' f" and EMB': "iso r' r'' f'"
-shows "embedS r r'' (f' o f)"
+shows "embedS r r'' (f' \<circ> f)"
 using assms unfolding iso_def
 by (auto simp add: embedS_comp_embed)
 
 lemma iso_comp_embedS:
 assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
         and  EMB: "iso r r' f" and EMB': "embedS r' r'' f'"
-shows "embedS r r'' (f' o f)"
+shows "embedS r r'' (f' \<circ> f)"
 using assms unfolding iso_def  using embed_comp_embedS
 by (auto simp add: embed_comp_embedS)
 
--- a/src/HOL/Basic_BNFs.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Basic_BNFs.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -19,8 +19,8 @@
   "s = Inr x \<Longrightarrow> x \<in> setr s"
 
 lemma sum_set_defs[code]:
-  "setl = (\<lambda>x. case x of Inl z => {z} | _ => {})"
-  "setr = (\<lambda>x. case x of Inr z => {z} | _ => {})"
+  "setl = (\<lambda>x. case x of Inl z \<Rightarrow> {z} | _ \<Rightarrow> {})"
+  "setr = (\<lambda>x. case x of Inr z \<Rightarrow> {z} | _ \<Rightarrow> {})"
   by (auto simp: fun_eq_iff intro: setl.intros setr.intros elim: setl.cases setr.cases split: sum.splits)
 
 lemma rel_sum_simps[code, simp]:
@@ -52,7 +52,7 @@
   show "map_sum id id = id" by (rule map_sum.id)
 next
   fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r"
-  show "map_sum (g1 o f1) (g2 o f2) = map_sum g1 g2 o map_sum f1 f2"
+  show "map_sum (g1 \<circ> f1) (g2 \<circ> f2) = map_sum g1 g2 \<circ> map_sum f1 f2"
     by (rule map_sum.comp[symmetric])
 next
   fix x and f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" and g1 g2
@@ -66,11 +66,11 @@
   qed
 next
   fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r"
-  show "setl o map_sum f1 f2 = image f1 o setl"
+  show "setl \<circ> map_sum f1 f2 = image f1 \<circ> setl"
     by (rule ext, unfold o_apply) (simp add: sum_set_defs(1) split: sum.split)
 next
   fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r"
-  show "setr o map_sum f1 f2 = image f2 o setr"
+  show "setr \<circ> map_sum f1 f2 = image f2 \<circ> setr"
     by (rule ext, unfold o_apply) (simp add: sum_set_defs(2) split: sum.split)
 next
   show "card_order natLeq" by (rule natLeq_card_order)
@@ -149,7 +149,7 @@
   show "map_prod id id = id" by (rule map_prod.id)
 next
   fix f1 f2 g1 g2
-  show "map_prod (g1 o f1) (g2 o f2) = map_prod g1 g2 o map_prod f1 f2"
+  show "map_prod (g1 \<circ> f1) (g2 \<circ> f2) = map_prod g1 g2 \<circ> map_prod f1 f2"
     by (rule map_prod.comp[symmetric])
 next
   fix x f1 f2 g1 g2
@@ -157,11 +157,11 @@
   thus "map_prod f1 f2 x = map_prod g1 g2 x" by (cases x) simp
 next
   fix f1 f2
-  show "(\<lambda>x. {fst x}) o map_prod f1 f2 = image f1 o (\<lambda>x. {fst x})"
+  show "(\<lambda>x. {fst x}) \<circ> map_prod f1 f2 = image f1 \<circ> (\<lambda>x. {fst x})"
     by (rule ext, unfold o_apply) simp
 next
   fix f1 f2
-  show "(\<lambda>x. {snd x}) o map_prod f1 f2 = image f2 o (\<lambda>x. {snd x})"
+  show "(\<lambda>x. {snd x}) \<circ> map_prod f1 f2 = image f2 \<circ> (\<lambda>x. {snd x})"
     by (rule ext, unfold o_apply) simp
 next
   show "card_order natLeq" by (rule natLeq_card_order)
--- a/src/HOL/Computational_Algebra/Polynomial.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Computational_Algebra/Polynomial.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -1094,7 +1094,7 @@
 lemma poly_prod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)"
   by (induct A rule: infinite_finite_induct) simp_all
 
-lemma degree_prod_sum_le: "finite S \<Longrightarrow> degree (prod f S) \<le> sum (degree o f) S"
+lemma degree_prod_sum_le: "finite S \<Longrightarrow> degree (prod f S) \<le> sum (degree \<circ> f) S"
 proof (induct S rule: finite_induct)
   case empty
   then show ?case by simp
--- a/src/HOL/Computational_Algebra/Primes.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Computational_Algebra/Primes.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -271,7 +271,7 @@
 subsubsection \<open>Make prime naively executable\<close>
 
 lemma prime_nat_iff':
-  "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)"
+  "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. \<not> n dvd p)"
 proof safe
   assume "p > 1" and *: "\<forall>n\<in>{2..<p}. \<not>n dvd p"
   show "prime p" unfolding prime_nat_iff
@@ -286,7 +286,7 @@
 qed (auto simp: prime_nat_iff)
 
 lemma prime_int_iff':
-  "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)" (is "?lhs = ?rhs")
+  "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. \<not> n dvd p)" (is "?lhs = ?rhs")
 proof
   assume "?lhs"
   thus "?rhs"
@@ -352,9 +352,9 @@
 lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
 proof
   assume "finite {(p::nat). prime p}"
-  with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
+  with Max_ge have "(\<exists>b. (\<forall>x \<in> {(p::nat). prime p}. x \<le> b))"
     by auto
-  then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
+  then obtain b where "\<forall>(x::nat). prime x \<longrightarrow> x \<le> b"
     by auto
   with bigger_prime [of b] show False
     by auto
--- a/src/HOL/Conditionally_Complete_Lattices.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Conditionally_Complete_Lattices.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -430,7 +430,7 @@
   assumes "a < b" and "P a" and "\<not> P b"
   shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
              (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
-proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
+proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x}"], auto)
   show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
     by (rule cSup_upper, auto simp: bdd_above_def)
        (metis \<open>a < b\<close> \<open>\<not> P b\<close> linear less_le)
@@ -439,7 +439,7 @@
     apply (rule cSup_least)
     apply auto
     apply (metis less_le_not_le)
-    apply (metis \<open>a<b\<close> \<open>~ P b\<close> linear less_le)
+    apply (metis \<open>a<b\<close> \<open>\<not> P b\<close> linear less_le)
     done
 next
   fix x
@@ -454,7 +454,7 @@
     assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
     thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
       by (rule_tac cSup_upper, auto simp: bdd_above_def)
-         (metis \<open>a<b\<close> \<open>~ P b\<close> linear less_le)
+         (metis \<open>a<b\<close> \<open>\<not> P b\<close> linear less_le)
 qed
 
 end
--- a/src/HOL/Divides.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Divides.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -673,7 +673,7 @@
      "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
   by (fact div_add1_eq)
 
-lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
+lemma zmod_eq_0_iff: "(m mod d = 0) = (\<exists>q::int. m = d*q)"
 by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
 
 (* REVISIT: should this be generalized to all semiring_div types? *)
@@ -769,20 +769,20 @@
 text\<open>The proofs of the two lemmas below are essentially identical\<close>
 
 lemma split_pos_lemma:
- "0<k ==>
-    P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)"
+ "0<k \<Longrightarrow>
+    P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i j)"
   by auto
 
 lemma split_neg_lemma:
- "k<0 ==>
-    P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)"
+ "k<0 \<Longrightarrow>
+    P(n div k :: int)(n mod k) = (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i j)"
   by auto
 
 lemma split_zdiv:
  "P(n div k :: int) =
-  ((k = 0 --> P 0) &
-   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) &
-   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))"
+  ((k = 0 \<longrightarrow> P 0) \<and>
+   (0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i)) \<and>
+   (k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i)))"
   apply (cases "k = 0")
   apply simp
 apply (simp only: linorder_neq_iff)
@@ -792,9 +792,9 @@
 
 lemma split_zmod:
  "P(n mod k :: int) =
-  ((k = 0 --> P n) &
-   (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) &
-   (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))"
+  ((k = 0 \<longrightarrow> P n) \<and>
+   (0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P j)) \<and>
+   (k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P j)))"
 apply (case_tac "k=0", simp)
 apply (simp only: linorder_neq_iff)
 apply (erule disjE)
@@ -955,7 +955,7 @@
 by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)
 
 lemma nonneg1_imp_zdiv_pos_iff:
-  "(0::int) <= a \<Longrightarrow> (a div b > 0) = (a >= b & b>0)"
+  "(0::int) \<le> a \<Longrightarrow> (a div b > 0) = (a \<ge> b \<and> b>0)"
 apply rule
  apply rule
   using div_pos_pos_trivial[of a b]apply arith
--- a/src/HOL/Enum.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Enum.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -88,7 +88,7 @@
   [code del]: "enum_the = The"
 
 lemma [code]:
-  "The P = (case filter P enum of [x] => x | _ => enum_the P)"
+  "The P = (case filter P enum of [x] \<Rightarrow> x | _ \<Rightarrow> enum_the P)"
 proof -
   {
     fix a
@@ -212,16 +212,16 @@
   then have "\<exists> n. x : bacc r n"
   proof (induct x arbitrary: rule: acc.induct)
     case (accI x)
-    then have "\<forall>y. \<exists> n. (y, x) \<in> r --> y : bacc r n" by simp
+    then have "\<forall>y. \<exists> n. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n" by simp
     from choice[OF this] obtain n where n: "\<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r (n y)" ..
-    obtain n where "\<And>y. (y, x) : r \<Longrightarrow> y : bacc r n"
+    obtain n where "\<And>y. (y, x) \<in> r \<Longrightarrow> y \<in> bacc r n"
     proof
-      fix y assume y: "(y, x) : r"
+      fix y assume y: "(y, x) \<in> r"
       with n have "y : bacc r (n y)" by auto
-      moreover have "n y <= Max ((%(y, x). n y) ` r)"
+      moreover have "n y <= Max ((\<lambda>(y, x). n y) ` r)"
         using y \<open>finite r\<close> by (auto intro!: Max_ge)
       note bacc_mono[OF this, of r]
-      ultimately show "y : bacc r (Max ((%(y, x). n y) ` r))" by auto
+      ultimately show "y : bacc r (Max ((\<lambda>(y, x). n y) ` r))" by auto
     qed
     then show ?case
       by (auto simp add: Let_def intro!: exI[of _ "Suc n"])
@@ -297,13 +297,13 @@
 begin
 
 definition
-  "enum = map (\<lambda>ys. the o map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)"
+  "enum = map (\<lambda>ys. the \<circ> map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)"
 
 definition
-  "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
+  "enum_all P = all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"
 
 definition
-  "enum_ex P = ex_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
+  "enum_ex P = ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum bs))) (length (enum :: 'a list))"
 
 instance proof
   show "UNIV = set (enum :: ('a \<Rightarrow> 'b) list)"
@@ -368,17 +368,17 @@
 end
 
 lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list)
-  in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
+  in map (\<lambda>ys. the \<circ> map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
   by (simp add: enum_fun_def Let_def)
 
 lemma enum_all_fun_code [code]:
   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
-   in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
+   in all_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"
   by (simp only: enum_all_fun_def Let_def)
 
 lemma enum_ex_fun_code [code]:
   "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
-   in ex_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))"
+   in ex_n_lists (\<lambda>bs. P (the \<circ> map_of (zip enum_a bs))) (length enum_a))"
   by (simp only: enum_ex_fun_def Let_def)
 
 instantiation set :: (enum) enum
--- a/src/HOL/Fields.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Fields.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -400,10 +400,10 @@
 lemma inverse_mult_distrib [simp]:
   "inverse (a * b) = inverse a * inverse b"
 proof cases
-  assume "a \<noteq> 0 & b \<noteq> 0"
+  assume "a \<noteq> 0 \<and> b \<noteq> 0"
   thus ?thesis by (simp add: nonzero_inverse_mult_distrib ac_simps)
 next
-  assume "~ (a \<noteq> 0 & b \<noteq> 0)"
+  assume "\<not> (a \<noteq> 0 \<and> b \<noteq> 0)"
   thus ?thesis by force
 qed
 
@@ -596,7 +596,7 @@
   assumes invle: "inverse a \<le> inverse b" and apos: "0 < a"
   shows "b \<le> a"
 proof (rule classical)
-  assume "~ b \<le> a"
+  assume "\<not> b \<le> a"
   hence "a < b"  by (simp add: linorder_not_le)
   hence bpos: "0 < b"  by (blast intro: apos less_trans)
   hence "a * inverse a \<le> a * inverse b"
@@ -649,9 +649,9 @@
   assumes less: "a < b" and apos:  "0 < a"
   shows "inverse b < inverse a"
 proof (rule ccontr)
-  assume "~ inverse b < inverse a"
+  assume "\<not> inverse b < inverse a"
   hence "inverse a \<le> inverse b" by simp
-  hence "~ (a < b)"
+  hence "\<not> (a < b)"
     by (simp add: not_less inverse_le_imp_le [OF _ apos])
   thus False by (rule notE [OF _ less])
 qed
@@ -1142,19 +1142,19 @@
 text\<open>Simplify quotients that are compared with the value 1.\<close>
 
 lemma le_divide_eq_1:
-  "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
+  "(1 \<le> b / a) = ((0 < a \<and> a \<le> b) \<or> (a < 0 \<and> b \<le> a))"
 by (auto simp add: le_divide_eq)
 
 lemma divide_le_eq_1:
-  "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
+  "(b / a \<le> 1) = ((0 < a \<and> b \<le> a) \<or> (a < 0 \<and> a \<le> b) \<or> a=0)"
 by (auto simp add: divide_le_eq)
 
 lemma less_divide_eq_1:
-  "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
+  "(1 < b / a) = ((0 < a \<and> a < b) \<or> (a < 0 \<and> b < a))"
 by (auto simp add: less_divide_eq)
 
 lemma divide_less_eq_1:
-  "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
+  "(b / a < 1) = ((0 < a \<and> b < a) \<or> (a < 0 \<and> a < b) \<or> a=0)"
 by (auto simp add: divide_less_eq)
 
 lemma divide_nonneg_nonneg [simp]:
@@ -1208,11 +1208,11 @@
 by (auto simp add: divide_less_eq)
 
 lemma eq_divide_eq_1 [simp]:
-  "(1 = b/a) = ((a \<noteq> 0 & a = b))"
+  "(1 = b/a) = ((a \<noteq> 0 \<and> a = b))"
 by (auto simp add: eq_divide_eq)
 
 lemma divide_eq_eq_1 [simp]:
-  "(b/a = 1) = ((a \<noteq> 0 & a = b))"
+  "(b/a = 1) = ((a \<noteq> 0 \<and> a = b))"
 by (auto simp add: divide_eq_eq)
 
 lemma abs_div_pos: "0 < y ==>
@@ -1221,10 +1221,10 @@
   apply (simp add: order_less_imp_le)
 done
 
-lemma zero_le_divide_abs_iff [simp]: "(0 \<le> a / \<bar>b\<bar>) = (0 \<le> a | b = 0)"
+lemma zero_le_divide_abs_iff [simp]: "(0 \<le> a / \<bar>b\<bar>) = (0 \<le> a \<or> b = 0)"
 by (auto simp: zero_le_divide_iff)
 
-lemma divide_le_0_abs_iff [simp]: "(a / \<bar>b\<bar> \<le> 0) = (a \<le> 0 | b = 0)"
+lemma divide_le_0_abs_iff [simp]: "(a / \<bar>b\<bar> \<le> 0) = (a \<le> 0 \<or> b = 0)"
 by (auto simp: divide_le_0_iff)
 
 lemma field_le_mult_one_interval:
--- a/src/HOL/Groebner_Basis.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Groebner_Basis.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -19,8 +19,8 @@
   by blast+
 
 lemma dnf:
-  "(P & (Q | R)) = ((P&Q) | (P&R))"
-  "((Q | R) & P) = ((Q&P) | (R&P))"
+  "(P \<and> (Q \<or> R)) = ((P\<and>Q) \<or> (P\<and>R))"
+  "((Q \<or> R) \<and> P) = ((Q\<and>P) \<or> (R\<and>P))"
   "(P \<and> Q) = (Q \<and> P)"
   "(P \<or> Q) = (Q \<or> P)"
   by blast+
--- a/src/HOL/HOL.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/HOL.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -1627,12 +1627,10 @@
 
 subsection \<open>Other simple lemmas and lemma duplicates\<close>
 
-lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
-    (ALL x. Q x \<longrightarrow> P x) = (ALL x. Q x \<longrightarrow> P' x)"
+lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> (\<forall>x. Q x \<longrightarrow> P x) = (\<forall>x. Q x \<longrightarrow> P' x)"
   by auto
 
-lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
-    (EX x. Q x \<and> P x) = (EX x. Q x \<and> P' x)"
+lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> (\<exists>x. Q x \<and> P x) = (\<exists>x. Q x \<and> P' x)"
   by auto
 
 lemma ex1_eq [iff]: "\<exists>!x. x = t" "\<exists>!x. t = x"
--- a/src/HOL/Library/BNF_Corec.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/BNF_Corec.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -31,17 +31,17 @@
   unfolding BNF_Def.Grp_def by auto
 
 lemma sum_comp_cases:
-  assumes "f o Inl = g o Inl" and "f o Inr = g o Inr"
+  assumes "f \<circ> Inl = g \<circ> Inl" and "f \<circ> Inr = g \<circ> Inr"
   shows "f = g"
 proof (rule ext)
   fix a show "f a = g a"
     using assms unfolding comp_def fun_eq_iff by (cases a) auto
 qed
 
-lemma case_sum_Inl_Inr_L: "case_sum (f o Inl) (f o Inr) = f"
+lemma case_sum_Inl_Inr_L: "case_sum (f \<circ> Inl) (f \<circ> Inr) = f"
   by (metis case_sum_expand_Inr')
 
-lemma eq_o_InrI: "\<lbrakk>g o Inl = h; case_sum h f = g\<rbrakk> \<Longrightarrow> f = g o Inr"
+lemma eq_o_InrI: "\<lbrakk>g \<circ> Inl = h; case_sum h f = g\<rbrakk> \<Longrightarrow> f = g \<circ> Inr"
   by (auto simp: fun_eq_iff split: sum.splits)
 
 lemma id_bnf_o: "BNF_Composition.id_bnf \<circ> f = f"
@@ -63,7 +63,7 @@
 
 subsection \<open>Coinduction\<close>
 
-lemma eq_comp_compI: "a o b = f o x \<Longrightarrow> x o c = id \<Longrightarrow> f = a o (b o c)"
+lemma eq_comp_compI: "a \<circ> b = f \<circ> x \<Longrightarrow> x \<circ> c = id \<Longrightarrow> f = a \<circ> (b \<circ> c)"
   unfolding fun_eq_iff by simp
 
 lemma self_bounded_weaken_left: "(a :: 'a :: semilattice_inf) \<le> inf a b \<Longrightarrow> a \<le> b"
@@ -159,7 +159,7 @@
   by (simp add: antisym cong_gen_cong gen_cong_minimal leq_gen_cong)
 
 lemma gen_cong_rho:
-  "\<rho> = eval o f \<Longrightarrow> rel (gen_cong R) (f x) (f y) \<Longrightarrow> gen_cong R (\<rho> x) (\<rho> y)"
+  "\<rho> = eval \<circ> f \<Longrightarrow> rel (gen_cong R) (f x) (f y) \<Longrightarrow> gen_cong R (\<rho> x) (\<rho> y)"
   by (simp add: gen_cong_eval)
 lemma coinduction:
   assumes coind: "\<forall>R. R \<le> retr R \<longrightarrow> R \<le> op ="
@@ -181,7 +181,7 @@
   fixes rel eval rel' eval' retr emb
   assumes base: "cong rel eval retr"
   and step: "cong rel' eval' retr"
-  and emb: "eval' o emb = eval"
+  and emb: "eval' \<circ> emb = eval"
   and emb_transfer: "rel_fun (rel R) (rel' R) emb emb"
 begin
 
--- a/src/HOL/Library/Cardinality.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Cardinality.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -92,7 +92,7 @@
       unfolding bs[symmetric] distinct_card[OF distb] ..
     have ca: "CARD('a) = length as"
       unfolding as[symmetric] distinct_card[OF dista] ..
-    let ?xs = "map (\<lambda>ys. the o map_of (zip as ys)) (List.n_lists (length as) bs)"
+    let ?xs = "map (\<lambda>ys. the \<circ> map_of (zip as ys)) (List.n_lists (length as) bs)"
     have "UNIV = set ?xs"
     proof(rule UNIV_eq_I)
       fix f :: "'a \<Rightarrow> 'b"
--- a/src/HOL/Library/DAList.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/DAList.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -22,7 +22,7 @@
 typedef ('key, 'value) alist = "{xs :: ('key \<times> 'value) list. (distinct \<circ> map fst) xs}"
   morphisms impl_of Alist
 proof
-  show "[] \<in> {xs. (distinct o map fst) xs}"
+  show "[] \<in> {xs. (distinct \<circ> map fst) xs}"
     by simp
 qed
 
--- a/src/HOL/Library/Diagonal_Subsequence.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Diagonal_Subsequence.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -8,22 +8,22 @@
 
 locale subseqs =
   fixes P::"nat\<Rightarrow>(nat\<Rightarrow>nat)\<Rightarrow>bool"
-  assumes ex_subseq: "\<And>n s. strict_mono (s::nat\<Rightarrow>nat) \<Longrightarrow> \<exists>r'. strict_mono r' \<and> P n (s o r')"
+  assumes ex_subseq: "\<And>n s. strict_mono (s::nat\<Rightarrow>nat) \<Longrightarrow> \<exists>r'. strict_mono r' \<and> P n (s \<circ> r')"
 begin
 
-definition reduce where "reduce s n = (SOME r'::nat\<Rightarrow>nat. strict_mono r' \<and> P n (s o r'))"
+definition reduce where "reduce s n = (SOME r'::nat\<Rightarrow>nat. strict_mono r' \<and> P n (s \<circ> r'))"
 
 lemma subseq_reduce[intro, simp]:
   "strict_mono s \<Longrightarrow> strict_mono (reduce s n)"
   unfolding reduce_def by (rule someI2_ex[OF ex_subseq]) auto
 
 lemma reduce_holds:
-  "strict_mono s \<Longrightarrow> P n (s o reduce s n)"
+  "strict_mono s \<Longrightarrow> P n (s \<circ> reduce s n)"
   unfolding reduce_def by (rule someI2_ex[OF ex_subseq]) (auto simp: o_def)
 
 primrec seqseq :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   "seqseq 0 = id"
-| "seqseq (Suc n) = seqseq n o reduce (seqseq n) n"
+| "seqseq (Suc n) = seqseq n \<circ> reduce (seqseq n) n"
 
 lemma subseq_seqseq[intro, simp]: "strict_mono (seqseq n)"
 proof (induct n)
@@ -35,7 +35,7 @@
 lemma seqseq_holds:
   "P n (seqseq (Suc n))"
 proof -
-  have "P n (seqseq n o reduce (seqseq n) n)"
+  have "P n (seqseq n \<circ> reduce (seqseq n) n)"
     by (intro reduce_holds subseq_seqseq)
   thus ?thesis by simp
 qed
@@ -57,14 +57,14 @@
 
 primrec fold_reduce where
   "fold_reduce n 0 = id"
-| "fold_reduce n (Suc k) = fold_reduce n k o reduce (seqseq (n + k)) (n + k)"
+| "fold_reduce n (Suc k) = fold_reduce n k \<circ> reduce (seqseq (n + k)) (n + k)"
 
 lemma subseq_fold_reduce[intro, simp]: "strict_mono (fold_reduce n k)"
 proof (induct k)
   case (Suc k) from strict_mono_o[OF this subseq_reduce] show ?case by (simp add: o_def)
 qed (simp add: strict_mono_def)
 
-lemma ex_subseq_reduce_index: "seqseq (n + k) = seqseq n o fold_reduce n k"
+lemma ex_subseq_reduce_index: "seqseq (n + k) = seqseq n \<circ> fold_reduce n k"
   by (induct k) simp_all
 
 lemma seqseq_fold_reduce: "seqseq n = fold_reduce 0 n"
@@ -73,20 +73,20 @@
 lemma diagseq_fold_reduce: "diagseq n = fold_reduce 0 n n"
   using seqseq_fold_reduce by (simp add: diagseq_def)
 
-lemma fold_reduce_add: "fold_reduce 0 (m + n) = fold_reduce 0 m o fold_reduce m n"
+lemma fold_reduce_add: "fold_reduce 0 (m + n) = fold_reduce 0 m \<circ> fold_reduce m n"
   by (induct n) simp_all
 
-lemma diagseq_add: "diagseq (k + n) = (seqseq k o (fold_reduce k n)) (k + n)"
+lemma diagseq_add: "diagseq (k + n) = (seqseq k \<circ> (fold_reduce k n)) (k + n)"
 proof -
   have "diagseq (k + n) = fold_reduce 0 (k + n) (k + n)"
     by (simp add: diagseq_fold_reduce)
-  also have "\<dots> = (seqseq k o fold_reduce k n) (k + n)"
+  also have "\<dots> = (seqseq k \<circ> fold_reduce k n) (k + n)"
     unfolding fold_reduce_add seqseq_fold_reduce ..
   finally show ?thesis .
 qed
 
 lemma diagseq_sub:
-  assumes "m \<le> n" shows "diagseq n = (seqseq m o (fold_reduce m (n - m))) n"
+  assumes "m \<le> n" shows "diagseq n = (seqseq m \<circ> (fold_reduce m (n - m))) n"
   using diagseq_add[of m "n - m"] assms by simp
 
 lemma subseq_diagonal_rest: "strict_mono (\<lambda>x. fold_reduce k x (k + x))"
@@ -100,12 +100,12 @@
   finally show "?lhs < \<dots>" .
 qed
 
-lemma diagseq_seqseq: "diagseq o (op + k) = (seqseq k o (\<lambda>x. fold_reduce k x (k + x)))"
+lemma diagseq_seqseq: "diagseq \<circ> (op + k) = (seqseq k \<circ> (\<lambda>x. fold_reduce k x (k + x)))"
   by (auto simp: o_def diagseq_add)
 
 lemma diagseq_holds:
-  assumes subseq_stable: "\<And>r s n. strict_mono r \<Longrightarrow> P n s \<Longrightarrow> P n (s o r)"
-  shows "P k (diagseq o (op + (Suc k)))"
+  assumes subseq_stable: "\<And>r s n. strict_mono r \<Longrightarrow> P n s \<Longrightarrow> P n (s \<circ> r)"
+  shows "P k (diagseq \<circ> (op + (Suc k)))"
   unfolding diagseq_seqseq by (intro subseq_stable subseq_diagonal_rest seqseq_holds)
 
 end
--- a/src/HOL/Library/Extended.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Extended.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -23,7 +23,7 @@
 case_of_simps less_eq_extended_case: less_eq_extended.simps
 
 definition less_extended :: "'a extended \<Rightarrow> 'a extended \<Rightarrow> bool" where
-"((x::'a extended) < y) = (x \<le> y & \<not> y \<le> x)"
+"((x::'a extended) < y) = (x \<le> y \<and> \<not> y \<le> x)"
 
 instance
   by intro_classes (auto simp: less_extended_def less_eq_extended_case split: extended.splits)
--- a/src/HOL/Library/Extended_Nat.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Extended_Nat.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -602,9 +602,9 @@
 by (induct n) auto
 
 lemma enat_less_induct:
-  assumes prem: "!!n. \<forall>m::enat. m < n --> P m ==> P n" shows "P n"
+  assumes prem: "\<And>n. \<forall>m::enat. m < n \<longrightarrow> P m \<Longrightarrow> P n" shows "P n"
 proof -
-  have P_enat: "!!k. P (enat k)"
+  have P_enat: "\<And>k. P (enat k)"
     apply (rule nat_less_induct)
     apply (rule prem, clarify)
     apply (erule less_enatE, simp)
--- a/src/HOL/Library/Extended_Real.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Extended_Real.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -1138,7 +1138,7 @@
   }
   moreover
   {
-    assume "y = -\<infinity> | y = \<infinity>"
+    assume "y = -\<infinity> \<or> y = \<infinity>"
     then have ?thesis
       using assms[rule_format, of 1] by (cases x) auto
   }
@@ -2845,7 +2845,7 @@
 proof -
   {
     fix x
-    have "(real_of_ereal o ereal) x = id x"
+    have "(real_of_ereal \<circ> ereal) x = id x"
       by auto
   }
   then show ?thesis
--- a/src/HOL/Library/FSet.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/FSet.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -510,7 +510,7 @@
   by transfer auto
 
 lemma pfsubset_ffilter:
-  "(\<And>x. x |\<in>| A \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| A & \<not> P x & Q x) \<Longrightarrow>
+  "(\<And>x. x |\<in>| A \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| A \<and> \<not> P x \<and> Q x) \<Longrightarrow>
     ffilter P A |\<subset>| ffilter Q A"
   unfolding less_fset_def by (auto simp add: subset_ffilter eq_ffilter)
 
--- a/src/HOL/Library/Old_Datatype.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Old_Datatype.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -15,7 +15,7 @@
 
 subsection \<open>The datatype universe\<close>
 
-definition "Node = {p. EX f x k. p = (f :: nat => 'b + nat, x ::'a + nat) & f k = Inr 0}"
+definition "Node = {p. \<exists>f x k. p = (f :: nat => 'b + nat, x ::'a + nat) \<and> f k = Inr 0}"
 
 typedef ('a, 'b) node = "Node :: ((nat => 'b + nat) * ('a + nat)) set"
   morphisms Rep_Node Abs_Node
@@ -44,9 +44,9 @@
 
 (*Leaf nodes, with arbitrary or nat labels*)
 definition Leaf :: "'a => ('a, 'b) dtree"
-  where "Leaf == Atom o Inl"
+  where "Leaf == Atom \<circ> Inl"
 definition Numb :: "nat => ('a, 'b) dtree"
-  where "Numb == Atom o Inr"
+  where "Numb == Atom \<circ> Inr"
 
 (*Injections of the "disjoint sum"*)
 definition In0 :: "('a, 'b) dtree => ('a, 'b) dtree"
@@ -56,13 +56,13 @@
 
 (*Function spaces*)
 definition Lim :: "('b => ('a, 'b) dtree) => ('a, 'b) dtree"
-  where "Lim f == \<Union>{z. ? x. z = Push_Node (Inl x) ` (f x)}"
+  where "Lim f == \<Union>{z. \<exists>x. z = Push_Node (Inl x) ` (f x)}"
 
 (*the set of nodes with depth less than k*)
 definition ndepth :: "('a, 'b) node => nat"
   where "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)"
 definition ntrunc :: "[nat, ('a, 'b) dtree] => ('a, 'b) dtree"
-  where "ntrunc k N == {n. n:N & ndepth(n)<k}"
+  where "ntrunc k N == {n. n:N \<and> ndepth(n)<k}"
 
 (*products and sums for the "universe"*)
 definition uprod :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set"
@@ -72,10 +72,10 @@
 
 (*the corresponding eliminators*)
 definition Split :: "[[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
-  where "Split c M == THE u. EX x y. M = Scons x y & u = c x y"
+  where "Split c M == THE u. \<exists>x y. M = Scons x y \<and> u = c x y"
 
 definition Case :: "[[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c"
-  where "Case c d M == THE u. (EX x . M = In0(x) & u = c(x)) | (EX y . M = In1(y) & u = d(y))"
+  where "Case c d M == THE u. (\<exists>x . M = In0(x) \<and> u = c(x)) \<or> (\<exists>y . M = In1(y) \<and> u = d(y))"
 
 
 (** equality for the "universe" **)
@@ -207,7 +207,7 @@
     "[| Scons M N = Scons M' N';  [| M=M';  N=N' |] ==> P |] ==> P"
 by (iprover dest: Scons_inject1 Scons_inject2)
 
-lemma Scons_Scons_eq [iff]: "(Scons M N = Scons M' N') = (M=M' & N=N')"
+lemma Scons_Scons_eq [iff]: "(Scons M N = Scons M' N') = (M=M' \<and> N=N')"
 by (blast elim!: Scons_inject)
 
 (*** Distinctness involving Leaf and Numb ***)
@@ -241,7 +241,7 @@
 by (simp add: ndepth_def  Node_K0_I [THEN Abs_Node_inverse] Least_equality)
 
 lemma ndepth_Push_Node_aux:
-     "case_nat (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"
+     "case_nat (Inr (Suc i)) f k = Inr 0 \<longrightarrow> Suc(LEAST x. f x = Inr 0) \<le> k"
 apply (induct_tac "k", auto)
 apply (erule Least_le)
 done
@@ -385,7 +385,7 @@
 done
 
 lemma ntrunc_o_equality: 
-    "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"
+    "[| !!k. (ntrunc(k) \<circ> h1) = (ntrunc(k) \<circ> h2) |] ==> h1=h2"
 apply (rule ntrunc_equality [THEN ext])
 apply (simp add: fun_eq_iff) 
 done
--- a/src/HOL/Library/Old_Recdef.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Old_Recdef.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -13,8 +13,8 @@
 
 subsection \<open>Lemmas for TFL\<close>
 
-lemma tfl_wf_induct: "ALL R. wf R -->
-       (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"
+lemma tfl_wf_induct: "\<forall>R. wf R \<longrightarrow>
+       (\<forall>P. (\<forall>x. (\<forall>y. (y,x)\<in>R \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> (\<forall>x. P x))"
 apply clarify
 apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast)
 done
@@ -22,39 +22,39 @@
 lemma tfl_cut_def: "cut f r x \<equiv> (\<lambda>y. if (y,x) \<in> r then f y else undefined)"
   unfolding cut_def .
 
-lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)"
+lemma tfl_cut_apply: "\<forall>f R. (x,a)\<in>R \<longrightarrow> (cut f R a)(x) = f(x)"
 apply clarify
 apply (rule cut_apply, assumption)
 done
 
 lemma tfl_wfrec:
-     "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)"
+     "\<forall>M R f. (f=wfrec R M) \<longrightarrow> wf R \<longrightarrow> (\<forall>x. f x = M (cut f R x) x)"
 apply clarify
 apply (erule wfrec)
 done
 
-lemma tfl_eq_True: "(x = True) --> x"
+lemma tfl_eq_True: "(x = True) \<longrightarrow> x"
   by blast
 
-lemma tfl_rev_eq_mp: "(x = y) --> y --> x"
+lemma tfl_rev_eq_mp: "(x = y) \<longrightarrow> y \<longrightarrow> x"
   by blast
 
-lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)"
+lemma tfl_simp_thm: "(x \<longrightarrow> y) \<longrightarrow> (x = x') \<longrightarrow> (x' \<longrightarrow> y)"
   by blast
 
-lemma tfl_P_imp_P_iff_True: "P ==> P = True"
+lemma tfl_P_imp_P_iff_True: "P \<Longrightarrow> P = True"
   by blast
 
-lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)"
+lemma tfl_imp_trans: "(A \<longrightarrow> B) \<Longrightarrow> (B \<longrightarrow> C) \<Longrightarrow> (A \<longrightarrow> C)"
   by blast
 
-lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)"
+lemma tfl_disj_assoc: "(a \<or> b) \<or> c \<equiv> a \<or> (b \<or> c)"
   by simp
 
-lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R"
+lemma tfl_disjE: "P \<or> Q \<Longrightarrow> P \<longrightarrow> R \<Longrightarrow> Q \<longrightarrow> R \<Longrightarrow> R"
   by blast
 
-lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q"
+lemma tfl_exE: "\<exists>x. P x \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q \<Longrightarrow> Q"
   by blast
 
 ML_file "old_recdef.ML"
--- a/src/HOL/Library/Option_ord.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Option_ord.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -101,12 +101,12 @@
   fix z :: "'a option"
   assume H: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x"
   have "P None" by (rule H) simp
-  then have P_Some [case_names Some]: "P z" if "\<And>x. z = Some x \<Longrightarrow> (P o Some) x" for z
+  then have P_Some [case_names Some]: "P z" if "\<And>x. z = Some x \<Longrightarrow> (P \<circ> Some) x" for z
     using \<open>P None\<close> that by (cases z) simp_all
   show "P z"
   proof (cases z rule: P_Some)
     case (Some w)
-    show "(P o Some) w"
+    show "(P \<circ> Some) w"
     proof (induct rule: less_induct)
       case (less x)
       have "P (Some x)"
@@ -117,7 +117,7 @@
         proof (cases y rule: P_Some)
           case (Some v)
           with \<open>y < Some x\<close> have "v < x" by simp
-          with less show "(P o Some) v" .
+          with less show "(P \<circ> Some) v" .
         qed
       qed
       then show ?case by simp
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -15,11 +15,11 @@
 declare less_bool_def[abs_def, code_pred_inline]
 declare le_bool_def[abs_def, code_pred_inline]
 
-lemma min_bool_eq [code_pred_inline]: "(min :: bool => bool => bool) == (op &)"
+lemma min_bool_eq [code_pred_inline]: "(min :: bool => bool => bool) == (op \<and>)"
 by (rule eq_reflection) (auto simp add: fun_eq_iff min_def)
 
 lemma [code_pred_inline]: 
-  "((A::bool) ~= (B::bool)) = ((A & ~ B) | (B & ~ A))"
+  "((A::bool) \<noteq> (B::bool)) = ((A \<and> \<not> B) \<or> (B \<and> \<not> A))"
 by fast
 
 setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name Let}]\<close>
@@ -45,7 +45,7 @@
   by (simp add: fun_eq_iff)
 
 lemma subset_eq[code_pred_inline]:
-  "(P :: 'a => bool) < (Q :: 'a => bool) == ((\<exists>x. Q x \<and> (\<not> P x)) \<and> (\<forall> x. P x --> Q x))"
+  "(P :: 'a \<Rightarrow> bool) < (Q :: 'a \<Rightarrow> bool) \<equiv> ((\<exists>x. Q x \<and> (\<not> P x)) \<and> (\<forall>x. P x \<longrightarrow> Q x))"
   by (rule eq_reflection) (auto simp add: less_fun_def le_fun_def)
 
 lemma set_equality[code_pred_inline]:
--- a/src/HOL/Library/Product_Order.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Product_Order.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -220,11 +220,11 @@
 of two complete lattices:\<close>
 
 lemma INF_prod_alt_def:
-  "INFIMUM A f = (INFIMUM A (fst o f), INFIMUM A (snd o f))"
+  "INFIMUM A f = (INFIMUM A (fst \<circ> f), INFIMUM A (snd \<circ> f))"
   unfolding Inf_prod_def by simp
 
 lemma SUP_prod_alt_def:
-  "SUPREMUM A f = (SUPREMUM A (fst o f), SUPREMUM A (snd o f))"
+  "SUPREMUM A f = (SUPREMUM A (fst \<circ> f), SUPREMUM A (snd \<circ> f))"
   unfolding Sup_prod_def by simp
 
 
--- a/src/HOL/Library/RBT_Set.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/RBT_Set.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -548,8 +548,8 @@
 
 lemma subset_Coset_empty_Set_empty [code]:
   "Coset t1 \<le> Set t2 \<longleftrightarrow> (case (RBT.impl_of t1, RBT.impl_of t2) of 
-    (rbt.Empty, rbt.Empty) => False |
-    (_, _) => Code.abort (STR ''non_empty_trees'') (\<lambda>_. Coset t1 \<le> Set t2))"
+    (rbt.Empty, rbt.Empty) \<Rightarrow> False |
+    (_, _) \<Rightarrow> Code.abort (STR ''non_empty_trees'') (\<lambda>_. Coset t1 \<le> Set t2))"
 proof -
   have *: "\<And>t. RBT.impl_of t = rbt.Empty \<Longrightarrow> t = RBT rbt.Empty"
     by (subst(asm) RBT_inverse[symmetric]) (auto simp: impl_of_inject)
@@ -568,7 +568,7 @@
   by (auto simp add: card.eq_fold intro: finite_fold_fold_keys comp_fun_commute_const)
 
 lemma sum_Set [code]:
-  "sum f (Set xs) = fold_keys (plus o f) xs 0"
+  "sum f (Set xs) = fold_keys (plus \<circ> f) xs 0"
 proof -
   have "comp_fun_commute (\<lambda>x. op + (f x))"
     by standard (auto simp: ac_simps)
--- a/src/HOL/Library/Ramsey.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Ramsey.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -27,7 +27,7 @@
   (is "\<exists>r\<ge>1. ?R m n r")
 proof (induct k \<equiv> "m + n" arbitrary: m n)
   case 0
-  show ?case (is "EX r. ?Q r")
+  show ?case (is "\<exists>r. ?Q r")
   proof
     from 0 show "?Q 1"
       by (clarsimp simp: indep_def) (metis card.empty emptyE empty_subsetI)
@@ -35,7 +35,7 @@
 next
   case (Suc k)
   consider "m = 0 \<or> n = 0" | "m \<noteq> 0" "n \<noteq> 0" by auto
-  then show ?case (is "EX r. ?Q r")
+  then show ?case (is "\<exists>r. ?Q r")
   proof cases
     case 1
     then have "?Q 1"
--- a/src/HOL/Library/Stream.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Stream.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -251,7 +251,7 @@
 
 lemma sdrop_while_sdrop_LEAST:
   assumes "\<exists>n. P (s !! n)"
-  shows "sdrop_while (Not o P) s = sdrop (LEAST n. P (s !! n)) s"
+  shows "sdrop_while (Not \<circ> P) s = sdrop (LEAST n. P (s !! n)) s"
 proof -
   from assms obtain m where "P (s !! m)" "\<And>n. P (s !! n) \<Longrightarrow> m \<le> n"
     and *: "(LEAST n. P (s !! n)) = m" by atomize_elim (auto intro: LeastI Least_le)
@@ -266,8 +266,8 @@
 qed
 
 primcorec sfilter where
-  "shd (sfilter P s) = shd (sdrop_while (Not o P) s)"
-| "stl (sfilter P s) = sfilter P (stl (sdrop_while (Not o P) s))"
+  "shd (sfilter P s) = shd (sdrop_while (Not \<circ> P) s)"
+| "stl (sfilter P s) = sfilter P (stl (sdrop_while (Not \<circ> P) s))"
 
 lemma sfilter_Stream: "sfilter P (x ## s) = (if P x then x ## sfilter P s else sfilter P s)"
 proof (cases "P x")
--- a/src/HOL/Library/Sublist.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Sublist.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -1018,7 +1018,7 @@
   "subseq (xs @ zs) (ys @ zs) \<longleftrightarrow> subseq xs ys" (is "?l = ?r")
 proof
   { fix xs' ys' xs ys zs :: "'a list" assume "subseq xs' ys'"
-    then have "xs' = xs @ zs & ys' = ys @ zs \<longrightarrow> subseq xs ys"
+    then have "xs' = xs @ zs \<and> ys' = ys @ zs \<longrightarrow> subseq xs ys"
     proof (induct arbitrary: xs ys zs)
       case list_emb_Nil show ?case by simp
     next
--- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -1018,8 +1018,8 @@
           addsimps [@{thm power_divide}]
         val th =
           Thm.instantiate' [] [SOME d, SOME (Thm.dest_arg P)]
-            (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto}
-             else @{lemma "(d=0 --> P) & (d ~= (0::real) --> P) ==> P" by blast})
+            (if ord then @{lemma "(d=0 \<longrightarrow> P) \<and> (d>0 \<longrightarrow> P) \<and> (d<(0::real) \<longrightarrow> P) \<Longrightarrow> P" by auto}
+             else @{lemma "(d=0 \<longrightarrow> P) \<and> (d \<noteq> (0::real) \<longrightarrow> P) \<Longrightarrow> P" by blast})
       in resolve_tac ctxt [th] i THEN Simplifier.asm_full_simp_tac simp_ctxt i end));
 
 fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
--- a/src/HOL/Library/While_Combinator.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/While_Combinator.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -12,8 +12,8 @@
 subsection \<open>Partial version\<close>
 
 definition while_option :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
-"while_option b c s = (if (\<exists>k. ~ b ((c ^^ k) s))
-   then Some ((c ^^ (LEAST k. ~ b ((c ^^ k) s))) s)
+"while_option b c s = (if (\<exists>k. \<not> b ((c ^^ k) s))
+   then Some ((c ^^ (LEAST k. \<not> b ((c ^^ k) s))) s)
    else None)"
 
 theorem while_option_unfold[code]:
@@ -21,42 +21,42 @@
 proof cases
   assume "b s"
   show ?thesis
-  proof (cases "\<exists>k. ~ b ((c ^^ k) s)")
+  proof (cases "\<exists>k. \<not> b ((c ^^ k) s)")
     case True
-    then obtain k where 1: "~ b ((c ^^ k) s)" ..
+    then obtain k where 1: "\<not> b ((c ^^ k) s)" ..
     with \<open>b s\<close> obtain l where "k = Suc l" by (cases k) auto
-    with 1 have "~ b ((c ^^ l) (c s))" by (auto simp: funpow_swap1)
-    then have 2: "\<exists>l. ~ b ((c ^^ l) (c s))" ..
+    with 1 have "\<not> b ((c ^^ l) (c s))" by (auto simp: funpow_swap1)
+    then have 2: "\<exists>l. \<not> b ((c ^^ l) (c s))" ..
     from 1
-    have "(LEAST k. ~ b ((c ^^ k) s)) = Suc (LEAST l. ~ b ((c ^^ Suc l) s))"
+    have "(LEAST k. \<not> b ((c ^^ k) s)) = Suc (LEAST l. \<not> b ((c ^^ Suc l) s))"
       by (rule Least_Suc) (simp add: \<open>b s\<close>)
-    also have "... = Suc (LEAST l. ~ b ((c ^^ l) (c s)))"
+    also have "... = Suc (LEAST l. \<not> b ((c ^^ l) (c s)))"
       by (simp add: funpow_swap1)
     finally
     show ?thesis 
       using True 2 \<open>b s\<close> by (simp add: funpow_swap1 while_option_def)
   next
     case False
-    then have "~ (\<exists>l. ~ b ((c ^^ Suc l) s))" by blast
-    then have "~ (\<exists>l. ~ b ((c ^^ l) (c s)))"
+    then have "\<not> (\<exists>l. \<not> b ((c ^^ Suc l) s))" by blast
+    then have "\<not> (\<exists>l. \<not> b ((c ^^ l) (c s)))"
       by (simp add: funpow_swap1)
     with False  \<open>b s\<close> show ?thesis by (simp add: while_option_def)
   qed
 next
-  assume [simp]: "~ b s"
-  have least: "(LEAST k. ~ b ((c ^^ k) s)) = 0"
+  assume [simp]: "\<not> b s"
+  have least: "(LEAST k. \<not> b ((c ^^ k) s)) = 0"
     by (rule Least_equality) auto
   moreover 
-  have "\<exists>k. ~ b ((c ^^ k) s)" by (rule exI[of _ "0::nat"]) auto
+  have "\<exists>k. \<not> b ((c ^^ k) s)" by (rule exI[of _ "0::nat"]) auto
   ultimately show ?thesis unfolding while_option_def by auto 
 qed
 
 lemma while_option_stop2:
- "while_option b c s = Some t \<Longrightarrow> EX k. t = (c^^k) s \<and> \<not> b t"
+ "while_option b c s = Some t \<Longrightarrow> \<exists>k. t = (c^^k) s \<and> \<not> b t"
 apply(simp add: while_option_def split: if_splits)
 by (metis (lifting) LeastI_ex)
 
-lemma while_option_stop: "while_option b c s = Some t \<Longrightarrow> ~ b t"
+lemma while_option_stop: "while_option b c s = Some t \<Longrightarrow> \<not> b t"
 by(metis while_option_stop2)
 
 theorem while_option_rule:
@@ -65,13 +65,13 @@
     and init: "P s"
   shows "P t"
 proof -
-  define k where "k = (LEAST k. ~ b ((c ^^ k) s))"
+  define k where "k = (LEAST k. \<not> b ((c ^^ k) s))"
   from assms have t: "t = (c ^^ k) s"
     by (simp add: while_option_def k_def split: if_splits)    
   have 1: "ALL i<k. b ((c ^^ i) s)"
     by (auto simp: k_def dest: not_less_Least)
 
-  { fix i assume "i <= k" then have "P ((c ^^ i) s)"
+  { fix i assume "i \<le> k" then have "P ((c ^^ i) s)"
       by (induct i) (auto simp: init step 1) }
   thus "P t" by (auto simp: t)
 qed
@@ -243,8 +243,8 @@
 
 theorem wf_while_option_Some:
   assumes "wf {(t, s). (P s \<and> b s) \<and> t = c s}"
-  and "!!s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s)" and "P s"
-  shows "EX t. while_option b c s = Some t"
+  and "\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s)" and "P s"
+  shows "\<exists>t. while_option b c s = Some t"
 using assms(1,3)
 proof (induction s)
   case less thus ?case using assms(2)
@@ -264,8 +264,8 @@
 qed
 
 theorem measure_while_option_Some: fixes f :: "'s \<Rightarrow> nat"
-shows "(!!s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s) \<and> f(c s) < f s)
-  \<Longrightarrow> P s \<Longrightarrow> EX t. while_option b c s = Some t"
+shows "(\<And>s. P s \<Longrightarrow> b s \<Longrightarrow> P(c s) \<and> f(c s) < f s)
+  \<Longrightarrow> P s \<Longrightarrow> \<exists>t. while_option b c s = Some t"
 by(blast intro: wf_while_option_Some[OF wf_if_measure, of P b f])
 
 text\<open>Kleene iteration starting from the empty set and assuming some finite
--- a/src/HOL/Library/positivstellensatz.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Library/positivstellensatz.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -177,12 +177,12 @@
   if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb)  (* FIXME !? *)
   then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb;
 
-val pth = @{lemma "(((x::real) < y) == (y - x > 0))" and "((x <= y) == (y - x >= 0))" and
-     "((x = y) == (x - y = 0))" and "((~(x < y)) == (x - y >= 0))" and
-     "((~(x <= y)) == (x - y > 0))" and "((~(x = y)) == (x - y > 0 | -(x - y) > 0))"
+val pth = @{lemma "(((x::real) < y) \<equiv> (y - x > 0))" and "((x \<le> y) \<equiv> (y - x \<ge> 0))" and
+     "((x = y) \<equiv> (x - y = 0))" and "((\<not>(x < y)) \<equiv> (x - y \<ge> 0))" and
+     "((\<not>(x \<le> y)) \<equiv> (x - y > 0))" and "((\<not>(x = y)) \<equiv> (x - y > 0 \<or> -(x - y) > 0))"
   by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)};
 
-val pth_final = @{lemma "(~p ==> False) ==> p" by blast}
+val pth_final = @{lemma "(\<not>p \<Longrightarrow> False) \<Longrightarrow> p" by blast}
 val pth_add =
   @{lemma "(x = (0::real) ==> y = 0 ==> x + y = 0 )" and "( x = 0 ==> y >= 0 ==> x + y >= 0)" and
     "(x = 0 ==> y > 0 ==> x + y > 0)" and "(x >= 0 ==> y = 0 ==> x + y >= 0)" and
@@ -204,8 +204,8 @@
 
 val weak_dnf_simps =
   List.take (@{thms simp_thms}, 34) @
-    @{lemma "((P & (Q | R)) = ((P&Q) | (P&R)))" and "((Q | R) & P) = ((Q&P) | (R&P))" and
-      "(P & Q) = (Q & P)" and "((P | Q) = (Q | P))" by blast+};
+    @{lemma "((P \<and> (Q \<or> R)) = ((P\<and>Q) \<or> (P\<and>R)))" and "((Q \<or> R) \<and> P) = ((Q\<and>P) \<or> (R\<and>P))" and
+      "(P \<and> Q) = (Q \<and> P)" and "((P \<or> Q) = (Q \<or> P))" by blast+};
 
 (*
 val nnfD_simps =
@@ -214,65 +214,65 @@
     "((~(P = Q)) = ((P & ~ Q) | (~P & Q)) )" and "((~ ~(P)) = P)" by blast+};
 *)
 
-val choice_iff = @{lemma "(ALL x. EX y. P x y) = (EX f. ALL x. P x (f x))" by metis};
+val choice_iff = @{lemma "(\<forall>x. \<exists>y. P x y) = (\<exists>f. \<forall>x. P x (f x))" by metis};
 val prenex_simps =
   map (fn th => th RS sym)
     ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @
       @{thms "HOL.all_simps"(1-4)} @ @{thms "ex_simps"(1-4)});
 
 val real_abs_thms1 = @{lemma
-  "((-1 * \<bar>x::real\<bar> >= r) = (-1 * x >= r & 1 * x >= r))" and
-  "((-1 * \<bar>x\<bar> + a >= r) = (a + -1 * x >= r & a + 1 * x >= r))" and
-  "((a + -1 * \<bar>x\<bar> >= r) = (a + -1 * x >= r & a + 1 * x >= r))" and
-  "((a + -1 * \<bar>x\<bar> + b >= r) = (a + -1 * x + b >= r & a + 1 * x + b >= r))" and
-  "((a + b + -1 * \<bar>x\<bar> >= r) = (a + b + -1 * x >= r & a + b + 1 * x >= r))" and
-  "((a + b + -1 * \<bar>x\<bar> + c >= r) = (a + b + -1 * x + c >= r & a + b + 1 * x + c >= r))" and
-  "((-1 * max x y >= r) = (-1 * x >= r & -1 * y >= r))" and
-  "((-1 * max x y + a >= r) = (a + -1 * x >= r & a + -1 * y >= r))" and
-  "((a + -1 * max x y >= r) = (a + -1 * x >= r & a + -1 * y >= r))" and
-  "((a + -1 * max x y + b >= r) = (a + -1 * x + b >= r & a + -1 * y  + b >= r))" and
-  "((a + b + -1 * max x y >= r) = (a + b + -1 * x >= r & a + b + -1 * y >= r))" and
-  "((a + b + -1 * max x y + c >= r) = (a + b + -1 * x + c >= r & a + b + -1 * y  + c >= r))" and
-  "((1 * min x y >= r) = (1 * x >= r & 1 * y >= r))" and
-  "((1 * min x y + a >= r) = (a + 1 * x >= r & a + 1 * y >= r))" and
-  "((a + 1 * min x y >= r) = (a + 1 * x >= r & a + 1 * y >= r))" and
-  "((a + 1 * min x y + b >= r) = (a + 1 * x + b >= r & a + 1 * y  + b >= r))" and
-  "((a + b + 1 * min x y >= r) = (a + b + 1 * x >= r & a + b + 1 * y >= r))" and
-  "((a + b + 1 * min x y + c >= r) = (a + b + 1 * x + c >= r & a + b + 1 * y  + c >= r))" and
-  "((min x y >= r) = (x >= r &  y >= r))" and
-  "((min x y + a >= r) = (a + x >= r & a + y >= r))" and
-  "((a + min x y >= r) = (a + x >= r & a + y >= r))" and
-  "((a + min x y + b >= r) = (a + x + b >= r & a + y  + b >= r))" and
-  "((a + b + min x y >= r) = (a + b + x >= r & a + b + y >= r))" and
-  "((a + b + min x y + c >= r) = (a + b + x + c >= r & a + b + y + c >= r))" and
-  "((-1 * \<bar>x\<bar> > r) = (-1 * x > r & 1 * x > r))" and
-  "((-1 * \<bar>x\<bar> + a > r) = (a + -1 * x > r & a + 1 * x > r))" and
-  "((a + -1 * \<bar>x\<bar> > r) = (a + -1 * x > r & a + 1 * x > r))" and
-  "((a + -1 * \<bar>x\<bar> + b > r) = (a + -1 * x + b > r & a + 1 * x + b > r))" and
-  "((a + b + -1 * \<bar>x\<bar> > r) = (a + b + -1 * x > r & a + b + 1 * x > r))" and
-  "((a + b + -1 * \<bar>x\<bar> + c > r) = (a + b + -1 * x + c > r & a + b + 1 * x + c > r))" and
-  "((-1 * max x y > r) = ((-1 * x > r) & -1 * y > r))" and
-  "((-1 * max x y + a > r) = (a + -1 * x > r & a + -1 * y > r))" and
-  "((a + -1 * max x y > r) = (a + -1 * x > r & a + -1 * y > r))" and
-  "((a + -1 * max x y + b > r) = (a + -1 * x + b > r & a + -1 * y  + b > r))" and
-  "((a + b + -1 * max x y > r) = (a + b + -1 * x > r & a + b + -1 * y > r))" and
-  "((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r & a + b + -1 * y  + c > r))" and
-  "((min x y > r) = (x > r &  y > r))" and
-  "((min x y + a > r) = (a + x > r & a + y > r))" and
-  "((a + min x y > r) = (a + x > r & a + y > r))" and
-  "((a + min x y + b > r) = (a + x + b > r & a + y  + b > r))" and
-  "((a + b + min x y > r) = (a + b + x > r & a + b + y > r))" and
-  "((a + b + min x y + c > r) = (a + b + x + c > r & a + b + y + c > r))"
+  "((-1 * \<bar>x::real\<bar> \<ge> r) = (-1 * x \<ge> r \<and> 1 * x \<ge> r))" and
+  "((-1 * \<bar>x\<bar> + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and
+  "((a + -1 * \<bar>x\<bar> \<ge> r) = (a + -1 * x \<ge> r \<and> a + 1 * x \<ge> r))" and
+  "((a + -1 * \<bar>x\<bar> + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + 1 * x + b \<ge> r))" and
+  "((a + b + -1 * \<bar>x\<bar> \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + 1 * x \<ge> r))" and
+  "((a + b + -1 * \<bar>x\<bar> + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + 1 * x + c \<ge> r))" and
+  "((-1 * max x y \<ge> r) = (-1 * x \<ge> r \<and> -1 * y \<ge> r))" and
+  "((-1 * max x y + a \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and
+  "((a + -1 * max x y \<ge> r) = (a + -1 * x \<ge> r \<and> a + -1 * y \<ge> r))" and
+  "((a + -1 * max x y + b \<ge> r) = (a + -1 * x + b \<ge> r \<and> a + -1 * y  + b \<ge> r))" and
+  "((a + b + -1 * max x y \<ge> r) = (a + b + -1 * x \<ge> r \<and> a + b + -1 * y \<ge> r))" and
+  "((a + b + -1 * max x y + c \<ge> r) = (a + b + -1 * x + c \<ge> r \<and> a + b + -1 * y  + c \<ge> r))" and
+  "((1 * min x y \<ge> r) = (1 * x \<ge> r \<and> 1 * y \<ge> r))" and
+  "((1 * min x y + a \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and
+  "((a + 1 * min x y \<ge> r) = (a + 1 * x \<ge> r \<and> a + 1 * y \<ge> r))" and
+  "((a + 1 * min x y + b \<ge> r) = (a + 1 * x + b \<ge> r \<and> a + 1 * y  + b \<ge> r))" and
+  "((a + b + 1 * min x y \<ge> r) = (a + b + 1 * x \<ge> r \<and> a + b + 1 * y \<ge> r))" and
+  "((a + b + 1 * min x y + c \<ge> r) = (a + b + 1 * x + c \<ge> r \<and> a + b + 1 * y  + c \<ge> r))" and
+  "((min x y \<ge> r) = (x \<ge> r \<and> y \<ge> r))" and
+  "((min x y + a \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and
+  "((a + min x y \<ge> r) = (a + x \<ge> r \<and> a + y \<ge> r))" and
+  "((a + min x y + b \<ge> r) = (a + x + b \<ge> r \<and> a + y  + b \<ge> r))" and
+  "((a + b + min x y \<ge> r) = (a + b + x \<ge> r \<and> a + b + y \<ge> r))" and
+  "((a + b + min x y + c \<ge> r) = (a + b + x + c \<ge> r \<and> a + b + y + c \<ge> r))" and
+  "((-1 * \<bar>x\<bar> > r) = (-1 * x > r \<and> 1 * x > r))" and
+  "((-1 * \<bar>x\<bar> + a > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and
+  "((a + -1 * \<bar>x\<bar> > r) = (a + -1 * x > r \<and> a + 1 * x > r))" and
+  "((a + -1 * \<bar>x\<bar> + b > r) = (a + -1 * x + b > r \<and> a + 1 * x + b > r))" and
+  "((a + b + -1 * \<bar>x\<bar> > r) = (a + b + -1 * x > r \<and> a + b + 1 * x > r))" and
+  "((a + b + -1 * \<bar>x\<bar> + c > r) = (a + b + -1 * x + c > r \<and> a + b + 1 * x + c > r))" and
+  "((-1 * max x y > r) = ((-1 * x > r) \<and> -1 * y > r))" and
+  "((-1 * max x y + a > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and
+  "((a + -1 * max x y > r) = (a + -1 * x > r \<and> a + -1 * y > r))" and
+  "((a + -1 * max x y + b > r) = (a + -1 * x + b > r \<and> a + -1 * y  + b > r))" and
+  "((a + b + -1 * max x y > r) = (a + b + -1 * x > r \<and> a + b + -1 * y > r))" and
+  "((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r \<and> a + b + -1 * y  + c > r))" and
+  "((min x y > r) = (x > r \<and> y > r))" and
+  "((min x y + a > r) = (a + x > r \<and> a + y > r))" and
+  "((a + min x y > r) = (a + x > r \<and> a + y > r))" and
+  "((a + min x y + b > r) = (a + x + b > r \<and> a + y  + b > r))" and
+  "((a + b + min x y > r) = (a + b + x > r \<and> a + b + y > r))" and
+  "((a + b + min x y + c > r) = (a + b + x + c > r \<and> a + b + y + c > r))"
   by auto};
 
-val abs_split' = @{lemma "P \<bar>x::'a::linordered_idom\<bar> == (x >= 0 & P x | x < 0 & P (-x))"
+val abs_split' = @{lemma "P \<bar>x::'a::linordered_idom\<bar> == (x \<ge> 0 \<and> P x \<or> x < 0 \<and> P (-x))"
   by (atomize (full)) (auto split: abs_split)};
 
-val max_split = @{lemma "P (max x y) == ((x::'a::linorder) <= y & P y | x > y & P x)"
-  by (atomize (full)) (cases "x <= y", auto simp add: max_def)};
+val max_split = @{lemma "P (max x y) \<equiv> ((x::'a::linorder) \<le> y \<and> P y \<or> x > y \<and> P x)"
+  by (atomize (full)) (cases "x \<le> y", auto simp add: max_def)};
 
-val min_split = @{lemma "P (min x y) == ((x::'a::linorder) <= y & P x | x > y & P y)"
-  by (atomize (full)) (cases "x <= y", auto simp add: min_def)};
+val min_split = @{lemma "P (min x y) \<equiv> ((x::'a::linorder) \<le> y \<and> P x \<or> x > y \<and> P y)"
+  by (atomize (full)) (cases "x \<le> y", auto simp add: min_def)};
 
 
          (* Miscellaneous *)
--- a/src/HOL/Limited_Sequence.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Limited_Sequence.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -171,8 +171,8 @@
 definition neg_not_seq :: "unit pos_dseq \<Rightarrow> unit neg_dseq"
 where
   "neg_not_seq x = (\<lambda>i. case Lazy_Sequence.yield (x i) of
-    None => Lazy_Sequence.hb_single ()
-  | Some ((), xq) => Lazy_Sequence.empty)"
+    None \<Rightarrow> Lazy_Sequence.hb_single ()
+  | Some ((), xq) \<Rightarrow> Lazy_Sequence.empty)"
 
 
 ML \<open>
--- a/src/HOL/Limits.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Limits.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -2365,7 +2365,7 @@
   
 lemma uniformly_continuous_imp_Cauchy_continuous:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space"
-  shows "\<lbrakk>uniformly_continuous_on S f; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f o \<sigma>)"
+  shows "\<lbrakk>uniformly_continuous_on S f; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)"
   by (simp add: uniformly_continuous_on_def Cauchy_def) meson
 
 lemma (in bounded_linear) isUCont: "isUCont f"
--- a/src/HOL/List.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/List.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -147,7 +147,7 @@
 primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
 "zip xs [] = []" |
 zip_Cons: "zip xs (y # ys) =
-  (case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
+  (case xs of [] \<Rightarrow> [] | z # zs \<Rightarrow> (z, y) # zip zs ys)"
   \<comment> \<open>Warning: simpset does not contain this definition, but separate
        theorems for \<open>xs = []\<close> and \<open>xs = z # zs\<close>\<close>
 
@@ -182,7 +182,7 @@
 "find P (x#xs) = (if P x then Some x else find P xs)"
 
 text \<open>In the context of multisets, \<open>count_list\<close> is equivalent to
-  @{term "count o mset"} and it it advisable to use the latter.\<close>
+  @{term "count \<circ> mset"} and it it advisable to use the latter.\<close>
 primrec count_list :: "'a list \<Rightarrow> 'a \<Rightarrow> nat" where
 "count_list [] y = 0" |
 "count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)"
@@ -190,9 +190,9 @@
 definition
    "extract" :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> ('a list * 'a * 'a list) option"
 where "extract P xs =
-  (case dropWhile (Not o P) xs of
+  (case dropWhile (Not \<circ> P) xs of
      [] \<Rightarrow> None |
-     y#ys \<Rightarrow> Some(takeWhile (Not o P) xs, y, ys))"
+     y#ys \<Rightarrow> Some(takeWhile (Not \<circ> P) xs, y, ys))"
 
 hide_const (open) "extract"
 
@@ -783,10 +783,10 @@
 lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
 by (induct xs) auto
 
-lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
+lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (\<exists>x. xs = [x])"
 by (cases xs) auto
 
-lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
+lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (\<exists>x. xs = [x])"
 by (cases xs) auto
 
 lemma length_induct:
@@ -987,7 +987,7 @@
 done
 
 lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
-  (EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
+  (\<exists>us. xs = zs @ us \<and> us @ ys = ts \<or> xs @ us = zs \<and> ys = us @ ts)"
 apply (induct xs arbitrary: ys zs ts)
  apply fastforce
 apply(case_tac zs)
@@ -1019,7 +1019,7 @@
 lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
 by (simp add: hd_append split: list.split)
 
-lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
+lemma tl_append: "tl (xs @ ys) = (case xs of [] \<Rightarrow> tl ys | z#zs \<Rightarrow> zs @ ys)"
 by (simp split: list.split)
 
 lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
@@ -1027,11 +1027,11 @@
 
 
 lemma Cons_eq_append_conv: "x#xs = ys@zs =
- (ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
+ (ys = [] \<and> x#xs = zs \<or> (\<exists>ys'. x#ys' = ys \<and> xs = ys'@zs))"
 by(cases ys) auto
 
 lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
- (ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))"
+ (ys = [] \<and> zs = x#xs \<or> (\<exists>ys'. ys = x#ys' \<and> ys'@zs = xs))"
 by(cases ys) auto
 
 lemma longest_common_prefix:
@@ -1112,7 +1112,7 @@
 lemma map_tl: "map f (tl xs) = tl (map f xs)"
 by (cases xs) simp_all
 
-lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
+lemma map_ext: "(\<And>x. x \<in> set xs \<longrightarrow> f x = g x) ==> map f xs = map g xs"
 by (induct xs) simp_all
 
 lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
@@ -1124,7 +1124,7 @@
 lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
 by (induct xs) auto
 
-lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)"
+lemma map_comp_map[simp]: "((map f) \<circ> (map g)) = map(f \<circ> g)"
 by (rule ext) simp
 
 lemma rev_map: "rev (map f xs) = map f (rev xs)"
@@ -1156,7 +1156,7 @@
 declare map_eq_Cons_D [dest!]  Cons_eq_map_D [dest!]
 
 lemma ex_map_conv:
-  "(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
+  "(\<exists>xs. ys = map f xs) = (\<forall>y \<in> set ys. \<exists>x. y = f x)"
 by(induct ys, auto simp add: Cons_eq_map_conv)
 
 lemma map_eq_imp_length_eq:
@@ -1380,7 +1380,7 @@
   "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
   by (auto dest!: split_list_last)
 
-lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs & P x"
+lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs \<and> P x"
 proof (induct xs)
   case Nil thus ?case by simp
 next
@@ -1450,7 +1450,7 @@
   by rule (erule split_list_last_prop, auto)
 
 
-lemma finite_list: "finite A ==> EX xs. set xs = A"
+lemma finite_list: "finite A \<Longrightarrow> \<exists>xs. set xs = A"
   by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2))
 
 lemma card_length: "card (set xs) \<le> length xs"
@@ -1481,7 +1481,7 @@
 by (induct xs) (auto simp add: le_SucI)
 
 lemma sum_length_filter_compl:
-  "length(filter P xs) + length(filter (%x. ~P x) xs) = length xs"
+  "length(filter P xs) + length(filter (\<lambda>x. \<not>P x) xs) = length xs"
 by(induct xs) simp_all
 
 lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
@@ -1500,18 +1500,18 @@
 apply simp
 done
 
-lemma filter_map: "filter P (map f xs) = map f (filter (P o f) xs)"
+lemma filter_map: "filter P (map f xs) = map f (filter (P \<circ> f) xs)"
 by (induct xs) simp_all
 
 lemma length_filter_map[simp]:
-  "length (filter P (map f xs)) = length(filter (P o f) xs)"
+  "length (filter P (map f xs)) = length(filter (P \<circ> f) xs)"
 by (simp add:filter_map)
 
 lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
 by auto
 
 lemma length_filter_less:
-  "\<lbrakk> x : set xs; ~ P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
+  "\<lbrakk> x \<in> set xs; \<not> P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
 proof (induct xs)
   case Nil thus ?case by simp
 next
@@ -1522,12 +1522,12 @@
 qed
 
 lemma length_filter_conv_card:
-  "length(filter p xs) = card{i. i < length xs & p(xs!i)}"
+  "length(filter p xs) = card{i. i < length xs \<and> p(xs!i)}"
 proof (induct xs)
   case Nil thus ?case by simp
 next
   case (Cons x xs)
-  let ?S = "{i. i < length xs & p(xs!i)}"
+  let ?S = "{i. i < length xs \<and> p(xs!i)}"
   have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
   show ?case (is "?l = card ?S'")
   proof (cases)
@@ -1621,7 +1621,7 @@
 lemma partition_filter1: "fst (partition P xs) = filter P xs"
 by (induct xs) (auto simp add: Let_def split_def)
 
-lemma partition_filter2: "snd (partition P xs) = filter (Not o P) xs"
+lemma partition_filter2: "snd (partition P xs) = filter (Not \<circ> P) xs"
 by (induct xs) (auto simp add: Let_def split_def)
 
 lemma partition_P:
@@ -1643,7 +1643,7 @@
 qed
 
 lemma partition_filter_conv[simp]:
-  "partition f xs = (filter f xs,filter (Not o f) xs)"
+  "partition f xs = (filter f xs,filter (Not \<circ> f) xs)"
 unfolding partition_filter2[symmetric]
 unfolding partition_filter1[symmetric] by simp
 
@@ -1784,7 +1784,7 @@
 by (auto simp add: set_conv_nth)
 
 lemma all_set_conv_all_nth:
-  "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
+  "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))"
 by (auto simp add: set_conv_nth)
 
 lemma rev_nth:
@@ -1808,18 +1808,18 @@
 qed
 
 lemma Skolem_list_nth:
-  "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
-  (is "_ = (EX xs. ?P k xs)")
+  "(\<forall>i<k. \<exists>x. P i x) = (\<exists>xs. size xs = k \<and> (\<forall>i<k. P i (xs!i)))"
+  (is "_ = (\<exists>xs. ?P k xs)")
 proof(induct k)
   case 0 show ?case by simp
 next
   case (Suc k)
-  show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
+  show ?case (is "?L = ?R" is "_ = (\<exists>xs. ?P' xs)")
   proof
     assume "?R" thus "?L" using Suc by auto
   next
     assume "?L"
-    with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
+    with Suc obtain x xs where "?P k xs \<and> P k x" by (metis less_Suc_eq)
     hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
     thus "?R" ..
   qed
@@ -1967,7 +1967,7 @@
 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))"
+  "x \<in> set (butlast xs) \<or> x \<in> set (butlast ys) \<Longrightarrow> x \<in> set (butlast (xs @ ys))"
 by (auto dest: in_set_butlastD simp add: butlast_append)
 
 lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
@@ -2004,7 +2004,7 @@
 by (induct xs) simp_all
 
 lemma snoc_eq_iff_butlast:
-  "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
+  "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] \<and> butlast ys = xs \<and> last ys = x)"
 by fastforce
 
 corollary longest_common_suffix:
@@ -2036,7 +2036,7 @@
 
 declare take_Cons [simp del] and drop_Cons [simp del]
 
-lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
+lemma take_Suc: "xs \<noteq> [] \<Longrightarrow> take (Suc n) xs = hd xs # take n (tl xs)"
 by(clarsimp simp add:neq_Nil_conv)
 
 lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
@@ -2285,7 +2285,7 @@
 by (induct xs) auto
 
 lemma takeWhile_append1 [simp]:
-  "[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
+  "\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> takeWhile P (xs @ ys) = takeWhile P xs"
 by (induct xs) auto
 
 lemma takeWhile_append2 [simp]:
@@ -2306,7 +2306,7 @@
 by (induct xs) auto
 
 lemma dropWhile_append1 [simp]:
-  "[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
+  "\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
 by (induct xs) auto
 
 lemma dropWhile_append2 [simp]:
@@ -2336,7 +2336,7 @@
 by(induct xs, auto)
 
 lemma dropWhile_eq_Cons_conv:
-  "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
+  "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \<and> \<not> P y)"
 by(induct xs, auto)
 
 lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
@@ -2693,7 +2693,7 @@
 
 lemma list_all2_append1:
   "list_all2 P (xs @ ys) zs =
-  (EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
+  (\<exists>us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
     list_all2 P xs us \<and> list_all2 P ys vs)"
 apply (simp add: list_all2_iff zip_append1)
 apply (rule iffI)
@@ -2705,7 +2705,7 @@
 
 lemma list_all2_append2:
   "list_all2 P xs (ys @ zs) =
-  (EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
+  (\<exists>us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
     list_all2 P us ys \<and> list_all2 P vs zs)"
 apply (simp add: list_all2_iff zip_append2)
 apply (rule iffI)
@@ -2904,7 +2904,7 @@
 lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \<circ> fold f xs"
 by (induct xs) simp_all
 
-lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g o f) xs"
+lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \<circ> f) xs"
 by (induct xs) simp_all
 
 lemma fold_filter:
@@ -3044,7 +3044,7 @@
 lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
 by (simp add: foldl_conv_fold)
 
-lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g o f) xs a"
+lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \<circ> f) xs a"
 by (simp add: foldr_conv_fold fold_map rev_map)
 
 lemma foldr_filter:
@@ -3075,7 +3075,7 @@
 by(induct j)simp_all
 
 lemma upt_eq_Cons_conv:
- "([i..<j] = x#xs) = (i < j & i = x & [i+1..<j] = xs)"
+ "([i..<j] = x#xs) = (i < j \<and> i = x \<and> [i+1..<j] = xs)"
 apply(induct j arbitrary: x xs)
  apply simp
 apply(clarsimp simp add: append_eq_Cons_conv)
@@ -3147,8 +3147,8 @@
 
 
 lemma nth_take_lemma:
-  "k <= length xs ==> k <= length ys ==>
-     (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys"
+  "k \<le> length xs \<Longrightarrow> k \<le> length ys \<Longrightarrow>
+     (\<And>i. i < k \<longrightarrow> xs!i = ys!i) \<Longrightarrow> 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 \<open>Both lists must be non-empty\<close>
@@ -3294,7 +3294,7 @@
 lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
 by (metis distinct_remdups distinct_remdups_id)
 
-lemma finite_distinct_list: "finite A \<Longrightarrow> EX xs. set xs = A & distinct xs"
+lemma finite_distinct_list: "finite A \<Longrightarrow> \<exists>xs. set xs = A \<and> distinct xs"
 by (metis distinct_remdups finite_list set_remdups)
 
 lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
@@ -3319,7 +3319,7 @@
 by (induct xs) auto
 
 lemma distinct_map:
-  "distinct(map f xs) = (distinct xs & inj_on f (set xs))"
+  "distinct(map f xs) = (distinct xs \<and> inj_on f (set xs))"
 by (induct xs) auto
 
 lemma distinct_map_filter:
@@ -3379,7 +3379,7 @@
 sometimes it is useful.\<close>
 
 lemma distinct_conv_nth:
-"distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j --> xs!i \<noteq> xs!j)"
+"distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j \<longrightarrow> xs!i \<noteq> xs!j)"
 apply (induct xs, simp, simp)
 apply (rule iffI, clarsimp)
  apply (case_tac i)
@@ -3451,7 +3451,7 @@
 lemma distinct_length_filter: "distinct xs \<Longrightarrow> length (filter P xs) = card ({x. P x} Int set xs)"
 by (induct xs) (auto)
 
-lemma not_distinct_decomp: "~ distinct ws ==> EX xs ys zs y. ws = xs@[y]@ys@[y]@zs"
+lemma not_distinct_decomp: "\<not> distinct ws \<Longrightarrow> \<exists>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)
@@ -3563,7 +3563,7 @@
 by force
 
 lemma remdups_adj_altdef: "(remdups_adj xs = ys) \<longleftrightarrow>
-  (\<exists>f::nat => nat. mono f & f ` {0 ..< size xs} = {0 ..< size ys}
+  (\<exists>f::nat => nat. mono f \<and> f ` {0 ..< size xs} = {0 ..< size ys}
     \<and> (\<forall>i < size xs. xs!i = ys!(f i))
     \<and> (\<forall>i. i + 1 < size xs \<longrightarrow> (xs!i = xs!(i+1) \<longleftrightarrow> f i = f(i+1))))" (is "?L \<longleftrightarrow> (\<exists>f. ?p f xs ys)")
 proof
@@ -3645,7 +3645,7 @@
     proof cases
       assume "x1 = x2"
 
-      let ?f' = "f o Suc"
+      let ?f' = "f \<circ> Suc"
 
       have "remdups_adj (x1 # xs) = ys"
       proof (intro "3.hyps" exI conjI impI allI)
@@ -4114,15 +4114,15 @@
 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
 by auto
 
-lemma in_set_replicate[simp]: "(x : set (replicate n y)) = (x = y & n \<noteq> 0)"
+lemma in_set_replicate[simp]: "(x \<in> set (replicate n y)) = (x = y \<and> n \<noteq> 0)"
 by (simp add: set_replicate_conv_if)
 
 lemma Ball_set_replicate[simp]:
-  "(ALL x : set(replicate n a). P x) = (P a | n=0)"
+  "(\<forall>x \<in> set(replicate n a). P x) = (P a \<or> n=0)"
 by(simp add: set_replicate_conv_if)
 
 lemma Bex_set_replicate[simp]:
-  "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
+  "(\<exists>x \<in> set(replicate n a). P x) = (P a \<and> n\<noteq>0)"
 by(simp add: set_replicate_conv_if)
 
 lemma replicate_append_same:
@@ -4144,7 +4144,7 @@
 by (induct n) auto
 
 lemma replicate_eq_replicate[simp]:
-  "(replicate m x = replicate n y) \<longleftrightarrow> (m=n & (m\<noteq>0 \<longrightarrow> x=y))"
+  "(replicate m x = replicate n y) \<longleftrightarrow> (m=n \<and> (m\<noteq>0 \<longrightarrow> x=y))"
 apply(induct m arbitrary: n)
  apply simp
 apply(induct_tac n)
@@ -4306,7 +4306,7 @@
 by(simp add:rotate_def)
 
 lemma rotate_add:
-  "rotate (m+n) = rotate m o rotate n"
+  "rotate (m+n) = rotate m \<circ> rotate n"
 by(simp add:rotate_def funpow_add)
 
 lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs"
@@ -4469,7 +4469,7 @@
   case Nil thus ?case by simp
 next
   case (Cons a xs)
-  then have "!x. x: set xs \<longrightarrow> x \<noteq> a" by auto
+  then have "\<forall>x. x \<in> set xs \<longrightarrow> x \<noteq> a" by auto
   with Cons show ?case by(simp add: nths_Cons cong:filter_cong)
 qed
 
@@ -4809,7 +4809,7 @@
 subsubsection \<open>(In)finiteness\<close>
 
 lemma finite_maxlen:
-  "finite (M::'a list set) ==> EX n. ALL s:M. size s < n"
+  "finite (M::'a list set) \<Longrightarrow> \<exists>n. \<forall>s\<in>M. size s < n"
 proof (induct rule: finite.induct)
   case emptyI show ?case by simp
 next
@@ -4891,7 +4891,7 @@
   from this show ?thesis by (rule card_lists_distinct_length_eq)
 qed
 
-lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)"
+lemma infinite_UNIV_listI: "\<not> finite(UNIV::'a list set)"
 apply (rule notI)
 apply (drule finite_maxlen)
 apply clarsimp
@@ -4977,7 +4977,7 @@
 by (cases xs) (simp_all add: sorted_Cons)
 
 lemma sorted_append:
-  "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
+  "sorted (xs@ys) = (sorted xs \<and> sorted ys \<and> (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
 by (induct xs) (auto simp add:sorted_Cons)
 
 lemma sorted_nth_mono:
@@ -5878,8 +5878,8 @@
 by (unfold lenlex_def) blast
 
 lemma lenlex_conv:
-    "lenlex r = {(xs,ys). length xs < length ys |
-                 length xs = length ys \<and> (xs, ys) : lex r}"
+    "lenlex r = {(xs,ys). length xs < length ys \<or>
+                 length xs = length ys \<and> (xs, ys) \<in> lex r}"
 by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def)
 
 lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
@@ -5889,8 +5889,8 @@
 by (simp add:lex_conv)
 
 lemma Cons_in_lex [simp]:
-    "((x # xs, y # ys) : lex r) =
-      ((x, y) : r \<and> length xs = length ys | x = y \<and> (xs, ys) : lex r)"
+    "((x # xs, y # ys) \<in> lex r) =
+      ((x, y) \<in> r \<and> length xs = length ys \<or> x = y \<and> (xs, ys) \<in> lex r)"
 apply (simp add: lex_conv)
 apply (rule iffI)
  prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
@@ -5944,7 +5944,7 @@
 by (unfold lexord_def, induct_tac x, auto)
 
 lemma lexord_cons_cons[simp]:
-     "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r | (a = b & (x,y)\<in> lexord r))"
+     "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r \<or> (a = b \<and> (x,y)\<in> lexord r))"
   apply (unfold lexord_def, safe, simp_all)
   apply (case_tac u, simp, simp)
   apply (case_tac u, simp, clarsimp, blast, blast, clarsimp)
@@ -5964,13 +5964,13 @@
 by (induct x, auto)
 
 lemma lexord_append_leftD:
-     "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (! a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
+     "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (\<forall>a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
 by (erule rev_mp, induct_tac x, auto)
 
 lemma lexord_take_index_conv:
-   "((x,y) : lexord r) =
+   "((x,y) \<in> lexord r) =
     ((length x < length y \<and> take (length x) y = x) \<or>
-     (\<exists>i. i < min(length x)(length y) & take i x = take i y & (x!i,y!i) \<in> r))"
+     (\<exists>i. i < min(length x)(length y) \<and> take i x = take i y \<and> (x!i,y!i) \<in> r))"
   apply (unfold lexord_def Let_def, clarsimp)
   apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
   apply auto
@@ -5992,7 +5992,7 @@
   apply (induct_tac x, clarsimp)
   by (clarify, case_tac x, simp, force)
 
-lemma lexord_irreflexive: "ALL x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
+lemma lexord_irreflexive: "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
 by (induct xs) auto
 
 text\<open>By Ren\'e Thiemann:\<close>
@@ -6033,7 +6033,7 @@
 lemma lexord_transI:  "trans r \<Longrightarrow> trans (lexord r)"
 by (rule transI, drule lexord_trans, blast)
 
-lemma lexord_linear: "(! a b. (a,b)\<in> r | a = b | (b,a) \<in> r) \<Longrightarrow> (x,y) : lexord r | x = y | (y,x) : lexord r"
+lemma lexord_linear: "(\<forall>a b. (a,b)\<in> r \<or> a = b \<or> (b,a) \<in> r) \<Longrightarrow> (x,y) \<in> lexord r \<or> x = y \<or> (y,x) \<in> lexord r"
   apply (rule_tac x = y in spec)
   apply (induct_tac x, rule allI)
   apply (case_tac x, simp, simp)
@@ -6427,8 +6427,8 @@
 lemma listrel_eq_len:  "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
 by(induct rule: listrel.induct) auto
 
-lemma listrel_iff_zip [code_unfold]: "(xs,ys) : listrel r \<longleftrightarrow>
-  length xs = length ys & (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
+lemma listrel_iff_zip [code_unfold]: "(xs,ys) \<in> listrel r \<longleftrightarrow>
+  length xs = length ys \<and> (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
 proof
   assume ?L thus ?R by induct (auto intro: listrel_eq_len)
 next
@@ -6437,8 +6437,8 @@
     by (induct rule: list_induct2) (auto intro: listrel.intros)
 qed
 
-lemma listrel_iff_nth: "(xs,ys) : listrel r \<longleftrightarrow>
-  length xs = length ys & (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
+lemma listrel_iff_nth: "(xs,ys) \<in> listrel r \<longleftrightarrow>
+  length xs = length ys \<and> (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
 by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
 
 
@@ -6579,7 +6579,7 @@
   "x \<in> set xs \<Longrightarrow> y \<le> f x \<Longrightarrow> y \<le> size_list f xs"
 by (induct xs) auto
 
-lemma size_list_map[simp]: "size_list f (map g xs) = size_list (f o g) xs"
+lemma size_list_map[simp]: "size_list f (map g xs) = size_list (f \<circ> g) xs"
 by (induct xs) auto
 
 lemma size_list_append[simp]: "size_list f (xs @ ys) = size_list f xs + size_list f ys"
@@ -6757,7 +6757,7 @@
 lemma [code]:
   "lexordp r xs [] = False"
   "lexordp r [] (y#ys) = True"
-  "lexordp r (x # xs) (y # ys) = (r x y | (x = y & lexordp r xs ys))"
+  "lexordp r (x # xs) (y # ys) = (r x y \<or> (x = y \<and> lexordp r xs ys))"
 unfolding lexordp_def by auto
 
 text \<open>Bounded quantification and summation over nats.\<close>
--- a/src/HOL/MacLaurin.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/MacLaurin.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -209,7 +209,7 @@
   fixes n :: nat and h :: real
   shows
     "h < 0 \<and> n > 0 \<and> diff 0 = f \<and>
-      (\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
+      (\<forall>m t. m < n \<and> h \<le> t \<and> t \<le> 0 \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t) \<longrightarrow>
     (\<exists>t. h < t \<and> t < 0 \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n)"
   by (blast intro: Maclaurin_minus)
 
@@ -373,7 +373,7 @@
 
 subsection \<open>Version for Sine Function\<close>
 
-lemma mod_exhaust_less_4: "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = 3"
+lemma mod_exhaust_less_4: "m mod 4 = 0 \<or> m mod 4 = 1 \<or> m mod 4 = 2 \<or> m mod 4 = 3"
   for m :: nat
   by auto
 
--- a/src/HOL/Map.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Map.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -258,11 +258,11 @@
 
 subsection \<open>@{const map_option} related\<close>
 
-lemma map_option_o_empty [simp]: "map_option f o empty = empty"
+lemma map_option_o_empty [simp]: "map_option f \<circ> empty = empty"
 by (rule ext) simp
 
 lemma map_option_o_map_upd [simp]:
-  "map_option f o m(a\<mapsto>b) = (map_option f o m)(a\<mapsto>f b)"
+  "map_option f \<circ> m(a\<mapsto>b) = (map_option f \<circ> m)(a\<mapsto>f b)"
 by (rule ext) simp
 
 
@@ -299,7 +299,7 @@
 by (rule ext) (simp add: map_add_def split: option.split)
 
 lemma map_add_Some_iff:
-  "((m ++ n) k = Some x) = (n k = Some x | n k = None & m k = Some x)"
+  "((m ++ n) k = Some x) = (n k = Some x \<or> n k = None \<and> m k = Some x)"
 by (simp add: map_add_def split: option.split)
 
 lemma map_add_SomeD [dest!]:
@@ -309,7 +309,7 @@
 lemma map_add_find_right [simp]: "n k = Some xx \<Longrightarrow> (m ++ n) k = Some xx"
 by (subst map_add_Some_iff) fast
 
-lemma map_add_None [iff]: "((m ++ n) k = None) = (n k = None & m k = None)"
+lemma map_add_None [iff]: "((m ++ n) k = None) = (n k = None \<and> m k = None)"
 by (simp add: map_add_def split: option.split)
 
 lemma map_add_upd[simp]: "f ++ g(x\<mapsto>y) = (f ++ g)(x\<mapsto>y)"
--- a/src/HOL/Meson.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Meson.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -15,21 +15,21 @@
 
 text \<open>de Morgan laws\<close>
 
-lemma not_conjD: "~(P&Q) ==> ~P | ~Q"
-  and not_disjD: "~(P|Q) ==> ~P & ~Q"
-  and not_notD: "~~P ==> P"
-  and not_allD: "!!P. ~(\<forall>x. P(x)) ==> \<exists>x. ~P(x)"
-  and not_exD: "!!P. ~(\<exists>x. P(x)) ==> \<forall>x. ~P(x)"
+lemma not_conjD: "\<not>(P\<and>Q) \<Longrightarrow> \<not>P \<or> \<not>Q"
+  and not_disjD: "\<not>(P\<or>Q) \<Longrightarrow> \<not>P \<and> \<not>Q"
+  and not_notD: "\<not>\<not>P \<Longrightarrow> P"
+  and not_allD: "\<And>P. \<not>(\<forall>x. P(x)) \<Longrightarrow> \<exists>x. \<not>P(x)"
+  and not_exD: "\<And>P. \<not>(\<exists>x. P(x)) \<Longrightarrow> \<forall>x. \<not>P(x)"
   by fast+
 
 text \<open>Removal of \<open>\<longrightarrow>\<close> and \<open>\<longleftrightarrow>\<close> (positive and negative occurrences)\<close>
 
-lemma imp_to_disjD: "P-->Q ==> ~P | Q"
-  and not_impD: "~(P-->Q) ==> P & ~Q"
-  and iff_to_disjD: "P=Q ==> (~P | Q) & (~Q | P)"
-  and not_iffD: "~(P=Q) ==> (P | Q) & (~P | ~Q)"
-    \<comment> \<open>Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF\<close>
-  and not_refl_disj_D: "x ~= x | P ==> P"
+lemma imp_to_disjD: "P\<longrightarrow>Q \<Longrightarrow> \<not>P \<or> Q"
+  and not_impD: "\<not>(P\<longrightarrow>Q) \<Longrightarrow> P \<and> \<not>Q"
+  and iff_to_disjD: "P=Q \<Longrightarrow> (\<not>P \<or> Q) \<and> (\<not>Q \<or> P)"
+  and not_iffD: "\<not>(P=Q) \<Longrightarrow> (P \<or> Q) \<and> (\<not>P \<or> \<not>Q)"
+    \<comment> \<open>Much more efficient than @{prop "(P \<and> \<not>Q) \<or> (Q \<and> \<not>P)"} for computing CNF\<close>
+  and not_refl_disj_D: "x \<noteq> x \<or> P \<Longrightarrow> P"
   by fast+
 
 
@@ -37,24 +37,24 @@
 
 text \<open>Conjunction\<close>
 
-lemma conj_exD1: "!!P Q. (\<exists>x. P(x)) & Q ==> \<exists>x. P(x) & Q"
-  and conj_exD2: "!!P Q. P & (\<exists>x. Q(x)) ==> \<exists>x. P & Q(x)"
+lemma conj_exD1: "\<And>P Q. (\<exists>x. P(x)) \<and> Q \<Longrightarrow> \<exists>x. P(x) \<and> Q"
+  and conj_exD2: "\<And>P Q. P \<and> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P \<and> Q(x)"
   by fast+
 
 
 text \<open>Disjunction\<close>
 
-lemma disj_exD: "!!P Q. (\<exists>x. P(x)) | (\<exists>x. Q(x)) ==> \<exists>x. P(x) | Q(x)"
+lemma disj_exD: "\<And>P Q. (\<exists>x. P(x)) \<or> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P(x) \<or> Q(x)"
   \<comment> \<open>DO NOT USE with forall-Skolemization: makes fewer schematic variables!!\<close>
   \<comment> \<open>With ex-Skolemization, makes fewer Skolem constants\<close>
-  and disj_exD1: "!!P Q. (\<exists>x. P(x)) | Q ==> \<exists>x. P(x) | Q"
-  and disj_exD2: "!!P Q. P | (\<exists>x. Q(x)) ==> \<exists>x. P | Q(x)"
+  and disj_exD1: "\<And>P Q. (\<exists>x. P(x)) \<or> Q \<Longrightarrow> \<exists>x. P(x) \<or> Q"
+  and disj_exD2: "\<And>P Q. P \<or> (\<exists>x. Q(x)) \<Longrightarrow> \<exists>x. P \<or> Q(x)"
   by fast+
 
-lemma disj_assoc: "(P|Q)|R ==> P|(Q|R)"
-  and disj_comm: "P|Q ==> Q|P"
-  and disj_FalseD1: "False|P ==> P"
-  and disj_FalseD2: "P|False ==> P"
+lemma disj_assoc: "(P\<or>Q)\<or>R \<Longrightarrow> P\<or>(Q\<or>R)"
+  and disj_comm: "P\<or>Q \<Longrightarrow> Q\<or>P"
+  and disj_FalseD1: "False\<or>P \<Longrightarrow> P"
+  and disj_FalseD2: "P\<or>False \<Longrightarrow> P"
   by fast+
 
 
@@ -63,15 +63,15 @@
 text\<open>Inserts negated disjunct after removing the negation; P is a literal.
   Model elimination requires assuming the negation of every attempted subgoal,
   hence the negated disjuncts.\<close>
-lemma make_neg_rule: "~P|Q ==> ((~P==>P) ==> Q)"
+lemma make_neg_rule: "\<not>P\<or>Q \<Longrightarrow> ((\<not>P\<Longrightarrow>P) \<Longrightarrow> Q)"
 by blast
 
 text\<open>Version for Plaisted's "Postive refinement" of the Meson procedure\<close>
-lemma make_refined_neg_rule: "~P|Q ==> (P ==> Q)"
+lemma make_refined_neg_rule: "\<not>P\<or>Q \<Longrightarrow> (P \<Longrightarrow> Q)"
 by blast
 
 text\<open>@{term P} should be a literal\<close>
-lemma make_pos_rule: "P|Q ==> ((P==>~P) ==> Q)"
+lemma make_pos_rule: "P\<or>Q \<Longrightarrow> ((P\<Longrightarrow>\<not>P) \<Longrightarrow> Q)"
 by blast
 
 text\<open>Versions of \<open>make_neg_rule\<close> and \<open>make_pos_rule\<close> that don't
@@ -79,15 +79,15 @@
 
 lemmas make_neg_rule' = make_refined_neg_rule
 
-lemma make_pos_rule': "[|P|Q; ~P|] ==> Q"
+lemma make_pos_rule': "\<lbrakk>P\<or>Q; \<not>P\<rbrakk> \<Longrightarrow> Q"
 by blast
 
 text\<open>Generation of a goal clause -- put away the final literal\<close>
 
-lemma make_neg_goal: "~P ==> ((~P==>P) ==> False)"
+lemma make_neg_goal: "\<not>P \<Longrightarrow> ((\<not>P\<Longrightarrow>P) \<Longrightarrow> False)"
 by blast
 
-lemma make_pos_goal: "P ==> ((P==>~P) ==> False)"
+lemma make_pos_goal: "P \<Longrightarrow> ((P\<Longrightarrow>\<not>P) \<Longrightarrow> False)"
 by blast
 
 
@@ -97,18 +97,17 @@
 
 (*NOTE: could handle conjunctions (faster?) by
     nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *)
-lemma conj_forward: "[| P'&Q';  P' ==> P;  Q' ==> Q |] ==> P&Q"
+lemma conj_forward: "\<lbrakk>P'\<and>Q';  P' \<Longrightarrow> P;  Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P\<and>Q"
 by blast
 
-lemma disj_forward: "[| P'|Q';  P' ==> P;  Q' ==> Q |] ==> P|Q"
+lemma disj_forward: "\<lbrakk>P'\<or>Q';  P' \<Longrightarrow> P;  Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P\<or>Q"
 by blast
 
-lemma imp_forward: "[| P' \<longrightarrow> Q';  P ==> P';  Q' ==> Q |] ==> P \<longrightarrow> Q"
+lemma imp_forward: "\<lbrakk>P' \<longrightarrow> Q';  P \<Longrightarrow> P';  Q' \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> P \<longrightarrow> Q"
 by blast
 
 (*Version of @{text disj_forward} for removal of duplicate literals*)
-lemma disj_forward2:
-    "[| P'|Q';  P' ==> P;  [| Q'; P==>False |] ==> Q |] ==> P|Q"
+lemma disj_forward2: "\<lbrakk> P'\<or>Q';  P' \<Longrightarrow> P;  \<lbrakk>Q'; P\<Longrightarrow>False\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> P\<or>Q"
 apply blast 
 done
 
--- a/src/HOL/Nat.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nat.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -360,10 +360,10 @@
   for m n :: nat
   by (cases m) simp_all
 
-lemma add_is_1: "m + n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = 0 | m = 0 \<and> n = Suc 0"
+lemma add_is_1: "m + n = Suc 0 \<longleftrightarrow> m = Suc 0 \<and> n = 0 \<or> m = 0 \<and> n = Suc 0"
   by (cases m) simp_all
 
-lemma one_is_add: "Suc 0 = m + n \<longleftrightarrow> m = Suc 0 \<and> n = 0 | m = 0 \<and> n = Suc 0"
+lemma one_is_add: "Suc 0 = m + n \<longleftrightarrow> m = Suc 0 \<and> n = 0 \<or> m = 0 \<and> n = Suc 0"
   by (rule trans, rule eq_commute, rule add_is_1)
 
 lemma add_eq_self_zero: "m + n = m \<Longrightarrow> n = 0"
@@ -735,7 +735,7 @@
   for m n :: nat
   unfolding less_le ..
 
-lemma nat_le_linear: "m \<le> n | n \<le> m"
+lemma nat_le_linear: "m \<le> n \<or> n \<le> m"
   for m n :: nat
   by (rule linear)
 
--- a/src/HOL/Nitpick.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nitpick.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -150,10 +150,10 @@
   "one_frac \<equiv> Abs_Frac (1, 1)"
 
 definition num :: "'a \<Rightarrow> int" where
-  "num \<equiv> fst o Rep_Frac"
+  "num \<equiv> fst \<circ> Rep_Frac"
 
 definition denom :: "'a \<Rightarrow> int" where
-  "denom \<equiv> snd o Rep_Frac"
+  "denom \<equiv> snd \<circ> Rep_Frac"
 
 function norm_frac :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
   "norm_frac a b =
--- a/src/HOL/Nonstandard_Analysis/HTranscendental.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HTranscendental.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -370,7 +370,7 @@
 done
 
 lemma starfun_exp_HInfinite_Infinitesimal_disj:
- "x \<in> HInfinite ==> ( *f* exp) x \<in> HInfinite | ( *f* exp) (x::hypreal) \<in> Infinitesimal"
+ "x \<in> HInfinite \<Longrightarrow> ( *f* exp) x \<in> HInfinite \<or> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
 apply (insert linorder_linear [of x 0]) 
 apply (auto intro: starfun_exp_HInfinite starfun_exp_Infinitesimal)
 done
--- a/src/HOL/Nonstandard_Analysis/HyperNat.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HyperNat.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -154,7 +154,7 @@
 lemma hypnat_of_nat_Suc [simp]: "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1"
   by transfer simp
 
-lemma of_nat_eq_add [rule_format]: "\<forall>d::hypnat. of_nat m = of_nat n + d --> d \<in> range of_nat"
+lemma of_nat_eq_add [rule_format]: "\<forall>d::hypnat. of_nat m = of_nat n + d \<longrightarrow> d \<in> range of_nat"
   apply (induct n)
    apply (auto simp add: add.assoc)
   apply (case_tac x)
--- a/src/HOL/Nonstandard_Analysis/NSCA.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/NSCA.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -16,7 +16,7 @@
 
 definition \<comment>\<open>standard part map\<close>
   stc :: "hcomplex => hcomplex" where 
-  "stc x = (SOME r. x \<in> HFinite & r:SComplex & r \<approx> x)"
+  "stc x = (SOME r. x \<in> HFinite \<and> r\<in>SComplex \<and> r \<approx> x)"
 
 
 subsection\<open>Closure Laws for SComplex, the Standard Complex Numbers\<close>
@@ -63,7 +63,7 @@
 
 lemma SComplex_SReal_dense:
      "[| x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y  
-      |] ==> \<exists>r \<in> Reals. hcmod x< r & r < hcmod y"
+      |] ==> \<exists>r \<in> Reals. hcmod x< r \<and> r < hcmod y"
 apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex)
 done
 
@@ -350,7 +350,7 @@
 apply (simp add: Standard_HComplex st_SReal [unfolded Reals_eq_Standard])
 done
 
-lemma stc_part_Ex1: "x:HFinite ==> \<exists>!t. t \<in> SComplex &  x \<approx> t"
+lemma stc_part_Ex1: "x\<in>HFinite \<Longrightarrow> \<exists>!t. t \<in> SComplex \<and> x \<approx> t"
 apply (drule stc_part_Ex, safe)
 apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)
 apply (auto intro!: approx_unique_complex)
--- a/src/HOL/Nonstandard_Analysis/NatStar.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/NatStar.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -143,9 +143,9 @@
   by transfer (rule refl)
 
 lemma starfunNat_inverse_real_of_nat_Infinitesimal [simp]:
-  "N \<in> HNatInfinite \<Longrightarrow> ( *f* (%x. inverse (real x))) N \<in> Infinitesimal"
+  "N \<in> HNatInfinite \<Longrightarrow> ( *f* (\<lambda>x. inverse (real x))) N \<in> Infinitesimal"
   apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
-  apply (subgoal_tac "hypreal_of_hypnat N ~= 0")
+  apply (subgoal_tac "hypreal_of_hypnat N \<noteq> 0")
    apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat)
   done
 
--- a/src/HOL/Number_Theory/Euler_Criterion.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Number_Theory/Euler_Criterion.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -78,7 +78,7 @@
 private lemma P_lemma: assumes "x \<in> S1"
   shows "\<exists>! y. P x y"
 proof -
-  have "~ p dvd x" using assms zdvd_not_zless S1_def by auto
+  have "\<not> p dvd x" using assms zdvd_not_zless S1_def by auto
   hence co_xp: "coprime x p" using p_prime prime_imp_coprime_int[of p x] 
     by (simp add: ac_simps)
   then obtain y' where y': "[x * y' = 1] (mod p)" using cong_solve_coprime_int by blast
@@ -110,11 +110,11 @@
 private lemma f_lemma_1: assumes "x \<in> S1"
   shows "f x = f (f_1 x)" using f_def f_1_lemma_2[of x] assms by auto
 
-private lemma l1: assumes "~ QuadRes p a" "x \<in> S1"
+private lemma l1: assumes "\<not> QuadRes p a" "x \<in> S1"
   shows "x \<noteq> f_1 x"
   using f_1_lemma_1[of x] assms unfolding P_def QuadRes_def power2_eq_square by fastforce
 
-private lemma l2: assumes "~ QuadRes p a" "x \<in> S1"
+private lemma l2: assumes "\<not> QuadRes p a" "x \<in> S1"
   shows "[\<Prod> (f x) = a] (mod p)"
   using assms l1 f_1_lemma_1 P_def f_def by auto
 
@@ -165,7 +165,7 @@
     by simp blast 
 qed
 
-private lemma l11: assumes "~ QuadRes p a"
+private lemma l11: assumes "\<not> QuadRes p a"
   shows "card S2 = (p - 1) div 2"
 proof -
   have "sum card S2 = 2 * card S2"
@@ -175,11 +175,11 @@
   ultimately show ?thesis by linarith
 qed
 
-private lemma l12: assumes "~ QuadRes p a"
+private lemma l12: assumes "\<not> QuadRes p a"
   shows "[prod Prod S2 = a ^ ((p - 1) div 2)] (mod p)"
   using assms l2 l10 l11 unfolding S2_def by blast
 
-private lemma E_2: assumes "~ QuadRes p a"
+private lemma E_2: assumes "\<not> QuadRes p a"
   shows "[a ^ ((p - 1) div 2) = -1] (mod p)" using l9 l12 cong_trans cong_sym assms by blast
 
 lemma euler_criterion_aux: "[(Legendre a p) = a ^ ((p - 1) div 2)] (mod p)"
--- a/src/HOL/Number_Theory/Gauss.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Number_Theory/Gauss.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -323,8 +323,8 @@
     done
   then have "\<forall>x \<in> E. [(p-x) mod p = - x](mod p)"
     by (metis cong_def minus_mod_self1 mod_mod_trivial)
-  then have "[prod ((\<lambda>x. x mod p) o (op - p)) E = prod (uminus) E](mod p)"
-    using finite_E p_ge_2 cong_prod [of E "(\<lambda>x. x mod p) o (op - p)" uminus p]
+  then have "[prod ((\<lambda>x. x mod p) \<circ> (op - p)) E = prod (uminus) E](mod p)"
+    using finite_E p_ge_2 cong_prod [of E "(\<lambda>x. x mod p) \<circ> (op - p)" uminus p]
     by auto
   then have two: "[prod id F = prod (uminus) E](mod p)"
     by (metis FE cong_cong_mod_int cong_refl cong_prod minus_mod_self1)
--- a/src/HOL/Number_Theory/Pocklington.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Number_Theory/Pocklington.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -212,7 +212,7 @@
   from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime_nat an1
   have an: "coprime a n" "coprime (a ^ (n - 1)) n"
     using \<open>n \<ge> 2\<close> by simp_all
-  have False if H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
+  have False if H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "\<exists>m. ?P m")
   proof -
     from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
       m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k"
--- a/src/HOL/Number_Theory/Residues.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Number_Theory/Residues.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -296,7 +296,7 @@
   fixes a m :: nat
   assumes "coprime a m"
   shows "[a ^ totient m = 1] (mod m)"
-proof (cases "m = 0 | m = 1")
+proof (cases "m = 0 \<or> m = 1")
   case True
   then show ?thesis by auto
 next
--- a/src/HOL/Numeral_Simprocs.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Numeral_Simprocs.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -77,7 +77,7 @@
 by auto
 
 lemma nat_mult_dvd_cancel_disj[simp]:
-  "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))"
+  "(k*m) dvd (k*n) = (k=0 \<or> m dvd (n::nat))"
 by (auto simp: dvd_eq_mod_eq_0 mod_mult_mult1)
 
 lemma nat_mult_dvd_cancel1: "0 < k \<Longrightarrow> (k*m) dvd (k*n::nat) = (m dvd n)"
--- a/src/HOL/Option.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Option.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -109,7 +109,7 @@
   by (simp add: map_option_case split: option.split)
 
 lemma map_option_o_case_sum [simp]:
-    "map_option f o case_sum g h = case_sum (map_option f o g) (map_option f o h)"
+    "map_option f \<circ> case_sum g h = case_sum (map_option f \<circ> g) (map_option f \<circ> h)"
   by (rule o_case_sum)
 
 lemma map_option_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> map_option f x = map_option g y"
--- a/src/HOL/Orderings.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Orderings.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -389,7 +389,7 @@
 done
 
 lemma not_less_iff_gr_or_eq:
- "\<not>(x < y) \<longleftrightarrow> (x > y | x = y)"
+ "\<not>(x < y) \<longleftrightarrow> (x > y \<or> x = y)"
 apply(simp add:not_less le_less)
 apply blast
 done
@@ -745,14 +745,14 @@
   "_Ex_less_eq" :: "[idt, 'a, bool] => bool"    ("(3? _<=_./ _)" [0, 0, 10] 10)
 
 translations
-  "ALL x<y. P"   =>  "ALL x. x < y \<longrightarrow> P"
-  "EX x<y. P"    =>  "EX x. x < y \<and> P"
-  "ALL x<=y. P"  =>  "ALL x. x <= y \<longrightarrow> P"
-  "EX x<=y. P"   =>  "EX x. x <= y \<and> P"
-  "ALL x>y. P"   =>  "ALL x. x > y \<longrightarrow> P"
-  "EX x>y. P"    =>  "EX x. x > y \<and> P"
-  "ALL x>=y. P"  =>  "ALL x. x >= y \<longrightarrow> P"
-  "EX x>=y. P"   =>  "EX x. x >= y \<and> P"
+  "\<forall>x<y. P" \<rightharpoonup> "\<forall>x. x < y \<longrightarrow> P"
+  "\<exists>x<y. P" \<rightharpoonup> "\<exists>x. x < y \<and> P"
+  "\<forall>x\<le>y. P" \<rightharpoonup> "\<forall>x. x \<le> y \<longrightarrow> P"
+  "\<exists>x\<le>y. P" \<rightharpoonup> "\<exists>x. x \<le> y \<and> P"
+  "\<forall>x>y. P" \<rightharpoonup> "\<forall>x. x > y \<longrightarrow> P"
+  "\<exists>x>y. P" \<rightharpoonup> "\<exists>x. x > y \<and> P"
+  "\<forall>x\<ge>y. P" \<rightharpoonup> "\<forall>x. x \<ge> y \<longrightarrow> P"
+  "\<exists>x\<ge>y. P" \<rightharpoonup> "\<exists>x. x \<ge> y \<and> P"
 
 print_translation \<open>
 let
@@ -1001,22 +1001,22 @@
     a >= b >= c ... in Isar proofs.\<close>
 
 lemma xt1 [no_atp]:
-  "a = b ==> b > c ==> a > c"
-  "a > b ==> b = c ==> a > c"
-  "a = b ==> b >= c ==> a >= c"
-  "a >= b ==> b = c ==> a >= c"
-  "(x::'a::order) >= y ==> y >= x ==> x = y"
-  "(x::'a::order) >= y ==> y >= z ==> x >= z"
-  "(x::'a::order) > y ==> y >= z ==> x > z"
-  "(x::'a::order) >= y ==> y > z ==> x > z"
-  "(a::'a::order) > b ==> b > a ==> P"
-  "(x::'a::order) > y ==> y > z ==> x > z"
-  "(a::'a::order) >= b ==> a ~= b ==> a > b"
-  "(a::'a::order) ~= b ==> a >= b ==> a > b"
-  "a = f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c"
-  "a > b ==> f b = c ==> (!!x y. x > y ==> f x > f y) ==> f a > c"
-  "a = f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c"
-  "a >= b ==> f b = c ==> (!! x y. x >= y ==> f x >= f y) ==> f a >= c"
+  "a = b \<Longrightarrow> b > c \<Longrightarrow> a > c"
+  "a > b \<Longrightarrow> b = c \<Longrightarrow> a > c"
+  "a = b \<Longrightarrow> b \<ge> c \<Longrightarrow> a \<ge> c"
+  "a \<ge> b \<Longrightarrow> b = c \<Longrightarrow> a \<ge> c"
+  "(x::'a::order) \<ge> y \<Longrightarrow> y \<ge> x \<Longrightarrow> x = y"
+  "(x::'a::order) \<ge> y \<Longrightarrow> y \<ge> z \<Longrightarrow> x \<ge> z"
+  "(x::'a::order) > y \<Longrightarrow> y \<ge> z \<Longrightarrow> x > z"
+  "(x::'a::order) \<ge> y \<Longrightarrow> y > z \<Longrightarrow> x > z"
+  "(a::'a::order) > b \<Longrightarrow> b > a \<Longrightarrow> P"
+  "(x::'a::order) > y \<Longrightarrow> y > z \<Longrightarrow> x > z"
+  "(a::'a::order) \<ge> b \<Longrightarrow> a \<noteq> b \<Longrightarrow> a > b"
+  "(a::'a::order) \<noteq> b \<Longrightarrow> a \<ge> b \<Longrightarrow> a > b"
+  "a = f b \<Longrightarrow> b > c \<Longrightarrow> (\<And>x y. x > y \<Longrightarrow> f x > f y) \<Longrightarrow> a > f c"
+  "a > b \<Longrightarrow> f b = c \<Longrightarrow> (\<And>x y. x > y \<Longrightarrow> f x > f y) \<Longrightarrow> f a > c"
+  "a = f b \<Longrightarrow> b \<ge> c \<Longrightarrow> (\<And>x y. x \<ge> y \<Longrightarrow> f x \<ge> f y) \<Longrightarrow> a \<ge> f c"
+  "a \<ge> b \<Longrightarrow> f b = c \<Longrightarrow> (\<And>x y. x \<ge> y \<Longrightarrow> f x \<ge> f y) \<Longrightarrow> f a \<ge> c"
   by auto
 
 lemma xt2 [no_atp]:
--- a/src/HOL/Partial_Function.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Partial_Function.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -340,7 +340,7 @@
 
 lemma admissible_image:
   assumes pfun: "partial_function_definitions le lub"
-  assumes adm: "ccpo.admissible lub le (P o g)"
+  assumes adm: "ccpo.admissible lub le (P \<circ> g)"
   assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y"
   assumes inv: "\<And>x. f (g x) = x"
   shows "ccpo.admissible (img_lub f g lub) (img_ord f le) P"
@@ -350,10 +350,10 @@
     by (auto simp: img_ord_def intro: chainI dest: chainD)
   assume "A \<noteq> {}"
   assume P_A: "\<forall>x\<in>A. P x"
-  have "(P o g) (lub (f ` A))" using adm ch'
+  have "(P \<circ> g) (lub (f ` A))" using adm ch'
   proof (rule ccpo.admissibleD)
     fix x assume "x \<in> f ` A"
-    with P_A show "(P o g) x" by (auto simp: inj[OF inv])
+    with P_A show "(P \<circ> g) x" by (auto simp: inj[OF inv])
   qed(simp add: \<open>A \<noteq> {}\<close>)
   thus "P (img_lub f g lub A)" unfolding img_lub_def by simp
 qed
--- a/src/HOL/Predicate.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Predicate.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -383,7 +383,7 @@
 by (rule unit_pred_cases) (auto elim: botE intro: singleI)
 
 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a pred \<Rightarrow> 'b pred" where
-  "map f P = P \<bind> (single o f)"
+  "map f P = P \<bind> (single \<circ> f)"
 
 lemma eval_map [simp]:
   "eval (map f P) = (\<Squnion>x\<in>{x. eval P x}. (\<lambda>y. f x = y))"
--- a/src/HOL/Presburger.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Presburger.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -180,8 +180,8 @@
 
 subsubsection\<open>First some trivial facts about periodic sets or predicates\<close>
 lemma periodic_finite_ex:
-  assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)"
-  shows "(EX x. P x) = (EX j : {1..d}. P j)"
+  assumes dpos: "(0::int) < d" and modd: "\<forall>x k. P x = P(x - k*d)"
+  shows "(\<exists>x. P x) = (\<exists>j \<in> {1..d}. P j)"
   (is "?LHS = ?RHS")
 proof
   assume ?LHS
@@ -219,7 +219,7 @@
 
 lemma decr_mult_lemma:
   assumes dpos: "(0::int) < d" and minus: "\<forall>x. P x \<longrightarrow> P(x - d)" and knneg: "0 <= k"
-  shows "ALL x. P x \<longrightarrow> P(x - k*d)"
+  shows "\<forall>x. P x \<longrightarrow> P(x - k*d)"
 using knneg
 proof (induct rule:int_ge_induct)
   case base thus ?case by simp
@@ -235,25 +235,25 @@
 
 lemma  minusinfinity:
   assumes dpos: "0 < d" and
-    P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and ePeqP1: "EX z::int. ALL x. x < z \<longrightarrow> (P x = P1 x)"
-  shows "(EX x. P1 x) \<longrightarrow> (EX x. P x)"
+    P1eqP1: "\<forall>x k. P1 x = P1(x - k*d)" and ePeqP1: "\<exists>z::int. \<forall>x. x < z \<longrightarrow> (P x = P1 x)"
+  shows "(\<exists>x. P1 x) \<longrightarrow> (\<exists>x. P x)"
 proof
-  assume eP1: "EX x. P1 x"
+  assume eP1: "\<exists>x. P1 x"
   then obtain x where P1: "P1 x" ..
-  from ePeqP1 obtain z where P1eqP: "ALL x. x < z \<longrightarrow> (P x = P1 x)" ..
+  from ePeqP1 obtain z where P1eqP: "\<forall>x. x < z \<longrightarrow> (P x = P1 x)" ..
   let ?w = "x - (\<bar>x - z\<bar> + 1) * d"
   from dpos have w: "?w < z" by(rule decr_lemma)
   have "P1 x = P1 ?w" using P1eqP1 by blast
   also have "\<dots> = P(?w)" using w P1eqP by blast
   finally have "P ?w" using P1 by blast
-  thus "EX x. P x" ..
+  thus "\<exists>x. P x" ..
 qed
 
 lemma cpmi: 
   assumes dp: "0 < D" and p1:"\<exists>z. \<forall> x< z. P x = P' x"
-  and nb:"\<forall>x.(\<forall> j\<in> {1..D}. \<forall>(b::int) \<in> B. x \<noteq> b+j) --> P (x) --> P (x - D)"
+  and nb:"\<forall>x.(\<forall> j\<in> {1..D}. \<forall>(b::int) \<in> B. x \<noteq> b+j) \<longrightarrow> P (x) \<longrightarrow> P (x - D)"
   and pd: "\<forall> x k. P' x = P' (x-k*D)"
-  shows "(\<exists>x. P x) = ((\<exists> j\<in> {1..D} . P' j) | (\<exists> j \<in> {1..D}.\<exists> b\<in> B. P (b+j)))" 
+  shows "(\<exists>x. P x) = ((\<exists>j \<in> {1..D} . P' j) \<or> (\<exists>j \<in> {1..D}. \<exists> b \<in> B. P (b+j)))"
          (is "?L = (?R1 \<or> ?R2)")
 proof-
  {assume "?R2" hence "?L"  by blast}
@@ -263,11 +263,11 @@
  { fix x
    assume P: "P x" and H: "\<not> ?R2"
    {fix y assume "\<not> (\<exists>j\<in>{1..D}. \<exists>b\<in>B. P (b + j))" and P: "P y"
-     hence "~(EX (j::int) : {1..D}. EX (b::int) : B. y = b+j)" by auto
+     hence "\<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. y = b+j)" by auto
      with nb P  have "P (y - D)" by auto }
-   hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D)" by blast
+   hence "\<forall>x. \<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> B. P(b+j)) \<longrightarrow> P (x) \<longrightarrow> P (x - D)" by blast
    with H P have th: " \<forall>x. P x \<longrightarrow> P (x - D)" by auto
-   from p1 obtain z where z: "ALL x. x < z --> (P x = P' x)" by blast
+   from p1 obtain z where z: "\<forall>x. x < z \<longrightarrow> (P x = P' x)" by blast
    let ?y = "x - (\<bar>x - z\<bar> + 1)*D"
    have zp: "0 <= (\<bar>x - z\<bar> + 1)" by arith
    from dp have yz: "?y < z" using decr_lemma[OF dp] by simp   
@@ -284,7 +284,7 @@
     P1eqP1: "\<forall>x k. P' x = P'(x - k*d)" and ePeqP1: "\<exists> z. \<forall> x>z. P x = P' x"
   shows "(\<exists> x. P' x) \<longrightarrow> (\<exists> x. P x)"
 proof
-  assume eP1: "EX x. P' x"
+  assume eP1: "\<exists>x. P' x"
   then obtain x where P1: "P' x" ..
   from ePeqP1 obtain z where P1eqP: "\<forall>x>z. P x = P' x" ..
   let ?w' = "x + (\<bar>x - z\<bar> + 1) * d"
@@ -294,12 +294,12 @@
   hence "P' x = P' ?w" using P1eqP1 by blast
   also have "\<dots> = P(?w)" using w P1eqP by blast
   finally have "P ?w" using P1 by blast
-  thus "EX x. P x" ..
+  thus "\<exists>x. P x" ..
 qed
 
 lemma incr_mult_lemma:
-  assumes dpos: "(0::int) < d" and plus: "ALL x::int. P x \<longrightarrow> P(x + d)" and knneg: "0 <= k"
-  shows "ALL x. P x \<longrightarrow> P(x + k*d)"
+  assumes dpos: "(0::int) < d" and plus: "\<forall>x::int. P x \<longrightarrow> P(x + d)" and knneg: "0 <= k"
+  shows "\<forall>x. P x \<longrightarrow> P(x + k*d)"
 using knneg
 proof (induct rule:int_ge_induct)
   case base thus ?case by simp
@@ -315,9 +315,9 @@
 
 lemma cppi: 
   assumes dp: "0 < D" and p1:"\<exists>z. \<forall> x> z. P x = P' x"
-  and nb:"\<forall>x.(\<forall> j\<in> {1..D}. \<forall>(b::int) \<in> A. x \<noteq> b - j) --> P (x) --> P (x + D)"
+  and nb:"\<forall>x.(\<forall> j\<in> {1..D}. \<forall>(b::int) \<in> A. x \<noteq> b - j) \<longrightarrow> P (x) \<longrightarrow> P (x + D)"
   and pd: "\<forall> x k. P' x= P' (x-k*D)"
-  shows "(\<exists>x. P x) = ((\<exists> j\<in> {1..D} . P' j) | (\<exists> j \<in> {1..D}.\<exists> b\<in> A. P (b - j)))" (is "?L = (?R1 \<or> ?R2)")
+  shows "(\<exists>x. P x) = ((\<exists>j \<in> {1..D} . P' j) \<or> (\<exists> j \<in> {1..D}. \<exists> b\<in> A. P (b - j)))" (is "?L = (?R1 \<or> ?R2)")
 proof-
  {assume "?R2" hence "?L"  by blast}
  moreover
@@ -326,11 +326,11 @@
  { fix x
    assume P: "P x" and H: "\<not> ?R2"
    {fix y assume "\<not> (\<exists>j\<in>{1..D}. \<exists>b\<in>A. P (b - j))" and P: "P y"
-     hence "~(EX (j::int) : {1..D}. EX (b::int) : A. y = b - j)" by auto
+     hence "\<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> A. y = b - j)" by auto
      with nb P  have "P (y + D)" by auto }
-   hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : A. P(b-j)) --> P (x) --> P (x + D)" by blast
+   hence "\<forall>x. \<not>(\<exists>(j::int) \<in> {1..D}. \<exists>(b::int) \<in> A. P(b-j)) \<longrightarrow> P (x) \<longrightarrow> P (x + D)" by blast
    with H P have th: " \<forall>x. P x \<longrightarrow> P (x + D)" by auto
-   from p1 obtain z where z: "ALL x. x > z --> (P x = P' x)" by blast
+   from p1 obtain z where z: "\<forall>x. x > z \<longrightarrow> (P x = P' x)" by blast
    let ?y = "x + (\<bar>x - z\<bar> + 1)*D"
    have zp: "0 <= (\<bar>x - z\<bar> + 1)" by arith
    from dp have yz: "?y > z" using incr_lemma[OF dp] by simp
--- a/src/HOL/Quickcheck_Exhaustive.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Quickcheck_Exhaustive.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -336,7 +336,7 @@
       enum = (Enum.enum :: 'a list)
     in
       check_all_n_lists
-        (\<lambda>(ys, yst). f (the o map_of (zip enum ys), mk_term yst))
+        (\<lambda>(ys, yst). f (the \<circ> map_of (zip enum ys), mk_term yst))
         (natural_of_nat (length enum)))"
 
 definition enum_term_of_fun :: "('a \<Rightarrow> 'b) itself \<Rightarrow> unit \<Rightarrow> term list"
--- a/src/HOL/Quickcheck_Narrowing.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Quickcheck_Narrowing.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -52,7 +52,7 @@
 
 primrec map_cons :: "('a => 'b) => 'a narrowing_cons => 'b narrowing_cons"
 where
-  "map_cons f (Narrowing_cons ty cs) = Narrowing_cons ty (map (\<lambda>c. f o c) cs)"
+  "map_cons f (Narrowing_cons ty cs) = Narrowing_cons ty (map (\<lambda>c. f \<circ> c) cs)"
 
 subsubsection \<open>From narrowing's deep representation of terms to @{theory Code_Evaluation}'s terms\<close>
 
@@ -111,8 +111,8 @@
 definition sum :: "'a narrowing => 'a narrowing => 'a narrowing"
 where
   "sum a b d =
-    (case a d of Narrowing_cons (Narrowing_sum_of_products ssa) ca => 
-      case b d of Narrowing_cons (Narrowing_sum_of_products ssb) cb =>
+    (case a d of Narrowing_cons (Narrowing_sum_of_products ssa) ca \<Rightarrow>
+      case b d of Narrowing_cons (Narrowing_sum_of_products ssb) cb \<Rightarrow>
       Narrowing_cons (Narrowing_sum_of_products (ssa @ ssb)) (ca @ cb))"
 
 lemma [fundef_cong]:
@@ -146,11 +146,11 @@
 (* FIXME: hard-wired maximal depth of 100 here *)
 definition exists :: "('a :: {narrowing, partial_term_of} => property) => property"
 where
-  "exists f = (case narrowing (100 :: integer) of Narrowing_cons ty cs => Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
+  "exists f = (case narrowing (100 :: integer) of Narrowing_cons ty cs \<Rightarrow> Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 
 definition "all" :: "('a :: {narrowing, partial_term_of} => property) => property"
 where
-  "all f = (case narrowing (100 :: integer) of Narrowing_cons ty cs => Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
+  "all f = (case narrowing (100 :: integer) of Narrowing_cons ty cs \<Rightarrow> Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 
 subsubsection \<open>class \<open>is_testable\<close>\<close>
 
--- a/src/HOL/Quotient.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Quotient.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -107,7 +107,7 @@
   by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp part_equivpI)
 
 lemma abs_o_rep:
-  "Abs o Rep = id"
+  "Abs \<circ> Rep = id"
   unfolding fun_eq_iff
   by (simp add: Quotient3_abs_rep)
 
@@ -280,25 +280,25 @@
 
 (* Next four lemmas are unused *)
 lemma all_reg:
-  assumes a: "!x :: 'a. (P x --> Q x)"
+  assumes a: "\<forall>x :: 'a. (P x \<longrightarrow> Q x)"
   and     b: "All P"
   shows "All Q"
   using a b by fast
 
 lemma ex_reg:
-  assumes a: "!x :: 'a. (P x --> Q x)"
+  assumes a: "\<forall>x :: 'a. (P x \<longrightarrow> Q x)"
   and     b: "Ex P"
   shows "Ex Q"
   using a b by fast
 
 lemma ball_reg:
-  assumes a: "!x :: 'a. (x \<in> R --> P x --> Q x)"
+  assumes a: "\<forall>x :: 'a. (x \<in> R \<longrightarrow> P x \<longrightarrow> Q x)"
   and     b: "Ball R P"
   shows "Ball R Q"
   using a b by fast
 
 lemma bex_reg:
-  assumes a: "!x :: 'a. (x \<in> R --> P x --> Q x)"
+  assumes a: "\<forall>x :: 'a. (x \<in> R \<longrightarrow> P x \<longrightarrow> Q x)"
   and     b: "Bex R P"
   shows "Bex R Q"
   using a b by fast
--- a/src/HOL/Random_Pred.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Random_Pred.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -59,10 +59,10 @@
    in if Predicate.eval P' () then (Orderings.bot, s') else (Predicate.single (), s'))"
 
 definition Random :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> 'a random_pred"
-  where "Random g = scomp g (Pair o (Predicate.single o fst))"
+  where "Random g = scomp g (Pair \<circ> (Predicate.single \<circ> fst))"
 
 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a random_pred \<Rightarrow> 'b random_pred"
-  where "map f P = bind P (single o f)"
+  where "map f P = bind P (single \<circ> f)"
 
 hide_const (open) iter' iter empty single bind union if_randompred
   iterate_upto not_randompred Random map
--- a/src/HOL/Random_Sequence.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Random_Sequence.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -42,7 +42,7 @@
 
 definition map :: "('a => 'b) => 'a random_dseq => 'b random_dseq"
 where
-  "map f P = bind P (single o f)"
+  "map f P = bind P (single \<circ> f)"
 
 fun Random :: "(natural \<Rightarrow> Random.seed \<Rightarrow> (('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)) \<Rightarrow> 'a random_dseq"
 where
@@ -82,7 +82,7 @@
 
 definition pos_map :: "('a => 'b) => 'a pos_random_dseq => 'b pos_random_dseq"
 where
-  "pos_map f P = pos_bind P (pos_single o f)"
+  "pos_map f P = pos_bind P (pos_single \<circ> f)"
 
 fun iter :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
   \<Rightarrow> natural \<Rightarrow> Random.seed \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
@@ -132,7 +132,7 @@
 
 definition neg_map :: "('a => 'b) => 'a neg_random_dseq => 'b neg_random_dseq"
 where
-  "neg_map f P = neg_bind P (neg_single o f)"
+  "neg_map f P = neg_bind P (neg_single \<circ> f)"
 
 definition pos_not_random_dseq :: "unit neg_random_dseq => unit pos_random_dseq"
 where
--- a/src/HOL/Record.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Record.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -224,70 +224,70 @@
 lemmas simps = Let_def repr_inv abst_inv repr_inj abst_inj
 
 lemma iso_tuple_access_update_fst_fst:
-  "f o h g = j o f \<Longrightarrow>
-    (f o iso_tuple_fst isom) o (iso_tuple_fst_update isom o h) g =
-      j o (f o iso_tuple_fst isom)"
+  "f \<circ> h g = j \<circ> f \<Longrightarrow>
+    (f \<circ> iso_tuple_fst isom) \<circ> (iso_tuple_fst_update isom \<circ> h) g =
+      j \<circ> (f \<circ> iso_tuple_fst isom)"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_fst_def simps
     fun_eq_iff)
 
 lemma iso_tuple_access_update_snd_snd:
-  "f o h g = j o f \<Longrightarrow>
-    (f o iso_tuple_snd isom) o (iso_tuple_snd_update isom o h) g =
-      j o (f o iso_tuple_snd isom)"
+  "f \<circ> h g = j \<circ> f \<Longrightarrow>
+    (f \<circ> iso_tuple_snd isom) \<circ> (iso_tuple_snd_update isom \<circ> h) g =
+      j \<circ> (f \<circ> iso_tuple_snd isom)"
   by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_snd_def simps
     fun_eq_iff)
 
 lemma iso_tuple_access_update_fst_snd:
-  "(f o iso_tuple_fst isom) o (iso_tuple_snd_update isom o h) g =
-    id o (f o iso_tuple_fst isom)"
+  "(f \<circ> iso_tuple_fst isom) \<circ> (iso_tuple_snd_update isom \<circ> h) g =
+    id \<circ> (f \<circ> iso_tuple_fst isom)"
   by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_fst_def simps
     fun_eq_iff)
 
 lemma iso_tuple_access_update_snd_fst:
-  "(f o iso_tuple_snd isom) o (iso_tuple_fst_update isom o h) g =
-    id o (f o iso_tuple_snd isom)"
+  "(f \<circ> iso_tuple_snd isom) \<circ> (iso_tuple_fst_update isom \<circ> h) g =
+    id \<circ> (f \<circ> iso_tuple_snd isom)"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_def simps
     fun_eq_iff)
 
 lemma iso_tuple_update_swap_fst_fst:
-  "h f o j g = j g o h f \<Longrightarrow>
-    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g =
-      (iso_tuple_fst_update isom o j) g o (iso_tuple_fst_update isom o h) f"
+  "h f \<circ> j g = j g \<circ> h f \<Longrightarrow>
+    (iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g =
+      (iso_tuple_fst_update isom \<circ> j) g \<circ> (iso_tuple_fst_update isom \<circ> h) f"
   by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose fun_eq_iff)
 
 lemma iso_tuple_update_swap_snd_snd:
-  "h f o j g = j g o h f \<Longrightarrow>
-    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g =
-      (iso_tuple_snd_update isom o j) g o (iso_tuple_snd_update isom o h) f"
+  "h f \<circ> j g = j g \<circ> h f \<Longrightarrow>
+    (iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g =
+      (iso_tuple_snd_update isom \<circ> j) g \<circ> (iso_tuple_snd_update isom \<circ> h) f"
   by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose fun_eq_iff)
 
 lemma iso_tuple_update_swap_fst_snd:
-  "(iso_tuple_snd_update isom o h) f o (iso_tuple_fst_update isom o j) g =
-    (iso_tuple_fst_update isom o j) g o (iso_tuple_snd_update isom o h) f"
+  "(iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g =
+    (iso_tuple_fst_update isom \<circ> j) g \<circ> (iso_tuple_snd_update isom \<circ> h) f"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def
     simps fun_eq_iff)
 
 lemma iso_tuple_update_swap_snd_fst:
-  "(iso_tuple_fst_update isom o h) f o (iso_tuple_snd_update isom o j) g =
-    (iso_tuple_snd_update isom o j) g o (iso_tuple_fst_update isom o h) f"
+  "(iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g =
+    (iso_tuple_snd_update isom \<circ> j) g \<circ> (iso_tuple_fst_update isom \<circ> h) f"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def simps
     fun_eq_iff)
 
 lemma iso_tuple_update_compose_fst_fst:
-  "h f o j g = k (f o g) \<Longrightarrow>
-    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g =
-      (iso_tuple_fst_update isom o k) (f o g)"
+  "h f \<circ> j g = k (f \<circ> g) \<Longrightarrow>
+    (iso_tuple_fst_update isom \<circ> h) f \<circ> (iso_tuple_fst_update isom \<circ> j) g =
+      (iso_tuple_fst_update isom \<circ> k) (f \<circ> g)"
   by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose fun_eq_iff)
 
 lemma iso_tuple_update_compose_snd_snd:
-  "h f o j g = k (f o g) \<Longrightarrow>
-    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g =
-      (iso_tuple_snd_update isom o k) (f o g)"
+  "h f \<circ> j g = k (f \<circ> g) \<Longrightarrow>
+    (iso_tuple_snd_update isom \<circ> h) f \<circ> (iso_tuple_snd_update isom \<circ> j) g =
+      (iso_tuple_snd_update isom \<circ> k) (f \<circ> g)"
   by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose fun_eq_iff)
 
 lemma iso_tuple_surjective_proof_assist_step:
-  "iso_tuple_surjective_proof_assist v a (iso_tuple_fst isom o f) \<Longrightarrow>
-    iso_tuple_surjective_proof_assist v b (iso_tuple_snd isom o f) \<Longrightarrow>
+  "iso_tuple_surjective_proof_assist v a (iso_tuple_fst isom \<circ> f) \<Longrightarrow>
+    iso_tuple_surjective_proof_assist v b (iso_tuple_snd isom \<circ> f) \<Longrightarrow>
     iso_tuple_surjective_proof_assist v (iso_tuple_cons isom a b) f"
   by (clarsimp simp: iso_tuple_surjective_proof_assist_def simps
     iso_tuple_fst_def iso_tuple_snd_def iso_tuple_cons_def)
@@ -295,7 +295,7 @@
 lemma iso_tuple_fst_update_accessor_cong_assist:
   assumes "iso_tuple_update_accessor_cong_assist f g"
   shows "iso_tuple_update_accessor_cong_assist
-    (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)"
+    (iso_tuple_fst_update isom \<circ> f) (g \<circ> iso_tuple_fst isom)"
 proof -
   from assms have "f id = id"
     by (rule iso_tuple_update_accessor_cong_assist_id)
@@ -307,7 +307,7 @@
 lemma iso_tuple_snd_update_accessor_cong_assist:
   assumes "iso_tuple_update_accessor_cong_assist f g"
   shows "iso_tuple_update_accessor_cong_assist
-    (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)"
+    (iso_tuple_snd_update isom \<circ> f) (g \<circ> iso_tuple_snd isom)"
 proof -
   from assms have "f id = id"
     by (rule iso_tuple_update_accessor_cong_assist_id)
@@ -319,7 +319,7 @@
 lemma iso_tuple_fst_update_accessor_eq_assist:
   assumes "iso_tuple_update_accessor_eq_assist f g a u a' v"
   shows "iso_tuple_update_accessor_eq_assist
-    (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)
+    (iso_tuple_fst_update isom \<circ> f) (g \<circ> iso_tuple_fst isom)
     (iso_tuple_cons isom a b) u (iso_tuple_cons isom a' b) v"
 proof -
   from assms have "f id = id"
@@ -334,7 +334,7 @@
 lemma iso_tuple_snd_update_accessor_eq_assist:
   assumes "iso_tuple_update_accessor_eq_assist f g b u b' v"
   shows "iso_tuple_update_accessor_eq_assist
-    (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)
+    (iso_tuple_snd_update isom \<circ> f) (g \<circ> iso_tuple_snd isom)
     (iso_tuple_cons isom a b) u (iso_tuple_cons isom a b') v"
 proof -
   from assms have "f id = id"
--- a/src/HOL/SMT.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/SMT.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -371,7 +371,7 @@
   by auto
 
 lemma [z3_rule]:
-  "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not> P)))"
+  "((P = Q) \<longrightarrow> R) = (R \<or> (Q = (\<not> P)))"
   by auto
 
 lemma [z3_rule]:
--- a/src/HOL/Set.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Set.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -837,7 +837,7 @@
 lemma subset_Diff_insert: "A \<subseteq> B - insert x C \<longleftrightarrow> A \<subseteq> B - C \<and> x \<notin> A"
   by blast
 
-lemma doubleton_eq_iff: "{a, b} = {c, d} \<longleftrightarrow> a = c \<and> b = d \<or> a = d & b = c"
+lemma doubleton_eq_iff: "{a, b} = {c, d} \<longleftrightarrow> a = c \<and> b = d \<or> a = d \<and> b = c"
   by (blast elim: equalityE)
 
 lemma Un_singleton_iff: "A \<union> B = {x} \<longleftrightarrow> A = {} \<and> B = {x} \<or> A = {x} \<and> B = {} \<or> A = {x} \<and> B = {x}"
@@ -1584,7 +1584,7 @@
   "\<And>A P Q. (\<exists>x\<in>A. P \<and> Q x) \<longleftrightarrow> (P \<and> (\<exists>x\<in>A. Q x))"
   "\<And>P. (\<exists>x\<in>{}. P x) \<longleftrightarrow> False"
   "\<And>P. (\<exists>x\<in>UNIV. P x) \<longleftrightarrow> (\<exists>x. P x)"
-  "\<And>a B P. (\<exists>x\<in>insert a B. P x) \<longleftrightarrow> (P a | (\<exists>x\<in>B. P x))"
+  "\<And>a B P. (\<exists>x\<in>insert a B. P x) \<longleftrightarrow> (P a \<or> (\<exists>x\<in>B. P x))"
   "\<And>P Q. (\<exists>x\<in>Collect Q. P x) \<longleftrightarrow> (\<exists>x. Q x \<and> P x)"
   "\<And>A P f. (\<exists>x\<in>f`A. P x) \<longleftrightarrow> (\<exists>x\<in>A. P (f x))"
   "\<And>A P. (\<not>(\<exists>x\<in>A. P x)) \<longleftrightarrow> (\<forall>x\<in>A. \<not> P x)"
--- a/src/HOL/Set_Interval.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Set_Interval.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -196,19 +196,19 @@
 begin
 
 lemma greaterThanLessThan_iff [simp]:
-  "(i : {l<..<u}) = (l < i & i < u)"
+  "(i \<in> {l<..<u}) = (l < i \<and> i < u)"
 by (simp add: greaterThanLessThan_def)
 
 lemma atLeastLessThan_iff [simp]:
-  "(i : {l..<u}) = (l <= i & i < u)"
+  "(i \<in> {l..<u}) = (l \<le> i \<and> i < u)"
 by (simp add: atLeastLessThan_def)
 
 lemma greaterThanAtMost_iff [simp]:
-  "(i : {l<..u}) = (l < i & i <= u)"
+  "(i \<in> {l<..u}) = (l < i \<and> i \<le> u)"
 by (simp add: greaterThanAtMost_def)
 
 lemma atLeastAtMost_iff [simp]:
-  "(i : {l..u}) = (l <= i & i <= u)"
+  "(i \<in> {l..u}) = (l \<le> i \<and> i \<le> u)"
 by (simp add: atLeastAtMost_def)
 
 text \<open>The above four lemmas could be declared as iffs. Unfortunately this
@@ -234,11 +234,11 @@
 by(auto simp: atLeastAtMost_def atLeast_def atMost_def)
 
 lemma atLeastatMost_empty_iff[simp]:
-  "{a..b} = {} \<longleftrightarrow> (~ a <= b)"
+  "{a..b} = {} \<longleftrightarrow> (\<not> a \<le> b)"
 by auto (blast intro: order_trans)
 
 lemma atLeastatMost_empty_iff2[simp]:
-  "{} = {a..b} \<longleftrightarrow> (~ a <= b)"
+  "{} = {a..b} \<longleftrightarrow> (\<not> a \<le> b)"
 by auto (blast intro: order_trans)
 
 lemma atLeastLessThan_empty[simp]:
@@ -246,20 +246,20 @@
 by(auto simp: atLeastLessThan_def)
 
 lemma atLeastLessThan_empty_iff[simp]:
-  "{a..<b} = {} \<longleftrightarrow> (~ a < b)"
+  "{a..<b} = {} \<longleftrightarrow> (\<not> a < b)"
 by auto (blast intro: le_less_trans)
 
 lemma atLeastLessThan_empty_iff2[simp]:
-  "{} = {a..<b} \<longleftrightarrow> (~ a < b)"
+  "{} = {a..<b} \<longleftrightarrow> (\<not> a < b)"
 by auto (blast intro: le_less_trans)
 
 lemma greaterThanAtMost_empty[simp]: "l \<le> k ==> {k<..l} = {}"
 by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def)
 
-lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \<longleftrightarrow> ~ k < l"
+lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \<longleftrightarrow> \<not> k < l"
 by auto (blast intro: less_le_trans)
 
-lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \<longleftrightarrow> ~ k < l"
+lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \<longleftrightarrow> \<not> k < l"
 by auto (blast intro: less_le_trans)
 
 lemma greaterThanLessThan_empty[simp]:"l \<le> k ==> {k<..<l} = {}"
@@ -271,13 +271,13 @@
 lemma atLeastAtMost_singleton': "a = b \<Longrightarrow> {a .. b} = {a}" by simp
 
 lemma atLeastatMost_subset_iff[simp]:
-  "{a..b} <= {c..d} \<longleftrightarrow> (~ a <= b) | c <= a & b <= d"
+  "{a..b} \<le> {c..d} \<longleftrightarrow> (\<not> a \<le> b) \<or> c \<le> a \<and> b \<le> d"
 unfolding atLeastAtMost_def atLeast_def atMost_def
 by (blast intro: order_trans)
 
 lemma atLeastatMost_psubset_iff:
   "{a..b} < {c..d} \<longleftrightarrow>
-   ((~ a <= b) | c <= a & b <= d & (c < a | b < d))  &  c <= d"
+   ((\<not> a \<le> b) \<or> c \<le> a \<and> b \<le> d \<and> (c < a \<or> b < d)) \<and> c \<le> d"
 by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans)
 
 lemma Icc_eq_Icc[simp]:
@@ -294,11 +294,11 @@
 qed simp
 
 lemma Icc_subset_Ici_iff[simp]:
-  "{l..h} \<subseteq> {l'..} = (~ l\<le>h \<or> l\<ge>l')"
+  "{l..h} \<subseteq> {l'..} = (\<not> l\<le>h \<or> l\<ge>l')"
 by(auto simp: subset_eq intro: order_trans)
 
 lemma Icc_subset_Iic_iff[simp]:
-  "{l..h} \<subseteq> {..h'} = (~ l\<le>h \<or> h\<le>h')"
+  "{l..h} \<subseteq> {..h'} = (\<not> l\<le>h \<or> h\<le>h')"
 by(auto simp: subset_eq intro: order_trans)
 
 lemma not_Ici_eq_empty[simp]: "{l..} \<noteq> {}"
@@ -462,7 +462,7 @@
 end
 
 lemma (in linorder) atLeastLessThan_subset_iff:
-  "{a..<b} <= {c..<d} \<Longrightarrow> b <= a | c<=a & b<=d"
+  "{a..<b} \<subseteq> {c..<d} \<Longrightarrow> b \<le> a \<or> c\<le>a \<and> b\<le>d"
 apply (auto simp:subset_eq Ball_def)
 apply(frule_tac x=a in spec)
 apply(erule_tac x=d in allE)
@@ -1092,7 +1092,7 @@
 
 text \<open>A set of natural numbers is finite iff it is bounded.\<close>
 lemma finite_nat_set_iff_bounded:
-  "finite(N::nat set) = (EX m. ALL n:N. n<m)" (is "?F = ?B")
+  "finite(N::nat set) = (\<exists>m. \<forall>n\<in>N. n<m)" (is "?F = ?B")
 proof
   assume f:?F  show ?B
     using Max_ge[OF \<open>?F\<close>, simplified less_Suc_eq_le[symmetric]] by blast
@@ -1101,7 +1101,7 @@
 qed
 
 lemma finite_nat_set_iff_bounded_le:
-  "finite(N::nat set) = (EX m. ALL n:N. n<=m)"
+  "finite(N::nat set) = (\<exists>m. \<forall>n\<in>N. n<=m)"
 apply(simp add:finite_nat_set_iff_bounded)
 apply(blast dest:less_imp_le_nat le_imp_less_Suc)
 done
@@ -1177,8 +1177,8 @@
   proof
     fix x assume "x : ?B"
     then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto
-    hence "i-k\<le>n & x : M((i-k)+k)" by auto
-    thus "x : ?A" by blast
+    hence "i-k\<le>n \<and> x \<in> M((i-k)+k)" by auto
+    thus "x \<in> ?A" by blast
   qed
 qed
 
@@ -1263,7 +1263,7 @@
 by (blast dest: ex_bij_betw_nat_finite bij_betw_inv)
 
 lemma finite_same_card_bij:
-  "finite A \<Longrightarrow> finite B \<Longrightarrow> card A = card B \<Longrightarrow> EX h. bij_betw h A B"
+  "finite A \<Longrightarrow> finite B \<Longrightarrow> card A = card B \<Longrightarrow> \<exists>h. bij_betw h A B"
 apply(drule ex_bij_betw_finite_nat)
 apply(drule ex_bij_betw_nat_finite)
 apply(auto intro!:bij_betw_trans)
@@ -1282,7 +1282,7 @@
     using assms ex_bij_betw_finite_nat by blast
   moreover obtain g where "bij_betw g {0 ..< card B} B"
     using assms ex_bij_betw_nat_finite by blast
-  ultimately have "bij_betw (g o f) A B"
+  ultimately have "bij_betw (g \<circ> f) A B"
     by (auto simp: bij_betw_trans)
   thus "(\<exists>f. bij_betw f A B)" by blast
 qed (auto simp: bij_betw_same_card)
@@ -1297,11 +1297,11 @@
   moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B"
   using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force
   ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force
-  hence "inj_on (g o f) A" using 1 comp_inj_on by blast
+  hence "inj_on (g \<circ> f) A" using 1 comp_inj_on by blast
   moreover
   {have "{0 ..< card A} \<le> {0 ..< card B}" using * by force
    with 2 have "f ` A  \<le> {0 ..< card B}" by blast
-   hence "(g o f) ` A \<le> B" unfolding comp_def using 3 by force
+   hence "(g \<circ> f) ` A \<le> B" unfolding comp_def using 3 by force
   }
   ultimately show "(\<exists>f. inj_on f A \<and> f ` A \<le> B)" by blast
 qed (insert assms, auto)
@@ -1541,7 +1541,7 @@
 subsubsection \<open>Some Subset Conditions\<close>
 
 lemma ivl_subset [simp]:
- "({i..<j} \<subseteq> {m..<n}) = (j \<le> i | m \<le> i & j \<le> (n::'a::linorder))"
+ "({i..<j} \<subseteq> {m..<n}) = (j \<le> i \<or> m \<le> i \<and> j \<le> (n::'a::linorder))"
 apply(auto simp:linorder_not_le)
 apply(rule ccontr)
 apply(insert linorder_le_less_linear[of i n])
@@ -1809,7 +1809,7 @@
 by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
 
 lemma sum_nat_group: "(\<Sum>m<n::nat. sum f {m * k ..< m*k + k}) = sum f {..< n * k}"
-  apply (subgoal_tac "k = 0 | 0 < k", auto)
+  apply (subgoal_tac "k = 0 \<or> 0 < k", auto)
   apply (induct "n")
   apply (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric])
   done
--- a/src/HOL/String.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/String.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -359,7 +359,7 @@
 subsection \<open>Dedicated conversion for generated computations\<close>
 
 definition char_of_num :: "num \<Rightarrow> char"
-  where "char_of_num = char_of_nat o nat_of_num"
+  where "char_of_num = char_of_nat \<circ> nat_of_num"
 
 lemma [code_computation_unfold]:
   "Char = char_of_num"
--- a/src/HOL/Tools/ATP/atp_problem_generate.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/ATP/atp_problem_generate.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -1598,7 +1598,7 @@
 
 (** Helper facts **)
 
-val not_ffalse = @{lemma "~ fFalse" by (unfold fFalse_def) fast}
+val not_ffalse = @{lemma "\<not> fFalse" by (unfold fFalse_def) fast}
 val ftrue = @{lemma "fTrue" by (unfold fTrue_def) fast}
 
 (* The Boolean indicates that a fairly sound type encoding is needed. *)
@@ -1624,13 +1624,13 @@
            fNot_def [THEN Meson.iff_to_disjD, THEN conjunct2]}
     |> map (pair Non_Rec_Def)),
    (("fconj", false),
-    @{lemma "~ P | ~ Q | fconj P Q" "~ fconj P Q | P" "~ fconj P Q | Q" by (unfold fconj_def) fast+}
+    @{lemma "\<not> P \<or> \<not> Q \<or> fconj P Q" "\<not> fconj P Q \<or> P" "\<not> fconj P Q \<or> Q" by (unfold fconj_def) fast+}
     |> map (pair General)),
    (("fdisj", false),
-    @{lemma "~ P | fdisj P Q" "~ Q | fdisj P Q" "~ fdisj P Q | P | Q" by (unfold fdisj_def) fast+}
+    @{lemma "\<not> P \<or> fdisj P Q" "\<not> Q \<or> fdisj P Q" "\<not> fdisj P Q \<or> P \<or> Q" by (unfold fdisj_def) fast+}
     |> map (pair General)),
    (("fimplies", false),
-    @{lemma "P | fimplies P Q" "~ Q | fimplies P Q" "~ fimplies P Q | ~ P | Q"
+    @{lemma "P \<or> fimplies P Q" "\<not> Q \<or> fimplies P Q" "\<not> fimplies P Q \<or> \<not> P \<or> Q"
         by (unfold fimplies_def) fast+}
     |> map (pair General)),
    (("fequal", true),
@@ -1642,8 +1642,8 @@
     |> map (pair General)),
    (* Partial characterization of "fAll" and "fEx". A complete characterization
       would require the axiom of choice for replay with Metis. *)
-   (("fAll", false), [(General, @{lemma "~ fAll P | P x" by (auto simp: fAll_def)})]),
-   (("fEx", false), [(General, @{lemma "~ P x | fEx P" by (auto simp: fEx_def)})])]
+   (("fAll", false), [(General, @{lemma "\<not> fAll P \<or> P x" by (auto simp: fAll_def)})]),
+   (("fEx", false), [(General, @{lemma "\<not> P x \<or> fEx P" by (auto simp: fEx_def)})])]
   |> map (apsnd (map (apsnd zero_var_indexes)))
 
 val completish_helper_table =
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -285,10 +285,10 @@
         else if s = tptp_and then HOLogic.conj
         else if s = tptp_implies then HOLogic.imp
         else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT
-        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "%P Q. Q ~= P"}
-        else if s = tptp_if then @{term "%P Q. Q --> P"}
-        else if s = tptp_not_and then @{term "%P Q. ~ (P & Q)"}
-        else if s = tptp_not_or then @{term "%P Q. ~ (P | Q)"}
+        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "\<lambda>P Q. Q \<noteq> P"}
+        else if s = tptp_if then @{term "\<lambda>P Q. Q \<longrightarrow> P"}
+        else if s = tptp_not_and then @{term "\<lambda>P Q. \<not> (P \<and> Q)"}
+        else if s = tptp_not_or then @{term "\<lambda>P Q. \<not> (P \<or> Q)"}
         else if s = tptp_not then HOLogic.Not
         else if s = tptp_ho_forall then HOLogic.all_const dummyT
         else if s = tptp_ho_exists then HOLogic.exists_const dummyT
--- a/src/HOL/Tools/Argo/argo_real.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Argo/argo_real.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -241,8 +241,8 @@
   by (simp add: max_def)}
 val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp}
 val sub_eq_thm = mk_rewr @{lemma "((x::real) = y) = (x - y = 0)" by simp}
-val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x <= y) & (y <= x))" by (rule order_eq_iff)}
-val add_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x + n <= y + n)" by simp}
+val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x \<le> y) \<and> (y \<le> x))" by (rule order_eq_iff)}
+val add_le_thm = mk_rewr @{lemma "((x::real) \<le> y) = (x + n \<le> y + n)" by simp}
 val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp}
 val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp}
 val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp}
@@ -250,8 +250,8 @@
 val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp}
 val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp}
 val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp}
-val not_le_thm = mk_rewr @{lemma "(~((x::real) <= y)) = (y < x)" by auto}
-val not_lt_thm = mk_rewr @{lemma "(~((x::real) < y)) = (y <= x)" by auto}
+val not_le_thm = mk_rewr @{lemma "(\<not>((x::real) \<le> y)) = (y < x)" by auto}
+val not_lt_thm = mk_rewr @{lemma "(\<not>((x::real) < y)) = (y \<le> x)" by auto}
 
 fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm
   | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps
--- a/src/HOL/Tools/Argo/argo_tactic.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Argo/argo_tactic.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -325,12 +325,12 @@
 fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)]
 fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts
 
-val iff_1_taut = @{lemma "P = Q | P | Q" by fast}
-val iff_2_taut = @{lemma "P = Q | (~P) | (~Q)" by fast}
-val iff_3_taut = @{lemma "~(P = Q) | (~P) | Q" by fast}
-val iff_4_taut = @{lemma "~(P = Q) | P | (~Q)" by fast}
-val ite_then_taut = @{lemma "~P | (if P then t else u) = t" by auto}
-val ite_else_taut = @{lemma "P | (if P then t else u) = u" by auto}
+val iff_1_taut = @{lemma "P = Q \<or> P \<or> Q" by fast}
+val iff_2_taut = @{lemma "P = Q \<or> (\<not>P) \<or> (\<not>Q)" by fast}
+val iff_3_taut = @{lemma "\<not>(P = Q) \<or> (\<not>P) \<or> Q" by fast}
+val iff_4_taut = @{lemma "\<not>(P = Q) \<or> P \<or> (\<not>Q)" by fast}
+val ite_then_taut = @{lemma "\<not>P \<or> (if P then t else u) = t" by auto}
+val ite_else_taut = @{lemma "P \<or> (if P then t else u) = u" by auto}
 
 fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term
   | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i)
@@ -364,8 +364,8 @@
   if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct
   else Conv.all_conv ct
 
-val flatten_and_thm = @{lemma "(P1 & P2) & P3 == P1 & P2 & P3" by simp}
-val flatten_or_thm = @{lemma "(P1 | P2) | P3 == P1 | P2 | P3" by simp}
+val flatten_and_thm = @{lemma "(P1 \<and> P2) \<and> P3 \<equiv> P1 \<and> P2 \<and> P3" by simp}
+val flatten_or_thm = @{lemma "(P1 \<or> P2) \<or> P3 \<equiv> P1 \<or> P2 \<or> P3" by simp}
 
 fun flatten_conv cv rule ct = (
   Conv.try_conv (Conv.arg_conv cv) then_conv
@@ -384,14 +384,14 @@
 fun explode rule1 rule2 thm =
   explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm]
 val explode_conj = explode @{thm conjunct1} @{thm conjunct2}
-val explode_ndis = explode @{lemma "~(P | Q) ==> ~P" by auto} @{lemma "~(P | Q) ==> ~Q" by auto}
+val explode_ndis = explode @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>P" by auto} @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>Q" by auto}
 
 fun pick_false i thms = nth thms i
 
 fun pick_dual rule (i1, i2) thms =
   rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1]
-val pick_dual_conj = pick_dual @{lemma "~P ==> P ==> False" by auto}
-val pick_dual_ndis = pick_dual @{lemma "~P ==> P ==> ~True" by auto}
+val pick_dual_conj = pick_dual @{lemma "\<not>P \<Longrightarrow> P \<Longrightarrow> False" by auto}
+val pick_dual_ndis = pick_dual @{lemma "\<not>P \<Longrightarrow> P \<Longrightarrow> \<not>True" by auto}
 
 fun join thm0 rule is thms =
   let
@@ -399,13 +399,13 @@
     val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is []
   in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end
 
-val join_conj = join @{lemma "True" by auto} @{lemma "P ==> Q ==> P & Q" by auto}
-val join_ndis = join @{lemma "~False" by auto} @{lemma "~P ==> ~Q ==> ~(P | Q)" by auto}
+val join_conj = join @{lemma "True" by auto} @{lemma "P \<Longrightarrow> Q \<Longrightarrow> P \<and> Q" by auto}
+val join_ndis = join @{lemma "\<not>False" by auto} @{lemma "\<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> \<not>(P \<or> Q)" by auto}
 
-val false_thm = @{lemma "False ==> P" by auto}
-val ntrue_thm = @{lemma "~True ==> P" by auto}
-val iff_conj_thm = @{lemma "(P ==> Q) ==> (Q ==> P) ==> P = Q" by auto}
-val iff_ndis_thm = @{lemma "(~P ==> ~Q) ==> (~Q ==> ~P) ==> P = Q" by auto}
+val false_thm = @{lemma "False \<Longrightarrow> P" by auto}
+val ntrue_thm = @{lemma "\<not>True \<Longrightarrow> P" by auto}
+val iff_conj_thm = @{lemma "(P \<Longrightarrow> Q) \<Longrightarrow> (Q \<Longrightarrow> P) \<Longrightarrow> P = Q" by auto}
+val iff_ndis_thm = @{lemma "(\<not>P \<Longrightarrow> \<not>Q) \<Longrightarrow> (\<not>Q \<Longrightarrow> \<not>P) \<Longrightarrow> P = Q" by auto}
 
 fun iff_intro rule lf rf ct =
   let
@@ -421,21 +421,21 @@
 val sort_conj = sort_nary with_conj join_conj explode_conj
 val sort_ndis = sort_nary with_ndis join_ndis explode_ndis 
 
-val not_true_thm = mk_rewr @{lemma "(~True) = False" by auto}
-val not_false_thm = mk_rewr @{lemma "(~False) = True" by auto}
-val not_not_thm = mk_rewr @{lemma "(~~P) = P" by auto}
-val not_and_thm = mk_rewr @{lemma "(~(P & Q)) = (~P | ~Q)" by auto}
-val not_or_thm = mk_rewr @{lemma "(~(P | Q)) = (~P & ~Q)" by auto}
+val not_true_thm = mk_rewr @{lemma "(\<not>True) = False" by auto}
+val not_false_thm = mk_rewr @{lemma "(\<not>False) = True" by auto}
+val not_not_thm = mk_rewr @{lemma "(\<not>\<not>P) = P" by auto}
+val not_and_thm = mk_rewr @{lemma "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" by auto}
+val not_or_thm = mk_rewr @{lemma "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" by auto}
 val not_iff_thms = map mk_rewr
-  @{lemma "(~((~P) = Q)) = (P = Q)" "(~(P = (~Q))) = (P = Q)" "(~(P = Q)) = ((~P) = Q)" by auto}
+  @{lemma "(\<not>((\<not>P) = Q)) = (P = Q)" "(\<not>(P = (\<not>Q))) = (P = Q)" "(\<not>(P = Q)) = ((\<not>P) = Q)" by auto}
 val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto}
-val iff_false_thms = map mk_rewr @{lemma "(False = P) = (~P)" "(P = False) = (~P)" by auto}
-val iff_not_not_thm = mk_rewr @{lemma "((~P) = (~Q)) = (P = Q)" by auto}
+val iff_false_thms = map mk_rewr @{lemma "(False = P) = (\<not>P)" "(P = False) = (\<not>P)" by auto}
+val iff_not_not_thm = mk_rewr @{lemma "((\<not>P) = (\<not>Q)) = (P = Q)" by auto}
 val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto}
 val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto}
-val iff_dual_thms = map mk_rewr @{lemma "(P = (~P)) = False" "((~P) = P) = False" by auto}
-val imp_thm = mk_rewr @{lemma "(P --> Q) = (~P | Q)" by auto}
-val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((~P | Q) & (P | R) & (Q | R))" by auto}
+val iff_dual_thms = map mk_rewr @{lemma "(P = (\<not>P)) = False" "((\<not>P) = P) = False" by auto}
+val imp_thm = mk_rewr @{lemma "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" by auto}
+val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((\<not>P \<or> Q) \<and> (P \<or> R) \<and> (Q \<or> R))" by auto}
 val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto}
 val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto}
 val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto}
@@ -494,8 +494,8 @@
 
 (* proof replay for clauses *)
 
-val prep_clause_rule = @{lemma "P ==> ~P ==> False" by fast}
-val extract_lit_rule = @{lemma "(~(P | Q) ==> False) ==> ~P ==> ~Q ==> False" by fast}
+val prep_clause_rule = @{lemma "P \<Longrightarrow> \<not>P \<Longrightarrow> False" by fast}
+val extract_lit_rule = @{lemma "(\<not>(P \<or> Q) \<Longrightarrow> False) \<Longrightarrow> \<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> False" by fast}
 
 fun add_lit i thm lits =
   let val ct = Thm.cprem_of thm 1
@@ -518,7 +518,7 @@
 
 (* proof replay for unit resolution *)
 
-val unit_rule = @{lemma "(P ==> False) ==> (~P ==> False) ==> False" by fast}
+val unit_rule = @{lemma "(P \<Longrightarrow> False) \<Longrightarrow> (\<not>P \<Longrightarrow> False) \<Longrightarrow> False" by fast}
 val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1))
 val bogus_ct = @{cterm HOL.True}
 
@@ -538,7 +538,7 @@
 
 (* proof replay for hypothesis *)
 
-val dneg_rule = @{lemma "~~P ==> P" by auto}
+val dneg_rule = @{lemma "\<not>\<not>P \<Longrightarrow> P" by auto}
 
 fun replay_hyp i ct =
   if i < 0 then (Thm.assume ct, [(~i, ct)])
@@ -562,7 +562,7 @@
 
 (* proof replay for symmetry *)
 
-val symm_rules = @{lemma "a = b ==> b = a" "~(a = b) ==> ~(b = a)" by simp_all}
+val symm_rules = @{lemma "a = b ==> b = a" "\<not>(a = b) \<Longrightarrow> \<not>(b = a)" by simp_all}
 
 fun replay_symm thm = hd (discharges thm symm_rules)
 
@@ -570,9 +570,9 @@
 (* proof replay for transitivity *)
 
 val trans_rules = @{lemma
-  "~(a = b) ==> b = c ==> ~(a = c)"
-  "a = b ==> ~(b = c) ==> ~(a = c)"
-  "a = b ==> b = c ==> a = c"
+  "\<not>(a = b) \<Longrightarrow> b = c \<Longrightarrow> \<not>(a = c)"
+  "a = b \<Longrightarrow> \<not>(b = c) \<Longrightarrow> \<not>(a = c)"
+  "a = b \<Longrightarrow> b = c ==> a = c"
   by simp_all}
 
 fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules))
@@ -585,8 +585,8 @@
 
 (* proof replay for substitution *)
 
-val subst_rule1 = @{lemma "~(p a) ==> p = q ==> a = b ==> ~(q b)" by simp}
-val subst_rule2 = @{lemma "p a ==> p = q ==> a = b ==> q b" by simp}
+val subst_rule1 = @{lemma "\<not>(p a) \<Longrightarrow> p = q \<Longrightarrow> a = b \<Longrightarrow> \<not>(q b)" by simp}
+val subst_rule2 = @{lemma "p a \<Longrightarrow> p = q \<Longrightarrow> a = b \<Longrightarrow> q b" by simp}
 
 fun replay_subst thm1 thm2 thm3 =
   subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3]
@@ -596,8 +596,8 @@
 
 structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord)
 
-val unclausify_rule1 = @{lemma "(~P ==> False) ==> P" by auto}
-val unclausify_rule2 = @{lemma "(P ==> False) ==> ~P" by auto}
+val unclausify_rule1 = @{lemma "(\<not>P \<Longrightarrow> False) \<Longrightarrow> P" by auto}
+val unclausify_rule2 = @{lemma "(P \<Longrightarrow> False) \<Longrightarrow> \<not>P" by auto}
 
 fun unclausify (thm, lits) ls =
   (case (Thm.prop_of thm, lits) of
--- a/src/HOL/Tools/BNF/bnf_comp.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -932,7 +932,7 @@
        rtac ctxt refl) 1;
     fun pred_set_tac ctxt =
       HEADGOAL (EVERY'
-        [rtac ctxt (pred_set_of_bnf bnf RS @{thm arg_cong[of _ _ "\<lambda>f. f o _"]} RS trans),
+        [rtac ctxt (pred_set_of_bnf bnf RS @{thm arg_cong[of _ _ "\<lambda>f. f \<circ> _"]} RS trans),
         SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})), rtac ctxt refl]);
 
     val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -69,8 +69,8 @@
   unfold_thms_tac ctxt xtor_un_fold_defs THEN
   HEADGOAL (REPEAT_DETERM o FIRST' [hyp_subst_tac_thin true ctxt, rtac ctxt refl,
     rtac ctxt conjI,
-    rtac ctxt @{thm arg_cong2[of _ _ _ _ "op o", OF refl]},
-    rtac ctxt @{thm arg_cong2[of _ _ _ _ "op o", OF _ refl]},
+    rtac ctxt @{thm arg_cong2[of _ _ _ _ "op \<circ>", OF refl]},
+    rtac ctxt @{thm arg_cong2[of _ _ _ _ "op \<circ>", OF _ refl]},
     resolve_tac ctxt map_arg_congs,
     resolve_tac ctxt fp_un_fold_uniques THEN_ALL_NEW case_fp fp (K all_tac) (rtac ctxt sym),
     SELECT_GOAL (CHANGED (unfold_thms_tac ctxt (flat [simp_thms, map_thms, map_defs,
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -517,7 +517,7 @@
       val (tfrees, _) = BNF_Util.mk_TFrees n @{context};
       val T = mk_tupleT_balanced tfrees;
     in
-      @{thm asm_rl[of "ALL x. P x --> Q x" for P Q]}
+      @{thm asm_rl[of "\<forall>x. P x \<longrightarrow> Q x" for P Q]}
       |> Thm.instantiate' [SOME (Thm.ctyp_of @{context} T)] []
       |> Raw_Simplifier.rewrite_goals_rule @{context} @{thms split_paired_All[THEN eq_reflection]}
       |> (fn thm => impI RS funpow n (fn th => allI RS th) thm)
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -1629,7 +1629,7 @@
         val theta =
           (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) dtors);
         val dtor_unfold_dtors = (dtor_unfold_unique_thm OF
-          map (fn thm => mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "op o", OF _ refl]})
+          map (fn thm => mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "op \<circ>", OF _ refl]})
             @{thm trans[OF id_o o_id[symmetric]]}) map_id0s)
           |> split_conj_thm |> map mk_sym;
       in
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -807,10 +807,10 @@
     EVERY' (map2 (fn map_thm => fn map_comp0 =>
       EVERY' (map (rtac ctxt)
         [@{thm comp_assoc[symmetric]} RS trans,
-        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans,
+        @{thm arg_cong2[of _ _ _ _ "op \<circ>"]} OF [map_thm, refl] RS trans,
         @{thm comp_assoc[symmetric]} RS sym RS trans, map_thm RS arg_cong RS trans,
         @{thm comp_assoc[symmetric]} RS trans,
-        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp0 RS sym, refl]]))
+        @{thm arg_cong2[of _ _ _ _ "op \<circ>"]} OF [map_comp0 RS sym, refl]]))
     maps map_comp0s)] 1;
 
 fun mk_mcong_tac ctxt m coinduct_tac map_comps dtor_maps map_cong0s set_map0ss set_Jsetss
--- a/src/HOL/Tools/BNF/bnf_lfp.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -1279,7 +1279,7 @@
           (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) ctors);
         val ctor_fold_ctors = (ctor_fold_unique_thm OF
           map (fn thm => mk_trans @{thm id_o} (mk_sym (thm RS
-            @{thm trans[OF arg_cong2[of _ _ _ _ "op o", OF refl] o_id]}))) map_id0s)
+            @{thm trans[OF arg_cong2[of _ _ _ _ "op \<circ>", OF refl] o_id]}))) map_id0s)
           |> split_conj_thm |> map mk_sym;
       in
         infer_instantiate lthy theta ctor_fold_unique_thm
--- a/src/HOL/Tools/BNF/bnf_lift.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lift.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -310,7 +310,7 @@
 
             fun pred_set_tac ctxt =
               HEADGOAL (EVERY'
-                [rtac ctxt (pred_set_of_bnf bnf RS @{thm arg_cong[of _ _ "\<lambda>f. f o _"]} RS trans),
+                [rtac ctxt (pred_set_of_bnf bnf RS @{thm arg_cong[of _ _ "\<lambda>f. f \<circ> _"]} RS trans),
                 SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})),
                 rtac ctxt refl]);
 
--- a/src/HOL/Tools/Meson/meson.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Meson/meson.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -770,8 +770,8 @@
 (*Rules to convert the head literal into a negated assumption. If the head
   literal is already negated, then using notEfalse instead of notEfalse'
   prevents a double negation.*)
-val notEfalse = @{lemma "~ P ==> P ==> False" by (rule notE)};
-val notEfalse' = @{lemma "P ==> ~ P ==> False" by (rule notE)};
+val notEfalse = @{lemma "\<not> P \<Longrightarrow> P \<Longrightarrow> False" by (rule notE)};
+val notEfalse' = @{lemma "P \<Longrightarrow> \<not> P \<Longrightarrow> False" by (rule notE)};
 
 fun negated_asm_of_head th =
     th RS notEfalse handle THM _ => th RS notEfalse';
--- a/src/HOL/Tools/Meson/meson_clausify.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Meson/meson_clausify.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -292,7 +292,7 @@
 fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths
 
 val cheat_choice =
-  @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"}
+  @{prop "\<forall>x. \<exists>y. Q x y \<Longrightarrow> \<exists>f. \<forall>x. Q x (f x)"}
   |> Logic.varify_global
   |> Skip_Proof.make_thm @{theory}
 
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -99,7 +99,7 @@
 
 (* INFERENCE RULE: ASSUME *)
 
-val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
+val EXCLUDED_MIDDLE = @{lemma "P \<Longrightarrow> \<not> P \<Longrightarrow> False" by (rule notE)}
 
 fun inst_excluded_middle ctxt i_atom =
   let
@@ -256,7 +256,7 @@
    don't use this trick in general because it makes the proof object uglier than
    necessary. FIXME. *)
 fun negate_head ctxt th =
-  if exists (fn t => t aconv @{prop "~ False"}) (Thm.prems_of th) then
+  if exists (fn t => t aconv @{prop "\<not> False"}) (Thm.prems_of th) then
     (th RS @{thm select_FalseI})
     |> fold (rewrite_rule ctxt o single) @{thms not_atomize_select atomize_not_select}
   else
@@ -298,7 +298,7 @@
 
 (* INFERENCE RULE: REFL *)
 
-val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
+val REFL_THM = Thm.incr_indexes 2 @{lemma "t \<noteq> t \<Longrightarrow> False" by simp}
 
 val refl_x = Thm.cterm_of @{context} (Var (hd (Term.add_vars (Thm.prop_of REFL_THM) [])));
 val refl_idx = 1 + Thm.maxidx_of REFL_THM;
@@ -312,8 +312,8 @@
 
 (* INFERENCE RULE: EQUALITY *)
 
-val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
-val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
+val subst_em = @{lemma "s = t \<Longrightarrow> P s \<Longrightarrow> \<not> P t \<Longrightarrow> False" by simp}
+val ssubst_em = @{lemma "s = t \<Longrightarrow> P t \<Longrightarrow> \<not> P s \<Longrightarrow> False" by simp}
 
 fun equality_inference ctxt type_enc concealed sym_tab (pos, atom) fp fr =
   let val thy = Proof_Context.theory_of ctxt
--- a/src/HOL/Tools/Old_Datatype/old_datatype.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_datatype.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -47,7 +47,7 @@
 val Sumr_inject = @{thm Sumr_inject};
 
 val datatype_injI =
-  @{lemma "(!!x. ALL y. f x = f y --> x = y) ==> inj f" by (simp add: inj_on_def)};
+  @{lemma "(\<And>x. \<forall>y. f x = f y \<longrightarrow> x = y) \<Longrightarrow> inj f" by (simp add: inj_on_def)};
 
 
 (** proof of characteristic theorems **)
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -142,7 +142,7 @@
   map_filter (fn th => if Thm.no_prems th then SOME (OF1 @{thm equivp_reflp} th) else NONE
     handle THM _ => NONE) (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv}))
 
-val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
+val eq_imp_rel = @{lemma "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b" by (simp add: equivp_reflp)}
 
 fun eq_imp_rel_get ctxt =
   map (OF1 eq_imp_rel) (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv}))
@@ -569,10 +569,10 @@
    the Quot_True premise in 2nd records the lifted theorem
 *)
 val procedure_thm =
-  @{lemma  "[|A;
-              A --> B;
-              Quot_True D ==> B = C;
-              C = D|] ==> D"
+  @{lemma  "\<lbrakk>A;
+              A \<longrightarrow> B;
+              Quot_True D \<Longrightarrow> B = C;
+              C = D\<rbrakk> \<Longrightarrow> D"
       by (simp add: Quot_True_def)}
 
 (* in case of partial equivalence relations, this form of the 
--- a/src/HOL/Tools/SMT/conj_disj_perm.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/SMT/conj_disj_perm.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -20,8 +20,8 @@
 
 fun add_lit thm = Termtab.update (HOLogic.dest_Trueprop (Thm.prop_of thm), thm)
 
-val ndisj1_rule = @{lemma "~(P | Q) ==> ~P" by auto}
-val ndisj2_rule = @{lemma "~(P | Q) ==> ~Q" by auto}
+val ndisj1_rule = @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>P" by auto}
+val ndisj2_rule = @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>Q" by auto}
 
 fun explode_thm thm =
   (case HOLogic.dest_Trueprop (Thm.prop_of thm) of
@@ -33,7 +33,7 @@
 and explode_conj_thm rule1 rule2 thm lits =
   explode_thm (thm RS rule1) (explode_thm (thm RS rule2) (add_lit thm lits))
 
-val not_false_rule = @{lemma "~False" by auto}
+val not_false_rule = @{lemma "\<not>False" by auto}
 fun explode thm = explode_thm thm (add_lit not_false_rule (add_lit @{thm TrueI} Termtab.empty))
 
 fun find_dual_lit lits (@{const HOL.Not} $ t, thm) = Termtab.lookup lits t |> Option.map (pair thm)
@@ -41,8 +41,8 @@
 
 fun find_dual_lits lits = Termtab.get_first (find_dual_lit lits) lits
 
-val not_not_rule = @{lemma "P ==> ~~P" by auto}
-val ndisj_rule = @{lemma "~P ==> ~Q ==> ~(P | Q)" by auto}
+val not_not_rule = @{lemma "P \<Longrightarrow> \<not>\<not>P" by auto}
+val ndisj_rule = @{lemma "\<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> \<not>(P \<or> Q)" by auto}
 
 fun join lits t =
   (case Termtab.lookup lits t of
@@ -63,8 +63,8 @@
     val thm2 = prove_conj_disj_imp crhs clhs
   in eq_from_impls thm1 thm2 end
 
-val not_not_false_rule = @{lemma "~~False ==> P" by auto}
-val not_true_rule = @{lemma "~True ==> P" by auto}
+val not_not_false_rule = @{lemma "\<not>\<not>False \<Longrightarrow> P" by auto}
+val not_true_rule = @{lemma "\<not>True \<Longrightarrow> P" by auto}
 
 fun prove_any_imp ct =
   (case Thm.term_of ct of
@@ -94,7 +94,7 @@
     val thm2 = if to_right then prove_any_imp crhs else prove_contradiction_imp crhs
   in eq_from_impls thm1 thm2 end
 
-val contrapos_rule = @{lemma "(~P) = (~Q) ==> P = Q" by auto}
+val contrapos_rule = @{lemma "(\<not>P) = (\<not>Q) \<Longrightarrow> P = Q" by auto}
 fun contrapos prove cp = contrapos_rule OF [prove (apply2 (Thm.apply @{cterm HOL.Not}) cp)]
 
 datatype kind = True | False | Conj | Disj | Other
--- a/src/HOL/Tools/SMT/z3_replay.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/SMT/z3_replay.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -83,7 +83,7 @@
   fun rewrite_conv _ [] = Conv.all_conv
     | rewrite_conv ctxt eqs = Simplifier.full_rewrite (empty_simpset ctxt addsimps eqs)
 
-  val rewrite_true_rule = @{lemma "True == ~ False" by simp}
+  val rewrite_true_rule = @{lemma "True \<equiv> \<not> False" by simp}
   val prep_rules = [@{thm Let_def}, remove_trigger, remove_fun_app, rewrite_true_rule]
 
   fun rewrite _ [] = I
@@ -142,8 +142,8 @@
 (* |- (EX x. P x) = P c     |- ~ (ALL x. P x) = ~ P c *)
 local
   val sk_rules = @{lemma
-    "c = (SOME x. P x) ==> (EX x. P x) = P c"
-    "c = (SOME x. ~ P x) ==> (~ (ALL x. P x)) = (~ P c)"
+    "c = (SOME x. P x) \<Longrightarrow> (\<exists>x. P x) = P c"
+    "c = (SOME x. \<not> P x) \<Longrightarrow> (\<not> (\<forall>x. P x)) = (\<not> P c)"
     by (metis someI_ex)+}
 in
 
@@ -155,12 +155,12 @@
 
 end
 
-val true_thm = @{lemma "~False" by simp}
+val true_thm = @{lemma "\<not>False" by simp}
 fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl}, @{thm reflexive}, true_thm]
 
 val intro_def_rules = @{lemma
-  "(~ P | P) & (P | ~ P)"
-  "(P | ~ P) & (~ P | P)"
+  "(\<not> P \<or> P) \<and> (P \<or> \<not> P)"
+  "(P \<or> \<not> P) \<and> (\<not> P \<or> P)"
   by fast+}
 
 fun discharge_assms_tac ctxt rules =
--- a/src/HOL/Tools/SMT/z3_replay_methods.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/SMT/z3_replay_methods.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -319,8 +319,8 @@
   end
 
 val cong_dest_rules = @{lemma
-  "(~ P | Q) & (P | ~ Q) ==> P = Q"
-  "(P | ~ Q) & (~ P | Q) ==> P = Q"
+  "(\<not> P \<or> Q) \<and> (P \<or> \<not> Q) \<Longrightarrow> P = Q"
+  "(P \<or> \<not> Q) \<and> (\<not> P \<or> Q) \<Longrightarrow> P = Q"
   by fast+}
 
 fun cong_full_core_tac ctxt =
@@ -342,10 +342,10 @@
 (* quantifier introduction *)
 
 val quant_intro_rules = @{lemma
-  "(!!x. P x = Q x) ==> (ALL x. P x) = (ALL x. Q x)"
-  "(!!x. P x = Q x) ==> (EX x. P x) = (EX x. Q x)"
-  "(!!x. (~ P x) = Q x) ==> (~ (EX x. P x)) = (ALL x. Q x)"
-  "(!!x. (~ P x) = Q x) ==> (~ (ALL x. P x)) = (EX x. Q x)"
+  "(\<And>x. P x = Q x) ==> (\<forall>x. P x) = (\<forall>x. Q x)"
+  "(\<And>x. P x = Q x) ==> (\<exists>x. P x) = (\<exists>x. Q x)"
+  "(!!x. (\<not> P x) = Q x) \<Longrightarrow> (\<not>(\<exists>x. P x)) = (\<forall>x. Q x)"
+  "(\<And>x. (\<not> P x) = Q x) ==> (\<not>(\<forall>x. P x)) = (\<exists>x. Q x)"
   by fast+}
 
 fun quant_intro ctxt [thm] t =
@@ -476,8 +476,8 @@
 
 (* elimination of unused bound variables *)
 
-val elim_all = @{lemma "P = Q ==> (ALL x. P) = Q" by fast}
-val elim_ex = @{lemma "P = Q ==> (EX x. P) = Q" by fast}
+val elim_all = @{lemma "P = Q \<Longrightarrow> (\<forall>x. P) = Q" by fast}
+val elim_ex = @{lemma "P = Q \<Longrightarrow> (\<exists>x. P) = Q" by fast}
 
 fun elim_unused_tac ctxt i st = (
   match_tac ctxt [@{thm refl}]
@@ -496,7 +496,7 @@
 
 (* quantifier instantiation *)
 
-val quant_inst_rule = @{lemma "~P x | Q ==> ~(ALL x. P x) | Q" by fast}
+val quant_inst_rule = @{lemma "\<not>P x \<or> Q ==> \<not>(\<forall>x. P x) \<or> Q" by fast}
 
 fun quant_inst ctxt _ t = prove ctxt t (fn _ =>
   REPEAT_ALL_NEW (resolve_tac ctxt [quant_inst_rule])
@@ -507,8 +507,8 @@
 
 exception LEMMA of unit
 
-val intro_hyp_rule1 = @{lemma "(~P ==> Q) ==> P | Q" by fast}
-val intro_hyp_rule2 = @{lemma "(P ==> Q) ==> ~P | Q" by fast}
+val intro_hyp_rule1 = @{lemma "(\<not>P \<Longrightarrow> Q) \<Longrightarrow> P \<or> Q" by fast}
+val intro_hyp_rule2 = @{lemma "(P \<Longrightarrow> Q) \<Longrightarrow> \<not>P \<or> Q" by fast}
 
 fun norm_lemma thm =
   (thm COMP_INCR intro_hyp_rule1)
@@ -568,7 +568,7 @@
 
 (* iff-false *)
 
-val iff_false_rule = @{lemma "~P ==> P = False" by fast}
+val iff_false_rule = @{lemma "\<not>P \<Longrightarrow> P = False" by fast}
 
 fun iff_false _ [thm] _ = thm RS iff_false_rule
   | iff_false ctxt thms t = replay_rule_error ctxt Z3_Proof.Iff_False thms t
--- a/src/HOL/Tools/cnf.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/cnf.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -63,15 +63,15 @@
 val conj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P \<and> Q) = (P' \<and> Q')" by auto};
 val disj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P \<or> Q) = (P' \<or> Q')" by auto};
 
-val make_nnf_imp         = @{lemma "[| (~P) = P'; Q = Q' |] ==> (P \<longrightarrow> Q) = (P' \<or> Q')" by auto};
-val make_nnf_iff         = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' \<or> NQ) \<and> (NP \<or> Q'))" by auto};
-val make_nnf_not_false   = @{lemma "(~False) = True" by auto};
-val make_nnf_not_true    = @{lemma "(~True) = False" by auto};
-val make_nnf_not_conj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P \<and> Q)) = (P' \<or> Q')" by auto};
-val make_nnf_not_disj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P \<or> Q)) = (P' \<and> Q')" by auto};
-val make_nnf_not_imp     = @{lemma "[| P = P'; (~Q) = Q' |] ==> (~(P \<longrightarrow> Q)) = (P' \<and> Q')" by auto};
-val make_nnf_not_iff     = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' \<or> Q') \<and> (NP \<or> NQ))" by auto};
-val make_nnf_not_not     = @{lemma "P = P' ==> (~~P) = P'" by auto};
+val make_nnf_imp         = @{lemma "[| (\<not>P) = P'; Q = Q' |] ==> (P \<longrightarrow> Q) = (P' \<or> Q')" by auto};
+val make_nnf_iff         = @{lemma "[| P = P'; (\<not>P) = NP; Q = Q'; (\<not>Q) = NQ |] ==> (P = Q) = ((P' \<or> NQ) \<and> (NP \<or> Q'))" by auto};
+val make_nnf_not_false   = @{lemma "(\<not>False) = True" by auto};
+val make_nnf_not_true    = @{lemma "(\<not>True) = False" by auto};
+val make_nnf_not_conj    = @{lemma "[| (\<not>P) = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<and> Q)) = (P' \<or> Q')" by auto};
+val make_nnf_not_disj    = @{lemma "[| (\<not>P) = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<or> Q)) = (P' \<and> Q')" by auto};
+val make_nnf_not_imp     = @{lemma "[| P = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<longrightarrow> Q)) = (P' \<and> Q')" by auto};
+val make_nnf_not_iff     = @{lemma "[| P = P'; (\<not>P) = NP; Q = Q'; (\<not>Q) = NQ |] ==> (\<not>(P = Q)) = ((P' \<or> Q') \<and> (NP \<or> NQ))" by auto};
+val make_nnf_not_not     = @{lemma "P = P' ==> (\<not>\<not>P) = P'" by auto};
 
 val simp_TF_conj_True_l  = @{lemma "[| P = True; Q = Q' |] ==> (P \<and> Q) = Q'" by auto};
 val simp_TF_conj_True_r  = @{lemma "[| P = P'; Q = True |] ==> (P \<and> Q) = P'" by auto};
@@ -85,10 +85,10 @@
 val make_cnf_disj_conj_l = @{lemma "[| (P \<or> R) = PR; (Q \<or> R) = QR |] ==> ((P \<and> Q) \<or> R) = (PR \<and> QR)" by auto};
 val make_cnf_disj_conj_r = @{lemma "[| (P \<or> Q) = PQ; (P \<or> R) = PR |] ==> (P \<or> (Q \<and> R)) = (PQ \<and> PR)" by auto};
 
-val make_cnfx_disj_ex_l  = @{lemma "((EX (x::bool). P x) \<or> Q) = (EX x. P x \<or> Q)" by auto};
-val make_cnfx_disj_ex_r  = @{lemma "(P \<or> (EX (x::bool). Q x)) = (EX x. P \<or> Q x)" by auto};
-val make_cnfx_newlit     = @{lemma "(P \<or> Q) = (EX x. (P \<or> x) \<and> (Q \<or> ~x))" by auto};
-val make_cnfx_ex_cong    = @{lemma "(ALL (x::bool). P x = Q x) \<Longrightarrow> (EX x. P x) = (EX x. Q x)" by auto};
+val make_cnfx_disj_ex_l  = @{lemma "((\<exists>(x::bool). P x) \<or> Q) = (\<exists>x. P x \<or> Q)" by auto};
+val make_cnfx_disj_ex_r  = @{lemma "(P \<or> (\<exists>(x::bool). Q x)) = (\<exists>x. P \<or> Q x)" by auto};
+val make_cnfx_newlit     = @{lemma "(P \<or> Q) = (\<exists>x. (P \<or> x) \<and> (Q \<or> \<not>x))" by auto};
+val make_cnfx_ex_cong    = @{lemma "(\<forall>(x::bool). P x = Q x) \<Longrightarrow> (\<exists>x. P x) = (\<exists>x. Q x)" by auto};
 
 val weakening_thm        = @{lemma "[| P; Q |] ==> Q" by auto};
 
--- a/src/HOL/Tools/inductive.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/inductive.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -108,9 +108,9 @@
 
 val simp_thms1 =
   map mk_meta_eq
-    @{lemma "(~ True) = False" "(~ False) = True"
-        "(True --> P) = P" "(False --> P) = True"
-        "(P & True) = P" "(True & P) = P"
+    @{lemma "(\<not> True) = False" "(\<not> False) = True"
+        "(True \<longrightarrow> P) = P" "(False \<longrightarrow> P) = True"
+        "(P \<and> True) = P" "(True \<and> P) = P"
       by (fact simp_thms)+};
 
 val simp_thms2 =
@@ -420,7 +420,7 @@
       (mono RS (fp_def RS
         (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold})));
 
-    val rules = [refl, TrueI, @{lemma "~ False" by (rule notI)}, exI, conjI];
+    val rules = [refl, TrueI, @{lemma "\<not> False" by (rule notI)}, exI, conjI];
 
     val intrs = map_index (fn (i, intr) =>
       Goal.prove_sorry ctxt [] [] intr (fn _ => EVERY
@@ -451,7 +451,7 @@
     val intrs = map dest_intr intr_ts ~~ intr_names;
 
     val rules1 = [disjE, exE, FalseE];
-    val rules2 = [conjE, FalseE, @{lemma "~ True ==> R" by (rule notE [OF _ TrueI])}];
+    val rules2 = [conjE, FalseE, @{lemma "\<not> True \<Longrightarrow> R" by (rule notE [OF _ TrueI])}];
 
     fun prove_elim c =
       let
--- a/src/HOL/Tools/rewrite_hol_proof.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/rewrite_hol_proof.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -118,7 +118,7 @@
    "(iffD1 % Ex P % Ex Q %% (cong % TYPE('T1) % TYPE('T2) % Ex % Ex % P % Q %% prfT1 %% prfT2 %%  \
  \    (HOL.refl % TYPE('T3) % x1 %% prfT3) %%  \
  \    (ext % TYPE('a) % TYPE(bool) % x2 % x3 %% prfa %% prfb %% prf)) %% prf') ==  \
- \  (exE % TYPE('a) % P % EX x. Q x %% prfa %% prf' %%  \
+ \  (exE % TYPE('a) % P % \<exists>x. Q x %% prfa %% prf' %%  \
  \    (Lam x H : P x.  \
  \        exI % TYPE('a) % Q % x %% prfa %%  \
  \         (iffD1 % P x % Q x %% (prf % x) %% H)))",
@@ -126,95 +126,95 @@
    "(iffD2 % Ex P % Ex Q %% (cong % TYPE('T1) % TYPE('T2) % Ex % Ex % P % Q %% prfT1 %% prfT2 %%  \
  \    (HOL.refl % TYPE('T3) % x1 %% prfT3) %%  \
  \    (ext % TYPE('a) % TYPE(bool) % x2 % x3 %% prfa %% prfb %% prf)) %% prf') ==  \
- \  (exE % TYPE('a) % Q % EX x. P x %% prfa %% prf' %%  \
+ \  (exE % TYPE('a) % Q % \<exists>x. P x %% prfa %% prf' %%  \
  \    (Lam x H : Q x.  \
  \        exI % TYPE('a) % P % x %% prfa %%  \
  \         (iffD2 % P x % Q x %% (prf % x) %% H)))",
 
-   (* & *)
+   (* \<and> *)
 
-   "(iffD1 % A & C % B & D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op & % op & % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op & %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+   "(iffD1 % A \<and> C % B \<and> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<and> % op \<and> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<and> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
  \  (conjI % B % D %%  \
  \    (iffD1 % A % B %% prf1 %% (conjunct1 % A % C %% prf3)) %%  \
  \    (iffD1 % C % D %% prf2 %% (conjunct2 % A % C %% prf3)))",
 
-   "(iffD2 % A & C % B & D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op & % op & % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op & %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+   "(iffD2 % A \<and> C % B \<and> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<and> % op \<and> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<and> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
  \  (conjI % A % C %%  \
  \    (iffD2 % A % B %% prf1 %% (conjunct1 % B % D %% prf3)) %%  \
  \    (iffD2 % C % D %% prf2 %% (conjunct2 % B % D %% prf3)))",
 
-   "(cong % TYPE(bool) % TYPE(bool) % op & A % op & A % B % C %% prfb %% prfb %%  \
- \    (HOL.refl % TYPE(bool=>bool) % op & A %% prfbb)) ==  \
- \  (cong % TYPE(bool) % TYPE(bool) % op & A % op & A % B % C %% prfb %% prfb %%  \
+   "(cong % TYPE(bool) % TYPE(bool) % op \<and> A % op \<and> A % B % C %% prfb %% prfb %%  \
+ \    (HOL.refl % TYPE(bool=>bool) % op \<and> A %% prfbb)) ==  \
+ \  (cong % TYPE(bool) % TYPE(bool) % op \<and> A % op \<and> A % B % C %% prfb %% prfb %%  \
  \    (cong % TYPE(bool) % TYPE(bool=>bool) %  \
- \      (op & :: bool=>bool=>bool) % (op & :: bool=>bool=>bool) % A % A %%  \
+ \      (op \<and> :: bool=>bool=>bool) % (op \<and> :: bool=>bool=>bool) % A % A %%  \
  \        prfb %% prfbb %%  \
- \        (HOL.refl % TYPE(bool=>bool=>bool) % (op & :: bool=>bool=>bool) %%  \
+ \        (HOL.refl % TYPE(bool=>bool=>bool) % (op \<and> :: bool=>bool=>bool) %%  \
  \           (OfClass type_class % TYPE(bool=>bool=>bool))) %%  \
  \        (HOL.refl % TYPE(bool) % A %% prfb)))",
 
-   (* | *)
+   (* \<or> *)
 
-   "(iffD1 % A | C % B | D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op | % op | % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op | %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
- \  (disjE % A % C % B | D %% prf3 %%  \
+   "(iffD1 % A \<or> C % B \<or> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<or> % op \<or> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<or> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+ \  (disjE % A % C % B \<or> D %% prf3 %%  \
  \    (Lam H : A. disjI1 % B % D %% (iffD1 % A % B %% prf1 %% H)) %%  \
  \    (Lam H : C. disjI2 % D % B %% (iffD1 % C % D %% prf2 %% H)))",
 
-   "(iffD2 % A | C % B | D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op | % op | % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op | %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
- \  (disjE % B % D % A | C %% prf3 %%  \
+   "(iffD2 % A \<or> C % B \<or> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<or> % op \<or> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<or> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+ \  (disjE % B % D % A \<or> C %% prf3 %%  \
  \    (Lam H : B. disjI1 % A % C %% (iffD2 % A % B %% prf1 %% H)) %%  \
  \    (Lam H : D. disjI2 % C % A %% (iffD2 % C % D %% prf2 %% H)))",
 
-   "(cong % TYPE(bool) % TYPE(bool) % op | A % op | A % B % C %% prfb %% prfb %%  \
- \    (HOL.refl % TYPE(bool=>bool) % op | A %% prfbb)) ==  \
- \  (cong % TYPE(bool) % TYPE(bool) % op | A % op | A % B % C %% prfb %% prfb %%  \
+   "(cong % TYPE(bool) % TYPE(bool) % op \<or> A % op \<or> A % B % C %% prfb %% prfb %%  \
+ \    (HOL.refl % TYPE(bool=>bool) % op \<or> A %% prfbb)) ==  \
+ \  (cong % TYPE(bool) % TYPE(bool) % op \<or> A % op \<or> A % B % C %% prfb %% prfb %%  \
  \    (cong % TYPE(bool) % TYPE(bool=>bool) %  \
- \      (op | :: bool=>bool=>bool) % (op | :: bool=>bool=>bool) % A % A %%  \
+ \      (op \<or> :: bool=>bool=>bool) % (op \<or> :: bool=>bool=>bool) % A % A %%  \
  \        prfb %% prfbb %%  \
- \        (HOL.refl % TYPE(bool=>bool=>bool) % (op | :: bool=>bool=>bool) %%  \
+ \        (HOL.refl % TYPE(bool=>bool=>bool) % (op \<or> :: bool=>bool=>bool) %%  \
  \           (OfClass type_class % TYPE(bool=>bool=>bool))) %%  \
  \        (HOL.refl % TYPE(bool) % A %% prfb)))",
 
-   (* --> *)
+   (* \<longrightarrow> *)
 
-   "(iffD1 % A --> C % B --> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op --> % op --> % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op --> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+   "(iffD1 % A \<longrightarrow> C % B \<longrightarrow> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<longrightarrow> % op \<longrightarrow> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<longrightarrow> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
  \  (impI % B % D %% (Lam H: B. iffD1 % C % D %% prf2 %%  \
  \    (mp % A % C %% prf3 %% (iffD2 % A % B %% prf1 %% H))))",
 
-   "(iffD2 % A --> C % B --> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
- \    (cong % TYPE('T3) % TYPE('T4) % op --> % op --> % A % B %% prfT3 %% prfT4 %%  \
- \      (HOL.refl % TYPE('T5) % op --> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
+   "(iffD2 % A \<longrightarrow> C % B \<longrightarrow> D %% (cong % TYPE('T1) % TYPE('T2) % x1 % x2 % C % D %% prfT1 %% prfT2 %%  \
+ \    (cong % TYPE('T3) % TYPE('T4) % op \<longrightarrow> % op \<longrightarrow> % A % B %% prfT3 %% prfT4 %%  \
+ \      (HOL.refl % TYPE('T5) % op \<longrightarrow> %% prfT5) %% prf1) %% prf2) %% prf3) ==  \
  \  (impI % A % C %% (Lam H: A. iffD2 % C % D %% prf2 %%  \
  \    (mp % B % D %% prf3 %% (iffD1 % A % B %% prf1 %% H))))",
 
-   "(cong % TYPE(bool) % TYPE(bool) % op --> A % op --> A % B % C %% prfb %% prfb %%  \
- \    (HOL.refl % TYPE(bool=>bool) % op --> A %% prfbb)) ==  \
- \  (cong % TYPE(bool) % TYPE(bool) % op --> A % op --> A % B % C %% prfb %% prfb %%  \
+   "(cong % TYPE(bool) % TYPE(bool) % op \<longrightarrow> A % op \<longrightarrow> A % B % C %% prfb %% prfb %%  \
+ \    (HOL.refl % TYPE(bool=>bool) % op \<longrightarrow> A %% prfbb)) ==  \
+ \  (cong % TYPE(bool) % TYPE(bool) % op \<longrightarrow> A % op \<longrightarrow> A % B % C %% prfb %% prfb %%  \
  \    (cong % TYPE(bool) % TYPE(bool=>bool) %  \
- \      (op --> :: bool=>bool=>bool) % (op --> :: bool=>bool=>bool) % A % A %%  \
+ \      (op \<longrightarrow> :: bool=>bool=>bool) % (op \<longrightarrow> :: bool=>bool=>bool) % A % A %%  \
  \        prfb %% prfbb %%  \
- \        (HOL.refl % TYPE(bool=>bool=>bool) % (op --> :: bool=>bool=>bool) %%  \
+ \        (HOL.refl % TYPE(bool=>bool=>bool) % (op \<longrightarrow> :: bool=>bool=>bool) %%  \
  \           (OfClass type_class % TYPE(bool=>bool=>bool))) %%  \
  \        (HOL.refl % TYPE(bool) % A %% prfb)))",
 
-   (* ~ *)
+   (* \<not> *)
 
-   "(iffD1 % ~ P % ~ Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% prfT1 %% prfT2 %%  \
+   "(iffD1 % \<not> P % \<not> Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% prfT1 %% prfT2 %%  \
  \    (HOL.refl % TYPE('T3) % Not %% prfT3) %% prf1) %% prf2) ==  \
  \  (notI % Q %% (Lam H: Q.  \
  \    notE % P % False %% prf2 %% (iffD2 % P % Q %% prf1 %% H)))",
 
-   "(iffD2 % ~ P % ~ Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% prfT1 %% prfT2 %%  \
+   "(iffD2 % \<not> P % \<not> Q %% (cong % TYPE('T1) % TYPE('T2) % Not % Not % P % Q %% prfT1 %% prfT2 %%  \
  \    (HOL.refl % TYPE('T3) % Not %% prfT3) %% prf1) %% prf2) ==  \
  \  (notI % P %% (Lam H: P.  \
  \    notE % Q % False %% prf2 %% (iffD1 % P % Q %% prf1 %% H)))",
--- a/src/HOL/Tools/sat.ML	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Tools/sat.ML	Sun Nov 26 21:08:32 2017 +0100
@@ -71,7 +71,7 @@
 val counter = Unsynchronized.ref 0;
 
 val resolution_thm =
-  @{lemma "(P ==> False) ==> (~ P ==> False) ==> False" by (rule case_split)}
+  @{lemma "(P \<Longrightarrow> False) \<Longrightarrow> (\<not> P \<Longrightarrow> False) \<Longrightarrow> False" by (rule case_split)}
 
 (* ------------------------------------------------------------------------- *)
 (* lit_ord: an order on integers that considers their absolute values only,  *)
--- a/src/HOL/Transcendental.thy	Sun Nov 26 13:19:52 2017 +0100
+++ b/src/HOL/Transcendental.thy	Sun Nov 26 21:08:32 2017 +0100
@@ -4367,7 +4367,7 @@
 lemma sin_times_pi_eq_0: "sin (x * pi) = 0 \<longleftrightarrow> x \<in> \<int>"
   by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int)
 
-lemma cos_one_2pi: "cos x = 1 \<longleftrightarrow> (\<exists>n::nat. x = n * 2 * pi) | (\<exists>n::nat. x = - (n * 2 * pi))"
+lemma cos_one_2pi: "cos x = 1 \<longleftrightarrow> (\<exists>n::nat. x = n * 2 * pi) \<or> (\<exists>n::nat. x = - (n * 2 * pi))"
   (is "?lhs = ?rhs")
 proof
   assume ?lhs