--- a/src/HOL/Analysis/Improper_Integral.thy Tue Aug 01 17:30:02 2017 +0200
+++ b/src/HOL/Analysis/Improper_Integral.thy Tue Aug 01 17:33:04 2017 +0200
@@ -1498,7 +1498,7 @@
using bounded_integrals_over_subintervals [OF int_gab] unfolding bounded_pos real_norm_def by blast
show "(\<lambda>x. f x \<bullet> j) absolutely_integrable_on cbox a b"
using g
- proof --\<open>A lot of duplication in the two proofs\<close>
+ proof \<comment>\<open>A lot of duplication in the two proofs\<close>
assume fg [rule_format]: "\<forall>x\<in>cbox a b. f x \<bullet> j \<le> g x"
have "(\<lambda>x. (f x \<bullet> j)) = (\<lambda>x. g x - (g x - (f x \<bullet> j)))"
by simp
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Aug 01 17:30:02 2017 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Aug 01 17:33:04 2017 +0200
@@ -5300,7 +5300,7 @@
"compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"
using compact_imp_seq_compact seq_compact_imp_heine_borel by blast
-lemma compact_def: --\<open>this is the definition of compactness in HOL Light\<close>
+lemma compact_def: \<comment>\<open>this is the definition of compactness in HOL Light\<close>
"compact (S :: 'a::metric_space set) \<longleftrightarrow>
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f \<circ> r) \<longlonglongrightarrow> l))"
unfolding compact_eq_seq_compact_metric seq_compact_def by auto
--- a/src/HOL/Analysis/Winding_Numbers.thy Tue Aug 01 17:30:02 2017 +0200
+++ b/src/HOL/Analysis/Winding_Numbers.thy Tue Aug 01 17:33:04 2017 +0200
@@ -41,7 +41,7 @@
by (meson mem_interior)
define z where "z \<equiv> - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \<i>"
have "z \<in> ball 0 e"
- using `e>0`
+ using \<open>e>0\<close>
apply (simp add: z_def dist_norm)
apply (rule le_less_trans [OF norm_triangle_ineq4])
apply (simp add: norm_mult abs_sgn_eq)
@@ -49,7 +49,7 @@
then have "z \<in> {z. Im z * Re b \<le> Im b * Re z}"
using e by blast
then show False
- using `e>0` `b \<noteq> 0`
+ using \<open>e>0\<close> \<open>b \<noteq> 0\<close>
apply (simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm)
apply (auto simp: algebra_simps)
apply (meson less_asym less_trans mult_pos_pos neg_less_0_iff_less)
--- a/src/HOL/Number_Theory/Residues.thy Tue Aug 01 17:30:02 2017 +0200
+++ b/src/HOL/Number_Theory/Residues.thy Tue Aug 01 17:33:04 2017 +0200
@@ -378,9 +378,9 @@
by (metis residues_prime.wilson_theorem1 residues_prime.intro le_eq_less_or_eq)
qed
-text {*
+text \<open>
This result can be transferred to the multiplicative group of
- $\mathbb{Z}/p\mathbb{Z}$ for $p$ prime. *}
+ $\mathbb{Z}/p\mathbb{Z}$ for $p$ prime.\<close>
lemma mod_nat_int_pow_eq:
fixes n :: nat and p a :: int
@@ -409,22 +409,22 @@
have **:"nat ` {1 .. int p - 1} = {1 .. p - 1}" (is "?L = ?R")
proof
{ fix n assume n: "n \<in> ?L"
- then have "n \<in> ?R" using `p\<ge>2` by force
+ then have "n \<in> ?R" using \<open>p\<ge>2\<close> by force
} thus "?L \<subseteq> ?R" by blast
{ fix n assume n: "n \<in> ?R"
- then have "n \<in> ?L" using `p\<ge>2` Set_Interval.transfer_nat_int_set_functions(2) by fastforce
+ then have "n \<in> ?L" using \<open>p\<ge>2\<close> Set_Interval.transfer_nat_int_set_functions(2) by fastforce
} thus "?R \<subseteq> ?L" by blast
qed
have "nat ` {a^i mod (int p) | i::nat. i \<in> UNIV} = {nat a^i mod p | i . i \<in> UNIV}" (is "?L = ?R")
proof
{ fix x assume x: "x \<in> ?L"
then obtain i where i:"x = nat (a^i mod (int p))" by blast
- hence "x = nat a ^ i mod p" using mod_nat_int_pow_eq[of a "int p" i] a `p\<ge>2` by auto
+ hence "x = nat a ^ i mod p" using mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> by auto
hence "x \<in> ?R" using i by blast
} thus "?L \<subseteq> ?R" by blast
{ fix x assume x: "x \<in> ?R"
then obtain i where i:"x = nat a^i mod p" by blast
- hence "x \<in> ?L" using mod_nat_int_pow_eq[of a "int p" i] a `p\<ge>2` by auto
+ hence "x \<in> ?L" using mod_nat_int_pow_eq[of a "int p" i] a \<open>p\<ge>2\<close> by auto
} thus "?R \<subseteq> ?L" by blast
qed
hence "{1 .. p - 1} = {nat a^i mod p | i. i \<in> UNIV}"