# HG changeset patch # User wenzelm # Date 1501601584 -7200 # Node ID cde6ceffcbc7a0517b4eb283bf64ebb59de0f90b # Parent 210dae34b8bc84ee95ec40e4ef0a51b7eec6b694 isabelle update_cartouches -c -t; diff -r 210dae34b8bc -r cde6ceffcbc7 src/HOL/Analysis/Improper_Integral.thy --- 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 "(\x. f x \ j) absolutely_integrable_on cbox a b" using g - proof --\A lot of duplication in the two proofs\ + proof \\A lot of duplication in the two proofs\ assume fg [rule_format]: "\x\cbox a b. f x \ j \ g x" have "(\x. (f x \ j)) = (\x. g x - (g x - (f x \ j)))" by simp diff -r 210dae34b8bc -r cde6ceffcbc7 src/HOL/Analysis/Topology_Euclidean_Space.thy --- 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) \ seq_compact s" using compact_imp_seq_compact seq_compact_imp_heine_borel by blast -lemma compact_def: --\this is the definition of compactness in HOL Light\ +lemma compact_def: \\this is the definition of compactness in HOL Light\ "compact (S :: 'a::metric_space set) \ (\f. (\n. f n \ S) \ (\l\S. \r. subseq r \ (f \ r) \ l))" unfolding compact_eq_seq_compact_metric seq_compact_def by auto diff -r 210dae34b8bc -r cde6ceffcbc7 src/HOL/Analysis/Winding_Numbers.thy --- 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 \ - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \" have "z \ ball 0 e" - using `e>0` + using \e>0\ 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 \ {z. Im z * Re b \ Im b * Re z}" using e by blast then show False - using `e>0` `b \ 0` + using \e>0\ \b \ 0\ 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) diff -r 210dae34b8bc -r cde6ceffcbc7 src/HOL/Number_Theory/Residues.thy --- 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 \ 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.\ 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 \ ?L" - then have "n \ ?R" using `p\2` by force + then have "n \ ?R" using \p\2\ by force } thus "?L \ ?R" by blast { fix n assume n: "n \ ?R" - then have "n \ ?L" using `p\2` Set_Interval.transfer_nat_int_set_functions(2) by fastforce + then have "n \ ?L" using \p\2\ Set_Interval.transfer_nat_int_set_functions(2) by fastforce } thus "?R \ ?L" by blast qed have "nat ` {a^i mod (int p) | i::nat. i \ UNIV} = {nat a^i mod p | i . i \ UNIV}" (is "?L = ?R") proof { fix x assume x: "x \ ?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\2` by auto + hence "x = nat a ^ i mod p" using mod_nat_int_pow_eq[of a "int p" i] a \p\2\ by auto hence "x \ ?R" using i by blast } thus "?L \ ?R" by blast { fix x assume x: "x \ ?R" then obtain i where i:"x = nat a^i mod p" by blast - hence "x \ ?L" using mod_nat_int_pow_eq[of a "int p" i] a `p\2` by auto + hence "x \ ?L" using mod_nat_int_pow_eq[of a "int p" i] a \p\2\ by auto } thus "?R \ ?L" by blast qed hence "{1 .. p - 1} = {nat a^i mod p | i. i \ UNIV}"