# HG changeset patch # User wenzelm # Date 1546372674 -3600 # Node ID 1daf07b653853f6a43c141105a3ff3f119d73d1b # Parent a59f7d07bf1748638c22533801b27430788c489d retain important whitespace after 'text' that is suppressed, but swallows adjacent whitespace; diff -r a59f7d07bf17 -r 1daf07b65385 src/HOL/Analysis/Continuum_Not_Denumerable.thy --- a/src/HOL/Analysis/Continuum_Not_Denumerable.thy Tue Jan 01 13:26:37 2019 +0100 +++ b/src/HOL/Analysis/Continuum_Not_Denumerable.thy Tue Jan 01 20:57:54 2019 +0100 @@ -32,7 +32,7 @@ range of \f\ by generating a sequence of closed intervals then using the Nested Interval Property. \ - +text%important \%whitespace\ theorem real_non_denum: "\f :: nat \ real. surj f" proof assume "\f::nat \ real. surj f" diff -r a59f7d07bf17 -r 1daf07b65385 src/HOL/Analysis/Infinite_Products.thy --- a/src/HOL/Analysis/Infinite_Products.thy Tue Jan 01 13:26:37 2019 +0100 +++ b/src/HOL/Analysis/Infinite_Products.thy Tue Jan 01 20:57:54 2019 +0100 @@ -58,6 +58,7 @@ where "raw_has_prod f M p \ (\n. \i\n. f (i+M)) \ p \ p \ 0" text\The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\ +text%important \%whitespace\ definition%important has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr "has'_prod" 80) where "f has_prod p \ raw_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ raw_has_prod f (Suc i) q)" diff -r a59f7d07bf17 -r 1daf07b65385 src/HOL/Analysis/Norm_Arith.thy --- a/src/HOL/Analysis/Norm_Arith.thy Tue Jan 01 13:26:37 2019 +0100 +++ b/src/HOL/Analysis/Norm_Arith.thy Tue Jan 01 20:57:54 2019 +0100 @@ -137,7 +137,7 @@ text \Hence more metric properties.\ - +text%important \%whitespace\ proposition dist_triangle_add: fixes x y x' y' :: "'a::real_normed_vector" shows "dist (x + y) (x' + y') \ dist x x' + dist y y'" diff -r a59f7d07bf17 -r 1daf07b65385 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Tue Jan 01 13:26:37 2019 +0100 +++ b/src/HOL/Analysis/Path_Connected.thy Tue Jan 01 20:57:54 2019 +0100 @@ -3401,12 +3401,12 @@ (\x. h(1, x) = q x) \ (\t \ {0..1}. P(\x. h(t, x))))" - text\$p, q$ are functions $X \to Y$, and the property $P$ restricts all intermediate maps. We often just want to require that $P$ fixes some subset, but to include the case of a loop homotopy, it is convenient to have a general property $P$.\ text \We often want to just localize the ending function equality or whatever.\ +text%important \%whitespace\ proposition homotopic_with: fixes X :: "'a::topological_space set" and Y :: "'b::topological_space set" assumes "\h k. (\x. x \ X \ h x = k x) \ (P h \ P k)" diff -r a59f7d07bf17 -r 1daf07b65385 src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Tue Jan 01 13:26:37 2019 +0100 +++ b/src/HOL/Analysis/Tagged_Division.thy Tue Jan 01 20:57:54 2019 +0100 @@ -1293,7 +1293,7 @@ @{typ bool}.\ paragraph%important \Using additivity of lifted function to encode definedness.\ - +text%important \%whitespace\ definition%important lift_option :: "('a \ 'b \ 'c) \ 'a option \ 'b option \ 'c option" where "lift_option f a' b' = Option.bind a' (\a. Option.bind b' (\b. Some (f a b)))" @@ -1334,7 +1334,7 @@ paragraph \Division points\ - +text%important \%whitespace\ definition%important "division_points (k::('a::euclidean_space) set) d = {(j,x). j \ Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}"