--- a/src/HOL/ATP.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/ATP.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* Automatic Theorem Provers (ATPs) *}
+section \<open>Automatic Theorem Provers (ATPs)\<close>
theory ATP
imports Meson
begin
-subsection {* ATP problems and proofs *}
+subsection \<open>ATP problems and proofs\<close>
ML_file "Tools/ATP/atp_util.ML"
ML_file "Tools/ATP/atp_problem.ML"
@@ -18,7 +18,7 @@
ML_file "Tools/ATP/atp_satallax.ML"
-subsection {* Higher-order reasoning helpers *}
+subsection \<open>Higher-order reasoning helpers\<close>
definition fFalse :: bool where
"fFalse \<longleftrightarrow> False"
@@ -131,7 +131,7 @@
unfolding fFalse_def fTrue_def fequal_def by auto
-subsection {* Waldmeister helpers *}
+subsection \<open>Waldmeister helpers\<close>
(* Has all needed simplification lemmas for logic. *)
lemma boolean_equality: "(P \<longleftrightarrow> P) = True"
@@ -144,7 +144,7 @@
simp_thms(1-5,7-8,11-25,27-33) disj_comms disj_assoc conj_comms conj_assoc
-subsection {* Basic connection between ATPs and HOL *}
+subsection \<open>Basic connection between ATPs and HOL\<close>
ML_file "Tools/lambda_lifting.ML"
ML_file "Tools/monomorph.ML"
--- a/src/HOL/Archimedean_Field.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Archimedean_Field.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,15 +2,15 @@
Author: Brian Huffman
*)
-section {* Archimedean Fields, Floor and Ceiling Functions *}
+section \<open>Archimedean Fields, Floor and Ceiling Functions\<close>
theory Archimedean_Field
imports Main
begin
-subsection {* Class of Archimedean fields *}
+subsection \<open>Class of Archimedean fields\<close>
-text {* Archimedean fields have no infinite elements. *}
+text \<open>Archimedean fields have no infinite elements.\<close>
class archimedean_field = linordered_field +
assumes ex_le_of_int: "\<exists>z. x \<le> of_int z"
@@ -48,53 +48,53 @@
then show ?thesis ..
qed
-text {* Archimedean fields have no infinitesimal elements. *}
+text \<open>Archimedean fields have no infinitesimal elements.\<close>
lemma ex_inverse_of_nat_Suc_less:
fixes x :: "'a::archimedean_field"
assumes "0 < x" shows "\<exists>n. inverse (of_nat (Suc n)) < x"
proof -
- from `0 < x` have "0 < inverse x"
+ from \<open>0 < x\<close> have "0 < inverse x"
by (rule positive_imp_inverse_positive)
obtain n where "inverse x < of_nat n"
using ex_less_of_nat ..
then obtain m where "inverse x < of_nat (Suc m)"
- using `0 < inverse x` by (cases n) (simp_all del: of_nat_Suc)
+ using \<open>0 < inverse x\<close> by (cases n) (simp_all del: of_nat_Suc)
then have "inverse (of_nat (Suc m)) < inverse (inverse x)"
- using `0 < inverse x` by (rule less_imp_inverse_less)
+ using \<open>0 < inverse x\<close> by (rule less_imp_inverse_less)
then have "inverse (of_nat (Suc m)) < x"
- using `0 < x` by (simp add: nonzero_inverse_inverse_eq)
+ using \<open>0 < x\<close> by (simp add: nonzero_inverse_inverse_eq)
then show ?thesis ..
qed
lemma ex_inverse_of_nat_less:
fixes x :: "'a::archimedean_field"
assumes "0 < x" shows "\<exists>n>0. inverse (of_nat n) < x"
- using ex_inverse_of_nat_Suc_less [OF `0 < x`] by auto
+ using ex_inverse_of_nat_Suc_less [OF \<open>0 < x\<close>] by auto
lemma ex_less_of_nat_mult:
fixes x :: "'a::archimedean_field"
assumes "0 < x" shows "\<exists>n. y < of_nat n * x"
proof -
obtain n where "y / x < of_nat n" using ex_less_of_nat ..
- with `0 < x` have "y < of_nat n * x" by (simp add: pos_divide_less_eq)
+ with \<open>0 < x\<close> have "y < of_nat n * x" by (simp add: pos_divide_less_eq)
then show ?thesis ..
qed
-subsection {* Existence and uniqueness of floor function *}
+subsection \<open>Existence and uniqueness of floor function\<close>
lemma exists_least_lemma:
assumes "\<not> P 0" and "\<exists>n. P n"
shows "\<exists>n. \<not> P n \<and> P (Suc n)"
proof -
- from `\<exists>n. P n` have "P (Least P)" by (rule LeastI_ex)
- with `\<not> P 0` obtain n where "Least P = Suc n"
+ from \<open>\<exists>n. P n\<close> have "P (Least P)" by (rule LeastI_ex)
+ with \<open>\<not> P 0\<close> obtain n where "Least P = Suc n"
by (cases "Least P") auto
then have "n < Least P" by simp
then have "\<not> P n" by (rule not_less_Least)
then have "\<not> P n \<and> P (Suc n)"
- using `P (Least P)` `Least P = Suc n` by simp
+ using \<open>P (Least P)\<close> \<open>Least P = Suc n\<close> by simp
then show ?thesis ..
qed
@@ -135,7 +135,7 @@
qed
-subsection {* Floor function *}
+subsection \<open>Floor function\<close>
class floor_ceiling = archimedean_field +
fixes floor :: "'a \<Rightarrow> int"
@@ -185,7 +185,7 @@
lemma floor_mono: assumes "x \<le> y" shows "floor x \<le> floor y"
proof -
have "of_int (floor x) \<le> x" by (rule of_int_floor_le)
- also note `x \<le> y`
+ also note \<open>x \<le> y\<close>
finally show ?thesis by (simp add: le_floor_iff)
qed
@@ -201,7 +201,7 @@
lemma le_floor_add: "floor x + floor y \<le> floor (x + y)"
by (simp only: le_floor_iff of_int_add add_mono of_int_floor_le)
-text {* Floor with numerals *}
+text \<open>Floor with numerals\<close>
lemma floor_zero [simp]: "floor 0 = 0"
using floor_of_int [of 0] by simp
@@ -271,7 +271,7 @@
"floor x < - numeral v \<longleftrightarrow> x < - numeral v"
by (simp add: floor_less_iff)
-text {* Addition and subtraction of integers *}
+text \<open>Addition and subtraction of integers\<close>
lemma floor_add_of_int [simp]: "floor (x + of_int z) = floor x + z"
using floor_correct [of x] by (simp add: floor_unique)
@@ -324,7 +324,7 @@
case False
obtain r where "r = - l" by blast
then have l: "l = - r" by simp
- moreover with `l \<noteq> 0` False have "r > 0" by simp
+ moreover with \<open>l \<noteq> 0\<close> False have "r > 0" by simp
ultimately show ?thesis using pos_mod_bound [of r]
by (auto simp add: zmod_zminus2_eq_if less_le field_simps intro: floor_unique)
qed
@@ -374,7 +374,7 @@
qed
-subsection {* Ceiling function *}
+subsection \<open>Ceiling function\<close>
definition
ceiling :: "'a::floor_ceiling \<Rightarrow> int" where
@@ -422,7 +422,7 @@
lemma ceiling_add_le: "ceiling (x + y) \<le> ceiling x + ceiling y"
by (simp only: ceiling_le_iff of_int_add add_mono le_of_int_ceiling)
-text {* Ceiling with numerals *}
+text \<open>Ceiling with numerals\<close>
lemma ceiling_zero [simp]: "ceiling 0 = 0"
using ceiling_of_int [of 0] by simp
@@ -492,7 +492,7 @@
"- numeral v < ceiling x \<longleftrightarrow> - numeral v < x"
by (simp add: less_ceiling_iff)
-text {* Addition and subtraction of integers *}
+text \<open>Addition and subtraction of integers\<close>
lemma ceiling_add_of_int [simp]: "ceiling (x + of_int z) = ceiling x + z"
using ceiling_correct [of x] by (simp add: ceiling_unique)
@@ -529,7 +529,7 @@
unfolding of_int_less_iff by simp
qed
-subsection {* Negation *}
+subsection \<open>Negation\<close>
lemma floor_minus: "floor (- x) = - ceiling x"
unfolding ceiling_def by simp
@@ -537,7 +537,7 @@
lemma ceiling_minus: "ceiling (- x) = - floor x"
unfolding ceiling_def by simp
-subsection {* Frac Function *}
+subsection \<open>Frac Function\<close>
definition frac :: "'a \<Rightarrow> 'a::floor_ceiling" where
--- a/src/HOL/BNF_Cardinal_Arithmetic.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Cardinal_Arithmetic.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,7 +5,7 @@
Cardinal arithmetic as needed by bounded natural functors.
*)
-section {* Cardinal Arithmetic as Needed by Bounded Natural Functors *}
+section \<open>Cardinal Arithmetic as Needed by Bounded Natural Functors\<close>
theory BNF_Cardinal_Arithmetic
imports BNF_Cardinal_Order_Relation
@@ -39,7 +39,7 @@
using card_order_on_Card_order[of UNIV r] by simp
-subsection {* Zero *}
+subsection \<open>Zero\<close>
definition czero where
"czero = card_of {}"
@@ -80,7 +80,7 @@
apply (simp only: card_of_empty3)
done
-subsection {* (In)finite cardinals *}
+subsection \<open>(In)finite cardinals\<close>
definition cinfinite where
"cinfinite r = (\<not> finite (Field r))"
@@ -127,7 +127,7 @@
unfolding cinfinite_def by (auto dest: card_of_ordLeq_infinite[OF card_of_mono2])
-subsection {* Binary sum *}
+subsection \<open>Binary sum\<close>
definition csum (infixr "+c" 65) where
"r1 +c r2 \<equiv> |Field r1 <+> Field r2|"
@@ -223,7 +223,7 @@
using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast
-subsection {* One *}
+subsection \<open>One\<close>
definition cone where
"cone = card_of {()}"
@@ -241,7 +241,7 @@
unfolding cone_def by (rule Card_order_singl_ordLeq) (auto intro: czeroI)
-subsection {* Two *}
+subsection \<open>Two\<close>
definition ctwo where
"ctwo = |UNIV :: bool set|"
@@ -257,7 +257,7 @@
by (simp add: ctwo_not_czero Card_order_ctwo)
-subsection {* Family sum *}
+subsection \<open>Family sum\<close>
definition Csum where
"Csum r rs \<equiv> |SIGMA i : Field r. Field (rs i)|"
@@ -278,7 +278,7 @@
This should make cardinal reasoning more direct and natural. *)
-subsection {* Product *}
+subsection \<open>Product\<close>
definition cprod (infixr "*c" 80) where
"r1 *c r2 = |Field r1 <*> Field r2|"
@@ -360,7 +360,7 @@
by (rule csum_absorb1') auto
-subsection {* Exponentiation *}
+subsection \<open>Exponentiation\<close>
definition cexp (infixr "^c" 90) where
"r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"
@@ -386,7 +386,7 @@
with n have "Field r2 = {}" .
hence "cone \<le>o r1 ^c r2" unfolding cone_def cexp_def Func_def
by (auto intro: card_of_ordLeqI[where f="\<lambda>_ _. undefined"])
- thus ?thesis using `p1 ^c p2 \<le>o cone` ordLeq_transitive by auto
+ thus ?thesis using \<open>p1 ^c p2 \<le>o cone\<close> ordLeq_transitive by auto
next
case False with True have "|Field (p1 ^c p2)| =o czero"
unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,13 +5,13 @@
Cardinal-order relations as needed by bounded natural functors.
*)
-section {* Cardinal-Order Relations as Needed by Bounded Natural Functors *}
+section \<open>Cardinal-Order Relations as Needed by Bounded Natural Functors\<close>
theory BNF_Cardinal_Order_Relation
imports Zorn BNF_Wellorder_Constructions
begin
-text{* In this section, we define cardinal-order relations to be minim well-orders
+text\<open>In this section, we define cardinal-order relations to be minim well-orders
on their field. Then we define the cardinal of a set to be {\em some} cardinal-order
relation on that set, which will be unique up to order isomorphism. Then we study
the connection between cardinals and:
@@ -31,14 +31,14 @@
most of the standard set-theoretic constructions (except for the powerset)
{\em do not increase cardinality}. In particular, e.g., the set of words/lists over
any infinite set has the same cardinality (hence, is in bijection) with that set.
-*}
+\<close>
-subsection {* Cardinal orders *}
+subsection \<open>Cardinal orders\<close>
-text{* A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the
+text\<open>A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the
order-embedding relation, @{text "\<le>o"} (which is the same as being {\em minimal} w.r.t. the
-strict order-embedding relation, @{text "<o"}), among all the well-orders on its field. *}
+strict order-embedding relation, @{text "<o"}), among all the well-orders on its field.\<close>
definition card_order_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
where
@@ -56,7 +56,7 @@
"card_order_on A r \<Longrightarrow> A = Field r \<and> Card_order r"
unfolding card_order_on_def using well_order_on_Field by blast
-text{* The existence of a cardinal relation on any given set (which will mean
+text\<open>The existence of a cardinal relation on any given set (which will mean
that any set has a cardinal) follows from two facts:
\begin{itemize}
\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}),
@@ -64,7 +64,7 @@
\item The well-founded-ness of @{text "<o"}, ensuring that then there exists a minimal
such well-order, i.e., a cardinal order.
\end{itemize}
-*}
+\<close>
theorem card_order_on: "\<exists>r. card_order_on A r"
proof-
@@ -111,11 +111,11 @@
using assms Card_order_ordIso ordIso_symmetric by blast
-subsection {* Cardinal of a set *}
+subsection \<open>Cardinal of a set\<close>
-text{* We define the cardinal of set to be {\em some} cardinal order on that set.
+text\<open>We define the cardinal of set to be {\em some} cardinal order on that set.
We shall prove that this notion is unique up to order isomorphism, meaning
-that order isomorphism shall be the true identity of cardinals. *}
+that order isomorphism shall be the true identity of cardinals.\<close>
definition card_of :: "'a set \<Rightarrow> 'a rel" ("|_|" )
where "card_of A = (SOME r. card_order_on A r)"
@@ -343,13 +343,13 @@
using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto
-subsection {* Cardinals versus set operations on arbitrary sets *}
+subsection \<open>Cardinals versus set operations on arbitrary sets\<close>
-text{* Here we embark in a long journey of simple results showing
+text\<open>Here we embark in a long journey of simple results showing
that the standard set-theoretic operations are well-behaved w.r.t. the notion of
cardinal -- essentially, this means that they preserve the ``cardinal identity"
@{text "=o"} and are monotonic w.r.t. @{text "\<le>o"}.
-*}
+\<close>
lemma card_of_empty: "|{}| \<le>o |A|"
using card_of_ordLeq inj_on_id by blast
@@ -819,21 +819,21 @@
using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast
-subsection {* Cardinals versus set operations involving infinite sets *}
+subsection \<open>Cardinals versus set operations involving infinite sets\<close>
-text{* Here we show that, for infinite sets, most set-theoretic constructions
+text\<open>Here we show that, for infinite sets, most set-theoretic constructions
do not increase the cardinality. The cornerstone for this is
theorem @{text "Card_order_Times_same_infinite"}, which states that self-product
does not increase cardinality -- the proof of this fact adapts a standard
set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11
-at page 47 in @{cite "card-book"}. Then everything else follows fairly easily. *}
+at page 47 in @{cite "card-book"}. Then everything else follows fairly easily.\<close>
lemma infinite_iff_card_of_nat:
"\<not> finite A \<longleftrightarrow> ( |UNIV::nat set| \<le>o |A| )"
unfolding infinite_iff_countable_subset card_of_ordLeq ..
-text{* The next two results correspond to the ZF fact that all infinite cardinals are
-limit ordinals: *}
+text\<open>The next two results correspond to the ZF fact that all infinite cardinals are
+limit ordinals:\<close>
lemma Card_order_infinite_not_under:
assumes CARD: "Card_order r" and INF: "\<not>finite (Field r)"
@@ -1140,13 +1140,13 @@
qed
-subsection {* The cardinal $\omega$ and the finite cardinals *}
+subsection \<open>The cardinal $\omega$ and the finite cardinals\<close>
-text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict
+text\<open>The cardinal $\omega$, of natural numbers, shall be the standard non-strict
order relation on
@{text "nat"}, that we abbreviate by @{text "natLeq"}. The finite cardinals
shall be the restrictions of these relations to the numbers smaller than
-fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}. *}
+fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}.\<close>
definition "(natLeq::(nat * nat) set) \<equiv> {(x,y). x \<le> y}"
definition "(natLess::(nat * nat) set) \<equiv> {(x,y). x < y}"
@@ -1160,12 +1160,12 @@
proof
assume "finite (A \<times> B)"
from assms(1) have "A \<noteq> {}" by auto
- with `finite (A \<times> B)` have "finite B" using finite_cartesian_productD2 by auto
+ with \<open>finite (A \<times> B)\<close> have "finite B" using finite_cartesian_productD2 by auto
with assms(2) show False by simp
qed
-subsubsection {* First as well-orders *}
+subsubsection \<open>First as well-orders\<close>
lemma Field_natLeq: "Field natLeq = (UNIV::nat set)"
by(unfold Field_def natLeq_def, auto)
@@ -1225,7 +1225,7 @@
unfolding wo_rel_def using natLeq_on_Well_order .
-subsubsection {* Then as cardinals *}
+subsubsection \<open>Then as cardinals\<close>
lemma natLeq_Card_order: "Card_order natLeq"
proof(auto simp add: natLeq_Well_order
@@ -1258,11 +1258,11 @@
card_of_Well_order natLeq_Well_order by blast
-subsection {* The successor of a cardinal *}
+subsection \<open>The successor of a cardinal\<close>
-text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"}
+text\<open>First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"}
being a successor cardinal of @{text "r"}. Although the definition does
-not require @{text "r"} to be a cardinal, only this case will be meaningful. *}
+not require @{text "r"} to be a cardinal, only this case will be meaningful.\<close>
definition isCardSuc :: "'a rel \<Rightarrow> 'a set rel \<Rightarrow> bool"
where
@@ -1270,9 +1270,9 @@
Card_order r' \<and> r <o r' \<and>
(\<forall>(r''::'a set rel). Card_order r'' \<and> r <o r'' \<longrightarrow> r' \<le>o r'')"
-text{* Now we introduce the cardinal-successor operator @{text "cardSuc"},
+text\<open>Now we introduce the cardinal-successor operator @{text "cardSuc"},
by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}.
-Again, the picked item shall be proved unique up to order-isomorphism. *}
+Again, the picked item shall be proved unique up to order-isomorphism.\<close>
definition cardSuc :: "'a rel \<Rightarrow> 'a set rel"
where
@@ -1312,14 +1312,14 @@
"Card_order r \<Longrightarrow> r \<le>o cardSuc r"
using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast
-text{* The minimality property of @{text "cardSuc"} originally present in its definition
-is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}: *}
+text\<open>The minimality property of @{text "cardSuc"} originally present in its definition
+is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}:\<close>
lemma cardSuc_least_aux:
"\<lbrakk>Card_order (r::'a rel); Card_order (r'::'a set rel); r <o r'\<rbrakk> \<Longrightarrow> cardSuc r \<le>o r'"
using cardSuc_isCardSuc unfolding isCardSuc_def by blast
-text{* But from this we can infer general minimality: *}
+text\<open>But from this we can infer general minimality:\<close>
lemma cardSuc_least:
assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r <o r'"
@@ -1497,7 +1497,7 @@
ordLeq_transitive by fast
-subsection {* Regular cardinals *}
+subsection \<open>Regular cardinals\<close>
definition cofinal where
"cofinal A r \<equiv>
@@ -1616,7 +1616,7 @@
qed
-subsection {* Others *}
+subsection \<open>Others\<close>
lemma card_of_Func_Times:
"|Func (A <*> B) C| =o |Func A (Func B C)|"
--- a/src/HOL/BNF_Composition.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Composition.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,7 +6,7 @@
Composition of bounded natural functors.
*)
-section {* Composition of Bounded Natural Functors *}
+section \<open>Composition of Bounded Natural Functors\<close>
theory BNF_Composition
imports BNF_Def
@@ -129,7 +129,7 @@
fix b assume "F b"
show "\<exists>b'. F (Rep b')"
proof (rule exI)
- from `F b` show "F (Rep (Abs b))" using type_definition.Abs_inverse[OF type_copy] by auto
+ from \<open>F b\<close> show "F (Rep (Abs b))" using type_definition.Abs_inverse[OF type_copy] by auto
qed
qed blast
--- a/src/HOL/BNF_Def.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Def.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,7 +6,7 @@
Definition of bounded natural functors.
*)
-section {* Definition of Bounded Natural Functors *}
+section \<open>Definition of Bounded Natural Functors\<close>
theory BNF_Def
imports BNF_Cardinal_Arithmetic Fun_Def_Base
--- a/src/HOL/BNF_Fixpoint_Base.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Fixpoint_Base.thy Sat Jul 18 22:58:50 2015 +0200
@@ -8,7 +8,7 @@
Shared fixpoint operations on bounded natural functors.
*)
-section {* Shared Fixpoint Operations on Bounded Natural Functors *}
+section \<open>Shared Fixpoint Operations on Bounded Natural Functors\<close>
theory BNF_Fixpoint_Base
imports BNF_Composition Basic_BNFs
--- a/src/HOL/BNF_Greatest_Fixpoint.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Greatest_Fixpoint.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,7 +7,7 @@
Greatest fixpoint (codatatype) operation on bounded natural functors.
*)
-section {* Greatest Fixpoint (Codatatype) Operation on Bounded Natural Functors *}
+section \<open>Greatest Fixpoint (Codatatype) Operation on Bounded Natural Functors\<close>
theory BNF_Greatest_Fixpoint
imports BNF_Fixpoint_Base String
@@ -17,7 +17,7 @@
"primcorec" :: thy_decl
begin
-setup {* Sign.const_alias @{binding proj} @{const_name Equiv_Relations.proj} *}
+setup \<open>Sign.const_alias @{binding proj} @{const_name Equiv_Relations.proj}\<close>
lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by simp
@@ -240,7 +240,7 @@
unfolding rel_fun_def image2p_def by auto
-subsection {* Equivalence relations, quotients, and Hilbert's choice *}
+subsection \<open>Equivalence relations, quotients, and Hilbert's choice\<close>
lemma equiv_Eps_in:
"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> Eps (%x. x \<in> X) \<in> X"
--- a/src/HOL/BNF_Least_Fixpoint.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Least_Fixpoint.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,7 +7,7 @@
Least fixpoint (datatype) operation on bounded natural functors.
*)
-section {* Least Fixpoint (Datatype) Operation on Bounded Natural Functors *}
+section \<open>Least Fixpoint (Datatype) Operation on Bounded Natural Functors\<close>
theory BNF_Least_Fixpoint
imports BNF_Fixpoint_Base
--- a/src/HOL/BNF_Wellorder_Constructions.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Wellorder_Constructions.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,23 +5,23 @@
Constructions on wellorders as needed by bounded natural functors.
*)
-section {* Constructions on Wellorders as Needed by Bounded Natural Functors *}
+section \<open>Constructions on Wellorders as Needed by Bounded Natural Functors\<close>
theory BNF_Wellorder_Constructions
imports BNF_Wellorder_Embedding
begin
-text {* In this section, we study basic constructions on well-orders, such as restriction to
+text \<open>In this section, we study basic constructions on well-orders, such as restriction to
a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders,
and bounded square. We also define between well-orders
the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"}),
@{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"}), and
@{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}). We study the
connections between these relations, order filters, and the aforementioned constructions.
-A main result of this section is that @{text "<o"} is well-founded. *}
+A main result of this section is that @{text "<o"} is well-founded.\<close>
-subsection {* Restriction to a set *}
+subsection \<open>Restriction to a set\<close>
abbreviation Restr :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a rel"
where "Restr r A \<equiv> r Int (A \<times> A)"
@@ -90,7 +90,7 @@
order_on_defs[of "Field r" r] by auto
-subsection {* Order filters versus restrictions and embeddings *}
+subsection \<open>Order filters versus restrictions and embeddings\<close>
lemma Field_Restr_ofilter:
"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
@@ -261,7 +261,7 @@
qed
-subsection {* The strict inclusion on proper ofilters is well-founded *}
+subsection \<open>The strict inclusion on proper ofilters is well-founded\<close>
definition ofilterIncl :: "'a rel \<Rightarrow> 'a set rel"
where
@@ -302,9 +302,9 @@
qed
-subsection {* Ordering the well-orders by existence of embeddings *}
+subsection \<open>Ordering the well-orders by existence of embeddings\<close>
-text {* We define three relations between well-orders:
+text \<open>We define three relations between well-orders:
\begin{itemize}
\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"});
\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"});
@@ -314,7 +314,7 @@
The prefix "ord" and the index "o" in these names stand for "ordinal-like".
These relations shall be proved to be inter-connected in a similar fashion as the trio
@{text "\<le>"}, @{text "<"}, @{text "="} associated to a total order on a set.
-*}
+\<close>
definition ordLeq :: "('a rel * 'a' rel) set"
where
@@ -347,10 +347,10 @@
shows "Well_order r \<and> Well_order r'"
using assms unfolding ordLeq_def by simp
-text{* Notice that the relations @{text "\<le>o"}, @{text "<o"}, @{text "=o"} connect well-orders
+text\<open>Notice that the relations @{text "\<le>o"}, @{text "<o"}, @{text "=o"} connect well-orders
on potentially {\em distinct} types. However, some of the lemmas below, including the next one,
restrict implicitly the type of these relations to @{text "(('a rel) * ('a rel)) set"} , i.e.,
-to @{text "'a rel rel"}. *}
+to @{text "'a rel rel"}.\<close>
lemma ordLeq_reflexive:
"Well_order r \<Longrightarrow> r \<le>o r"
@@ -822,13 +822,13 @@
ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast
qed
-subsection{* @{text "<o"} is well-founded *}
+subsection\<open>@{text "<o"} is well-founded\<close>
-text {* Of course, it only makes sense to state that the @{text "<o"} is well-founded
+text \<open>Of course, it only makes sense to state that the @{text "<o"} is well-founded
on the restricted type @{text "'a rel rel"}. We prove this by first showing that, for any set
of well-orders all embedded in a fixed well-order, the function mapping each well-order
in the set to an order filter of the fixed well-order is compatible w.r.t. to @{text "<o"} versus
-{\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded. *}
+{\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded.\<close>
definition ord_to_filter :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a set"
where "ord_to_filter r0 r \<equiv> (SOME f. embed r r0 f) ` (Field r)"
@@ -883,11 +883,11 @@
qed
-subsection {* Copy via direct images *}
+subsection \<open>Copy via direct images\<close>
-text{* The direct image operator is the dual of the inverse image operator @{text "inv_image"}
+text\<open>The direct image operator is the dual of the inverse image operator @{text "inv_image"}
from @{text "Relation.thy"}. It is useful for transporting a well-order between
-different types. *}
+different types.\<close>
definition dir_image :: "'a rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> 'a' rel"
where
@@ -1043,9 +1043,9 @@
qed
-subsection {* Bounded square *}
+subsection \<open>Bounded square\<close>
-text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
+text\<open>This construction essentially defines, for an order relation @{text "r"}, a lexicographic
order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
following criteria (in this order):
\begin{itemize}
@@ -1058,7 +1058,7 @@
at proving that the square of an infinite set has the same cardinal
as that set. The essential property required there (and which is ensured by this
construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
-in a product of proper filters on the original relation (assumed to be a well-order). *}
+in a product of proper filters on the original relation (assumed to be a well-order).\<close>
definition bsqr :: "'a rel => ('a * 'a)rel"
where
@@ -1639,9 +1639,9 @@
show ?thesis
proof (cases "h b2 = undefined")
case True
- hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
+ hence b1: "h b2 \<in> f1 ` A1" using h \<open>b2 \<in> B2\<close> unfolding B1 Func_def by auto
show ?thesis using A2 f_inv_into_f[OF b1]
- unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
+ unfolding True g_def Func_map_def j1_def j2[OF \<open>b2 \<in> B2\<close>] by auto
qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
auto intro: f_inv_into_f)
qed(insert h, unfold Func_def Func_map_def, auto)
--- a/src/HOL/BNF_Wellorder_Embedding.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Wellorder_Embedding.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,22 +5,22 @@
Well-order embeddings as needed by bounded natural functors.
*)
-section {* Well-Order Embeddings as Needed by Bounded Natural Functors *}
+section \<open>Well-Order Embeddings as Needed by Bounded Natural Functors\<close>
theory BNF_Wellorder_Embedding
imports Hilbert_Choice BNF_Wellorder_Relation
begin
-text{* In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and
+text\<open>In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and
prove their basic properties. The notion of embedding is considered from the point
of view of the theory of ordinals, and therefore requires the source to be injected
as an {\em initial segment} (i.e., {\em order filter}) of the target. A main result
of this section is the existence of embeddings (in one direction or another) between
any two well-orders, having as a consequence the fact that, given any two sets on
-any two types, one is smaller than (i.e., can be injected into) the other. *}
+any two types, one is smaller than (i.e., can be injected into) the other.\<close>
-subsection {* Auxiliaries *}
+subsection \<open>Auxiliaries\<close>
lemma UNION_inj_on_ofilter:
assumes WELL: "Well_order r" and
@@ -55,10 +55,10 @@
qed
-subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
-functions *}
+subsection \<open>(Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
+functions\<close>
-text{* Standardly, a function is an embedding of a well-order in another if it injectively and
+text\<open>Standardly, a function is an embedding of a well-order in another if it injectively and
order-compatibly maps the former into an order filter of the latter.
Here we opt for a more succinct definition (operator @{text "embed"}),
asking that, for any element in the source, the function should be a bijection
@@ -66,7 +66,7 @@
and the set of strict lower bounds of its image. (Later we prove equivalence with
the standard definition -- lemma @{text "embed_iff_compat_inj_on_ofilter"}.)
A {\em strict embedding} (operator @{text "embedS"}) is a non-bijective embedding
-and an isomorphism (operator @{text "iso"}) is a bijective embedding. *}
+and an isomorphism (operator @{text "iso"}) is a bijective embedding.\<close>
definition embed :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
where
@@ -74,7 +74,7 @@
lemmas embed_defs = embed_def embed_def[abs_def]
-text {* Strict embeddings: *}
+text \<open>Strict embeddings:\<close>
definition embedS :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
where
@@ -151,7 +151,7 @@
using assms unfolding iso_def
by (auto simp add: comp_embed bij_betw_trans)
-text{* That @{text "embedS"} is also preserved by function composition shall be proved only later. *}
+text\<open>That @{text "embedS"} is also preserved by function composition shall be proved only later.\<close>
lemma embed_Field:
"\<lbrakk>Well_order r; embed r r' f\<rbrakk> \<Longrightarrow> f`(Field r) \<le> Field r'"
@@ -451,9 +451,9 @@
qed
-subsection {* Given any two well-orders, one can be embedded in the other *}
+subsection \<open>Given any two well-orders, one can be embedded in the other\<close>
-text{* Here is an overview of the proof of of this fact, stated in theorem
+text\<open>Here is an overview of the proof of of this fact, stated in theorem
@{text "wellorders_totally_ordered"}:
Fix the well-orders @{text "r::'a rel"} and @{text "r'::'a' rel"}.
@@ -470,7 +470,7 @@
Otherwise, it means that @{text "Field r'"} is the smaller one, and the inverse of
(the "good" segment of) @{text "f"} is the desired embedding from @{text "r'"} to @{text "r"}
(lemma @{text "wellorders_totally_ordered_aux2"}).
-*}
+\<close>
lemma wellorders_totally_ordered_aux:
fixes r ::"'a rel" and r'::"'a' rel" and
@@ -768,12 +768,12 @@
qed
-subsection {* Uniqueness of embeddings *}
+subsection \<open>Uniqueness of embeddings\<close>
-text{* Here we show a fact complementary to the one from the previous subsection -- namely,
+text\<open>Here we show a fact complementary to the one from the previous subsection -- namely,
that between any two well-orders there is {\em at most} one embedding, and is the one
definable by the expected well-order recursive equation. As a consequence, any two
-embeddings of opposite directions are mutually inverse. *}
+embeddings of opposite directions are mutually inverse.\<close>
lemma embed_determined:
assumes WELL: "Well_order r" and WELL': "Well_order r'" and
@@ -864,7 +864,7 @@
unfolding iso_def using assms by (auto simp add: embed_bothWays_bij_betw)
-subsection {* More properties of embeddings, strict embeddings and isomorphisms *}
+subsection \<open>More properties of embeddings, strict embeddings and isomorphisms\<close>
lemma embed_bothWays_Field_bij_betw:
assumes WELL: "Well_order r" and WELL': "Well_order r'" and
--- a/src/HOL/BNF_Wellorder_Relation.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/BNF_Wellorder_Relation.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,24 +5,24 @@
Well-order relations as needed by bounded natural functors.
*)
-section {* Well-Order Relations as Needed by Bounded Natural Functors *}
+section \<open>Well-Order Relations as Needed by Bounded Natural Functors\<close>
theory BNF_Wellorder_Relation
imports Order_Relation
begin
-text{* In this section, we develop basic concepts and results pertaining
+text\<open>In this section, we develop basic concepts and results pertaining
to well-order relations. Note that we consider well-order relations
as {\em non-strict relations},
-i.e., as containing the diagonals of their fields. *}
+i.e., as containing the diagonals of their fields.\<close>
locale wo_rel =
fixes r :: "'a rel"
assumes WELL: "Well_order r"
begin
-text{* The following context encompasses all this section. In other words,
-for the whole section, we consider a fixed well-order relation @{term "r"}. *}
+text\<open>The following context encompasses all this section. In other words,
+for the whole section, we consider a fixed well-order relation @{term "r"}.\<close>
(* context wo_rel *)
@@ -38,7 +38,7 @@
lemmas ofilter_def = Order_Relation.ofilter_def[of r]
-subsection {* Auxiliaries *}
+subsection \<open>Auxiliaries\<close>
lemma REFL: "Refl r"
using WELL order_on_defs[of _ r] by auto
@@ -72,13 +72,13 @@
using TOTALS by auto
-subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *}
+subsection \<open>Well-founded induction and recursion adapted to non-strict well-order relations\<close>
-text{* Here we provide induction and recursion principles specific to {\em non-strict}
+text\<open>Here we provide induction and recursion principles specific to {\em non-strict}
well-order relations.
Although minor variations of those for well-founded relations, they will be useful
for doing away with the tediousness of
-having to take out the diagonal each time in order to switch to a well-founded relation. *}
+having to take out the diagonal each time in order to switch to a well-founded relation.\<close>
lemma well_order_induct:
assumes IND: "\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> P y \<Longrightarrow> P x"
@@ -113,9 +113,9 @@
qed
-subsection {* The notions of maximum, minimum, supremum, successor and order filter *}
+subsection \<open>The notions of maximum, minimum, supremum, successor and order filter\<close>
-text{*
+text\<open>
We define the successor {\em of a set}, and not of an element (the latter is of course
a particular case). Also, we define the maximum {\em of two elements}, @{text "max2"},
and the minimum {\em of a set}, @{text "minim"} -- we chose these variants since we
@@ -124,7 +124,7 @@
defined in terms of minimum as expected.
The minimum is only meaningful for non-empty sets, and the successor is only
meaningful for sets for which strict upper bounds exist.
-Order filters for well-orders are also known as ``initial segments". *}
+Order filters for well-orders are also known as ``initial segments".\<close>
definition max2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
where "max2 a b \<equiv> if (a,b) \<in> r then b else a"
@@ -142,7 +142,7 @@
where "suc A \<equiv> minim (AboveS A)"
-subsubsection {* Properties of max2 *}
+subsubsection \<open>Properties of max2\<close>
lemma max2_greater_among:
assumes "a \<in> Field r" and "b \<in> Field r"
@@ -191,7 +191,7 @@
unfolding max2_def by auto
-subsubsection {* Existence and uniqueness for isMinim and well-definedness of minim *}
+subsubsection \<open>Existence and uniqueness for isMinim and well-definedness of minim\<close>
lemma isMinim_unique:
assumes MINIM: "isMinim B a" and MINIM': "isMinim B a'"
@@ -254,7 +254,7 @@
unfolding minim_def using theI[of ?phi b] by blast
qed
-subsubsection{* Properties of minim *}
+subsubsection\<open>Properties of minim\<close>
lemma minim_in:
assumes "B \<le> Field r" and "B \<noteq> {}"
@@ -294,7 +294,7 @@
using isMinim_unique by auto
qed
-subsubsection{* Properties of successor *}
+subsubsection\<open>Properties of successor\<close>
lemma suc_AboveS:
assumes SUB: "B \<le> Field r" and ABOVES: "AboveS B \<noteq> {}"
@@ -388,7 +388,7 @@
qed
-subsubsection {* Properties of order filters *}
+subsubsection \<open>Properties of order filters\<close>
lemma under_ofilter:
"ofilter (under a)"
@@ -494,7 +494,7 @@
thus "A \<le> (\<Union>a \<in> A. under a)" by blast
qed
-subsubsection{* Other properties *}
+subsubsection\<open>Other properties\<close>
lemma ofilter_linord:
assumes OF1: "ofilter A" and OF2: "ofilter B"
--- a/src/HOL/Basic_BNFs.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Basic_BNFs.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,7 +7,7 @@
Registration of basic types as bounded natural functors.
*)
-section {* Registration of Basic Types as Bounded Natural Functors *}
+section \<open>Registration of Basic Types as Bounded Natural Functors\<close>
theory Basic_BNFs
imports BNF_Def
--- a/src/HOL/Binomial.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Binomial.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,13 +5,13 @@
The integer version of factorial and other additions by Jeremy Avigad.
*)
-section{*Factorial Function, Binomial Coefficients and Binomial Theorem*}
+section\<open>Factorial Function, Binomial Coefficients and Binomial Theorem\<close>
theory Binomial
imports Main
begin
-subsection {* Factorial *}
+subsection \<open>Factorial\<close>
fun fact :: "nat \<Rightarrow> ('a::semiring_char_0)"
where "fact 0 = 1" | "fact (Suc n) = of_nat (Suc n) * fact n"
@@ -74,7 +74,7 @@
end
-text{*Note that @{term "fact 0 = fact 1"}*}
+text\<open>Note that @{term "fact 0 = fact 1"}\<close>
lemma fact_less_mono_nat: "\<lbrakk>0 < m; m < n\<rbrakk> \<Longrightarrow> fact m < (fact n :: nat)"
by (induct n) (auto simp: less_Suc_eq)
@@ -121,7 +121,7 @@
by (simp add: atLeastAtMostSuc_conv)
finally show ?case .
qed
- from this `m = n + d` show ?thesis by simp
+ from this \<open>m = n + d\<close> show ?thesis by simp
qed
lemma fact_num_eq_if:
@@ -141,15 +141,15 @@
by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono)
qed
-lemma fact_numeral: --{*Evaluation for specific numerals*}
+lemma fact_numeral: --\<open>Evaluation for specific numerals\<close>
"fact (numeral k) = (numeral k) * (fact (pred_numeral k))"
by (metis fact.simps(2) numeral_eq_Suc of_nat_numeral)
-text {* This development is based on the work of Andy Gordon and
- Florian Kammueller. *}
+text \<open>This development is based on the work of Andy Gordon and
+ Florian Kammueller.\<close>
-subsection {* Basic definitions and lemmas *}
+subsection \<open>Basic definitions and lemmas\<close>
primrec binomial :: "nat \<Rightarrow> nat \<Rightarrow> nat" (infixl "choose" 65)
where
@@ -207,27 +207,27 @@
apply (auto simp: power_Suc)
by (simp add: add_le_mono mult_2)
-text{*The absorption property*}
+text\<open>The absorption property\<close>
lemma Suc_times_binomial:
"Suc k * (Suc n choose Suc k) = Suc n * (n choose k)"
using Suc_times_binomial_eq by auto
-text{*This is the well-known version of absorption, but it's harder to use because of the
- need to reason about division.*}
+text\<open>This is the well-known version of absorption, but it's harder to use because of the
+ need to reason about division.\<close>
lemma binomial_Suc_Suc_eq_times:
"(Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k"
by (simp add: Suc_times_binomial_eq del: mult_Suc mult_Suc_right)
-text{*Another version of absorption, with -1 instead of Suc.*}
+text\<open>Another version of absorption, with -1 instead of Suc.\<close>
lemma times_binomial_minus1_eq:
"0 < k \<Longrightarrow> k * (n choose k) = n * ((n - 1) choose (k - 1))"
using Suc_times_binomial_eq [where n = "n - 1" and k = "k - 1"]
by (auto split add: nat_diff_split)
-subsection {* Combinatorial theorems involving @{text "choose"} *}
+subsection \<open>Combinatorial theorems involving @{text "choose"}\<close>
-text {*By Florian Kamm\"uller, tidied by LCP.*}
+text \<open>By Florian Kamm\"uller, tidied by LCP.\<close>
lemma card_s_0_eq_empty: "finite A \<Longrightarrow> card {B. B \<subseteq> A & card B = 0} = 1"
by (simp cong add: conj_cong add: finite_subset [THEN card_0_eq])
@@ -246,9 +246,9 @@
shows "finite {x. \<exists>A \<subseteq> B. P x A}"
by (metis (no_types) assms finite_Collect_bounded_ex finite_Collect_subsets)
-text{*There are as many subsets of @{term A} having cardinality @{term k}
+text\<open>There are as many subsets of @{term A} having cardinality @{term k}
as there are sets obtained from the former by inserting a fixed element
- @{term x} into each.*}
+ @{term x} into each.\<close>
lemma constr_bij:
"finite A \<Longrightarrow> x \<notin> A \<Longrightarrow>
card {B. \<exists>C. C \<subseteq> A \<and> card C = k \<and> B = insert x C} =
@@ -258,16 +258,16 @@
apply (metis card_Diff_singleton_if finite_subset in_mono)
done
-text {*
+text \<open>
Main theorem: combinatorial statement about number of subsets of a set.
-*}
+\<close>
theorem n_subsets: "finite A \<Longrightarrow> card {B. B \<subseteq> A \<and> card B = k} = (card A choose k)"
proof (induct k arbitrary: A)
case 0 then show ?case by (simp add: card_s_0_eq_empty)
next
case (Suc k)
- show ?case using `finite A`
+ show ?case using \<open>finite A\<close>
proof (induct A)
case empty show ?case by (simp add: card_s_0_eq_empty)
next
@@ -285,9 +285,9 @@
qed
-subsection {* The binomial theorem (courtesy of Tobias Nipkow): *}
+subsection \<open>The binomial theorem (courtesy of Tobias Nipkow):\<close>
-text{* Avigad's version, generalized to any commutative ring *}
+text\<open>Avigad's version, generalized to any commutative ring\<close>
theorem binomial_ring: "(a+b::'a::{comm_ring_1,power})^n =
(\<Sum>k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
proof (induct n)
@@ -324,7 +324,7 @@
finally show "?P (Suc n)" by simp
qed
-text{* Original version for the naturals *}
+text\<open>Original version for the naturals\<close>
corollary binomial: "(a+b::nat)^n = (\<Sum>k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))"
using binomial_ring [of "int a" "int b" n]
by (simp only: of_nat_add [symmetric] of_nat_mult [symmetric] of_nat_power [symmetric]
@@ -386,7 +386,7 @@
shows "(\<And>k. m \<le> k \<Longrightarrow> k \<le> n \<Longrightarrow> g k = f (m + n - k)) \<Longrightarrow> (\<Sum>k=m..n. f k) = (\<Sum>k=m..n. g k)"
by (rule setsum.reindex_bij_witness[where i="\<lambda>k. m+n-k" and j="\<lambda>k. m+n-k"]) auto
-text{*NW diagonal sum property*}
+text\<open>NW diagonal sum property\<close>
lemma sum_choose_diagonal:
assumes "m\<le>n" shows "(\<Sum>k=0..m. (n-k) choose (m-k)) = Suc n choose m"
proof -
@@ -399,9 +399,9 @@
finally show ?thesis .
qed
-subsection{* Pochhammer's symbol : generalized rising factorial *}
+subsection\<open>Pochhammer's symbol : generalized rising factorial\<close>
-text {* See @{url "http://en.wikipedia.org/wiki/Pochhammer_symbol"} *}
+text \<open>See @{url "http://en.wikipedia.org/wiki/Pochhammer_symbol"}\<close>
definition "pochhammer (a::'a::comm_semiring_1) n =
(if n = 0 then 1 else setprod (\<lambda>n. a + of_nat n) {0 .. n - 1})"
@@ -548,7 +548,7 @@
by (simp add: of_nat_diff pochhammer_fact)
-subsection{* Generalized binomial coefficients *}
+subsection\<open>Generalized binomial coefficients\<close>
definition gbinomial :: "'a::field_char_0 \<Rightarrow> nat \<Rightarrow> 'a" (infixl "gchoose" 65)
where "a gchoose n =
@@ -729,8 +729,8 @@
then show ?thesis using kn by simp
qed
-text{*Contributed by Manuel Eberl, generalised by LCP.
- Alternative definition of the binomial coefficient as @{term "\<Prod>i<k. (n - i) / (k - i)"} *}
+text\<open>Contributed by Manuel Eberl, generalised by LCP.
+ Alternative definition of the binomial coefficient as @{term "\<Prod>i<k. (n - i) / (k - i)"}\<close>
lemma gbinomial_altdef_of_nat:
fixes k :: nat
and x :: "'a :: {field_char_0,field}"
@@ -769,15 +769,15 @@
then have "x * of_nat (k - i) \<le> (x - of_nat i) * (of_nat k :: 'a)"
unfolding of_nat_mult[symmetric] of_nat_le_iff .
with assms show "x / of_nat k \<le> (x - of_nat i) / (of_nat (k - i) :: 'a)"
- using `i < k` by (simp add: field_simps)
+ using \<open>i < k\<close> by (simp add: field_simps)
qed (simp add: x zero_le_divide_iff)
finally show ?thesis .
qed
-text{*Versions of the theorems above for the natural-number version of "choose"*}
+text\<open>Versions of the theorems above for the natural-number version of "choose"\<close>
lemma binomial_altdef_of_nat:
fixes n k :: nat
- and x :: "'a :: {field_char_0,field}" --{*the point is to constrain @{typ 'a}*}
+ and x :: "'a :: {field_char_0,field}" --\<open>the point is to constrain @{typ 'a}\<close>
assumes "k \<le> n"
shows "of_nat (n choose k) = (\<Prod>i<k. of_nat (n - i) / of_nat (k - i) :: 'a)"
using assms
@@ -795,7 +795,7 @@
shows "n choose r \<le> n ^ r"
proof -
have "n choose r \<le> fact n div fact (n - r)"
- using `r \<le> n` by (subst binomial_fact_lemma[symmetric]) auto
+ using \<open>r \<le> n\<close> by (subst binomial_fact_lemma[symmetric]) auto
with fact_div_fact_le_pow [OF assms] show ?thesis by auto
qed
@@ -838,7 +838,7 @@
by (simp add: binomial_altdef_nat mult.commute)
qed
-text{*The "Subset of a Subset" identity*}
+text\<open>The "Subset of a Subset" identity\<close>
lemma choose_mult:
assumes "k\<le>m" "m\<le>n"
shows "(n choose m) * (m choose k) = (n choose k) * ((n-k) choose (m-k))"
@@ -846,7 +846,7 @@
by simp
-subsection {* Binomial coefficients *}
+subsection \<open>Binomial coefficients\<close>
lemma choose_one: "(n::nat) choose 1 = n"
by simp
@@ -885,7 +885,7 @@
fix x
assume x: "x \<in> \<Union>A"
def K \<equiv> "{X \<in> A. x \<in> X}"
- with `finite A` have K: "finite K" by auto
+ with \<open>finite A\<close> have K: "finite K" by auto
let ?I = "\<lambda>i. {I. I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I}"
have "inj_on snd (SIGMA i:{1..card A}. ?I i)"
using assms by(auto intro!: inj_onI)
@@ -901,7 +901,7 @@
by(subst setsum_right_distrib) simp
also have "\<dots> = (\<Sum>i=1..card K. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> K \<and> card I = i. 1))" (is "_ = ?rhs")
proof(rule setsum.mono_neutral_cong_right[rule_format])
- show "{1..card K} \<subseteq> {1..card A}" using `finite A`
+ show "{1..card K} \<subseteq> {1..card A}" using \<open>finite A\<close>
by(auto simp add: K_def intro: card_mono)
next
fix i
@@ -909,7 +909,7 @@
hence i: "i \<le> card A" "card K < i" by auto
have "{I. I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I} = {I. I \<subseteq> K \<and> card I = i}"
by(auto simp add: K_def)
- also have "\<dots> = {}" using `finite A` i
+ also have "\<dots> = {}" using \<open>finite A\<close> i
by(auto simp add: K_def dest: card_mono[rotated 1])
finally show "(- 1) ^ (i + 1) * (\<Sum>I | I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1 :: int) = 0"
by(simp only:) simp
@@ -937,8 +937,8 @@
finally show ?thesis ..
qed
-text{* The number of nat lists of length @{text m} summing to @{text N} is
-@{term "(N + m - 1) choose N"}: *}
+text\<open>The number of nat lists of length @{text m} summing to @{text N} is
+@{term "(N + m - 1) choose N"}:\<close>
lemma card_length_listsum_rec:
assumes "m\<ge>1"
@@ -1007,7 +1007,7 @@
proof cases
assume "m = 1"
with Suc.hyps have "N\<ge>1" by auto
- with `m = 1` show ?thesis by (simp add: binomial_eq_0)
+ with \<open>m = 1\<close> show ?thesis by (simp add: binomial_eq_0)
next
assume "m \<noteq> 1" thus ?thesis using Suc by fastforce
qed
--- a/src/HOL/Code_Evaluation.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Code_Evaluation.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,16 +2,16 @@
Author: Florian Haftmann, TU Muenchen
*)
-section {* Term evaluation using the generic code generator *}
+section \<open>Term evaluation using the generic code generator\<close>
theory Code_Evaluation
imports Typerep Limited_Sequence
keywords "value" :: diag
begin
-subsection {* Term representation *}
+subsection \<open>Term representation\<close>
-subsubsection {* Terms and class @{text term_of} *}
+subsubsection \<open>Terms and class @{text term_of}\<close>
datatype (plugins only: code extraction) "term" = dummy_term
@@ -44,7 +44,7 @@
by (simp only: valapp_def fst_conv snd_conv)
-subsubsection {* Syntax *}
+subsubsection \<open>Syntax\<close>
definition termify :: "'a \<Rightarrow> term" where
[code del]: "termify x = dummy_term"
@@ -66,7 +66,7 @@
and valapp (infixl "{\<cdot>}" 70)
-subsection {* Tools setup and evaluation *}
+subsection \<open>Tools setup and evaluation\<close>
lemma eq_eq_TrueD:
assumes "(x \<equiv> y) \<equiv> Trueprop True"
@@ -87,7 +87,7 @@
ML_file "~~/src/HOL/Tools/value.ML"
-subsection {* @{text term_of} instances *}
+subsection \<open>@{text term_of} instances\<close>
instantiation "fun" :: (typerep, typerep) term_of
begin
@@ -138,12 +138,12 @@
code_reserved Eval HOLogic
-subsection {* Generic reification *}
+subsection \<open>Generic reification\<close>
ML_file "~~/src/HOL/Tools/reification.ML"
-subsection {* Diagnostic *}
+subsection \<open>Diagnostic\<close>
definition tracing :: "String.literal \<Rightarrow> 'a \<Rightarrow> 'a" where
[code del]: "tracing s x = x"
--- a/src/HOL/Code_Numeral.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Code_Numeral.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Author: Florian Haftmann, TU Muenchen
*)
-section {* Numeric types for code generation onto target language numerals only *}
+section \<open>Numeric types for code generation onto target language numerals only\<close>
theory Code_Numeral
imports Nat_Transfer Divides Lifting
begin
-subsection {* Type of target language integers *}
+subsection \<open>Type of target language integers\<close>
typedef integer = "UNIV \<Colon> int set"
morphisms int_of_integer integer_of_int ..
@@ -238,9 +238,9 @@
"integer_of_nat (numeral n) = numeral n"
by transfer simp
-subsection {* Code theorems for target language integers *}
+subsection \<open>Code theorems for target language integers\<close>
-text {* Constructors *}
+text \<open>Constructors\<close>
definition Pos :: "num \<Rightarrow> integer"
where
@@ -261,7 +261,7 @@
code_datatype "0::integer" Pos Neg
-text {* Auxiliary operations *}
+text \<open>Auxiliary operations\<close>
lift_definition dup :: "integer \<Rightarrow> integer"
is "\<lambda>k::int. k + k"
@@ -290,7 +290,7 @@
by (transfer, simp add: dbl_def dbl_inc_def dbl_dec_def)+
-text {* Implementations *}
+text \<open>Implementations\<close>
lemma one_integer_code [code, code_unfold]:
"1 = Pos Num.One"
@@ -520,7 +520,7 @@
hide_const (open) Pos Neg sub dup divmod_abs
-subsection {* Serializer setup for target language integers *}
+subsection \<open>Serializer setup for target language integers\<close>
code_reserved Eval int Integer abs
@@ -541,12 +541,12 @@
and (Haskell) "!(0/ ::/ Integer)"
and (Scala) "BigInt(0)"
-setup {*
+setup \<open>
fold (fn target =>
Numeral.add_code @{const_name Code_Numeral.Pos} I Code_Printer.literal_numeral target
#> Numeral.add_code @{const_name Code_Numeral.Neg} (op ~) Code_Printer.literal_numeral target)
["SML", "OCaml", "Haskell", "Scala"]
-*}
+\<close>
code_printing
constant "plus :: integer \<Rightarrow> _ \<Rightarrow> _" \<rightharpoonup>
@@ -613,7 +613,7 @@
code_module Code_Numeral \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
-subsection {* Type of target language naturals *}
+subsection \<open>Type of target language naturals\<close>
typedef natural = "UNIV \<Colon> nat set"
morphisms nat_of_natural natural_of_nat ..
@@ -787,7 +787,7 @@
by (rule is_measure_trivial)
-subsection {* Inductive representation of target language naturals *}
+subsection \<open>Inductive representation of target language naturals\<close>
lift_definition Suc :: "natural \<Rightarrow> natural"
is Nat.Suc
@@ -831,7 +831,7 @@
hide_const (open) Suc
-subsection {* Code refinement for target language naturals *}
+subsection \<open>Code refinement for target language naturals\<close>
lift_definition Nat :: "integer \<Rightarrow> natural"
is nat
--- a/src/HOL/Coinduction.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Coinduction.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,7 +6,7 @@
Coinduction method that avoids some boilerplate compared to coinduct.
*)
-section {* Coinduction Method *}
+section \<open>Coinduction Method\<close>
theory Coinduction
imports Ctr_Sugar
--- a/src/HOL/Complete_Lattices.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Complete_Lattices.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,6 +1,6 @@
(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *)
-section {* Complete lattices *}
+section \<open>Complete lattices\<close>
theory Complete_Lattices
imports Fun
@@ -11,7 +11,7 @@
less (infix "\<sqsubset>" 50)
-subsection {* Syntactic infimum and supremum operations *}
+subsection \<open>Syntactic infimum and supremum operations\<close>
class Inf =
fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
@@ -79,11 +79,11 @@
end
-text {*
+text \<open>
Note: must use names @{const INFIMUM} and @{const SUPREMUM} here instead of
@{text INF} and @{text SUP} to allow the following syntax coexist
with the plain constant names.
-*}
+\<close>
syntax
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3INF _./ _)" [0, 10] 10)
@@ -107,17 +107,17 @@
"SUP x. B" == "SUP x:CONST UNIV. B"
"SUP x:A. B" == "CONST SUPREMUM A (%x. B)"
-print_translation {*
+print_translation \<open>
[Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFIMUM} @{syntax_const "_INF"},
Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPREMUM} @{syntax_const "_SUP"}]
-*} -- {* to avoid eta-contraction of body *}
+\<close> -- \<open>to avoid eta-contraction of body\<close>
-subsection {* Abstract complete lattices *}
+subsection \<open>Abstract complete lattices\<close>
-text {* A complete lattice always has a bottom and a top,
+text \<open>A complete lattice always has a bottom and a top,
so we include them into the following type class,
along with assumptions that define bottom and top
-in terms of infimum and supremum. *}
+in terms of infimum and supremum.\<close>
class complete_lattice = lattice + Inf + Sup + bot + top +
assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
@@ -250,8 +250,8 @@
proof (rule Inf_greatest)
fix b assume "b \<in> B"
with assms obtain a where "a \<in> A" and "a \<sqsubseteq> b" by blast
- from `a \<in> A` have "\<Sqinter>A \<sqsubseteq> a" by (rule Inf_lower)
- with `a \<sqsubseteq> b` show "\<Sqinter>A \<sqsubseteq> b" by auto
+ from \<open>a \<in> A\<close> have "\<Sqinter>A \<sqsubseteq> a" by (rule Inf_lower)
+ with \<open>a \<sqsubseteq> b\<close> show "\<Sqinter>A \<sqsubseteq> b" by auto
qed
lemma INF_mono:
@@ -264,8 +264,8 @@
proof (rule Sup_least)
fix a assume "a \<in> A"
with assms obtain b where "b \<in> B" and "a \<sqsubseteq> b" by blast
- from `b \<in> B` have "b \<sqsubseteq> \<Squnion>B" by (rule Sup_upper)
- with `a \<sqsubseteq> b` show "a \<sqsubseteq> \<Squnion>B" by auto
+ from \<open>b \<in> B\<close> have "b \<sqsubseteq> \<Squnion>B" by (rule Sup_upper)
+ with \<open>a \<sqsubseteq> b\<close> show "a \<sqsubseteq> \<Squnion>B" by auto
qed
lemma SUP_mono:
@@ -274,7 +274,7 @@
lemma INF_superset_mono:
"B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
- -- {* The last inclusion is POSITIVE! *}
+ -- \<open>The last inclusion is POSITIVE!\<close>
by (blast intro: INF_mono dest: subsetD)
lemma SUP_subset_mono:
@@ -286,8 +286,8 @@
and "A \<noteq> {}"
shows "\<Sqinter>A \<sqsubseteq> u"
proof -
- from `A \<noteq> {}` obtain v where "v \<in> A" by blast
- moreover from `v \<in> A` assms(1) have "v \<sqsubseteq> u" by blast
+ from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
+ moreover from \<open>v \<in> A\<close> assms(1) have "v \<sqsubseteq> u" by blast
ultimately show ?thesis by (rule Inf_lower2)
qed
@@ -296,8 +296,8 @@
and "A \<noteq> {}"
shows "u \<sqsubseteq> \<Squnion>A"
proof -
- from `A \<noteq> {}` obtain v where "v \<in> A" by blast
- moreover from `v \<in> A` assms(1) have "u \<sqsubseteq> v" by blast
+ from \<open>A \<noteq> {}\<close> obtain v where "v \<in> A" by blast
+ moreover from \<open>v \<in> A\<close> assms(1) have "u \<sqsubseteq> v" by blast
ultimately show ?thesis by (rule Sup_upper2)
qed
@@ -359,7 +359,7 @@
assume "\<not> (\<forall>x\<in>A. x = \<top>)"
then obtain x where "x \<in> A" and "x \<noteq> \<top>" by blast
then obtain B where "A = insert x B" by blast
- with `\<Sqinter>A = \<top>` `x \<noteq> \<top>` show False by simp
+ with \<open>\<Sqinter>A = \<top>\<close> \<open>x \<noteq> \<top>\<close> show False by simp
qed
qed
then show "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)" by auto
@@ -435,8 +435,8 @@
lemma less_INF_D:
assumes "y < (\<Sqinter>i\<in>A. f i)" "i \<in> A" shows "y < f i"
proof -
- note `y < (\<Sqinter>i\<in>A. f i)`
- also have "(\<Sqinter>i\<in>A. f i) \<le> f i" using `i \<in> A`
+ note \<open>y < (\<Sqinter>i\<in>A. f i)\<close>
+ also have "(\<Sqinter>i\<in>A. f i) \<le> f i" using \<open>i \<in> A\<close>
by (rule INF_lower)
finally show "y < f i" .
qed
@@ -444,9 +444,9 @@
lemma SUP_lessD:
assumes "(\<Squnion>i\<in>A. f i) < y" "i \<in> A" shows "f i < y"
proof -
- have "f i \<le> (\<Squnion>i\<in>A. f i)" using `i \<in> A`
+ have "f i \<le> (\<Squnion>i\<in>A. f i)" using \<open>i \<in> A\<close>
by (rule SUP_upper)
- also note `(\<Squnion>i\<in>A. f i) < y`
+ also note \<open>(\<Squnion>i\<in>A. f i) < y\<close>
finally show "f i < y" .
qed
@@ -550,19 +550,19 @@
lemma mono_Inf:
shows "f (\<Sqinter>A) \<le> (\<Sqinter>x\<in>A. f x)"
- using `mono f` by (auto intro: complete_lattice_class.INF_greatest Inf_lower dest: monoD)
+ using \<open>mono f\<close> by (auto intro: complete_lattice_class.INF_greatest Inf_lower dest: monoD)
lemma mono_Sup:
shows "(\<Squnion>x\<in>A. f x) \<le> f (\<Squnion>A)"
- using `mono f` by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD)
+ using \<open>mono f\<close> by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD)
lemma mono_INF:
"f (INF i : I. A i) \<le> (INF x : I. f (A x))"
- by (intro complete_lattice_class.INF_greatest monoD[OF `mono f`] INF_lower)
+ by (intro complete_lattice_class.INF_greatest monoD[OF \<open>mono f\<close>] INF_lower)
lemma mono_SUP:
"(SUP x : I. f (A x)) \<le> f (SUP i : I. A i)"
- by (intro complete_lattice_class.SUP_least monoD[OF `mono f`] SUP_upper)
+ by (intro complete_lattice_class.SUP_least monoD[OF \<open>mono f\<close>] SUP_upper)
end
@@ -698,7 +698,7 @@
end
-subsection {* Complete lattice on @{typ bool} *}
+subsection \<open>Complete lattice on @{typ bool}\<close>
instantiation bool :: complete_lattice
begin
@@ -734,7 +734,7 @@
qed (auto intro: bool_induct)
-subsection {* Complete lattice on @{typ "_ \<Rightarrow> _"} *}
+subsection \<open>Complete lattice on @{typ "_ \<Rightarrow> _"}\<close>
instantiation "fun" :: (type, Inf) Inf
begin
@@ -787,7 +787,7 @@
instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra ..
-subsection {* Complete lattice on unary and binary predicates *}
+subsection \<open>Complete lattice on unary and binary predicates\<close>
lemma Inf1_I:
"(\<And>P. P \<in> A \<Longrightarrow> P a) \<Longrightarrow> (\<Sqinter>A) a"
@@ -878,7 +878,7 @@
using assms by auto
-subsection {* Complete lattice on @{typ "_ set"} *}
+subsection \<open>Complete lattice on @{typ "_ set"}\<close>
instantiation "set" :: (type) complete_lattice
begin
@@ -899,7 +899,7 @@
qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def)
-subsubsection {* Inter *}
+subsubsection \<open>Inter\<close>
abbreviation Inter :: "'a set set \<Rightarrow> 'a set" where
"Inter S \<equiv> \<Sqinter>S"
@@ -923,18 +923,18 @@
lemma InterI [intro!]: "(\<And>X. X \<in> C \<Longrightarrow> A \<in> X) \<Longrightarrow> A \<in> \<Inter>C"
by (simp add: Inter_eq)
-text {*
+text \<open>
\medskip A ``destruct'' rule -- every @{term X} in @{term C}
contains @{term A} as an element, but @{prop "A \<in> X"} can hold when
@{prop "X \<in> C"} does not! This rule is analogous to @{text spec}.
-*}
+\<close>
lemma InterD [elim, Pure.elim]: "A \<in> \<Inter>C \<Longrightarrow> X \<in> C \<Longrightarrow> A \<in> X"
by auto
lemma InterE [elim]: "A \<in> \<Inter>C \<Longrightarrow> (X \<notin> C \<Longrightarrow> R) \<Longrightarrow> (A \<in> X \<Longrightarrow> R) \<Longrightarrow> R"
- -- {* ``Classical'' elimination rule -- does not require proving
- @{prop "X \<in> C"}. *}
+ -- \<open>``Classical'' elimination rule -- does not require proving
+ @{prop "X \<in> C"}.\<close>
by (unfold Inter_eq) blast
lemma Inter_lower: "B \<in> A \<Longrightarrow> \<Inter>A \<subseteq> B"
@@ -971,15 +971,15 @@
by (fact Inf_superset_mono)
-subsubsection {* Intersections of families *}
+subsubsection \<open>Intersections of families\<close>
abbreviation INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
"INTER \<equiv> INFIMUM"
-text {*
+text \<open>
Note: must use name @{const INTER} here instead of @{text INT}
to allow the following syntax coexist with the plain constant name.
-*}
+\<close>
syntax
"_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10)
@@ -999,9 +999,9 @@
"INT x. B" == "INT x:CONST UNIV. B"
"INT x:A. B" == "CONST INTER A (%x. B)"
-print_translation {*
+print_translation \<open>
[Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}]
-*} -- {* to avoid eta-contraction of body *}
+\<close> -- \<open>to avoid eta-contraction of body\<close>
lemma INTER_eq:
"(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
@@ -1021,7 +1021,7 @@
by auto
lemma INT_E [elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> (b \<in> B a \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
- -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}. *}
+ -- \<open>"Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}.\<close>
by (auto simp add: INF_def image_def)
lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
@@ -1068,7 +1068,7 @@
lemma INT_anti_mono:
"A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow> (\<Inter>x\<in>B. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
- -- {* The last inclusion is POSITIVE! *}
+ -- \<open>The last inclusion is POSITIVE!\<close>
by (fact INF_superset_mono)
lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
@@ -1078,7 +1078,7 @@
by blast
-subsubsection {* Union *}
+subsubsection \<open>Union\<close>
abbreviation Union :: "'a set set \<Rightarrow> 'a set" where
"Union S \<equiv> \<Squnion>S"
@@ -1102,8 +1102,8 @@
lemma UnionI [intro]:
"X \<in> C \<Longrightarrow> A \<in> X \<Longrightarrow> A \<in> \<Union>C"
- -- {* The order of the premises presupposes that @{term C} is rigid;
- @{term A} may be flexible. *}
+ -- \<open>The order of the premises presupposes that @{term C} is rigid;
+ @{term A} may be flexible.\<close>
by auto
lemma UnionE [elim!]:
@@ -1147,15 +1147,15 @@
by (fact Sup_subset_mono)
-subsubsection {* Unions of families *}
+subsubsection \<open>Unions of families\<close>
abbreviation UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
"UNION \<equiv> SUPREMUM"
-text {*
+text \<open>
Note: must use name @{const UNION} here instead of @{text UN}
to allow the following syntax coexist with the plain constant name.
-*}
+\<close>
syntax
"_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10)
@@ -1175,18 +1175,18 @@
"UN x. B" == "UN x:CONST UNIV. B"
"UN x:A. B" == "CONST UNION A (%x. B)"
-text {*
+text \<open>
Note the difference between ordinary xsymbol syntax of indexed
unions and intersections (e.g.\ @{text"\<Union>a\<^sub>1\<in>A\<^sub>1. B"})
and their \LaTeX\ rendition: @{term"\<Union>a\<^sub>1\<in>A\<^sub>1. B"}. The
former does not make the index expression a subscript of the
union/intersection symbol because this leads to problems with nested
subscripts in Proof General.
-*}
+\<close>
-print_translation {*
+print_translation \<open>
[Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}]
-*} -- {* to avoid eta-contraction of body *}
+\<close> -- \<open>to avoid eta-contraction of body\<close>
lemma UNION_eq:
"(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
@@ -1211,8 +1211,8 @@
using Union_iff [of _ "B ` A"] by simp
lemma UN_I [intro]: "a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> b \<in> (\<Union>x\<in>A. B x)"
- -- {* The order of the premises presupposes that @{term A} is rigid;
- @{term b} may be flexible. *}
+ -- \<open>The order of the premises presupposes that @{term A} is rigid;
+ @{term b} may be flexible.\<close>
by auto
lemma UN_E [elim!]: "b \<in> (\<Union>x\<in>A. B x) \<Longrightarrow> (\<And>x. x\<in>A \<Longrightarrow> b \<in> B x \<Longrightarrow> R) \<Longrightarrow> R"
@@ -1295,7 +1295,7 @@
by blast
lemma vimage_eq_UN: "f -` B = (\<Union>y\<in>B. f -` {y})"
- -- {* NOT suitable for rewriting *}
+ -- \<open>NOT suitable for rewriting\<close>
by blast
lemma image_UN: "f ` UNION A B = (\<Union>x\<in>A. f ` B x)"
@@ -1305,7 +1305,7 @@
by blast
-subsubsection {* Distributive laws *}
+subsubsection \<open>Distributive laws\<close>
lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
by (fact inf_Sup)
@@ -1322,19 +1322,19 @@
lemma UN_Un_distrib: "(\<Union>i\<in>I. A i \<union> B i) = (\<Union>i\<in>I. A i) \<union> (\<Union>i\<in>I. B i)"
by (rule sym) (rule SUP_sup_distrib)
-lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A ` C) \<inter> \<Inter>(B ` C)" -- {* FIXME drop *}
+lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A ` C) \<inter> \<Inter>(B ` C)" -- \<open>FIXME drop\<close>
by (simp add: INT_Int_distrib)
-lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A ` C) \<union> \<Union>(B ` C)" -- {* FIXME drop *}
- -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *}
- -- {* Union of a family of unions *}
+lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A ` C) \<union> \<Union>(B ` C)" -- \<open>FIXME drop\<close>
+ -- \<open>Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:\<close>
+ -- \<open>Union of a family of unions\<close>
by (simp add: UN_Un_distrib)
lemma Un_INT_distrib: "B \<union> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. B \<union> A i)"
by (fact sup_INF)
lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. B \<inter> A i)"
- -- {* Halmos, Naive Set Theory, page 35. *}
+ -- \<open>Halmos, Naive Set Theory, page 35.\<close>
by (fact inf_SUP)
lemma Int_UN_distrib2: "(\<Union>i\<in>I. A i) \<inter> (\<Union>j\<in>J. B j) = (\<Union>i\<in>I. \<Union>j\<in>J. A i \<inter> B j)"
@@ -1347,7 +1347,7 @@
by (fact Sup_inf_eq_bot_iff)
-subsection {* Injections and bijections *}
+subsection \<open>Injections and bijections\<close>
lemma inj_on_Inter:
"S \<noteq> {} \<Longrightarrow> (\<And>A. A \<in> S \<Longrightarrow> inj_on f A) \<Longrightarrow> inj_on f (\<Inter>S)"
@@ -1429,7 +1429,7 @@
by (auto split: if_splits)
-subsubsection {* Complement *}
+subsubsection \<open>Complement\<close>
lemma Compl_INT [simp]: "- (\<Inter>x\<in>A. B x) = (\<Union>x\<in>A. -B x)"
by (fact uminus_INF)
@@ -1438,10 +1438,10 @@
by (fact uminus_SUP)
-subsubsection {* Miniscoping and maxiscoping *}
+subsubsection \<open>Miniscoping and maxiscoping\<close>
-text {* \medskip Miniscoping: pushing in quantifiers and big Unions
- and Intersections. *}
+text \<open>\medskip Miniscoping: pushing in quantifiers and big Unions
+ and Intersections.\<close>
lemma UN_simps [simp]:
"\<And>a B C. (\<Union>x\<in>C. insert a (B x)) = (if C={} then {} else insert a (\<Union>x\<in>C. B x))"
@@ -1477,7 +1477,7 @@
by auto
-text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *}
+text \<open>\medskip Maxiscoping: pulling out big Unions and Intersections.\<close>
lemma UN_extend_simps:
"\<And>a B C. insert a (\<Union>x\<in>C. B x) = (if C={} then {a} else (\<Union>x\<in>C. insert a (B x)))"
@@ -1505,7 +1505,7 @@
"\<And>A B f. (\<Inter>a\<in>A. B (f a)) = (\<Inter>x\<in>f`A. B x)"
by auto
-text {* Finally *}
+text \<open>Finally\<close>
no_notation
less_eq (infix "\<sqsubseteq>" 50) and
@@ -1514,7 +1514,7 @@
lemmas mem_simps =
insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
- -- {* Each of these has ALREADY been added @{text "[simp]"} above. *}
+ -- \<open>Each of these has ALREADY been added @{text "[simp]"} above.\<close>
end
--- a/src/HOL/Complete_Partial_Order.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Complete_Partial_Order.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,15 +3,15 @@
Author: Alexander Krauss, TU Muenchen
*)
-section {* Chain-complete partial orders and their fixpoints *}
+section \<open>Chain-complete partial orders and their fixpoints\<close>
theory Complete_Partial_Order
imports Product_Type
begin
-subsection {* Monotone functions *}
+subsection \<open>Monotone functions\<close>
-text {* Dictionary-passing version of @{const Orderings.mono}. *}
+text \<open>Dictionary-passing version of @{const Orderings.mono}.\<close>
definition monotone :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))"
@@ -24,11 +24,11 @@
unfolding monotone_def by iprover
-subsection {* Chains *}
+subsection \<open>Chains\<close>
-text {* A chain is a totally-ordered set. Chains are parameterized over
+text \<open>A chain is a totally-ordered set. Chains are parameterized over
the order for maximal flexibility, since type classes are not enough.
-*}
+\<close>
definition
chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
@@ -67,12 +67,12 @@
shows "chain le_b (f ` Y)"
by(blast intro: chainI dest: chainD[OF chain] mono)
-subsection {* Chain-complete partial orders *}
+subsection \<open>Chain-complete partial orders\<close>
-text {*
+text \<open>
A ccpo has a least upper bound for any chain. In particular, the
empty set is a chain, so every ccpo must have a bottom element.
-*}
+\<close>
class ccpo = order + Sup +
assumes ccpo_Sup_upper: "\<lbrakk>chain (op \<le>) A; x \<in> A\<rbrakk> \<Longrightarrow> x \<le> Sup A"
@@ -85,7 +85,7 @@
lemma ccpo_Sup_singleton [simp]: "\<Squnion>{x} = x"
by(rule antisym)(auto intro: ccpo_Sup_least ccpo_Sup_upper simp add: chain_singleton)
-subsection {* Transfinite iteration of a function *}
+subsection \<open>Transfinite iteration of a function\<close>
inductive_set iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set"
for f :: "'a \<Rightarrow> 'a"
@@ -145,7 +145,7 @@
lemma bot_in_iterates: "Sup {} \<in> iterates f"
by(auto intro: iterates.Sup simp add: chain_empty)
-subsection {* Fixpoint combinator *}
+subsection \<open>Fixpoint combinator\<close>
definition
fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
@@ -183,9 +183,9 @@
end
-subsection {* Fixpoint induction *}
+subsection \<open>Fixpoint induction\<close>
-setup {* Sign.map_naming (Name_Space.mandatory_path "ccpo") *}
+setup \<open>Sign.map_naming (Name_Space.mandatory_path "ccpo")\<close>
definition admissible :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
where "admissible lub ord P = (\<forall>A. chain ord A \<longrightarrow> (A \<noteq> {}) \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (lub A))"
@@ -203,7 +203,7 @@
shows "P (lub A)"
using assms by (auto simp: ccpo.admissible_def)
-setup {* Sign.map_naming Name_Space.parent_path *}
+setup \<open>Sign.map_naming Name_Space.parent_path\<close>
lemma (in ccpo) fixp_induct:
assumes adm: "ccpo.admissible Sup (op \<le>) P"
--- a/src/HOL/Complex.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Complex.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,17 +4,17 @@
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
*)
-section {* Complex Numbers: Rectangular and Polar Representations *}
+section \<open>Complex Numbers: Rectangular and Polar Representations\<close>
theory Complex
imports Transcendental
begin
-text {*
+text \<open>
We use the @{text codatatype} command to define the type of complex numbers. This allows us to use
@{text primcorec} to define complex functions by defining their real and imaginary result
separately.
-*}
+\<close>
codatatype complex = Complex (Re: real) (Im: real)
@@ -27,7 +27,7 @@
lemma complex_eq_iff: "x = y \<longleftrightarrow> Re x = Re y \<and> Im x = Im y"
by (auto intro: complex.expand)
-subsection {* Addition and Subtraction *}
+subsection \<open>Addition and Subtraction\<close>
instantiation complex :: ab_group_add
begin
@@ -53,7 +53,7 @@
end
-subsection {* Multiplication and Division *}
+subsection \<open>Multiplication and Division\<close>
instantiation complex :: field
begin
@@ -99,7 +99,7 @@
lemma Im_power_real [simp]: "Im x = 0 \<Longrightarrow> Im (x ^ n) = 0"
by (induct n) simp_all
-subsection {* Scalar Multiplication *}
+subsection \<open>Scalar Multiplication\<close>
instantiation complex :: real_field
begin
@@ -127,7 +127,7 @@
end
-subsection {* Numerals, Arithmetic, and Embedding from Reals *}
+subsection \<open>Numerals, Arithmetic, and Embedding from Reals\<close>
abbreviation complex_of_real :: "real \<Rightarrow> complex"
where "complex_of_real \<equiv> of_real"
@@ -171,7 +171,7 @@
"z \<in> \<real> \<Longrightarrow> of_real (Re z) = z"
by (auto simp: Reals_def)
-subsection {* The Complex Number $i$ *}
+subsection \<open>The Complex Number $i$\<close>
primcorec "ii" :: complex ("\<i>") where
"Re ii = 0"
@@ -231,7 +231,7 @@
lemma divide_numeral_i [simp]: "z / (numeral n * ii) = -(ii*z) / numeral n"
by (metis divide_divide_eq_left divide_i mult.commute mult_minus_right)
-subsection {* Vector Norm *}
+subsection \<open>Vector Norm\<close>
instantiation complex :: real_normed_field
begin
@@ -329,7 +329,7 @@
by (simp add: norm_complex_def divide_simps complex_eq_iff)
-text {* Properties of complex signum. *}
+text \<open>Properties of complex signum.\<close>
lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)"
by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult.commute)
@@ -341,7 +341,7 @@
by (simp add: complex_sgn_def divide_inverse)
-subsection {* Completeness of the Complexes *}
+subsection \<open>Completeness of the Complexes\<close>
lemma bounded_linear_Re: "bounded_linear Re"
by (rule bounded_linear_intro [where K=1], simp_all add: norm_complex_def)
@@ -407,7 +407,7 @@
declare
DERIV_power[where 'a=complex, unfolded of_nat_def[symmetric], derivative_intros]
-subsection {* Complex Conjugation *}
+subsection \<open>Complex Conjugation\<close>
primcorec cnj :: "complex \<Rightarrow> complex" where
"Re (cnj z) = Re z"
@@ -514,7 +514,7 @@
by (simp add: sums_def lim_cnj cnj_setsum [symmetric] del: cnj_setsum)
-subsection{*Basic Lemmas*}
+subsection\<open>Basic Lemmas\<close>
lemma complex_eq_0: "z=0 \<longleftrightarrow> (Re z)\<^sup>2 + (Im z)\<^sup>2 = 0"
by (metis zero_complex.sel complex_eqI sum_power2_eq_zero_iff)
@@ -616,13 +616,13 @@
done
qed
-subsection{*Polar Form for Complex Numbers*}
+subsection\<open>Polar Form for Complex Numbers\<close>
lemma complex_unimodular_polar: "(norm z = 1) \<Longrightarrow> \<exists>x. z = Complex (cos x) (sin x)"
using sincos_total_2pi [of "Re z" "Im z"]
by auto (metis cmod_power2 complex_eq power_one)
-subsubsection {* $\cos \theta + i \sin \theta$ *}
+subsubsection \<open>$\cos \theta + i \sin \theta$\<close>
primcorec cis :: "real \<Rightarrow> complex" where
"Re (cis a) = cos a"
@@ -661,7 +661,7 @@
lemma cis_pi: "cis pi = -1"
by (simp add: complex_eq_iff)
-subsubsection {* $r(\cos \theta + i \sin \theta)$ *}
+subsubsection \<open>$r(\cos \theta + i \sin \theta)$\<close>
definition rcis :: "real \<Rightarrow> real \<Rightarrow> complex" where
"rcis r a = complex_of_real r * cis a"
@@ -702,7 +702,7 @@
lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)"
by (simp add: rcis_def cis_divide [symmetric])
-subsubsection {* Complex exponential *}
+subsubsection \<open>Complex exponential\<close>
abbreviation Exp :: "complex \<Rightarrow> complex"
where "Exp \<equiv> exp"
@@ -747,7 +747,7 @@
lemma Exp_two_pi_i [simp]: "Exp((2::complex) * complex_of_real pi * ii) = 1"
by (simp add: Exp_eq_polar complex_eq_iff)
-subsubsection {* Complex argument *}
+subsubsection \<open>Complex argument\<close>
definition arg :: "complex \<Rightarrow> real" where
"arg z = (if z = 0 then 0 else (SOME a. sgn z = cis a \<and> -pi < a \<and> a \<le> pi))"
@@ -775,7 +775,7 @@
by (auto elim!: evenE dest!: less_2_cases)
thus "a = x" unfolding d_def by simp
qed (simp add: assms del: Re_sgn Im_sgn)
- with `z \<noteq> 0` show "arg z = x"
+ with \<open>z \<noteq> 0\<close> show "arg z = x"
unfolding arg_def by simp
qed
@@ -786,7 +786,7 @@
with assms have "r \<noteq> 0" by auto
def b \<equiv> "if 0 < r then a else a + pi"
have b: "sgn z = cis b"
- unfolding z b_def rcis_def using `r \<noteq> 0`
+ unfolding z b_def rcis_def using \<open>r \<noteq> 0\<close>
by (simp add: of_real_def sgn_scaleR sgn_if complex_eq_iff)
have cis_2pi_nat: "\<And>n. cis (2 * pi * real_of_nat n) = 1"
by (induct_tac n) (simp_all add: distrib_left cis_mult [symmetric] complex_eq_iff)
@@ -815,7 +815,7 @@
lemma cos_arg_i_mult_zero [simp]: "y \<noteq> 0 \<Longrightarrow> Re y = 0 \<Longrightarrow> cos (arg y) = 0"
using cis_arg [of y] by (simp add: complex_eq_iff)
-subsection {* Square root of complex numbers *}
+subsection \<open>Square root of complex numbers\<close>
primcorec csqrt :: "complex \<Rightarrow> complex" where
"Re (csqrt z) = sqrt ((cmod z + Re z) / 2)"
@@ -902,7 +902,7 @@
finally show ?thesis .
qed
-text {* Legacy theorem names *}
+text \<open>Legacy theorem names\<close>
lemmas expand_complex_eq = complex_eq_iff
lemmas complex_Re_Im_cancel_iff = complex_eq_iff
--- a/src/HOL/Complex_Main.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Complex_Main.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,4 +1,4 @@
-section {* Comprehensive Complex Theory *}
+section \<open>Comprehensive Complex Theory\<close>
theory Complex_Main
imports
--- a/src/HOL/Conditionally_Complete_Lattices.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Conditionally_Complete_Lattices.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Author: Luke S. Serafin, Carnegie Mellon University
*)
-section {* Conditionally-complete Lattices *}
+section \<open>Conditionally-complete Lattices\<close>
theory Conditionally_Complete_Lattices
imports Main
@@ -166,12 +166,12 @@
end
-text {*
+text \<open>
To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
@{const Inf} in theorem names with c.
-*}
+\<close>
class conditionally_complete_lattice = lattice + Sup + Inf +
assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
@@ -426,13 +426,13 @@
proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> 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 `a < b` `\<not> P b` linear less_le)
+ (metis \<open>a < b\<close> \<open>\<not> P b\<close> linear less_le)
next
show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
apply (rule cSup_least)
apply auto
apply (metis less_le_not_le)
- apply (metis `a<b` `~ P b` linear less_le)
+ apply (metis \<open>a<b\<close> \<open>~ P b\<close> linear less_le)
done
next
fix x
@@ -447,7 +447,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 `a<b` `~ P b` linear less_le)
+ (metis \<open>a<b\<close> \<open>~ P b\<close> linear less_le)
qed
end
@@ -538,13 +538,13 @@
{ fix z assume "z \<in> X"
have "z \<le> Max (X \<inter> {x..y})"
proof cases
- assume "x \<le> z" with `z \<in> X` `X \<subseteq> {..y}` *(1) show ?thesis
+ assume "x \<le> z" with \<open>z \<in> X\<close> \<open>X \<subseteq> {..y}\<close> *(1) show ?thesis
by (auto intro!: Max_ge)
next
assume "\<not> x \<le> z"
then have "z < x" by simp
also have "x \<le> Max (X \<inter> {x..y})"
- using `x \<in> X` *(1) `x \<le> y` by (intro Max_ge) auto
+ using \<open>x \<in> X\<close> *(1) \<open>x \<le> y\<close> by (intro Max_ge) auto
finally show ?thesis by simp
qed }
note le = this
--- a/src/HOL/Ctr_Sugar.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Ctr_Sugar.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,7 +6,7 @@
Wrapping existing freely generated type's constructors.
*)
-section {* Wrapping Existing Freely Generated Type's Constructors *}
+section \<open>Wrapping Existing Freely Generated Type's Constructors\<close>
theory Ctr_Sugar
imports HOL
--- a/src/HOL/Deriv.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Deriv.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,13 +6,13 @@
GMVT by Benjamin Porter, 2005
*)
-section{* Differentiation *}
+section\<open>Differentiation\<close>
theory Deriv
imports Limits
begin
-subsection {* Frechet derivative *}
+subsection \<open>Frechet derivative\<close>
definition
has_derivative :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> bool"
@@ -22,12 +22,12 @@
(bounded_linear f' \<and>
((\<lambda>y. ((f y - f (Lim F (\<lambda>x. x))) - f' (y - Lim F (\<lambda>x. x))) /\<^sub>R norm (y - Lim F (\<lambda>x. x))) ---> 0) F)"
-text {*
+text \<open>
Usually the filter @{term F} is @{term "at x within s"}. @{term "(f has_derivative D)
(at x within s)"} means: @{term D} is the derivative of function @{term f} at point @{term x}
within the set @{term s}. Where @{term s} is used to express left or right sided derivatives. In
most cases @{term s} is either a variable or @{term UNIV}.
-*}
+\<close>
lemma has_derivative_eq_rhs: "(f has_derivative f') F \<Longrightarrow> f' = g' \<Longrightarrow> (f has_derivative g') F"
by simp
@@ -51,7 +51,7 @@
by simp
named_theorems derivative_intros "structural introduction rules for derivatives"
-setup {*
+setup \<open>
let
val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs}
fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms
@@ -62,11 +62,11 @@
Named_Theorems.get (Context.proof_of context) @{named_theorems derivative_intros}
|> map_filter eq_rule)
end;
-*}
+\<close>
-text {*
+text \<open>
The following syntax is only used as a legacy syntax.
-*}
+\<close>
abbreviation (input)
FDERIV :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
@@ -189,7 +189,7 @@
lemmas has_derivative_within_subset = has_derivative_subset
-subsection {* Continuity *}
+subsection \<open>Continuity\<close>
lemma has_derivative_continuous:
assumes f: "(f has_derivative f') (at x within s)"
@@ -214,7 +214,7 @@
by (simp add: continuous_within)
qed
-subsection {* Composition *}
+subsection \<open>Composition\<close>
lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \<Longrightarrow> (f ---> y) (at x within s) \<longleftrightarrow> (f ---> y) (inf (nhds x) (principal s))"
unfolding tendsto_def eventually_inf_principal eventually_at_filter
@@ -390,7 +390,7 @@
then have "y \<noteq> 0"
by (auto simp: norm_conv_dist dist_commute)
have "norm (?inv y - ?inv x - ?f (y -x)) / norm (y - x) = norm ((?inv y - ?inv x) * (y - x) * ?inv x) / norm (y - x)"
- apply (subst inverse_diff_inverse [OF `y \<noteq> 0` x])
+ apply (subst inverse_diff_inverse [OF \<open>y \<noteq> 0\<close> x])
apply (subst minus_diff_minus)
apply (subst norm_minus_cancel)
apply (simp add: left_diff_distrib)
@@ -422,7 +422,7 @@
using has_derivative_mult[OF f has_derivative_inverse[OF x g]]
by (simp add: field_simps)
-text{*Conventional form requires mult-AC laws. Types real and complex only.*}
+text\<open>Conventional form requires mult-AC laws. Types real and complex only.\<close>
lemma has_derivative_divide'[derivative_intros]:
fixes f :: "_ \<Rightarrow> 'a::real_normed_field"
@@ -439,14 +439,14 @@
by simp
qed
-subsection {* Uniqueness *}
+subsection \<open>Uniqueness\<close>
-text {*
+text \<open>
This can not generally shown for @{const has_derivative}, as we need to approach the point from
all directions. There is a proof in @{text Multivariate_Analysis} for @{text euclidean_space}.
-*}
+\<close>
lemma has_derivative_zero_unique:
assumes "((\<lambda>x. 0) has_derivative F) (at x)" shows "F = (\<lambda>h. 0)"
@@ -485,7 +485,7 @@
unfolding fun_eq_iff right_minus_eq .
qed
-subsection {* Differentiability predicate *}
+subsection \<open>Differentiability predicate\<close>
definition
differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
@@ -609,7 +609,7 @@
lemma mult_commute_abs: "(\<lambda>x. x * c) = op * (c::'a::ab_semigroup_mult)"
by (simp add: fun_eq_iff mult.commute)
-subsection {* Vector derivative *}
+subsection \<open>Vector derivative\<close>
lemma has_field_derivative_iff_has_vector_derivative:
"(f has_field_derivative y) F \<longleftrightarrow> (f has_vector_derivative y) F"
@@ -687,7 +687,7 @@
by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left])
-subsection {* Derivatives *}
+subsection \<open>Derivatives\<close>
lemma DERIV_D: "DERIV f x :> D \<Longrightarrow> (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D"
by (simp add: DERIV_def)
@@ -748,7 +748,7 @@
by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult])
(auto simp: field_simps dest: has_field_derivative_imp_has_derivative)
-text {* Derivative of linear multiplication *}
+text \<open>Derivative of linear multiplication\<close>
lemma DERIV_cmult:
"(f has_field_derivative D) (at x within s) ==> ((\<lambda>x. c * f x) has_field_derivative c * D) (at x within s)"
@@ -784,24 +784,24 @@
by (rule arg_cong [of "\<lambda>x. x * D"]) (simp add: fun_eq_iff)
with assms have "(f has_derivative (\<lambda>x. x * D)) (at x within s)"
by (auto dest!: has_field_derivative_imp_has_derivative)
- then show ?thesis using `f x \<noteq> 0`
+ then show ?thesis using \<open>f x \<noteq> 0\<close>
by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse)
qed
-text {* Power of @{text "-1"} *}
+text \<open>Power of @{text "-1"}\<close>
lemma DERIV_inverse:
"x \<noteq> 0 \<Longrightarrow> ((\<lambda>x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)"
by (drule DERIV_inverse' [OF DERIV_ident]) simp
-text {* Derivative of inverse *}
+text \<open>Derivative of inverse\<close>
lemma DERIV_inverse_fun:
"(f has_field_derivative d) (at x within s) \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow>
((\<lambda>x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)"
by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib)
-text {* Derivative of quotient *}
+text \<open>Derivative of quotient\<close>
lemma DERIV_divide[derivative_intros]:
"(f has_field_derivative D) (at x within s) \<Longrightarrow>
@@ -842,7 +842,7 @@
((\<lambda>x. f (g x)) has_field_derivative Da * Db) (at x within s)"
by (rule DERIV_chain')
-text {* Standard version *}
+text \<open>Standard version\<close>
lemma DERIV_chain:
"DERIV f (g x) :> Da \<Longrightarrow> (g has_field_derivative Db) (at x within s) \<Longrightarrow>
@@ -872,7 +872,7 @@
declare
DERIV_power[where 'a=real, unfolded real_of_nat_def[symmetric], derivative_intros]
-text{*Alternative definition for differentiability*}
+text\<open>Alternative definition for differentiability\<close>
lemma DERIV_LIM_iff:
fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
@@ -908,7 +908,7 @@
by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right
tendsto_minus_cancel_left field_simps conj_commute)
-text {* Caratheodory formulation of derivative at a point *}
+text \<open>Caratheodory formulation of derivative at a point\<close>
lemma CARAT_DERIV: (*FIXME: SUPERSEDED BY THE ONE IN Deriv.thy. But still used by NSA/HDeriv.thy*)
"(DERIV f x :> l) \<longleftrightarrow> (\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> isCont g x \<and> g x = l)"
@@ -932,9 +932,9 @@
qed
-subsection {* Local extrema *}
+subsection \<open>Local extrema\<close>
-text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
+text\<open>If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right\<close>
lemma DERIV_pos_inc_right:
fixes f :: "real => real"
@@ -1037,23 +1037,23 @@
qed
-text{*Similar theorem for a local minimum*}
+text\<open>Similar theorem for a local minimum\<close>
lemma DERIV_local_min:
fixes f :: "real => real"
shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0"
by (drule DERIV_minus [THEN DERIV_local_max], auto)
-text{*In particular, if a function is locally flat*}
+text\<open>In particular, if a function is locally flat\<close>
lemma DERIV_local_const:
fixes f :: "real => real"
shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0"
by (auto dest!: DERIV_local_max)
-subsection {* Rolle's Theorem *}
+subsection \<open>Rolle's Theorem\<close>
-text{*Lemma about introducing open ball in open interval*}
+text\<open>Lemma about introducing open ball in open interval\<close>
lemma lemma_interval_lt:
"[| a < x; x < b |]
==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)"
@@ -1070,11 +1070,11 @@
apply force
done
-text{*Rolle's Theorem.
+text\<open>Rolle's Theorem.
If @{term f} is defined and continuous on the closed interval
@{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
and @{term "f(a) = f(b)"},
- then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
+ then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}\<close>
theorem Rolle:
assumes lt: "a < b"
and eq: "f(a) = f(b)"
@@ -1094,7 +1094,7 @@
show ?thesis
proof cases
assume axb: "a < x & x < b"
- --{*@{term f} attains its maximum within the interval*}
+ --\<open>@{term f} attains its maximum within the interval\<close>
hence ax: "a<x" and xb: "x<b" by arith +
from lemma_interval [OF ax xb]
obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
@@ -1104,7 +1104,7 @@
from differentiableD [OF dif [OF axb]]
obtain l where der: "DERIV f x :> l" ..
have "l=0" by (rule DERIV_local_max [OF der d bound'])
- --{*the derivative at a local maximum is zero*}
+ --\<open>the derivative at a local maximum is zero\<close>
thus ?thesis using ax xb der by auto
next
assume notaxb: "~ (a < x & x < b)"
@@ -1113,7 +1113,7 @@
show ?thesis
proof cases
assume ax'b: "a < x' & x' < b"
- --{*@{term f} attains its minimum within the interval*}
+ --\<open>@{term f} attains its minimum within the interval\<close>
hence ax': "a<x'" and x'b: "x'<b" by arith+
from lemma_interval [OF ax' x'b]
obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
@@ -1123,11 +1123,11 @@
from differentiableD [OF dif [OF ax'b]]
obtain l where der: "DERIV f x' :> l" ..
have "l=0" by (rule DERIV_local_min [OF der d bound'])
- --{*the derivative at a local minimum is zero*}
+ --\<open>the derivative at a local minimum is zero\<close>
thus ?thesis using ax' x'b der by auto
next
assume notax'b: "~ (a < x' & x' < b)"
- --{*@{term f} is constant througout the interval*}
+ --\<open>@{term f} is constant througout the interval\<close>
hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith
hence fb_eq_fx': "f b = f x'" by (auto simp add: eq)
from dense [OF lt]
@@ -1155,14 +1155,14 @@
from differentiableD [OF dif [OF conjI [OF ar rb]]]
obtain l where der: "DERIV f r :> l" ..
have "l=0" by (rule DERIV_local_const [OF der d bound'])
- --{*the derivative of a constant function is zero*}
+ --\<open>the derivative of a constant function is zero\<close>
thus ?thesis using ar rb der by auto
qed
qed
qed
-subsection{*Mean Value Theorem*}
+subsection\<open>Mean Value Theorem\<close>
lemma lemma_MVT:
"f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)"
@@ -1215,7 +1215,7 @@
done
-text{*A function is constant if its derivative is 0 over an interval.*}
+text\<open>A function is constant if its derivative is 0 over an interval.\<close>
lemma DERIV_isconst_end:
fixes f :: "real => real"
@@ -1261,14 +1261,14 @@
have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0"
proof (rule allI, rule impI)
fix z :: real assume "?a \<le> z \<and> z \<le> ?b"
- hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
+ hence "a < z" and "z < b" using \<open>x \<in> {a <..< b}\<close> and \<open>y \<in> {a <..< b}\<close> by auto
hence "z \<in> {a<..<b}" by auto
thus "DERIV f z :> 0" by (rule derivable)
qed
hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z"
and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto
- have "?a < ?b" using `x \<noteq> y` by auto
+ have "?a < ?b" using \<open>x \<noteq> y\<close> by auto
from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y]
show ?thesis by auto
qed auto
@@ -1302,7 +1302,7 @@
lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)"
by (simp)
-text{*Gallileo's "trick": average velocity = av. of end velocities*}
+text\<open>Gallileo's "trick": average velocity = av. of end velocities\<close>
lemma DERIV_const_average:
fixes v :: "real => real"
@@ -1455,7 +1455,7 @@
apply (metis filterlim_at_top_mirror lim)
done
-text {* Derivative of inverse function *}
+text \<open>Derivative of inverse function\<close>
lemma DERIV_inverse_function:
fixes f g :: "real \<Rightarrow> real"
@@ -1504,7 +1504,7 @@
using neq by (rule tendsto_inverse)
qed
-subsection {* Generalized Mean Value Theorem *}
+subsection \<open>Generalized Mean Value Theorem\<close>
theorem GMVT:
fixes a b :: real
@@ -1581,7 +1581,7 @@
qed
-subsection {* L'Hopitals rule *}
+subsection \<open>L'Hopitals rule\<close>
lemma isCont_If_ge:
fixes a :: "'a :: linorder_topology"
@@ -1645,21 +1645,21 @@
proof (rule bchoice, rule)
fix x assume "x \<in> {0 <..< a}"
then have x[arith]: "0 < x" "x < a" by auto
- with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x"
+ with g'_neq_0 g_neq_0 \<open>g 0 = 0\<close> have g': "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x"
by auto
have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont f x"
- using `isCont f 0` f by (auto intro: DERIV_isCont simp: le_less)
+ using \<open>isCont f 0\<close> f by (auto intro: DERIV_isCont simp: le_less)
moreover have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont g x"
- using `isCont g 0` g by (auto intro: DERIV_isCont simp: le_less)
+ using \<open>isCont g 0\<close> g by (auto intro: DERIV_isCont simp: le_less)
ultimately have "\<exists>c. 0 < c \<and> c < x \<and> (f x - f 0) * g' c = (g x - g 0) * f' c"
- using f g `x < a` by (intro GMVT') auto
+ using f g \<open>x < a\<close> by (intro GMVT') auto
then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c"
by blast
moreover
from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c"
by (simp add: field_simps)
ultimately show "\<exists>y. 0 < y \<and> y < x \<and> f x / g x = f' y / g' y"
- using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c])
+ using \<open>f 0 = 0\<close> \<open>g 0 = 0\<close> by (auto intro!: exI[of _ c])
qed
then obtain \<zeta> where "\<forall>x\<in>{0 <..< a}. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)" ..
then have \<zeta>: "eventually (\<lambda>x. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)) (at_right 0)"
@@ -1765,7 +1765,7 @@
by (intro tendsto_intros)
then have "((\<lambda>t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)"
by (simp add: inverse_eq_divide)
- from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e`
+ from this[unfolded tendsto_iff, rule_format, of "e / 2"] \<open>0 < e\<close>
have "eventually (\<lambda>t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)"
by (auto simp: dist_real_def)
@@ -1780,7 +1780,7 @@
and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y"
by blast
from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y"
- using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps)
+ using \<open>g a < g t\<close> g'_neq_0[of y] by (auto simp add: field_simps)
have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t"
by (simp add: field_simps)
@@ -1790,7 +1790,7 @@
also have "\<dots> = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)"
by (simp add: abs_mult D_eq dist_real_def)
also have "\<dots> < (e / 4) * 2 + e / 2"
- using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto
+ using ineq Df[of y] \<open>0 < e\<close> by (intro add_le_less_mono mult_mono) auto
finally show "dist (f t / g t) x < e"
by (simp add: dist_real_def)
qed
--- a/src/HOL/Divides.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Divides.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1999 University of Cambridge
*)
-section {* The division operators div and mod *}
+section \<open>The division operators div and mod\<close>
theory Divides
imports Parity
begin
-subsection {* Abstract division in commutative semirings. *}
+subsection \<open>Abstract division in commutative semirings.\<close>
class div = dvd + divide +
fixes mod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "mod" 70)
@@ -38,7 +38,7 @@
"n \<noteq> 0 \<Longrightarrow> a ^ n = 0 \<longleftrightarrow> a = 0"
using power_not_zero [of a n] by (auto simp add: zero_power)
-text {* @{const divide} and @{const mod} *}
+text \<open>@{const divide} and @{const mod}\<close>
lemma mod_div_equality2: "b * (a div b) + a mod b = a"
unfolding mult.commute [of b]
@@ -232,7 +232,7 @@
then show ?thesis by (simp add: mod_div_equality)
qed
-text {* Addition respects modular equivalence. *}
+text \<open>Addition respects modular equivalence.\<close>
lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c"
proof -
@@ -274,7 +274,7 @@
\<Longrightarrow> (x + y) div z = x div z + y div z"
by (cases "z = 0", simp, unfold dvd_def, auto simp add: algebra_simps)
-text {* Multiplication respects modular equivalence. *}
+text \<open>Multiplication respects modular equivalence.\<close>
lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c"
proof -
@@ -312,7 +312,7 @@
by (simp only: mod_mult_eq [symmetric])
qed
-text {* Exponentiation respects modular equivalence. *}
+text \<open>Exponentiation respects modular equivalence.\<close>
lemma power_mod: "(a mod b)^n mod b = a^n mod b"
apply (induct n, simp_all)
@@ -325,10 +325,10 @@
assumes "c dvd b"
shows "a mod b mod c = a mod c"
proof -
- from `c dvd b` obtain k where "b = c * k"
+ from \<open>c dvd b\<close> obtain k where "b = c * k"
by (rule dvdE)
have "a mod b mod c = a mod (c * k) mod c"
- by (simp only: `b = c * k`)
+ by (simp only: \<open>b = c * k\<close>)
also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
by (simp only: mod_mult_self1)
also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
@@ -400,7 +400,7 @@
next
assume "b div a = c"
then have "b div a * a = c * a" by simp
- moreover from `a dvd b` have "b div a * a = b" by simp
+ moreover from \<open>a dvd b\<close> have "b div a * a = b" by simp
ultimately show "b = c * a" by simp
qed
@@ -416,7 +416,7 @@
subclass idom_divide ..
-text {* Negation respects modular equivalence. *}
+text \<open>Negation respects modular equivalence.\<close>
lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b"
proof -
@@ -439,7 +439,7 @@
by (simp only: mod_minus_eq [symmetric])
qed
-text {* Subtraction respects modular equivalence. *}
+text \<open>Subtraction respects modular equivalence.\<close>
lemma mod_diff_left_eq:
"(a - b) mod c = (a mod c - b) mod c"
@@ -512,7 +512,7 @@
end
-subsubsection {* Parity and division *}
+subsubsection \<open>Parity and division\<close>
class semiring_div_parity = semiring_div + comm_semiring_1_cancel + numeral +
assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
@@ -594,15 +594,15 @@
end
-subsection {* Generic numeral division with a pragmatic type class *}
-
-text {*
+subsection \<open>Generic numeral division with a pragmatic type class\<close>
+
+text \<open>
The following type class contains everything necessary to formulate
a division algorithm in ring structures with numerals, restricted
to its positive segments. This is its primary motiviation, and it
could surely be formulated using a more fine-grained, more algebraic
and less technical class hierarchy.
-*}
+\<close>
class semiring_numeral_div = semiring_div + comm_semiring_1_cancel + linordered_semidom +
assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
@@ -636,11 +636,11 @@
then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
have "0 < 2" by simp
with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
- with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
+ with \<open>a mod 2 \<noteq> 0\<close> have "0 < a mod 2" by simp
with discrete have "1 \<le> a mod 2" by simp
- with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
+ with \<open>a mod 2 \<noteq> 1\<close> have "1 < a mod 2" by simp
with discrete have "2 \<le> a mod 2" by simp
- with `a mod 2 < 2` show False by simp
+ with \<open>a mod 2 < 2\<close> show False by simp
qed
next
show "1 mod 2 = 1"
@@ -657,9 +657,9 @@
proof -
from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a"
by (auto intro: trans)
- with `0 < b` have "0 < a div b" by (auto intro: div_positive)
+ with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive)
then have [simp]: "1 \<le> a div b" by (simp add: discrete)
- with `0 < b` have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
+ with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
have mod_w: "a mod (2 * b) = a mod b + b * w"
by (simp add: w_def mod_mult2_eq ac_simps)
@@ -668,7 +668,7 @@
with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp
have "2 * (a div (2 * b)) = a div b - w"
by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps)
- with `w = 1` have div: "2 * (a div (2 * b)) = a div b - 1" by simp
+ with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp
then show ?P and ?Q
by (simp_all add: div mod add_implies_diff [symmetric] le_add_diff_inverse2)
qed
@@ -683,7 +683,7 @@
by (simp add: w_def mod_mult2_eq ac_simps)
moreover have "b \<le> a mod b + b"
proof -
- from `0 < b` pos_mod_sign have "0 \<le> a mod b" by blast
+ from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast
then have "0 + b \<le> a mod b + b" by (rule add_right_mono)
then show ?thesis by simp
qed
@@ -692,7 +692,7 @@
with mod_w have mod: "a mod (2 * b) = a mod b" by simp
have "2 * (a div (2 * b)) = a div b - w"
by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps)
- with `w = 0` have div: "2 * (a div (2 * b)) = a div b" by simp
+ with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp
then show ?P and ?Q
by (simp_all add: div mod)
qed
@@ -715,12 +715,12 @@
in if r \<ge> numeral l then (2 * q + 1, r - numeral l)
else (2 * q, r))"
-text {*
+text \<open>
This is a formulation of one step (referring to one digit position)
in school-method division: compare the dividend at the current
digit position with the remainder from previous division steps
and evaluate accordingly.
-*}
+\<close>
lemma divmod_step_eq [code]:
"divmod_step l (q, r) = (if numeral l \<le> r
@@ -732,13 +732,13 @@
"numeral l \<le> r \<Longrightarrow> divmod_step l (q, r) = (2 * q + 1, r - numeral l)"
by (auto simp add: divmod_step_eq not_le)
-text {*
+text \<open>
This is a formulation of school-method division.
If the divisor is smaller than the dividend, terminate.
If not, shift the dividend to the right until termination
occurs and then reiterate single division steps in the
opposite direction.
-*}
+\<close>
lemma divmod_divmod_step [code]:
"divmod m n = (if m < n then (0, numeral m)
@@ -803,7 +803,7 @@
div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] add.commute del: numeral_times_numeral)
qed
-text {* Special case: divisibility *}
+text \<open>Special case: divisibility\<close>
definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"
where
@@ -820,21 +820,21 @@
end
-subsection {* Division on @{typ nat} *}
-
-text {*
+subsection \<open>Division on @{typ nat}\<close>
+
+text \<open>
We define @{const divide} and @{const mod} on @{typ nat} by means
of a characteristic relation with two input arguments
@{term "m\<Colon>nat"}, @{term "n\<Colon>nat"} and two output arguments
@{term "q\<Colon>nat"}(uotient) and @{term "r\<Colon>nat"}(emainder).
-*}
+\<close>
definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where
"divmod_nat_rel m n qr \<longleftrightarrow>
m = fst qr * n + snd qr \<and>
(if n = 0 then fst qr = 0 else if n > 0 then 0 \<le> snd qr \<and> snd qr < n else n < snd qr \<and> snd qr \<le> 0)"
-text {* @{const divmod_nat_rel} is total: *}
+text \<open>@{const divmod_nat_rel} is total:\<close>
lemma divmod_nat_rel_ex:
obtains q r where "divmod_nat_rel m n (q, r)"
@@ -845,7 +845,7 @@
case False
have "\<exists>q r. m = q * n + r \<and> r < n"
proof (induct m)
- case 0 with `n \<noteq> 0`
+ case 0 with \<open>n \<noteq> 0\<close>
have "(0\<Colon>nat) = 0 * n + 0 \<and> 0 < n" by simp
then show ?case by blast
next
@@ -860,14 +860,14 @@
moreover from n have "Suc r' \<le> n" by auto
ultimately have "n = Suc r'" by auto
with m have "Suc m = Suc q' * n + 0" by simp
- with `n \<noteq> 0` show ?thesis by blast
+ with \<open>n \<noteq> 0\<close> show ?thesis by blast
qed
qed
with that show thesis
- using `n \<noteq> 0` by (auto simp add: divmod_nat_rel_def)
+ using \<open>n \<noteq> 0\<close> by (auto simp add: divmod_nat_rel_def)
qed
-text {* @{const divmod_nat_rel} is injective: *}
+text \<open>@{const divmod_nat_rel} is injective:\<close>
lemma divmod_nat_rel_unique:
assumes "divmod_nat_rel m n qr"
@@ -884,17 +884,17 @@
apply (subst less_iff_Suc_add)
apply (auto simp add: add_mult_distrib)
done
- from `n \<noteq> 0` assms have *: "fst qr = fst qr'"
+ from \<open>n \<noteq> 0\<close> assms have *: "fst qr = fst qr'"
by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym)
with assms have "snd qr = snd qr'"
by (simp add: divmod_nat_rel_def)
with * show ?thesis by (cases qr, cases qr') simp
qed
-text {*
+text \<open>
We instantiate divisibility on the natural numbers by
means of @{const divmod_nat_rel}:
-*}
+\<close>
definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
"divmod_nat m n = (THE qr. divmod_nat_rel m n qr)"
@@ -966,7 +966,7 @@
unfolding divmod_nat_rel_def using assms by auto
qed
-text {* The ''recursion'' equations for @{const divide} and @{const mod} *}
+text \<open>The ''recursion'' equations for @{const divide} and @{const mod}\<close>
lemma div_less [simp]:
fixes m n :: nat
@@ -1043,11 +1043,11 @@
let (q, r) = divmod_nat (m - n) n in (Suc q, r))"
by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
-text {* Simproc for cancelling @{const divide} and @{const mod} *}
+text \<open>Simproc for cancelling @{const divide} and @{const mod}\<close>
ML_file "~~/src/Provers/Arith/cancel_div_mod.ML"
-ML {*
+ML \<open>
structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
(
val div_name = @{const_name divide};
@@ -1073,12 +1073,12 @@
val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
(@{thm add_0_left} :: @{thm add_0_right} :: @{thms ac_simps}))
)
-*}
-
-simproc_setup cancel_div_mod_nat ("(m::nat) + n") = {* K Cancel_Div_Mod_Nat.proc *}
-
-
-subsubsection {* Quotient *}
+\<close>
+
+simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \<open>K Cancel_Div_Mod_Nat.proc\<close>
+
+
+subsubsection \<open>Quotient\<close>
lemma div_geq: "0 < n \<Longrightarrow> \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)"
by (simp add: le_div_geq linorder_not_less)
@@ -1098,15 +1098,15 @@
assumes "m \<ge> n"
shows "m div n > 0"
proof -
- from `m \<ge> n` obtain q where "m = n + q"
+ from \<open>m \<ge> n\<close> obtain q where "m = n + q"
by (auto simp add: le_iff_add)
- with `n > 0` show ?thesis by simp
+ with \<open>n > 0\<close> show ?thesis by simp
qed
lemma div_eq_0_iff: "(a div b::nat) = 0 \<longleftrightarrow> a < b \<or> b = 0"
by (metis div_less div_positive div_by_0 gr0I less_numeral_extra(3) not_less)
-subsubsection {* Remainder *}
+subsubsection \<open>Remainder\<close>
lemma mod_less_divisor [simp]:
fixes m n :: nat
@@ -1144,7 +1144,7 @@
apply simp
done
-subsubsection {* Quotient and Remainder *}
+subsubsection \<open>Quotient and Remainder\<close>
lemma divmod_nat_rel_mult1_eq:
"divmod_nat_rel b c (q, r)
@@ -1192,7 +1192,7 @@
by intro_classes (auto intro: div_positive simp add: mult_div_cancel mod_mult2_eq div_mult2_eq)
-subsubsection {* Further Facts about Quotient and Remainder *}
+subsubsection \<open>Further Facts about Quotient and Remainder\<close>
lemma div_1 [simp]:
"m div Suc 0 = m"
@@ -1254,7 +1254,7 @@
apply (simp_all)
done
-text{*A fact for the mutilated chess board*}
+text\<open>A fact for the mutilated chess board\<close>
lemma mod_Suc: "Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))"
apply (case_tac "n=0", simp)
apply (induct "m" rule: nat_less_induct)
@@ -1400,7 +1400,7 @@
done
-subsubsection {* An ``induction'' law for modulus arithmetic. *}
+subsubsection \<open>An ``induction'' law for modulus arithmetic.\<close>
lemma mod_induct_0:
assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)"
@@ -1505,9 +1505,9 @@
then show ?thesis by auto
qed
-text{*These lemmas collapse some needless occurrences of Suc:
+text\<open>These lemmas collapse some needless occurrences of Suc:
at least three Sucs, since two and fewer are rewritten back to Suc again!
- We already have some rules to simplify operands smaller than 3.*}
+ We already have some rules to simplify operands smaller than 3.\<close>
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
by (simp add: Suc3_eq_add_3)
@@ -1611,27 +1611,27 @@
qed
-subsection {* Division on @{typ int} *}
+subsection \<open>Division on @{typ int}\<close>
definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where
- --{*definition of quotient and remainder*}
+ --\<open>definition of quotient and remainder\<close>
"divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and>
(if 0 < b then 0 \<le> r \<and> r < b else if b < 0 then b < r \<and> r \<le> 0 else q = 0))"
-text {*
+text \<open>
The following algorithmic devlopment actually echos what has already
been developed in class @{class semiring_numeral_div}. In the long
run it seems better to derive division on @{typ int} just from
division on @{typ nat} and instantiate @{class semiring_numeral_div}
accordingly.
-*}
+\<close>
definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where
- --{*for the division algorithm*}
+ --\<open>for the division algorithm\<close>
"adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b)
else (2 * q, r))"
-text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
+text\<open>algorithm for the case @{text "a\<ge>0, b>0"}\<close>
function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
"posDivAlg a b = (if a < b \<or> b \<le> 0 then (0, a)
else adjust b (posDivAlg a (2 * b)))"
@@ -1639,7 +1639,7 @@
termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))")
(auto simp add: mult_2)
-text{*algorithm for the case @{text "a<0, b>0"}*}
+text\<open>algorithm for the case @{text "a<0, b>0"}\<close>
function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
"negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0 then (-1, a + b)
else adjust b (negDivAlg a (2 * b)))"
@@ -1647,12 +1647,12 @@
termination by (relation "measure (\<lambda>(a, b). nat (- a - b))")
(auto simp add: mult_2)
-text{*algorithm for the general case @{term "b\<noteq>0"}*}
+text\<open>algorithm for the general case @{term "b\<noteq>0"}\<close>
definition divmod_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
- --{*The full division algorithm considers all possible signs for a, b
+ --\<open>The full division algorithm considers all possible signs for a, b
including the special case @{text "a=0, b<0"} because
- @{term negDivAlg} requires @{term "a<0"}.*}
+ @{term negDivAlg} requires @{term "a<0"}.\<close>
"divmod_int a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b
else if a = 0 then (0, 0)
else apsnd uminus (negDivAlg (-a) (-b))
@@ -1681,7 +1681,7 @@
"divmod_int p q = (p div q, p mod q)"
by (simp add: prod_eq_iff)
-text{*
+text\<open>
Here is the division algorithm in ML:
\begin{verbatim}
@@ -1707,10 +1707,10 @@
if 0<b then negDivAlg (a,b)
else negateSnd (posDivAlg (~a,~b));
\end{verbatim}
-*}
-
-
-subsubsection {* Uniqueness and Monotonicity of Quotients and Remainders *}
+\<close>
+
+
+subsubsection \<open>Uniqueness and Monotonicity of Quotients and Remainders\<close>
lemma unique_quotient_lemma:
"[| b*q' + r' \<le> b*q + r; 0 \<le> r'; r' < b; r < b |]
@@ -1750,9 +1750,9 @@
done
-subsubsection {* Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends *}
-
-text{*And positive divisors*}
+subsubsection \<open>Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends\<close>
+
+text\<open>And positive divisors\<close>
lemma adjust_eq [simp]:
"adjust b (q, r) =
@@ -1763,13 +1763,13 @@
declare posDivAlg.simps [simp del]
-text{*use with a simproc to avoid repeatedly proving the premise*}
+text\<open>use with a simproc to avoid repeatedly proving the premise\<close>
lemma posDivAlg_eqn:
"0 < b ==>
posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))"
by (rule posDivAlg.simps [THEN trans], simp)
-text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
+text\<open>Correctness of @{term posDivAlg}: it computes quotients correctly\<close>
theorem posDivAlg_correct:
assumes "0 \<le> a" and "0 < b"
shows "divmod_int_rel a b (posDivAlg a b)"
@@ -1785,13 +1785,13 @@
done
-subsubsection {* Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends *}
-
-text{*And positive divisors*}
+subsubsection \<open>Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends\<close>
+
+text\<open>And positive divisors\<close>
declare negDivAlg.simps [simp del]
-text{*use with a simproc to avoid repeatedly proving the premise*}
+text\<open>use with a simproc to avoid repeatedly proving the premise\<close>
lemma negDivAlg_eqn:
"0 < b ==>
negDivAlg a b =
@@ -1815,7 +1815,7 @@
done
-subsubsection {* Existence Shown by Proving the Division Algorithm to be Correct *}
+subsubsection \<open>Existence Shown by Proving the Division Algorithm to be Correct\<close>
(*the case a=0*)
lemma divmod_int_rel_0: "divmod_int_rel 0 b (0, 0)"
@@ -1916,14 +1916,14 @@
end
-text{*Basic laws about division and remainder*}
+text\<open>Basic laws about division and remainder\<close>
lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)"
by (fact mod_div_equality2 [symmetric])
-text {* Tool setup *}
-
-ML {*
+text \<open>Tool setup\<close>
+
+ML \<open>
structure Cancel_Div_Mod_Int = Cancel_Div_Mod
(
val div_name = @{const_name Rings.divide};
@@ -1937,9 +1937,9 @@
val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
(@{thm diff_conv_add_uminus} :: @{thms add_0_left add_0_right} @ @{thms ac_simps}))
)
-*}
-
-simproc_setup cancel_div_mod_int ("(k::int) + l") = {* K Cancel_Div_Mod_Int.proc *}
+\<close>
+
+simproc_setup cancel_div_mod_int ("(k::int) + l") = \<open>K Cancel_Div_Mod_Int.proc\<close>
lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b"
using divmod_int_correct [of a b]
@@ -1956,7 +1956,7 @@
and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2]
-subsubsection {* General Properties of div and mod *}
+subsubsection \<open>General Properties of div and mod\<close>
lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0"
apply (rule div_int_unique)
@@ -1990,10 +1990,10 @@
apply (auto simp add: divmod_int_rel_def)
done
-text{*There is no @{text mod_neg_pos_trivial}.*}
-
-
-subsubsection {* Laws for div and mod with Unary Minus *}
+text\<open>There is no @{text mod_neg_pos_trivial}.\<close>
+
+
+subsubsection \<open>Laws for div and mod with Unary Minus\<close>
lemma zminus1_lemma:
"divmod_int_rel a b (q, r) ==> b \<noteq> 0
@@ -2035,7 +2035,7 @@
unfolding zmod_zminus2_eq_if by auto
-subsubsection {* Computation of Division and Remainder *}
+subsubsection \<open>Computation of Division and Remainder\<close>
lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1"
by (simp add: div_int_def divmod_int_def)
@@ -2043,7 +2043,7 @@
lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1"
by (simp add: mod_int_def divmod_int_def)
-text{*a positive, b positive *}
+text\<open>a positive, b positive\<close>
lemma div_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a div b = fst (posDivAlg a b)"
by (simp add: div_int_def divmod_int_def)
@@ -2051,7 +2051,7 @@
lemma mod_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a mod b = snd (posDivAlg a b)"
by (simp add: mod_int_def divmod_int_def)
-text{*a negative, b positive *}
+text\<open>a negative, b positive\<close>
lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)"
by (simp add: div_int_def divmod_int_def)
@@ -2059,7 +2059,7 @@
lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)"
by (simp add: mod_int_def divmod_int_def)
-text{*a positive, b negative *}
+text\<open>a positive, b negative\<close>
lemma div_pos_neg:
"[| 0 < a; b < 0 |] ==> a div b = fst (apsnd uminus (negDivAlg (-a) (-b)))"
@@ -2069,7 +2069,7 @@
"[| 0 < a; b < 0 |] ==> a mod b = snd (apsnd uminus (negDivAlg (-a) (-b)))"
by (simp add: mod_int_def divmod_int_def)
-text{*a negative, b negative *}
+text\<open>a negative, b negative\<close>
lemma div_neg_neg:
"[| a < 0; b \<le> 0 |] ==> a div b = fst (apsnd uminus (posDivAlg (-a) (-b)))"
@@ -2079,7 +2079,7 @@
"[| a < 0; b \<le> 0 |] ==> a mod b = snd (apsnd uminus (posDivAlg (-a) (-b)))"
by (simp add: mod_int_def divmod_int_def)
-text {*Simplify expresions in which div and mod combine numerical constants*}
+text \<open>Simplify expresions in which div and mod combine numerical constants\<close>
lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q"
by (rule div_int_unique [of a b q r]) (simp add: divmod_int_rel_def)
@@ -2096,12 +2096,12 @@
by (rule mod_int_unique [of a b q r],
simp add: divmod_int_rel_def)
-text {*
+text \<open>
numeral simprocs -- high chance that these can be replaced
by divmod algorithm from @{class semiring_numeral_div}
-*}
-
-ML {*
+\<close>
+
+ML \<open>
local
val mk_number = HOLogic.mk_number HOLogic.intT
val plus = @{term "plus :: int \<Rightarrow> int \<Rightarrow> int"}
@@ -2133,21 +2133,21 @@
else (le $ r $ zero, less $ u $ r, negrule RS eq_reflection)
in SOME (rule OF map (prove ctxt) [goal1, goal2, goal3]) end)
end
-*}
+\<close>
simproc_setup binary_int_div
("numeral m div numeral n :: int" |
"numeral m div - numeral n :: int" |
"- numeral m div numeral n :: int" |
"- numeral m div - numeral n :: int") =
- {* K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq}) *}
+ \<open>K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq})\<close>
simproc_setup binary_int_mod
("numeral m mod numeral n :: int" |
"numeral m mod - numeral n :: int" |
"- numeral m mod numeral n :: int" |
"- numeral m mod - numeral n :: int") =
- {* K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq}) *}
+ \<open>K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq})\<close>
lemmas posDivAlg_eqn_numeral [simp] =
posDivAlg_eqn [of "numeral v" "numeral w", OF zero_less_numeral] for v w
@@ -2156,7 +2156,7 @@
negDivAlg_eqn [of "numeral v" "- numeral w", OF zero_less_numeral] for v w
-text {* Special-case simplification: @{text "\<plusminus>1 div z"} and @{text "\<plusminus>1 mod z"} *}
+text \<open>Special-case simplification: @{text "\<plusminus>1 div z"} and @{text "\<plusminus>1 mod z"}\<close>
lemma [simp]:
shows div_one_bit0: "1 div numeral (Num.Bit0 v) = (0 :: int)"
@@ -2178,7 +2178,7 @@
by (simp_all add: div_eq_minus1 zmod_minus1)
-subsubsection {* Monotonicity in the First Argument (Dividend) *}
+subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close>
lemma zdiv_mono1: "[| a \<le> a'; 0 < (b::int) |] ==> a div b \<le> a' div b"
apply (cut_tac a = a and b = b in zmod_zdiv_equality)
@@ -2197,7 +2197,7 @@
done
-subsubsection {* Monotonicity in the Second Argument (Divisor) *}
+subsubsection \<open>Monotonicity in the Second Argument (Divisor)\<close>
lemma q_pos_lemma:
"[| 0 \<le> b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \<le> (q'::int)"
@@ -2259,9 +2259,9 @@
done
-subsubsection {* More Algebraic Laws for div and mod *}
-
-text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
+subsubsection \<open>More Algebraic Laws for div and mod\<close>
+
+text\<open>proving (a*b) div c = a * (b div c) + a * (b mod c)\<close>
lemma zmult1_lemma:
"[| divmod_int_rel b c (q, r) |]
@@ -2273,7 +2273,7 @@
apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN div_int_unique])
done
-text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
+text\<open>proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c)\<close>
lemma zadd1_lemma:
"[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br) |]
@@ -2325,13 +2325,13 @@
using mod_div_equality [of m n] by arith
-subsubsection {* Proving @{term "a div (b * c) = (a div b) div c"} *}
+subsubsection \<open>Proving @{term "a div (b * c) = (a div b) div c"}\<close>
(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but
7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems
to cause particular problems.*)
-text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
+text\<open>first, four lemmas to bound the remainder for the cases b<0 and b>0\<close>
lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * c < b * (q mod c) + r"
apply (subgoal_tac "b * (c - q mod c) < r * 1")
@@ -2409,9 +2409,9 @@
qed
-subsubsection {* Splitting Rules for div and mod *}
-
-text{*The proofs of the two lemmas below are essentially identical*}
+subsubsection \<open>Splitting Rules for div and mod\<close>
+
+text\<open>The proofs of the two lemmas below are essentially identical\<close>
lemma split_pos_lemma:
"0<k ==>
@@ -2421,7 +2421,7 @@
apply (subst mod_add_eq)
apply (subst zdiv_zadd1_eq)
apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial)
-txt{*converse direction*}
+txt\<open>converse direction\<close>
apply (drule_tac x = "n div k" in spec)
apply (drule_tac x = "n mod k" in spec, simp)
done
@@ -2434,7 +2434,7 @@
apply (subst mod_add_eq)
apply (subst zdiv_zadd1_eq)
apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial)
-txt{*converse direction*}
+txt\<open>converse direction\<close>
apply (drule_tac x = "n div k" in spec)
apply (drule_tac x = "n mod k" in spec, simp)
done
@@ -2463,14 +2463,14 @@
split_neg_lemma [of concl: "%x y. P y"])
done
-text {* Enable (lin)arith to deal with @{const divide} and @{const mod}
+text \<open>Enable (lin)arith to deal with @{const divide} and @{const mod}
when these are applied to some constant that is of the form
- @{term "numeral k"}: *}
+ @{term "numeral k"}:\<close>
declare split_zdiv [of _ _ "numeral k", arith_split] for k
declare split_zmod [of _ _ "numeral k", arith_split] for k
-subsubsection {* Computing @{text "div"} and @{text "mod"} with shifting *}
+subsubsection \<open>Computing @{text "div"} and @{text "mod"} with shifting\<close>
lemma pos_divmod_int_rel_mult_2:
assumes "0 \<le> b"
@@ -2478,7 +2478,7 @@
shows "divmod_int_rel (1 + 2*a) (2*b) (q, 1 + 2*r)"
using assms unfolding divmod_int_rel_def by auto
-declaration {* K (Lin_Arith.add_simps @{thms uminus_numeral_One}) *}
+declaration \<open>K (Lin_Arith.add_simps @{thms uminus_numeral_One})\<close>
lemma neg_divmod_int_rel_mult_2:
assumes "b \<le> 0"
@@ -2486,7 +2486,7 @@
shows "divmod_int_rel (1 + 2*a) (2*b) (q, 2*r - 1)"
using assms unfolding divmod_int_rel_def by auto
-text{*computing div by shifting *}
+text\<open>computing div by shifting\<close>
lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"
using pos_divmod_int_rel_mult_2 [OF _ divmod_int_rel_div_mod]
@@ -2544,14 +2544,14 @@
proof
assume ?L
have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp
- with `?L` show ?R by blast
+ with \<open>?L\<close> show ?R by blast
next
assume ?R thus ?L
by(auto simp: div_pos_pos_trivial div_neg_neg_trivial)
qed
-subsubsection {* Quotients of Signs *}
+subsubsection \<open>Quotients of Signs\<close>
lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0"
apply (subgoal_tac "a div b \<le> -1", force)
@@ -2566,9 +2566,9 @@
lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0"
by (drule zdiv_mono1, auto)
-text{* Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
+text\<open>Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
conditional upon the sign of @{text a} or @{text b}. There are many more.
-They should all be simp rules unless that causes too much search. *}
+They should all be simp rules unless that causes too much search.\<close>
lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)"
apply auto
@@ -2613,14 +2613,14 @@
done
-subsubsection {* The Divides Relation *}
+subsubsection \<open>The Divides Relation\<close>
lemma dvd_eq_mod_eq_0_numeral:
"numeral x dvd (numeral y :: 'a) \<longleftrightarrow> numeral y mod numeral x = (0 :: 'a::semiring_div)"
by (fact dvd_eq_mod_eq_0)
-subsubsection {* Further properties *}
+subsubsection \<open>Further properties\<close>
lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
using zmod_zdiv_equality[where a="m" and b="n"]
@@ -2650,7 +2650,7 @@
lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y"
by (unfold dvd_def, cases "y=0", auto simp add: abs_mult)
-text{*Suggested by Matthias Daum*}
+text\<open>Suggested by Matthias Daum\<close>
lemma int_power_div_base:
"\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)")
@@ -2659,7 +2659,7 @@
apply simp_all
done
-text {* by Brian Huffman *}
+text \<open>by Brian Huffman\<close>
lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m"
by (rule mod_minus_eq [symmetric])
@@ -2677,7 +2677,7 @@
power_mod
zminus_zmod zdiff_zmod_left zdiff_zmod_right
-text {* Distributive laws for function @{text nat}. *}
+text \<open>Distributive laws for function @{text nat}.\<close>
lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y"
apply (rule linorder_cases [of y 0])
@@ -2693,7 +2693,7 @@
apply (simp add: nat_eq_iff zmod_int)
done
-text {* transfer setup *}
+text \<open>transfer setup\<close>
lemma transfer_nat_int_functions:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
@@ -2729,7 +2729,7 @@
transfer_int_nat_function_closures
]
-text{*Suggested by Matthias Daum*}
+text\<open>Suggested by Matthias Daum\<close>
lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"
apply (subgoal_tac "nat x div nat k < nat x")
apply (simp add: nat_div_distrib [symmetric])
@@ -2781,11 +2781,11 @@
thus ?lhs by simp
qed
-text {*
+text \<open>
This re-embedding of natural division on integers goes back to the
time when numerals had been signed numerals. It should
now be replaced by the algorithm developed in @{class semiring_numeral_div}.
-*}
+\<close>
lemma div_nat_numeral [simp]:
"(numeral v :: nat) div numeral v' = nat (numeral v div numeral v')"
@@ -2804,14 +2804,14 @@
by (subst nat_mod_distrib) simp_all
-subsubsection {* Tools setup *}
-
-text {* Nitpick *}
+subsubsection \<open>Tools setup\<close>
+
+text \<open>Nitpick\<close>
lemmas [nitpick_unfold] = dvd_eq_mod_eq_0 mod_div_equality' zmod_zdiv_equality'
-subsubsection {* Code generation *}
+subsubsection \<open>Code generation\<close>
definition divmod_abs :: "int \<Rightarrow> int \<Rightarrow> int \<times> int"
where
--- a/src/HOL/Enum.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Enum.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,12 +1,12 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Finite types as explicit enumerations *}
+section \<open>Finite types as explicit enumerations\<close>
theory Enum
imports Map Groups_List
begin
-subsection {* Class @{text enum} *}
+subsection \<open>Class @{text enum}\<close>
class enum =
fixes enum :: "'a list"
@@ -16,7 +16,7 @@
and enum_distinct: "distinct enum"
assumes enum_all_UNIV: "enum_all P \<longleftrightarrow> Ball UNIV P"
assumes enum_ex_UNIV: "enum_ex P \<longleftrightarrow> Bex UNIV P"
- -- {* tailored towards simple instantiation *}
+ -- \<open>tailored towards simple instantiation\<close>
begin
subclass finite proof
@@ -52,9 +52,9 @@
end
-subsection {* Implementations using @{class enum} *}
+subsection \<open>Implementations using @{class enum}\<close>
-subsubsection {* Unbounded operations and quantifiers *}
+subsubsection \<open>Unbounded operations and quantifiers\<close>
lemma Collect_code [code]:
"Collect P = set (filter P enum)"
@@ -82,7 +82,7 @@
by (auto simp add: list_ex1_iff enum_UNIV)
-subsubsection {* An executable choice operator *}
+subsubsection \<open>An executable choice operator\<close>
definition
[code del]: "enum_the = The"
@@ -106,7 +106,7 @@
and "\<forall> x \<in> set vs. \<not> P x"
and "P a"
by (auto simp add: filter_eq_Cons_iff) (simp only: filter_empty_conv[symmetric])
- with `P x` in_enum[of x, unfolded enum_eq] `x \<noteq> a` show "False" by auto
+ with \<open>P x\<close> in_enum[of x, unfolded enum_eq] \<open>x \<noteq> a\<close> show "False" by auto
qed
next
from filter_enum show "P a" by (auto simp add: filter_eq_Cons_iff)
@@ -122,7 +122,7 @@
constant enum_the \<rightharpoonup> (Eval) "(fn '_ => raise Match)"
-subsubsection {* Equality and order on functions *}
+subsubsection \<open>Equality and order on functions\<close>
instantiation "fun" :: (enum, equal) equal
begin
@@ -150,7 +150,7 @@
by (simp_all add: fun_eq_iff le_fun_def order_less_le)
-subsubsection {* Operations on relations *}
+subsubsection \<open>Operations on relations\<close>
lemma [code]:
"Id = image (\<lambda>x. (x, x)) (set Enum.enum)"
@@ -177,7 +177,7 @@
by (auto simp add: mlex_prod_def)
-subsubsection {* Bounded accessible part *}
+subsubsection \<open>Bounded accessible part\<close>
primrec bacc :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> 'a set"
where
@@ -219,7 +219,7 @@
fix y assume y: "(y, x) : r"
with n have "y : bacc r (n y)" by auto
moreover have "n y <= Max ((%(y, x). n y) ` r)"
- using y `finite r` by (auto intro!: Max_ge)
+ 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
qed
@@ -241,7 +241,7 @@
by (simp add: card_UNIV_def acc_bacc_eq)
-subsection {* Default instances for @{class enum} *}
+subsection \<open>Default instances for @{class enum}\<close>
lemma map_of_zip_enum_is_Some:
assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
@@ -329,7 +329,7 @@
fix f :: "'a \<Rightarrow> 'b"
have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
- from `enum_all P` have "P (the \<circ> map_of (zip enum (map f enum)))"
+ from \<open>enum_all P\<close> have "P (the \<circ> map_of (zip enum (map f enum)))"
unfolding enum_all_fun_def all_n_lists_def
apply (simp add: set_n_lists)
apply (erule_tac x="map f enum" in allE)
@@ -354,7 +354,7 @@
from this obtain f where "P f" ..
have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
- from `P f` this have "P (the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum)))"
+ from \<open>P f\<close> this have "P (the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum)))"
by auto
from this show "enum_ex P"
unfolding enum_ex_fun_def ex_n_lists_def
@@ -489,9 +489,9 @@
end
-subsection {* Small finite types *}
+subsection \<open>Small finite types\<close>
-text {* We define small finite types for use in Quickcheck *}
+text \<open>We define small finite types for use in Quickcheck\<close>
datatype (plugins only: code "quickcheck" extraction) finite_1 =
a\<^sub>1
@@ -562,12 +562,12 @@
lemma finite_1_eq: "x = a\<^sub>1"
by(cases x) simp
-simproc_setup finite_1_eq ("x::finite_1") = {*
+simproc_setup finite_1_eq ("x::finite_1") = \<open>
fn _ => fn _ => fn ct =>
(case Thm.term_of ct of
Const (@{const_name a\<^sub>1}, _) => NONE
| _ => SOME (mk_meta_eq @{thm finite_1_eq}))
-*}
+\<close>
instantiation finite_1 :: complete_boolean_algebra
begin
@@ -864,8 +864,8 @@
instantiation finite_4 :: complete_lattice begin
-text {* @{term a\<^sub>1} $<$ @{term a\<^sub>2},@{term a\<^sub>3} $<$ @{term a\<^sub>4},
- but @{term a\<^sub>2} and @{term a\<^sub>3} are incomparable. *}
+text \<open>@{term a\<^sub>1} $<$ @{term a\<^sub>2},@{term a\<^sub>3} $<$ @{term a\<^sub>4},
+ but @{term a\<^sub>2} and @{term a\<^sub>3} are incomparable.\<close>
definition
"x < y \<longleftrightarrow> (case (x, y) of
@@ -968,7 +968,7 @@
instantiation finite_5 :: complete_lattice
begin
-text {* The non-distributive pentagon lattice $N_5$ *}
+text \<open>The non-distributive pentagon lattice $N_5$\<close>
definition
"x < y \<longleftrightarrow> (case (x, y) of
@@ -1034,7 +1034,7 @@
hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5
-subsection {* Closing up *}
+subsection \<open>Closing up\<close>
hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5
hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl
--- a/src/HOL/Equiv_Relations.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Equiv_Relations.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Copyright 1996 University of Cambridge
*)
-section {* Equivalence Relations in Higher-Order Set Theory *}
+section \<open>Equivalence Relations in Higher-Order Set Theory\<close>
theory Equiv_Relations
imports Groups_Big Relation
begin
-subsection {* Equivalence relations -- set version *}
+subsection \<open>Equivalence relations -- set version\<close>
definition equiv :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where
"equiv A r \<longleftrightarrow> refl_on A r \<and> sym r \<and> trans r"
@@ -22,12 +22,12 @@
obtains "refl_on A r" and "sym r" and "trans r"
using assms by (simp add: equiv_def)
-text {*
+text \<open>
Suppes, Theorem 70: @{text r} is an equiv relation iff @{text "r\<inverse> O
r = r"}.
First half: @{text "equiv A r ==> r\<inverse> O r = r"}.
-*}
+\<close>
lemma sym_trans_comp_subset:
"sym r ==> trans r ==> r\<inverse> O r \<subseteq> r"
@@ -43,7 +43,7 @@
apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+
done
-text {* Second half. *}
+text \<open>Second half.\<close>
lemma comp_equivI:
"r\<inverse> O r = r ==> Domain r = A ==> equiv A r"
@@ -55,11 +55,11 @@
done
-subsection {* Equivalence classes *}
+subsection \<open>Equivalence classes\<close>
lemma equiv_class_subset:
"equiv A r ==> (a, b) \<in> r ==> r``{a} \<subseteq> r``{b}"
- -- {* lemma for the next result *}
+ -- \<open>lemma for the next result\<close>
by (unfold equiv_def trans_def sym_def) blast
theorem equiv_class_eq: "equiv A r ==> (a, b) \<in> r ==> r``{a} = r``{b}"
@@ -73,7 +73,7 @@
lemma subset_equiv_class:
"equiv A r ==> r``{b} \<subseteq> r``{a} ==> b \<in> A ==> (a,b) \<in> r"
- -- {* lemma for the next result *}
+ -- \<open>lemma for the next result\<close>
by (unfold equiv_def refl_on_def) blast
lemma eq_equiv_class:
@@ -96,10 +96,10 @@
by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type)
-subsection {* Quotients *}
+subsection \<open>Quotients\<close>
definition quotient :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set set" (infixl "'/'/" 90) where
- "A//r = (\<Union>x \<in> A. {r``{x}})" -- {* set of equiv classes *}
+ "A//r = (\<Union>x \<in> A. {r``{x}})" -- \<open>set of equiv classes\<close>
lemma quotientI: "x \<in> A ==> r``{x} \<in> A//r"
by (unfold quotient_def) blast
@@ -160,7 +160,7 @@
apply blast
done
-subsection {* Refinement of one equivalence relation WRT another *}
+subsection \<open>Refinement of one equivalence relation WRT another\<close>
lemma refines_equiv_class_eq:
"\<lbrakk>R \<subseteq> S; equiv A R; equiv A S\<rbrakk> \<Longrightarrow> R``(S``{a}) = S``{a}"
@@ -187,9 +187,9 @@
done
-subsection {* Defining unary operations upon equivalence classes *}
+subsection \<open>Defining unary operations upon equivalence classes\<close>
-text{*A congruence-preserving function*}
+text\<open>A congruence-preserving function\<close>
definition congruent :: "('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
"congruent r f \<longleftrightarrow> (\<forall>(y, z) \<in> r. f y = f z)"
@@ -209,13 +209,13 @@
lemma UN_constant_eq: "a \<in> A ==> \<forall>y \<in> A. f y = c ==> (\<Union>y \<in> A. f(y))=c"
- -- {* lemma required to prove @{text UN_equiv_class} *}
+ -- \<open>lemma required to prove @{text UN_equiv_class}\<close>
by auto
lemma UN_equiv_class:
"equiv A r ==> f respects r ==> a \<in> A
==> (\<Union>x \<in> r``{a}. f x) = f a"
- -- {* Conversion rule *}
+ -- \<open>Conversion rule\<close>
apply (rule equiv_class_self [THEN UN_constant_eq], assumption+)
apply (unfold equiv_def congruent_def sym_def)
apply (blast del: equalityI)
@@ -230,11 +230,11 @@
apply auto
done
-text {*
+text \<open>
Sufficient conditions for injectiveness. Could weaken premises!
major premise could be an inclusion; bcong could be @{text "!!y. y \<in>
A ==> f y \<in> B"}.
-*}
+\<close>
lemma UN_equiv_class_inject:
"equiv A r ==> f respects r ==>
@@ -252,9 +252,9 @@
done
-subsection {* Defining binary operations upon equivalence classes *}
+subsection \<open>Defining binary operations upon equivalence classes\<close>
-text{*A congruence-preserving function of two arguments*}
+text\<open>A congruence-preserving function of two arguments\<close>
definition congruent2 :: "('a \<times> 'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> bool" where
"congruent2 r1 r2 f \<longleftrightarrow> (\<forall>(y1, z1) \<in> r1. \<forall>(y2, z2) \<in> r2. f y1 y2 = f z1 z2)"
@@ -268,7 +268,7 @@
"congruent2 r1 r2 f \<Longrightarrow> (y1, z1) \<in> r1 \<Longrightarrow> (y2, z2) \<in> r2 \<Longrightarrow> f y1 y2 = f z1 z2"
using assms by (auto simp add: congruent2_def)
-text{*Abbreviation for the common case where the relations are identical*}
+text\<open>Abbreviation for the common case where the relations are identical\<close>
abbreviation
RESPECTS2:: "['a => 'a => 'b, ('a * 'a) set] => bool"
(infixr "respects2" 80) where
@@ -310,8 +310,8 @@
lemma UN_UN_split_split_eq:
"(\<Union>(x1, x2) \<in> X. \<Union>(y1, y2) \<in> Y. A x1 x2 y1 y2) =
(\<Union>x \<in> X. \<Union>y \<in> Y. (\<lambda>(x1, x2). (\<lambda>(y1, y2). A x1 x2 y1 y2) y) x)"
- -- {* Allows a natural expression of binary operators, *}
- -- {* without explicit calls to @{text split} *}
+ -- \<open>Allows a natural expression of binary operators,\<close>
+ -- \<open>without explicit calls to @{text split}\<close>
by auto
lemma congruent2I:
@@ -319,8 +319,8 @@
==> (!!y z w. w \<in> A2 ==> (y,z) \<in> r1 ==> f y w = f z w)
==> (!!y z w. w \<in> A1 ==> (y,z) \<in> r2 ==> f w y = f w z)
==> congruent2 r1 r2 f"
- -- {* Suggested by John Harrison -- the two subproofs may be *}
- -- {* \emph{much} simpler than the direct proof. *}
+ -- \<open>Suggested by John Harrison -- the two subproofs may be\<close>
+ -- \<open>\emph{much} simpler than the direct proof.\<close>
apply (unfold congruent2_def equiv_def refl_on_def)
apply clarify
apply (blast intro: trans)
@@ -340,12 +340,12 @@
done
-subsection {* Quotients and finiteness *}
+subsection \<open>Quotients and finiteness\<close>
-text {*Suggested by Florian Kammüller*}
+text \<open>Suggested by Florian Kammüller\<close>
lemma finite_quotient: "finite A ==> r \<subseteq> A \<times> A ==> finite (A//r)"
- -- {* recall @{thm equiv_type} *}
+ -- \<open>recall @{thm equiv_type}\<close>
apply (rule finite_subset)
apply (erule_tac [2] finite_Pow_iff [THEN iffD2])
apply (unfold quotient_def)
@@ -381,7 +381,7 @@
done
-subsection {* Projection *}
+subsection \<open>Projection\<close>
definition proj where "proj r x = r `` {x}"
@@ -429,13 +429,13 @@
using assms in_quotient_imp_in_rel equiv_type by fastforce
-subsection {* Equivalence relations -- predicate version *}
+subsection \<open>Equivalence relations -- predicate version\<close>
-text {* Partial equivalences *}
+text \<open>Partial equivalences\<close>
definition part_equivp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
"part_equivp R \<longleftrightarrow> (\<exists>x. R x x) \<and> (\<forall>x y. R x y \<longleftrightarrow> R x x \<and> R y y \<and> R x = R y)"
- -- {* John-Harrison-style characterization *}
+ -- \<open>John-Harrison-style characterization\<close>
lemma part_equivpI:
"(\<exists>x. R x x) \<Longrightarrow> symp R \<Longrightarrow> transp R \<Longrightarrow> part_equivp R"
@@ -481,10 +481,10 @@
by (auto elim: part_equivpE)
-text {* Total equivalences *}
+text \<open>Total equivalences\<close>
definition equivp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
- "equivp R \<longleftrightarrow> (\<forall>x y. R x y = (R x = R y))" -- {* John-Harrison-style characterization *}
+ "equivp R \<longleftrightarrow> (\<forall>x y. R x y = (R x = R y))" -- \<open>John-Harrison-style characterization\<close>
lemma equivpI:
"reflp R \<Longrightarrow> symp R \<Longrightarrow> transp R \<Longrightarrow> equivp R"
--- a/src/HOL/Extraction.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Extraction.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Stefan Berghofer, TU Muenchen
*)
-section {* Program extraction for HOL *}
+section \<open>Program extraction for HOL\<close>
theory Extraction
imports Option
@@ -10,9 +10,9 @@
ML_file "Tools/rewrite_hol_proof.ML"
-subsection {* Setup *}
+subsection \<open>Setup\<close>
-setup {*
+setup \<open>
Extraction.add_types
[("bool", ([], NONE))] #>
Extraction.set_preprocessor (fn thy =>
@@ -22,7 +22,7 @@
(RewriteHOLProof.rews,
ProofRewriteRules.rprocs true @ [ProofRewriteRules.expand_of_class thy]) o
ProofRewriteRules.elim_vars (curry Const @{const_name default}))
-*}
+\<close>
lemmas [extraction_expand] =
meta_spec atomize_eq atomize_all atomize_imp atomize_conj
@@ -40,7 +40,7 @@
datatype (plugins only: code extraction) sumbool = Left | Right
-subsection {* Type of extracted program *}
+subsection \<open>Type of extracted program\<close>
extract_type
"typeof (Trueprop P) \<equiv> typeof P"
@@ -90,7 +90,7 @@
"typeof (x \<in> P) \<equiv> typeof P"
-subsection {* Realizability *}
+subsection \<open>Realizability\<close>
realizability
"(realizes t (Trueprop P)) \<equiv> (Trueprop (realizes t P))"
@@ -150,7 +150,7 @@
"(realizes t (P = Q)) \<equiv> (realizes t ((P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)))"
-subsection {* Computational content of basic inference rules *}
+subsection \<open>Computational content of basic inference rules\<close>
theorem disjE_realizer:
assumes r: "case x of Inl p \<Rightarrow> P p | Inr q \<Rightarrow> Q q"
--- a/src/HOL/Fields.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Fields.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,17 +7,17 @@
Author: Jeremy Avigad
*)
-section {* Fields *}
+section \<open>Fields\<close>
theory Fields
imports Rings
begin
-subsection {* Division rings *}
+subsection \<open>Division rings\<close>
-text {*
+text \<open>
A division ring is like a field, but without the commutativity requirement.
-*}
+\<close>
class inverse = divide +
fixes inverse :: "'a \<Rightarrow> 'a"
@@ -29,7 +29,7 @@
end
-text{* Lemmas @{text divide_simps} move division to the outside and eliminates them on (in)equalities. *}
+text\<open>Lemmas @{text divide_simps} move division to the outside and eliminates them on (in)equalities.\<close>
named_theorems divide_simps "rewrite rules to eliminate divisions"
@@ -94,9 +94,9 @@
assumes "inverse a = inverse b" and "a \<noteq> 0" and "b \<noteq> 0"
shows "a = b"
proof -
- from `inverse a = inverse b`
+ from \<open>inverse a = inverse b\<close>
have "inverse (inverse a) = inverse (inverse b)" by (rule arg_cong)
- with `a \<noteq> 0` and `b \<noteq> 0` show "a = b"
+ with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> show "a = b"
by (simp add: nonzero_inverse_inverse_eq)
qed
@@ -272,7 +272,7 @@
end
-subsection {* Fields *}
+subsection \<open>Fields\<close>
class field = comm_ring_1 + inverse +
assumes field_inverse: "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
@@ -306,7 +306,7 @@
by (simp add: divide_inverse)
qed
-text{*There is no slick version using division by zero.*}
+text\<open>There is no slick version using division by zero.\<close>
lemma inverse_add:
"a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a + inverse b = (a + b) * inverse a * inverse b"
by (simp add: division_ring_inverse_add ac_simps)
@@ -345,7 +345,7 @@
by (simp only: mult.commute)
qed
-text{*Special Cancellation Simprules for Division*}
+text\<open>Special Cancellation Simprules for Division\<close>
lemma nonzero_divide_mult_cancel_right [simp]:
"b \<noteq> 0 \<Longrightarrow> b / (a * b) = 1 / a"
@@ -374,8 +374,8 @@
lemma divide_minus1 [simp]: "x / - 1 = - x"
using nonzero_minus_divide_right [of "1" x] by simp
-text{*This version builds in division by zero while also re-orienting
- the right-hand side.*}
+text\<open>This version builds in division by zero while also re-orienting
+ the right-hand side.\<close>
lemma inverse_mult_distrib [simp]:
"inverse (a * b) = inverse a * inverse b"
proof cases
@@ -391,11 +391,11 @@
by (simp add: divide_inverse mult.commute)
-text {* Calculations with fractions *}
+text \<open>Calculations with fractions\<close>
-text{* There is a whole bunch of simp-rules just for class @{text
+text\<open>There is a whole bunch of simp-rules just for class @{text
field} but none for class @{text field} and @{text nonzero_divides}
-because the latter are covered by a simproc. *}
+because the latter are covered by a simproc.\<close>
lemma mult_divide_mult_cancel_left:
"c \<noteq> 0 \<Longrightarrow> (c * a) / (c * b) = a / b"
@@ -421,14 +421,14 @@
"(x / y) / (z / w) = (x * w) / (y * z)"
by simp
-text {*Special Cancellation Simprules for Division*}
+text \<open>Special Cancellation Simprules for Division\<close>
lemma mult_divide_mult_cancel_left_if [simp]:
shows "(c * a) / (c * b) = (if c = 0 then 0 else a / b)"
by simp
-text {* Division and Unary Minus *}
+text \<open>Division and Unary Minus\<close>
lemma minus_divide_right:
"- (a / b) = a / - b"
@@ -489,7 +489,7 @@
end
-subsection {* Ordered fields *}
+subsection \<open>Ordered fields\<close>
class linordered_field = field + linordered_idom
begin
@@ -579,7 +579,7 @@
apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq)
done
-text{*Both premises are essential. Consider -1 and 1.*}
+text\<open>Both premises are essential. Consider -1 and 1.\<close>
lemma inverse_less_iff_less [simp]:
"0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less)
@@ -593,8 +593,8 @@
by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le)
-text{*These results refer to both operands being negative. The opposite-sign
-case is trivial, since inverse preserves signs.*}
+text\<open>These results refer to both operands being negative. The opposite-sign
+case is trivial, since inverse preserves signs.\<close>
lemma inverse_le_imp_le_neg:
"inverse a \<le> inverse b \<Longrightarrow> b < 0 \<Longrightarrow> b \<le> a"
apply (rule classical)
@@ -733,8 +733,8 @@
finally show ?thesis .
qed
-text{* The following @{text field_simps} rules are necessary, as minus is always moved atop of
-division but we want to get rid of division. *}
+text\<open>The following @{text field_simps} rules are necessary, as minus is always moved atop of
+division but we want to get rid of division.\<close>
lemma pos_le_minus_divide_eq [field_simps]: "0 < c \<Longrightarrow> a \<le> - (b / c) \<longleftrightarrow> a * c \<le> - b"
unfolding minus_divide_left by (rule pos_le_divide_eq)
@@ -768,10 +768,10 @@
"y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y \<le> w / z \<longleftrightarrow> (x * z - w * y) / (y * z) \<le> 0"
by (subst le_iff_diff_le_0) (simp add: diff_frac_eq )
-text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
+text\<open>Lemmas @{text sign_simps} is a first attempt to automate proofs
of positivity/negativity needed for @{text field_simps}. Have not added @{text
sign_simps} to @{text field_simps} because the former can lead to case
-explosions. *}
+explosions.\<close>
lemmas sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff
@@ -835,8 +835,8 @@
apply (simp add: less_imp_not_eq nonzero_minus_divide_right [symmetric])
done
-text{*The last premise ensures that @{term a} and @{term b}
- have the same sign*}
+text\<open>The last premise ensures that @{term a} and @{term b}
+ have the same sign\<close>
lemma divide_strict_left_mono:
"[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / b"
by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono)
@@ -984,7 +984,7 @@
and minus_divide_less_eq: "- (b / c) < a \<longleftrightarrow> (if 0 < c then - b < a * c else if c < 0 then a * c < - b else 0 < a)"
by (auto simp: field_simps not_less dest: antisym)
-text {*Division and Signs*}
+text \<open>Division and Signs\<close>
lemma
shows zero_less_divide_iff: "0 < a / b \<longleftrightarrow> 0 < a \<and> 0 < b \<or> a < 0 \<and> b < 0"
@@ -993,9 +993,9 @@
and divide_le_0_iff: "a / b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
by (auto simp add: divide_simps)
-text {* Division and the Number One *}
+text \<open>Division and the Number One\<close>
-text{*Simplify expressions equated with 1*}
+text\<open>Simplify expressions equated with 1\<close>
lemma zero_eq_1_divide_iff [simp]: "0 = 1 / a \<longleftrightarrow> a = 0"
by (cases "a = 0") (auto simp: field_simps)
@@ -1003,7 +1003,7 @@
lemma one_divide_eq_0_iff [simp]: "1 / a = 0 \<longleftrightarrow> a = 0"
using zero_eq_1_divide_iff[of a] by simp
-text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
+text\<open>Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}\<close>
lemma zero_le_divide_1_iff [simp]:
"0 \<le> 1 / a \<longleftrightarrow> 0 \<le> a"
@@ -1050,7 +1050,7 @@
lemma divide_less_cancel: "a / c < b / c \<longleftrightarrow> (0 < c \<longrightarrow> a < b) \<and> (c < 0 \<longrightarrow> b < a) \<and> c \<noteq> 0"
by (auto simp add: divide_inverse mult_less_cancel_right)
-text{*Simplify quotients that are compared with the value 1.*}
+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))"
@@ -1084,7 +1084,7 @@
"x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x / y \<le> 0"
by (auto simp add: divide_simps)
-text {*Conditional Simplification Rules: No Case Splits*}
+text \<open>Conditional Simplification Rules: No Case Splits\<close>
lemma le_divide_eq_1_pos [simp]:
"0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
@@ -1158,11 +1158,11 @@
assume "0 < x"
thus ?thesis
using dense_le_bounded[of 0 1 "y/x"] *
- unfolding le_divide_eq if_P[OF `0 < x`] by simp
+ unfolding le_divide_eq if_P[OF \<open>0 < x\<close>] by simp
next
assume "\<not>0 < x" hence "x \<le> 0" by simp
obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1\<Colon>'a"] by auto
- hence "x \<le> s * x" using mult_le_cancel_right[of 1 x s] `x \<le> 0` by auto
+ hence "x \<le> s * x" using mult_le_cancel_right[of 1 x s] \<open>x \<le> 0\<close> by auto
also note *[OF s]
finally show ?thesis .
qed
--- a/src/HOL/Filter.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Filter.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,17 +3,17 @@
Author: Johannes Hölzl
*)
-section {* Filters on predicates *}
+section \<open>Filters on predicates\<close>
theory Filter
imports Set_Interval Lifting_Set
begin
-subsection {* Filters *}
+subsection \<open>Filters\<close>
-text {*
+text \<open>
This definition also allows non-proper filters.
-*}
+\<close>
locale is_filter =
fixes F :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
@@ -34,7 +34,7 @@
using assms by (simp add: Abs_filter_inverse)
-subsubsection {* Eventually *}
+subsubsection \<open>Eventually\<close>
definition eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a filter \<Rightarrow> bool"
where "eventually P F \<longleftrightarrow> Rep_filter F P"
@@ -237,7 +237,7 @@
frequently_all
frequently_imp_iff
-ML {*
+ML \<open>
fun eventually_elim_tac ctxt facts = SUBGOAL_CASES (fn (goal, i) =>
let
val mp_thms = facts RL @{thms eventually_rev_mp}
@@ -252,16 +252,16 @@
in
CASES cases (resolve_tac ctxt [raw_elim_thm] i)
end)
-*}
+\<close>
-method_setup eventually_elim = {*
+method_setup eventually_elim = \<open>
Scan.succeed (fn ctxt => METHOD_CASES (HEADGOAL o eventually_elim_tac ctxt))
-*} "elimination of eventually quantifiers"
+\<close> "elimination of eventually quantifiers"
-subsubsection {* Finer-than relation *}
+subsubsection \<open>Finer-than relation\<close>
-text {* @{term "F \<le> F'"} means that filter @{term F} is finer than
-filter @{term F'}. *}
+text \<open>@{term "F \<le> F'"} means that filter @{term F} is finer than
+filter @{term F'}.\<close>
instantiation filter :: (type) complete_lattice
begin
@@ -457,12 +457,12 @@
then have "\<exists>b\<in>B. \<forall>x\<in>X. b \<le> x"
proof induct
case empty then show ?case
- using `B \<noteq> {}` by auto
+ using \<open>B \<noteq> {}\<close> by auto
next
case (insert x X)
then obtain b where "b \<in> B" "\<And>x. x \<in> X \<Longrightarrow> b \<le> x"
by auto
- with `insert x X \<subseteq> B` base[of b x] show ?case
+ with \<open>insert x X \<subseteq> B\<close> base[of b x] show ?case
by (auto intro: order_trans)
qed
then obtain b where "b \<in> B" "b \<le> Inf X"
@@ -477,7 +477,7 @@
unfolding INF_def by (subst eventually_Inf_base) auto
-subsubsection {* Map function for filters *}
+subsubsection \<open>Map function for filters\<close>
definition filtermap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> 'b filter"
where "filtermap f F = Abs_filter (\<lambda>P. eventually (\<lambda>x. P (f x)) F)"
@@ -526,7 +526,7 @@
unfolding le_filter_def eventually_filtermap
by (subst (1 2) eventually_INF) auto
qed
-subsubsection {* Standard filters *}
+subsubsection \<open>Standard filters\<close>
definition principal :: "'a set \<Rightarrow> 'a filter" where
"principal S = Abs_filter (\<lambda>P. \<forall>x\<in>S. P x)"
@@ -576,7 +576,7 @@
lemma filtermap_principal[simp]: "filtermap f (principal A) = principal (f ` A)"
unfolding filter_eq_iff eventually_filtermap eventually_principal by simp
-subsubsection {* Order filters *}
+subsubsection \<open>Order filters\<close>
definition at_top :: "('a::order) filter"
where "at_top = (INF k. principal {k ..})"
@@ -648,7 +648,7 @@
unfolding trivial_limit_def
by (metis eventually_at_top_linorder order_refl)
-subsection {* Sequentially *}
+subsection \<open>Sequentially\<close>
abbreviation sequentially :: "nat filter"
where "sequentially \<equiv> at_top"
@@ -734,7 +734,7 @@
by (blast intro: finite_subset)
qed
-subsection {* Limits *}
+subsection \<open>Limits\<close>
definition filterlim :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b filter \<Rightarrow> 'a filter \<Rightarrow> bool" where
"filterlim f F2 F1 \<longleftrightarrow> filtermap f F1 \<le> F2"
@@ -828,7 +828,7 @@
fix i j assume "i \<in> I" "j \<in> I"
with chain[OF this] show "\<exists>x\<in>I. principal (F x) \<le> inf (principal (F i)) (principal (F j))"
by auto
-qed (auto simp: eventually_principal `I \<noteq> {}`)
+qed (auto simp: eventually_principal \<open>I \<noteq> {}\<close>)
lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\<lambda>x. f (g x)) F1 F2"
unfolding filterlim_def filtermap_filtermap ..
@@ -850,7 +850,7 @@
LIM x F. if P x then f x else g x :> G"
unfolding filterlim_iff eventually_inf_principal by (auto simp: eventually_conj_iff)
-subsection {* Limits to @{const at_top} and @{const at_bot} *}
+subsection \<open>Limits to @{const at_top} and @{const at_bot}\<close>
lemma filterlim_at_top:
fixes f :: "'a \<Rightarrow> ('b::linorder)"
@@ -890,7 +890,7 @@
have "eventually (\<lambda>x. g z \<le> x) at_top"
by (rule eventually_ge_at_top)
with Q show "eventually (\<lambda>x. z \<le> f x) at_top"
- by eventually_elim (metis mono bij `P z`)
+ by eventually_elim (metis mono bij \<open>P z\<close>)
qed
qed
@@ -937,7 +937,7 @@
by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans)
-subsection {* Setup @{typ "'a filter"} for lifting and transfer *}
+subsection \<open>Setup @{typ "'a filter"} for lifting and transfer\<close>
context begin interpretation lifting_syntax .
@@ -1011,7 +1011,7 @@
fix P' Q'
assume "G P'" "G Q'"
moreover
- from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def]
+ from bi_total_fun[OF \<open>bi_unique A\<close> bi_total_eq, unfolded bi_total_def]
obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast
have "F P = G P'" "F Q = G Q'" by transfer_prover+
ultimately have "F (\<lambda>x. P x \<and> Q x)" by(simp add: conj)
@@ -1021,7 +1021,7 @@
fix P' Q'
assume "\<forall>x. P' x \<longrightarrow> Q' x" "G P'"
moreover
- from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def]
+ from bi_total_fun[OF \<open>bi_unique A\<close> bi_total_eq, unfolded bi_total_def]
obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast
have "F P = G P'" by transfer_prover
moreover have "(\<forall>x. P x \<longrightarrow> Q x) \<longleftrightarrow> (\<forall>x. P' x \<longrightarrow> Q' x)" by transfer_prover
@@ -1046,7 +1046,7 @@
shows "left_total (rel_filter A)"
proof(rule left_totalI)
fix F :: "'a filter"
- from bi_total_fun[OF bi_unique_fun[OF `bi_total A` bi_unique_eq] bi_total_eq]
+ from bi_total_fun[OF bi_unique_fun[OF \<open>bi_total A\<close> bi_unique_eq] bi_total_eq]
obtain G where [transfer_rule]: "((A ===> op =) ===> op =) (\<lambda>P. eventually P F) G"
unfolding bi_total_def by blast
moreover have "is_filter (\<lambda>P. eventually P F) \<longleftrightarrow> is_filter G" by transfer_prover
--- a/src/HOL/Finite_Set.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Finite_Set.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,30 +3,30 @@
with contributions by Jeremy Avigad and Andrei Popescu
*)
-section {* Finite sets *}
+section \<open>Finite sets\<close>
theory Finite_Set
imports Product_Type Sum_Type Nat
begin
-subsection {* Predicate for finite sets *}
+subsection \<open>Predicate for finite sets\<close>
inductive finite :: "'a set \<Rightarrow> bool"
where
emptyI [simp, intro!]: "finite {}"
| insertI [simp, intro!]: "finite A \<Longrightarrow> finite (insert a A)"
-simproc_setup finite_Collect ("finite (Collect P)") = {* K Set_Comprehension_Pointfree.simproc *}
+simproc_setup finite_Collect ("finite (Collect P)") = \<open>K Set_Comprehension_Pointfree.simproc\<close>
declare [[simproc del: finite_Collect]]
lemma finite_induct [case_names empty insert, induct set: finite]:
- -- {* Discharging @{text "x \<notin> F"} entails extra work. *}
+ -- \<open>Discharging @{text "x \<notin> F"} entails extra work.\<close>
assumes "finite F"
assumes "P {}"
and insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
shows "P F"
-using `finite F`
+using \<open>finite F\<close>
proof induct
show "P {}" by fact
fix x F assume F: "finite F" and P: "P F"
@@ -53,7 +53,7 @@
qed
-subsubsection {* Choice principles *}
+subsubsection \<open>Choice principles\<close>
lemma ex_new_if_finite: -- "does not depend on def of finite at all"
assumes "\<not> finite (UNIV :: 'a set)" and "finite A"
@@ -63,7 +63,7 @@
then show ?thesis by blast
qed
-text {* A finite choice principle. Does not need the SOME choice operator. *}
+text \<open>A finite choice principle. Does not need the SOME choice operator.\<close>
lemma finite_set_choice:
"finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
@@ -79,7 +79,7 @@
qed
-subsubsection {* Finite sets are the images of initial segments of natural numbers *}
+subsubsection \<open>Finite sets are the images of initial segments of natural numbers\<close>
lemma finite_imp_nat_seg_image_inj_on:
assumes "finite A"
@@ -130,7 +130,7 @@
assumes "finite A"
shows "\<exists>f n::nat. f ` A = {i. i < n} \<and> inj_on f A"
proof -
- from finite_imp_nat_seg_image_inj_on[OF `finite A`]
+ from finite_imp_nat_seg_image_inj_on[OF \<open>finite A\<close>]
obtain f and n::nat where bij: "bij_betw f {i. i<n} A"
by (auto simp:bij_betw_def)
let ?f = "the_inv_into {i. i<n} f"
@@ -148,7 +148,7 @@
by (simp add: le_eq_less_or_eq Collect_disj_eq)
-subsubsection {* Finiteness and common set operations *}
+subsubsection \<open>Finiteness and common set operations\<close>
lemma rev_finite_subset:
"finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> finite A"
@@ -215,7 +215,7 @@
shows "finite (A - B) \<longleftrightarrow> finite A"
proof -
have "finite A \<longleftrightarrow> finite((A - B) \<union> (A \<inter> B))" by (simp add: Un_Diff_Int)
- also have "\<dots> \<longleftrightarrow> finite (A - B)" using `finite B` by simp
+ also have "\<dots> \<longleftrightarrow> finite (A - B)" using \<open>finite B\<close> by simp
finally show ?thesis ..
qed
@@ -277,10 +277,10 @@
case (insert x B)
then have B_A: "insert x B = f ` A" by simp
then obtain y where "x = f y" and "y \<in> A" by blast
- from B_A `x \<notin> B` have "B = f ` A - {x}" by blast
- with B_A `x \<notin> B` `x = f y` `inj_on f A` `y \<in> A` have "B = f ` (A - {y})"
+ from B_A \<open>x \<notin> B\<close> have "B = f ` A - {x}" by blast
+ with B_A \<open>x \<notin> B\<close> \<open>x = f y\<close> \<open>inj_on f A\<close> \<open>y \<in> A\<close> have "B = f ` (A - {y})"
by (simp add: inj_on_image_set_diff Set.Diff_subset)
- moreover from `inj_on f A` have "inj_on f (A - {y})" by (rule inj_on_diff)
+ moreover from \<open>inj_on f A\<close> have "inj_on f (A - {y})" by (rule inj_on_diff)
ultimately have "finite (A - {y})" by (rule insert.hyps)
then show "finite A" by simp
qed
@@ -397,7 +397,7 @@
from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
by (auto simp add: finite_conv_nat_seg_image)
then have "fst ` (A \<times> B) = fst ` f ` {i::nat. i < n}" by simp
- with `B \<noteq> {}` have "A = (fst \<circ> f) ` {i::nat. i < n}"
+ with \<open>B \<noteq> {}\<close> have "A = (fst \<circ> f) ` {i::nat. i < n}"
by (simp add: image_comp)
then have "\<exists>n f. A = f ` {i::nat. i < n}" by blast
then show ?thesis
@@ -411,7 +411,7 @@
from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
by (auto simp add: finite_conv_nat_seg_image)
then have "snd ` (A \<times> B) = snd ` f ` {i::nat. i < n}" by simp
- with `A \<noteq> {}` have "B = (snd \<circ> f) ` {i::nat. i < n}"
+ with \<open>A \<noteq> {}\<close> have "B = (snd \<circ> f) ` {i::nat. i < n}"
by (simp add: image_comp)
then have "\<exists>n f. B = f ` {i::nat. i < n}" by blast
then show ?thesis
@@ -468,7 +468,7 @@
qed
-subsubsection {* Further induction rules on finite sets *}
+subsubsection \<open>Further induction rules on finite sets\<close>
lemma finite_ne_induct [case_names singleton insert, consumes 2]:
assumes "finite F" and "F \<noteq> {}"
@@ -487,7 +487,7 @@
assumes empty: "P {}"
and insert: "\<And>a F. finite F \<Longrightarrow> a \<in> A \<Longrightarrow> a \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert a F)"
shows "P F"
-using `finite F` `F \<subseteq> A`
+using \<open>finite F\<close> \<open>F \<subseteq> A\<close>
proof induct
show "P {}" by fact
next
@@ -514,16 +514,16 @@
proof -
fix B :: "'a set"
assume "B \<subseteq> A"
- with `finite A` have "finite B" by (rule rev_finite_subset)
- from this `B \<subseteq> A` show "P (A - B)"
+ with \<open>finite A\<close> have "finite B" by (rule rev_finite_subset)
+ from this \<open>B \<subseteq> A\<close> show "P (A - B)"
proof induct
case empty
- from `P A` show ?case by simp
+ from \<open>P A\<close> show ?case by simp
next
case (insert b B)
have "P (A - B - {b})"
proof (rule remove)
- from `finite A` show "finite (A - B)" by induct auto
+ from \<open>finite A\<close> show "finite (A - B)" by induct auto
from insert show "b \<in> A - B" by simp
from insert show "P (A - B)" by simp
qed
@@ -546,19 +546,19 @@
case (insert a A)
then have "A = {a'. (f(a := c)) a' \<noteq> c}" and "f a \<noteq> c"
by auto
- with `finite A` have "finite {a'. (f(a := c)) a' \<noteq> c}"
+ with \<open>finite A\<close> have "finite {a'. (f(a := c)) a' \<noteq> c}"
by simp
have "(f(a := c)) a = c"
by simp
- from insert `A = {a'. (f(a := c)) a' \<noteq> c}` have "P (f(a := c))"
+ from insert \<open>A = {a'. (f(a := c)) a' \<noteq> c}\<close> have "P (f(a := c))"
by simp
- with `finite {a'. (f(a := c)) a' \<noteq> c}` `(f(a := c)) a = c` `f a \<noteq> c` have "P ((f(a := c))(a := f a))"
+ with \<open>finite {a'. (f(a := c)) a' \<noteq> c}\<close> \<open>(f(a := c)) a = c\<close> \<open>f a \<noteq> c\<close> have "P ((f(a := c))(a := f a))"
by (rule update)
then show ?case by simp
qed
-subsection {* Class @{text finite} *}
+subsection \<open>Class @{text finite}\<close>
class finite =
assumes finite_UNIV: "finite (UNIV \<Colon> 'a set)"
@@ -605,12 +605,12 @@
by default (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
-subsection {* A basic fold functional for finite sets *}
+subsection \<open>A basic fold functional for finite sets\<close>
-text {* The intended behaviour is
+text \<open>The intended behaviour is
@{text "fold f z {x\<^sub>1, ..., x\<^sub>n} = f x\<^sub>1 (\<dots> (f x\<^sub>n z)\<dots>)"}
if @{text f} is ``left-commutative'':
-*}
+\<close>
locale comp_fun_commute =
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
@@ -637,17 +637,17 @@
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" where
"fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)"
-text{*A tempting alternative for the definiens is
+text\<open>A tempting alternative for the definiens is
@{term "if finite A then THE y. fold_graph f z A y else e"}.
It allows the removal of finiteness assumptions from the theorems
@{text fold_comm}, @{text fold_reindex} and @{text fold_distrib}.
-The proofs become ugly. It is not worth the effort. (???) *}
+The proofs become ugly. It is not worth the effort. (???)\<close>
lemma finite_imp_fold_graph: "finite A \<Longrightarrow> \<exists>x. fold_graph f z A x"
by (induct rule: finite_induct) auto
-subsubsection{*From @{const fold_graph} to @{term fold}*}
+subsubsection\<open>From @{const fold_graph} to @{term fold}\<close>
context comp_fun_commute
begin
@@ -670,7 +670,7 @@
have "f x y = f a (f x y')"
unfolding y by (rule fun_left_comm)
moreover have "fold_graph f z (insert x A - {a}) (f x y')"
- using y' and `x \<noteq> a` and `x \<notin> A`
+ using y' and \<open>x \<noteq> a\<close> and \<open>x \<notin> A\<close>
by (simp add: insert_Diff_if fold_graph.insertI)
ultimately show ?case by fast
qed
@@ -685,11 +685,11 @@
"fold_graph f z A x \<Longrightarrow> fold_graph f z A y \<Longrightarrow> y = x"
proof (induct arbitrary: y set: fold_graph)
case (insertI x A y v)
- from `fold_graph f z (insert x A) v` and `x \<notin> A`
+ from \<open>fold_graph f z (insert x A) v\<close> and \<open>x \<notin> A\<close>
obtain y' where "v = f x y'" and "fold_graph f z A y'"
by (rule fold_graph_insertE)
- from `fold_graph f z A y'` have "y' = y" by (rule insertI)
- with `v = f x y'` show "v = f x y" by simp
+ from \<open>fold_graph f z A y'\<close> have "y' = y" by (rule insertI)
+ with \<open>v = f x y'\<close> show "v = f x y" by simp
qed fast
lemma fold_equality:
@@ -707,7 +707,7 @@
with assms show ?thesis by (simp add: fold_def)
qed
-text {* The base case for @{text fold}: *}
+text \<open>The base case for @{text fold}:\<close>
lemma (in -) fold_infinite [simp]:
assumes "\<not> finite A"
@@ -718,20 +718,20 @@
"fold f z {} = z"
by (auto simp add: fold_def)
-text{* The various recursion equations for @{const fold}: *}
+text\<open>The various recursion equations for @{const fold}:\<close>
lemma fold_insert [simp]:
assumes "finite A" and "x \<notin> A"
shows "fold f z (insert x A) = f x (fold f z A)"
proof (rule fold_equality)
fix z
- from `finite A` have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold)
- with `x \<notin> A` have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI)
+ from \<open>finite A\<close> have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold)
+ with \<open>x \<notin> A\<close> have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI)
then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp
qed
declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del]
- -- {* No more proofs involve these. *}
+ -- \<open>No more proofs involve these.\<close>
lemma fold_fun_left_comm:
"finite A \<Longrightarrow> f x (fold f z A) = fold f (f x z) A"
@@ -750,10 +750,10 @@
assumes "finite A" and "x \<in> A"
shows "fold f z A = f x (fold f z (A - {x}))"
proof -
- have A: "A = insert x (A - {x})" using `x \<in> A` by blast
+ have A: "A = insert x (A - {x})" using \<open>x \<in> A\<close> by blast
then have "fold f z A = fold f z (insert x (A - {x}))" by simp
also have "\<dots> = f x (fold f z (A - {x}))"
- by (rule fold_insert) (simp add: `finite A`)+
+ by (rule fold_insert) (simp add: \<open>finite A\<close>)+
finally show ?thesis .
qed
@@ -761,7 +761,7 @@
assumes "finite A"
shows "fold f z (insert x A) = f x (fold f z (A - {x}))"
proof -
- from `finite A` have "finite (insert x A)" by auto
+ from \<open>finite A\<close> have "finite (insert x A)" by auto
moreover have "x \<in> insert x A" by auto
ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))"
by (rule fold_rec)
@@ -775,7 +775,7 @@
end
-text{* Other properties of @{const fold}: *}
+text\<open>Other properties of @{const fold}:\<close>
lemma fold_image:
assumes "inj_on g A"
@@ -794,12 +794,12 @@
case emptyI then show ?case by (auto intro: fold_graph.emptyI)
next
case (insertI x A r B)
- from `inj_on g B` `x \<notin> A` `insert x A = image g B` obtain x' A' where
+ from \<open>inj_on g B\<close> \<open>x \<notin> A\<close> \<open>insert x A = image g B\<close> obtain x' A' where
"x' \<notin> A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'"
by (rule inj_img_insertE)
from insertI.prems have "fold_graph (f o g) z A' r"
by (auto intro: insertI.hyps)
- with `x' \<notin> A'` have "fold_graph (f \<circ> g) z (insert x' A') ((f \<circ> g) x' r)"
+ with \<open>x' \<notin> A'\<close> have "fold_graph (f \<circ> g) z (insert x' A') ((f \<circ> g) x' r)"
by (rule fold_graph.insertI)
then show ?case by simp
qed
@@ -809,7 +809,7 @@
case emptyI thus ?case by (auto intro: fold_graph.emptyI)
next
case (insertI x A r)
- from `x \<notin> A` insertI.prems have "g x \<notin> g ` A" by auto
+ from \<open>x \<notin> A\<close> insertI.prems have "g x \<notin> g ` A" by auto
moreover from insertI have "fold_graph f z (g ` A) r" by simp
ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)"
by (rule fold_graph.insertI)
@@ -827,19 +827,19 @@
shows "fold f s A = fold g t B"
proof -
have "fold f s A = fold g s A"
- using `finite A` cong proof (induct A)
+ using \<open>finite A\<close> cong proof (induct A)
case empty then show ?case by simp
next
case (insert x A)
- interpret f: comp_fun_commute f by (fact `comp_fun_commute f`)
- interpret g: comp_fun_commute g by (fact `comp_fun_commute g`)
+ interpret f: comp_fun_commute f by (fact \<open>comp_fun_commute f\<close>)
+ interpret g: comp_fun_commute g by (fact \<open>comp_fun_commute g\<close>)
from insert show ?case by simp
qed
with assms show ?thesis by simp
qed
-text {* A simplified version for idempotent functions: *}
+text \<open>A simplified version for idempotent functions:\<close>
locale comp_fun_idem = comp_fun_commute +
assumes comp_fun_idem: "f x \<circ> f x = f x"
@@ -868,7 +868,7 @@
end
-subsubsection {* Liftings to @{text comp_fun_commute} etc. *}
+subsubsection \<open>Liftings to @{text comp_fun_commute} etc.\<close>
lemma (in comp_fun_commute) comp_comp_fun_commute:
"comp_fun_commute (f \<circ> g)"
@@ -919,7 +919,7 @@
qed
-subsubsection {* Expressing set operations via @{const fold} *}
+subsubsection \<open>Expressing set operations via @{const fold}\<close>
lemma comp_fun_commute_const:
"comp_fun_commute (\<lambda>_. f)"
@@ -951,7 +951,7 @@
shows "A \<union> B = fold insert B A"
proof -
interpret comp_fun_idem insert by (fact comp_fun_idem_insert)
- from `finite A` show ?thesis by (induct A arbitrary: B) simp_all
+ from \<open>finite A\<close> show ?thesis by (induct A arbitrary: B) simp_all
qed
lemma minus_fold_remove:
@@ -959,7 +959,7 @@
shows "B - A = fold Set.remove B A"
proof -
interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove)
- from `finite A` have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto
+ from \<open>finite A\<close> have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto
then show ?thesis ..
qed
@@ -1053,7 +1053,7 @@
shows "inf (Inf A) B = fold inf B A"
proof -
interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
- from `finite A` fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
+ from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
(simp_all add: inf_commute fun_eq_iff)
qed
@@ -1062,7 +1062,7 @@
shows "sup (Sup A) B = fold sup B A"
proof -
interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
- from `finite A` fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
+ from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
(simp_all add: sup_commute fun_eq_iff)
qed
@@ -1082,7 +1082,7 @@
proof (rule sym)
interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
interpret comp_fun_idem "inf \<circ> f" by (fact comp_comp_fun_idem)
- from `finite A` show "?fold = ?inf"
+ from \<open>finite A\<close> show "?fold = ?inf"
by (induct A arbitrary: B)
(simp_all add: inf_left_commute)
qed
@@ -1093,7 +1093,7 @@
proof (rule sym)
interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
interpret comp_fun_idem "sup \<circ> f" by (fact comp_comp_fun_idem)
- from `finite A` show "?fold = ?sup"
+ from \<open>finite A\<close> show "?fold = ?sup"
by (induct A arbitrary: B)
(simp_all add: sup_left_commute)
qed
@@ -1111,9 +1111,9 @@
end
-subsection {* Locales as mini-packages for fold operations *}
+subsection \<open>Locales as mini-packages for fold operations\<close>
-subsubsection {* The natural case *}
+subsubsection \<open>The natural case\<close>
locale folding =
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
@@ -1142,16 +1142,16 @@
proof -
from fold_insert assms
have "fold f z (insert x A) = f x (fold f z A)" by simp
- with `finite A` show ?thesis by (simp add: eq_fold fun_eq_iff)
+ with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
qed
lemma remove:
assumes "finite A" and "x \<in> A"
shows "F A = f x (F (A - {x}))"
proof -
- from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ from \<open>x \<in> A\<close> obtain B where A: "A = insert x B" and "x \<notin> B"
by (auto dest: mk_disjoint_insert)
- moreover from `finite A` A have "finite B" by simp
+ moreover from \<open>finite A\<close> A have "finite B" by simp
ultimately show ?thesis by simp
qed
@@ -1163,7 +1163,7 @@
end
-subsubsection {* With idempotency *}
+subsubsection \<open>With idempotency\<close>
locale folding_idem = folding +
assumes comp_fun_idem: "f x \<circ> f x = f x"
@@ -1180,20 +1180,20 @@
proof -
from fold_insert_idem assms
have "fold f z (insert x A) = f x (fold f z A)" by simp
- with `finite A` show ?thesis by (simp add: eq_fold fun_eq_iff)
+ with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
qed
end
-subsection {* Finite cardinality *}
+subsection \<open>Finite cardinality\<close>
-text {*
+text \<open>
The traditional definition
@{prop "card A \<equiv> LEAST n. EX f. A = {f i | i. i < n}"}
is ugly to work with.
But now that we have @{const fold} things are easy:
-*}
+\<close>
definition card :: "'a set \<Rightarrow> nat" where
"card = folding.F (\<lambda>_. Suc) 0"
@@ -1289,7 +1289,7 @@
then have "x \<in> B" by simp
from insert have "A \<subseteq> B - {x}" and "finite (B - {x})" by auto
with insert.hyps have "card A \<le> card (B - {x})" by auto
- with `finite A` `x \<notin> A` `finite B` `x \<in> B` show ?case by simp (simp only: card.remove)
+ with \<open>finite A\<close> \<open>x \<notin> A\<close> \<open>finite B\<close> \<open>x \<in> B\<close> show ?case by simp (simp only: card.remove)
qed
qed
@@ -1467,24 +1467,24 @@
have "0 \<le> card S" by simp
then have "\<exists>T \<subseteq> S. card T = card S \<and> P T"
proof (induct rule: dec_induct)
- case base with `P {}` show ?case
+ case base with \<open>P {}\<close> show ?case
by (intro exI[of _ "{}"]) auto
next
case (step n)
then obtain T where T: "T \<subseteq> S" "card T = n" "P T"
by auto
- with `n < card S` have "T \<subset> S" "P T"
+ with \<open>n < card S\<close> have "T \<subset> S" "P T"
by auto
with select[of T] obtain s where "s \<in> S" "s \<notin> T" "P (insert s T)"
by auto
- with step(2) T `finite S` show ?case
+ with step(2) T \<open>finite S\<close> show ?case
by (intro exI[of _ "insert s T"]) (auto dest: finite_subset)
qed
- with `finite S` show "P S"
+ with \<open>finite S\<close> show "P S"
by (auto dest: card_subset_eq)
qed
-text{* main cardinality theorem *}
+text\<open>main cardinality theorem\<close>
lemma card_partition [rule_format]:
"finite C ==>
finite (\<Union>C) -->
@@ -1515,7 +1515,7 @@
qed
qed
-text{*The form of a finite set of given cardinality*}
+text\<open>The form of a finite set of given cardinality\<close>
lemma card_eq_SucD:
assumes "card A = Suc k"
@@ -1568,7 +1568,7 @@
next
case (Suc n)
then guess B .. note B = this
- with `\<not> finite A` have "A \<noteq> B" by auto
+ with \<open>\<not> finite A\<close> have "A \<noteq> B" by auto
with B have "B \<subset> A" by auto
hence "\<exists>x. x \<in> A - B" by (elim psubset_imp_ex_mem)
then guess x .. note x = this
@@ -1577,7 +1577,7 @@
thus "\<exists>B. finite B \<and> card B = Suc n \<and> B \<subseteq> A" ..
qed
-subsubsection {* Cardinality of image *}
+subsubsection \<open>Cardinality of image\<close>
lemma card_image_le: "finite A ==> card (f ` A) \<le> card A"
by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if)
@@ -1643,7 +1643,7 @@
lemma card_vimage_inj: "\<lbrakk> inj f; A \<subseteq> range f \<rbrakk> \<Longrightarrow> card (f -` A) = card A"
by(auto 4 3 simp add: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on])
-subsubsection {* Pigeonhole Principles *}
+subsubsection \<open>Pigeonhole Principles\<close>
lemma pigeonhole: "card A > card(f ` A) \<Longrightarrow> ~ inj_on f A "
by (auto dest: card_image less_irrefl_nat)
@@ -1660,7 +1660,7 @@
show ?case
proof cases
assume "finite{a:A. f a = b}"
- hence "~ finite(A - {a:A. f a = b})" using `\<not> finite A` by simp
+ hence "~ finite(A - {a:A. f a = b})" using \<open>\<not> finite A\<close> by simp
also have "A - {a:A. f a = b} = {a:A. f a \<noteq> b}" by blast
finally have "~ finite({a:A. f a \<noteq> b})" .
from insert(3)[OF _ this]
@@ -1679,20 +1679,20 @@
shows "EX b:B. ~finite{a:A. R a b}"
proof -
let ?F = "%a. {b:B. R a b}"
- from finite_Pow_iff[THEN iffD2, OF `finite B`]
+ from finite_Pow_iff[THEN iffD2, OF \<open>finite B\<close>]
have "finite(?F ` A)" by(blast intro: rev_finite_subset)
from pigeonhole_infinite[where f = ?F, OF assms(1) this]
obtain a0 where "a0\<in>A" and 1: "\<not> finite {a\<in>A. ?F a = ?F a0}" ..
- obtain b0 where "b0 : B" and "R a0 b0" using `a0:A` assms(3) by blast
+ obtain b0 where "b0 : B" and "R a0 b0" using \<open>a0:A\<close> assms(3) by blast
{ assume "finite{a:A. R a b0}"
then have "finite {a\<in>A. ?F a = ?F a0}"
- using `b0 : B` `R a0 b0` by(blast intro: rev_finite_subset)
+ using \<open>b0 : B\<close> \<open>R a0 b0\<close> by(blast intro: rev_finite_subset)
}
- with 1 `b0 : B` show ?thesis by blast
+ with 1 \<open>b0 : B\<close> show ?thesis by blast
qed
-subsubsection {* Cardinality of sums *}
+subsubsection \<open>Cardinality of sums\<close>
lemma card_Plus:
assumes "finite A" and "finite B"
@@ -1708,7 +1708,7 @@
"card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
by (auto simp add: card_Plus)
-text {* Relates to equivalence classes. Based on a theorem of F. Kamm\"uller. *}
+text \<open>Relates to equivalence classes. Based on a theorem of F. Kamm\"uller.\<close>
lemma dvd_partition:
assumes f: "finite (\<Union>C)" and "\<forall>c\<in>C. k dvd card c" "\<forall>c1\<in>C. \<forall>c2\<in>C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {}"
@@ -1729,7 +1729,7 @@
qed
qed
-subsubsection {* Relating injectivity and surjectivity *}
+subsubsection \<open>Relating injectivity and surjectivity\<close>
lemma finite_surj_inj: assumes "finite A" "A \<subseteq> f ` A" shows "inj_on f A"
proof -
--- a/src/HOL/Fun.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Fun.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Copyright 1994, 2012
*)
-section {* Notions about functions *}
+section \<open>Notions about functions\<close>
theory Fun
imports Set
@@ -15,14 +15,14 @@
"f x = u \<Longrightarrow> (\<And>x. P x \<Longrightarrow> g (f x) = x) \<Longrightarrow> P x \<Longrightarrow> x = g u"
by auto
-text{*Uniqueness, so NOT the axiom of choice.*}
+text\<open>Uniqueness, so NOT the axiom of choice.\<close>
lemma uniq_choice: "\<forall>x. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x. Q x (f x)"
by (force intro: theI')
lemma b_uniq_choice: "\<forall>x\<in>S. \<exists>!y. Q x y \<Longrightarrow> \<exists>f. \<forall>x\<in>S. Q x (f x)"
by (force intro: theI')
-subsection {* The Identity Function @{text id} *}
+subsection \<open>The Identity Function @{text id}\<close>
definition id :: "'a \<Rightarrow> 'a" where
"id = (\<lambda>x. x)"
@@ -40,7 +40,7 @@
constant id \<rightharpoonup> (Haskell) "id"
-subsection {* The Composition Operator @{text "f \<circ> g"} *}
+subsection \<open>The Composition Operator @{text "f \<circ> g"}\<close>
definition comp :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "o" 55) where
"f o g = (\<lambda>x. f (g x))"
@@ -98,7 +98,7 @@
constant comp \<rightharpoonup> (SML) infixl 5 "o" and (Haskell) infixr 9 "."
-subsection {* The Forward Composition Operator @{text fcomp} *}
+subsection \<open>The Forward Composition Operator @{text fcomp}\<close>
definition fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c" (infixl "\<circ>>" 60) where
"f \<circ>> g = (\<lambda>x. g (f x))"
@@ -121,7 +121,7 @@
no_notation fcomp (infixl "\<circ>>" 60)
-subsection {* Mapping functions *}
+subsection \<open>Mapping functions\<close>
definition map_fun :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'c \<Rightarrow> 'd" where
"map_fun f g h = g \<circ> h \<circ> f"
@@ -131,7 +131,7 @@
by (simp add: map_fun_def)
-subsection {* Injectivity and Bijectivity *}
+subsection \<open>Injectivity and Bijectivity\<close>
definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool" where -- "injective"
"inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)"
@@ -139,8 +139,8 @@
definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool" where -- "bijective"
"bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B"
-text{*A common special case: functions injective, surjective or bijective over
-the entire domain type.*}
+text\<open>A common special case: functions injective, surjective or bijective over
+the entire domain type.\<close>
abbreviation
"inj f \<equiv> inj_on f UNIV"
@@ -151,7 +151,7 @@
abbreviation
"bij f \<equiv> bij_betw f UNIV UNIV"
-text{* The negated case: *}
+text\<open>The negated case:\<close>
translations
"\<not> CONST surj f" <= "CONST range f \<noteq> CONST UNIV"
@@ -278,14 +278,14 @@
with assms * have "B = f ` (A - {x'})"
by (auto dest: inj_on_contraD)
have "x' \<notin> A - {x'}" by simp
- from `x' \<notin> A - {x'}` `A = insert x' (A - {x'})` `x = f x'` `B = image f (A - {x'})`
+ from \<open>x' \<notin> A - {x'}\<close> \<open>A = insert x' (A - {x'})\<close> \<open>x = f x'\<close> \<open>B = image f (A - {x'})\<close>
show ?thesis ..
qed
lemma linorder_injI:
assumes hyp: "\<And>x y. x < (y::'a::linorder) \<Longrightarrow> f x \<noteq> f y"
shows "inj f"
- -- {* Courtesy of Stephan Merz *}
+ -- \<open>Courtesy of Stephan Merz\<close>
proof (rule inj_onI)
fix x y
assume f_eq: "f x = f y"
@@ -395,9 +395,9 @@
have "inj_on ?g B"
proof(rule inj_onI)
fix x y assume "x:B" "y:B" "?g x = ?g y"
- from s `x:B` obtain a1 where a1: "?P x a1" by blast
- from s `y:B` obtain a2 where a2: "?P y a2" by blast
- from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp
+ from s \<open>x:B\<close> obtain a1 where a1: "?P x a1" by blast
+ from s \<open>y:B\<close> obtain a2 where a2: "?P y a2" by blast
+ from g[OF a1] a1 g[OF a2] a2 \<open>?g x = ?g y\<close> show "x=y" by simp
qed
moreover have "?g ` B = A"
proof(auto simp: image_def)
@@ -453,7 +453,7 @@
lemma surj_vimage_empty:
assumes "surj f" shows "f -` A = {} \<longleftrightarrow> A = {}"
- using surj_image_vimage_eq[OF `surj f`, of A]
+ using surj_image_vimage_eq[OF \<open>surj f\<close>, of A]
by (intro iffI) fastforce+
lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A"
@@ -515,8 +515,8 @@
done
lemma inj_vimage_singleton: "inj f \<Longrightarrow> f -` {a} \<subseteq> {THE x. f x = a}"
- -- {* The inverse image of a singleton under an injective function
- is included in a singleton. *}
+ -- \<open>The inverse image of a singleton under an injective function
+ is included in a singleton.\<close>
apply (auto simp add: inj_on_def)
apply (blast intro: the_equality [symmetric])
done
@@ -594,7 +594,7 @@
qed
-subsection{*Function Updating*}
+subsection\<open>Function Updating\<close>
definition fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where
"fun_upd f a b == % x. if x=a then b else f x"
@@ -658,7 +658,7 @@
by auto
-subsection {* @{text override_on} *}
+subsection \<open>@{text override_on}\<close>
definition override_on :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b" where
"override_on f g A = (\<lambda>a. if a \<in> A then g a else f a)"
@@ -673,7 +673,7 @@
by(simp add:override_on_def)
-subsection {* @{text swap} *}
+subsection \<open>@{text swap}\<close>
definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)"
where
@@ -750,7 +750,7 @@
hide_const (open) swap
-subsection {* Inversion of injective functions *}
+subsection \<open>Inversion of injective functions\<close>
definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
"the_inv_into A f == %x. THE y. y : A & f y = x"
@@ -813,7 +813,7 @@
by (rule the_inv_into_f_f)
-subsection {* Cantor's Paradox *}
+subsection \<open>Cantor's Paradox\<close>
lemma Cantors_paradox:
"\<not>(\<exists>f. f ` A = Pow A)"
@@ -825,14 +825,14 @@
thus False by best
qed
-subsection {* Setup *}
+subsection \<open>Setup\<close>
-subsubsection {* Proof tools *}
+subsubsection \<open>Proof tools\<close>
-text {* simplifies terms of the form
- f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *}
+text \<open>simplifies terms of the form
+ f(...,x:=y,...,x:=z,...) to f(...,x:=z,...)\<close>
-simproc_setup fun_upd2 ("f(v := w, x := y)") = {* fn _ =>
+simproc_setup fun_upd2 ("f(v := w, x := y)") = \<open>fn _ =>
let
fun gen_fun_upd NONE T _ _ = NONE
| gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y)
@@ -860,10 +860,10 @@
simp_tac (put_simpset ss ctxt) 1))
end
in proc end
-*}
+\<close>
-subsubsection {* Functorial structure of types *}
+subsubsection \<open>Functorial structure of types\<close>
ML_file "Tools/functor.ML"
@@ -873,7 +873,7 @@
functor vimage
by (simp_all add: fun_eq_iff vimage_comp)
-text {* Legacy theorem names *}
+text \<open>Legacy theorem names\<close>
lemmas o_def = comp_def
lemmas o_apply = comp_apply
--- a/src/HOL/Fun_Def.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Fun_Def.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,14 +2,14 @@
Author: Alexander Krauss, TU Muenchen
*)
-section {* Function Definitions and Termination Proofs *}
+section \<open>Function Definitions and Termination Proofs\<close>
theory Fun_Def
imports Basic_BNF_LFPs Partial_Function SAT
keywords "function" "termination" :: thy_goal and "fun" "fun_cases" :: thy_decl
begin
-subsection {* Definitions with default value *}
+subsection \<open>Definitions with default value\<close>
definition
THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where
@@ -66,7 +66,7 @@
proof
assume "\<exists>y. G x y"
hence "D x" using graph ..
- with `\<not> D x` show False ..
+ with \<open>\<not> D x\<close> show False ..
qed
hence "\<not>(\<exists>!y. G x y)" by blast
@@ -88,26 +88,26 @@
ML_file "Tools/Function/relation.ML"
ML_file "Tools/Function/function_elims.ML"
-method_setup relation = {*
+method_setup relation = \<open>
Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t))
-*} "prove termination using a user-specified wellfounded relation"
+\<close> "prove termination using a user-specified wellfounded relation"
ML_file "Tools/Function/function.ML"
ML_file "Tools/Function/pat_completeness.ML"
-method_setup pat_completeness = {*
+method_setup pat_completeness = \<open>
Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac)
-*} "prove completeness of (co)datatype patterns"
+\<close> "prove completeness of (co)datatype patterns"
ML_file "Tools/Function/fun.ML"
ML_file "Tools/Function/induction_schema.ML"
-method_setup induction_schema = {*
+method_setup induction_schema = \<open>
Scan.succeed (EMPTY_CASES oo Induction_Schema.induction_schema_tac)
-*} "prove an induction principle"
+\<close> "prove an induction principle"
-subsection {* Measure functions *}
+subsection \<open>Measure functions\<close>
inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
where is_measure_trivial: "is_measure f"
@@ -125,13 +125,13 @@
ML_file "Tools/Function/lexicographic_order.ML"
-method_setup lexicographic_order = {*
+method_setup lexicographic_order = \<open>
Method.sections clasimp_modifiers >>
(K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false))
-*} "termination prover for lexicographic orderings"
+\<close> "termination prover for lexicographic orderings"
-subsection {* Congruence rules *}
+subsection \<open>Congruence rules\<close>
lemma let_cong [fundef_cong]:
"M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g"
@@ -151,7 +151,7 @@
unfolding o_apply .
-subsection {* Simp rules for termination proofs *}
+subsection \<open>Simp rules for termination proofs\<close>
declare
trans_less_add1[termination_simp]
@@ -166,7 +166,7 @@
by (induct p) auto
-subsection {* Decomposition *}
+subsection \<open>Decomposition\<close>
lemma less_by_empty:
"A = {} \<Longrightarrow> A \<subseteq> B"
@@ -179,7 +179,7 @@
by (auto simp add: wf_comp_self[of R])
-subsection {* Reduction pairs *}
+subsection \<open>Reduction pairs\<close>
definition
"reduction_pair P = (wf (fst P) \<and> fst P O snd P \<subseteq> fst P)"
@@ -194,11 +194,11 @@
assumes "wf S"
shows "wf (R \<union> S)"
proof -
- from rp `S \<subseteq> snd P` have "wf (fst P)" "fst P O S \<subseteq> fst P"
+ from rp \<open>S \<subseteq> snd P\<close> have "wf (fst P)" "fst P O S \<subseteq> fst P"
unfolding reduction_pair_def by auto
- with `wf S` have "wf (fst P \<union> S)"
+ with \<open>wf S\<close> have "wf (fst P \<union> S)"
by (auto intro: wf_union_compatible)
- moreover from `R \<subseteq> fst P` have "R \<union> S \<subseteq> fst P \<union> S" by auto
+ moreover from \<open>R \<subseteq> fst P\<close> have "R \<union> S \<subseteq> fst P \<union> S" by auto
ultimately show ?thesis by (rule wf_subset)
qed
@@ -211,7 +211,7 @@
by force
-subsection {* Concrete orders for SCNP termination proofs *}
+subsection \<open>Concrete orders for SCNP termination proofs\<close>
definition "pair_less = less_than <*lex*> less_than"
definition "pair_leq = pair_less^="
@@ -223,14 +223,14 @@
lemma wf_pair_less[simp]: "wf pair_less"
by (auto simp: pair_less_def)
-text {* Introduction rules for @{text pair_less}/@{text pair_leq} *}
+text \<open>Introduction rules for @{text pair_less}/@{text pair_leq}\<close>
lemma pair_leqI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
and pair_leqI2: "a \<le> b \<Longrightarrow> s \<le> t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
and pair_lessI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
and pair_lessI2: "a \<le> b \<Longrightarrow> s < t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
unfolding pair_leq_def pair_less_def by auto
-text {* Introduction rules for max *}
+text \<open>Introduction rules for max\<close>
lemma smax_emptyI:
"finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> ({}, Y) \<in> max_strict"
and smax_insertI:
@@ -241,7 +241,7 @@
"\<lbrakk>y \<in> YS; (x, y) \<in> pair_leq; (XS, YS) \<in> max_weak\<rbrakk> \<Longrightarrow> (insert x XS, YS) \<in> max_weak"
unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases)
-text {* Introduction rules for min *}
+text \<open>Introduction rules for min\<close>
lemma smin_emptyI:
"X \<noteq> {} \<Longrightarrow> (X, {}) \<in> min_strict"
and smin_insertI:
@@ -252,7 +252,7 @@
"\<lbrakk>x \<in> XS; (x, y) \<in> pair_leq; (XS, YS) \<in> min_weak\<rbrakk> \<Longrightarrow> (XS, insert y YS) \<in> min_weak"
by (auto simp: min_strict_def min_weak_def min_ext_def)
-text {* Reduction Pairs *}
+text \<open>Reduction Pairs\<close>
lemma max_ext_compat:
assumes "R O S \<subseteq> R"
@@ -294,7 +294,7 @@
by (auto simp: pair_less_def pair_leq_def)
-subsection {* Tool setup *}
+subsection \<open>Tool setup\<close>
ML_file "Tools/Function/termination.ML"
ML_file "Tools/Function/scnp_solve.ML"
@@ -302,9 +302,9 @@
ML_file "Tools/Function/fun_cases.ML"
ML_val -- "setup inactive"
-{*
+\<open>
Context.theory_map (Function_Common.set_termination_prover
(K (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])))
-*}
+\<close>
end
--- a/src/HOL/Fun_Def_Base.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Fun_Def_Base.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Alexander Krauss, TU Muenchen
*)
-section {* Function Definition Base *}
+section \<open>Function Definition Base\<close>
theory Fun_Def_Base
imports Ctr_Sugar Set Wellfounded
--- a/src/HOL/GCD.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/GCD.thy Sat Jul 18 22:58:50 2015 +0200
@@ -25,13 +25,13 @@
*)
-section {* Greatest common divisor and least common multiple *}
+section \<open>Greatest common divisor and least common multiple\<close>
theory GCD
imports Main
begin
-subsection {* GCD and LCM definitions *}
+subsection \<open>GCD and LCM definitions\<close>
class gcd = zero + one + dvd +
fixes gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -592,7 +592,7 @@
end
-subsection {* Transfer setup *}
+subsection \<open>Transfer setup\<close>
lemma transfer_nat_int_gcd:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> gcd (nat x) (nat y) = nat (gcd x y)"
@@ -622,7 +622,7 @@
transfer_int_nat_gcd transfer_int_nat_gcd_closures]
-subsection {* GCD properties *}
+subsection \<open>GCD properties\<close>
(* was gcd_induct *)
lemma gcd_nat_induct:
@@ -744,10 +744,10 @@
declare gcd_nat.simps [simp del]
-text {*
+text \<open>
\medskip @{term "gcd m n"} divides @{text m} and @{text n}. The
conjunctions don't seem provable separately.
-*}
+\<close>
instance nat :: semiring_gcd
proof
@@ -868,12 +868,12 @@
lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \<Longrightarrow> gcd (x::int) y = abs y"
by (metis gcd_proj1_if_dvd_int gcd_commute_int)
-text {*
+text \<open>
\medskip Multiplication laws
-*}
+\<close>
lemma gcd_mult_distrib_nat: "(k::nat) * gcd m n = gcd (k * m) (k * n)"
- -- {* @{cite \<open>page 27\<close> davenport92} *}
+ -- \<open>@{cite \<open>page 27\<close> davenport92}\<close>
apply (induct m n rule: gcd_nat_induct)
apply simp
apply (case_tac "k = 0")
@@ -944,16 +944,16 @@
assume ?rhs then show ?lhs by simp
next
assume ?lhs
- from `?lhs` have "a dvd b * d" by (auto intro: dvdI dest: sym)
- with `coprime a d` have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat)
- from `?lhs` have "b dvd a * c" by (auto intro: dvdI dest: sym)
- with `coprime b c` have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat)
- from `?lhs` have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult.commute)
- with `coprime b c` have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
- from `?lhs` have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult.commute)
- with `coprime a d` have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
- from `a dvd b` `b dvd a` have "a = b" by (rule Nat.dvd.antisym)
- moreover from `c dvd d` `d dvd c` have "c = d" by (rule Nat.dvd.antisym)
+ from \<open>?lhs\<close> have "a dvd b * d" by (auto intro: dvdI dest: sym)
+ with \<open>coprime a d\<close> have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat)
+ from \<open>?lhs\<close> have "b dvd a * c" by (auto intro: dvdI dest: sym)
+ with \<open>coprime b c\<close> have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat)
+ from \<open>?lhs\<close> have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult.commute)
+ with \<open>coprime b c\<close> have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
+ from \<open>?lhs\<close> have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult.commute)
+ with \<open>coprime a d\<close> have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat)
+ from \<open>a dvd b\<close> \<open>b dvd a\<close> have "a = b" by (rule Nat.dvd.antisym)
+ moreover from \<open>c dvd d\<close> \<open>d dvd c\<close> have "c = d" by (rule Nat.dvd.antisym)
ultimately show ?rhs ..
qed
@@ -963,7 +963,7 @@
shows "\<bar>a\<bar> * \<bar>c\<bar> = \<bar>b\<bar> * \<bar>d\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>b\<bar> \<and> \<bar>c\<bar> = \<bar>d\<bar>"
using assms by (intro coprime_crossproduct_nat [transferred]) auto
-text {* \medskip Addition laws *}
+text \<open>\medskip Addition laws\<close>
lemma gcd_add1_nat [simp]: "gcd ((m::nat) + n) n = gcd m n"
apply (case_tac "n = 0")
@@ -1085,7 +1085,7 @@
by (simp add: gcd_int_def nat_mod_distrib gcd_non_0_nat)
-subsection {* Coprimality *}
+subsection \<open>Coprimality\<close>
context semiring_gcd
begin
@@ -1504,7 +1504,7 @@
by (metis coprime_lmult_int gcd_1_int gcd_commute_int gcd_red_int)
-subsection {* Bezout's theorem *}
+subsection \<open>Bezout's theorem\<close>
(* Function bezw returns a pair of witnesses to Bezout's theorem --
see the theorems that follow the definition. *)
@@ -1592,7 +1592,7 @@
ultimately show ?thesis by blast
qed
-text {* versions of Bezout for nat, by Amine Chaieb *}
+text \<open>versions of Bezout for nat, by Amine Chaieb\<close>
lemma ind_euclid:
assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
@@ -1744,7 +1744,7 @@
qed
-subsection {* LCM properties *}
+subsection \<open>LCM properties\<close>
lemma lcm_altdef_int [code]: "lcm (a::int) b = (abs a) * (abs b) div gcd a b"
by (simp add: lcm_int_def lcm_nat_def zdiv_int
@@ -1906,7 +1906,7 @@
by (auto simp add: abs_mult_self trans [OF lcm_unique_int eq_commute, symmetric] zmult_eq_1_iff)
-subsection {* The complete divisibility lattice *}
+subsection \<open>The complete divisibility lattice\<close>
interpretation gcd_semilattice_nat: semilattice_inf gcd "op dvd" "(\<lambda>m n::nat. m dvd n \<and> \<not> n dvd m)"
by standard simp_all
@@ -1916,9 +1916,9 @@
interpretation gcd_lcm_lattice_nat: lattice gcd "op dvd" "(\<lambda>m n::nat. m dvd n & ~ n dvd m)" lcm ..
-text{* Lifting gcd and lcm to sets (Gcd/Lcm).
+text\<open>Lifting gcd and lcm to sets (Gcd/Lcm).
Gcd is defined via Lcm to facilitate the proof that we have a complete lattice.
-*}
+\<close>
instantiation nat :: Gcd
begin
@@ -2034,7 +2034,7 @@
simp add: unit_factor_Gcd uf)
qed
-text{* Alternative characterizations of Gcd: *}
+text\<open>Alternative characterizations of Gcd:\<close>
lemma Gcd_eq_Max: "finite(M::nat set) \<Longrightarrow> M \<noteq> {} \<Longrightarrow> 0 \<notin> M \<Longrightarrow> Gcd M = Max(\<Inter>m\<in>M. {d. d dvd m})"
apply(rule antisym)
@@ -2097,7 +2097,7 @@
dvd.neq_le_trans dvd_triv_right mult.commute)
done
-text{* Nitpick: *}
+text\<open>Nitpick:\<close>
lemma gcd_eq_nitpick_gcd [nitpick_unfold]: "gcd x y = Nitpick.nat_gcd x y"
by (induct x y rule: nat_gcd.induct)
@@ -2107,7 +2107,7 @@
by (simp only: lcm_nat_def Nitpick.nat_lcm_def gcd_eq_nitpick_gcd)
-subsubsection {* Setwise gcd and lcm for integers *}
+subsubsection \<open>Setwise gcd and lcm for integers\<close>
instantiation int :: Gcd
begin
--- a/src/HOL/Groebner_Basis.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Groebner_Basis.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,17 +2,17 @@
Author: Amine Chaieb, TU Muenchen
*)
-section {* Groebner bases *}
+section \<open>Groebner bases\<close>
theory Groebner_Basis
imports Semiring_Normalization Parity
begin
-subsection {* Groebner Bases *}
+subsection \<open>Groebner Bases\<close>
-lemmas bool_simps = simp_thms(1-34) -- {* FIXME move to @{theory HOL} *}
+lemmas bool_simps = simp_thms(1-34) -- \<open>FIXME move to @{theory HOL}\<close>
-lemma nnf_simps: -- {* FIXME shadows fact binding in @{theory HOL} *}
+lemma nnf_simps: -- \<open>FIXME shadows fact binding in @{theory HOL}\<close>
"(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)"
"(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
@@ -35,7 +35,7 @@
named_theorems algebra "pre-simplification rules for algebraic methods"
ML_file "Tools/groebner.ML"
-method_setup algebra = {*
+method_setup algebra = \<open>
let
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
val addN = "add"
@@ -48,7 +48,7 @@
(fn (add_ths, del_ths) => fn ctxt =>
SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))
end
-*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
+\<close> "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
declare dvd_def[algebra]
declare dvd_eq_mod_eq_0[symmetric, algebra]
--- a/src/HOL/Groups.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Groups.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,44 +2,44 @@
Author: Gertrud Bauer, Steven Obua, Lawrence C Paulson, Markus Wenzel, Jeremy Avigad
*)
-section {* Groups, also combined with orderings *}
+section \<open>Groups, also combined with orderings\<close>
theory Groups
imports Orderings
begin
-subsection {* Dynamic facts *}
+subsection \<open>Dynamic facts\<close>
named_theorems ac_simps "associativity and commutativity simplification rules"
-text{* The rewrites accumulated in @{text algebra_simps} deal with the
+text\<open>The rewrites accumulated in @{text algebra_simps} deal with the
classical algebraic structures of groups, rings and family. They simplify
terms by multiplying everything out (in case of a ring) and bringing sums and
products into a canonical form (by ordered rewriting). As a result it decides
group and ring equalities but also helps with inequalities.
Of course it also works for fields, but it knows nothing about multiplicative
-inverses or division. This is catered for by @{text field_simps}. *}
+inverses or division. This is catered for by @{text field_simps}.\<close>
named_theorems algebra_simps "algebra simplification rules"
-text{* Lemmas @{text field_simps} multiply with denominators in (in)equations
+text\<open>Lemmas @{text field_simps} multiply with denominators in (in)equations
if they can be proved to be non-zero (for equations) or positive/negative
(for inequations). Can be too aggressive and is therefore separate from the
-more benign @{text algebra_simps}. *}
+more benign @{text algebra_simps}.\<close>
named_theorems field_simps "algebra simplification rules for fields"
-subsection {* Abstract structures *}
+subsection \<open>Abstract structures\<close>
-text {*
+text \<open>
These locales provide basic structures for interpretation into
bigger structures; extensions require careful thinking, otherwise
undesired effects may occur due to interpretation.
-*}
+\<close>
locale semigroup =
fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
@@ -76,7 +76,7 @@
end
-subsection {* Generic operations *}
+subsection \<open>Generic operations\<close>
class zero =
fixes zero :: 'a ("0")
@@ -92,17 +92,17 @@
lemma Let_1 [simp]: "Let 1 f = f 1"
unfolding Let_def ..
-setup {*
+setup \<open>
Reorient_Proc.add
(fn Const(@{const_name Groups.zero}, _) => true
| Const(@{const_name Groups.one}, _) => true
| _ => false)
-*}
+\<close>
simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc
simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc
-typed_print_translation {*
+typed_print_translation \<open>
let
fun tr' c = (c, fn ctxt => fn T => fn ts =>
if null ts andalso Printer.type_emphasis ctxt T then
@@ -110,7 +110,7 @@
Syntax_Phases.term_of_typ ctxt T
else raise Match);
in map tr' [@{const_syntax Groups.one}, @{const_syntax Groups.zero}] end;
-*} -- {* show types that are presumably too general *}
+\<close> -- \<open>show types that are presumably too general\<close>
class plus =
fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "+" 65)
@@ -125,7 +125,7 @@
fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
-subsection {* Semigroups and Monoids *}
+subsection \<open>Semigroups and Monoids\<close>
class semigroup_add = plus +
assumes add_assoc [algebra_simps, field_simps]: "(a + b) + c = a + (b + c)"
@@ -327,7 +327,7 @@
end
-subsection {* Groups *}
+subsection \<open>Groups\<close>
class group_add = minus + uminus + monoid_add +
assumes left_minus [simp]: "- a + a = 0"
@@ -413,7 +413,7 @@
proof
assume "a - b = 0"
have "a = (a - b) + b" by (simp add: add.assoc)
- also have "\<dots> = b" using `a - b = 0` by simp
+ also have "\<dots> = b" using \<open>a - b = 0\<close> by simp
finally show "a = b" .
next
assume "a = b" thus "a - b = 0" by simp
@@ -454,7 +454,7 @@
"0 = - a \<longleftrightarrow> 0 = a"
by (subst neg_equal_iff_equal [symmetric]) simp
-text{*The next two equations can make the simplifier loop!*}
+text\<open>The next two equations can make the simplifier loop!\<close>
lemma equation_minus_iff:
"a = - b \<longleftrightarrow> b = - a"
@@ -557,9 +557,9 @@
end
-subsection {* (Partially) Ordered Groups *}
+subsection \<open>(Partially) Ordered Groups\<close>
-text {*
+text \<open>
The theory of partially ordered groups is taken from the books:
\begin{itemize}
\item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
@@ -570,7 +570,7 @@
\item @{url "http://www.mathworld.com"} by Eric Weisstein et. al.
\item \emph{Algebra I} by van der Waerden, Springer.
\end{itemize}
-*}
+\<close>
class ordered_ab_semigroup_add = order + ab_semigroup_add +
assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
@@ -580,7 +580,7 @@
"a \<le> b \<Longrightarrow> a + c \<le> b + c"
by (simp add: add.commute [of _ c] add_left_mono)
-text {* non-strict, in both arguments *}
+text \<open>non-strict, in both arguments\<close>
lemma add_mono:
"a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a + c \<le> b + d"
apply (erule add_right_mono [THEN order_trans])
@@ -601,7 +601,7 @@
"a < b \<Longrightarrow> a + c < b + c"
by (simp add: add.commute [of _ c] add_strict_left_mono)
-text{*Strict monotonicity in both arguments*}
+text\<open>Strict monotonicity in both arguments\<close>
lemma add_strict_mono:
"a < b \<Longrightarrow> c < d \<Longrightarrow> a + c < b + d"
apply (erule add_strict_right_mono [THEN less_trans])
@@ -700,23 +700,23 @@
lemma add_diff_inverse:
"a + (b - a) = b"
- using `a \<le> b` by (auto simp add: le_iff_add)
+ using \<open>a \<le> b\<close> by (auto simp add: le_iff_add)
lemma add_diff_assoc:
"c + (b - a) = c + b - a"
- using `a \<le> b` by (auto simp add: le_iff_add add.left_commute [of c])
+ using \<open>a \<le> b\<close> by (auto simp add: le_iff_add add.left_commute [of c])
lemma add_diff_assoc2:
"b - a + c = b + c - a"
- using `a \<le> b` by (auto simp add: le_iff_add add.assoc)
+ using \<open>a \<le> b\<close> by (auto simp add: le_iff_add add.assoc)
lemma diff_add_assoc:
"c + b - a = c + (b - a)"
- using `a \<le> b` by (simp add: add.commute add_diff_assoc)
+ using \<open>a \<le> b\<close> by (simp add: add.commute add_diff_assoc)
lemma diff_add_assoc2:
"b + c - a = b - a + c"
- using `a \<le> b`by (simp add: add.commute add_diff_assoc)
+ using \<open>a \<le> b\<close>by (simp add: add.commute add_diff_assoc)
lemma diff_diff_right:
"c - (b - a) = c + a - b"
@@ -751,7 +751,7 @@
end
-subsection {* Support for reasoning about signs *}
+subsection \<open>Support for reasoning about signs\<close>
class ordered_comm_monoid_add =
ordered_cancel_ab_semigroup_add + comm_monoid_add
@@ -915,7 +915,7 @@
lemma le_imp_neg_le:
assumes "a \<le> b" shows "-b \<le> -a"
proof -
- have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono)
+ have "-a+a \<le> -a+b" using \<open>a \<le> b\<close> by (rule add_left_mono)
then have "0 \<le> -a+b" by simp
then have "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)
then show ?thesis by (simp add: algebra_simps)
@@ -946,7 +946,7 @@
lemma neg_0_less_iff_less [simp]: "0 < - a \<longleftrightarrow> a < 0"
by (subst neg_less_iff_less [symmetric], simp)
-text{*The next several equations can make the simplifier loop!*}
+text\<open>The next several equations can make the simplifier loop!\<close>
lemma less_minus_iff: "a < - b \<longleftrightarrow> b < - a"
proof -
@@ -1042,19 +1042,19 @@
ML_file "Tools/group_cancel.ML"
simproc_setup group_cancel_add ("a + b::'a::ab_group_add") =
- {* fn phi => fn ss => try Group_Cancel.cancel_add_conv *}
+ \<open>fn phi => fn ss => try Group_Cancel.cancel_add_conv\<close>
simproc_setup group_cancel_diff ("a - b::'a::ab_group_add") =
- {* fn phi => fn ss => try Group_Cancel.cancel_diff_conv *}
+ \<open>fn phi => fn ss => try Group_Cancel.cancel_diff_conv\<close>
simproc_setup group_cancel_eq ("a = (b::'a::ab_group_add)") =
- {* fn phi => fn ss => try Group_Cancel.cancel_eq_conv *}
+ \<open>fn phi => fn ss => try Group_Cancel.cancel_eq_conv\<close>
simproc_setup group_cancel_le ("a \<le> (b::'a::ordered_ab_group_add)") =
- {* fn phi => fn ss => try Group_Cancel.cancel_le_conv *}
+ \<open>fn phi => fn ss => try Group_Cancel.cancel_le_conv\<close>
simproc_setup group_cancel_less ("a < (b::'a::ordered_ab_group_add)") =
- {* fn phi => fn ss => try Group_Cancel.cancel_less_conv *}
+ \<open>fn phi => fn ss => try Group_Cancel.cancel_less_conv\<close>
class linordered_ab_semigroup_add =
linorder + ordered_ab_semigroup_add
@@ -1380,7 +1380,7 @@
lemmas diff_diff_eq = diff_diff_add -- \<open>FIXME duplicate\<close>
-subsection {* Tools setup *}
+subsection \<open>Tools setup\<close>
lemma add_mono_thms_linordered_semiring:
fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
--- a/src/HOL/Groups_Big.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Groups_Big.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
with contributions by Jeremy Avigad
*)
-section {* Big sum and product over finite (non-empty) sets *}
+section \<open>Big sum and product over finite (non-empty) sets\<close>
theory Groups_Big
imports Finite_Set
begin
-subsection {* Generic monoid operation over a set *}
+subsection \<open>Generic monoid operation over a set\<close>
no_notation times (infixl "*" 70)
no_notation Groups.one ("1")
@@ -44,9 +44,9 @@
assumes "finite A" and "x \<in> A"
shows "F g A = g x * F g (A - {x})"
proof -
- from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ from \<open>x \<in> A\<close> obtain B where A: "A = insert x B" and "x \<notin> B"
by (auto dest: mk_disjoint_insert)
- moreover from `finite A` A have "finite B" by simp
+ moreover from \<open>finite A\<close> A have "finite B" by simp
ultimately show ?thesis by simp
qed
@@ -67,7 +67,7 @@
lemma union_inter:
assumes "finite A" and "finite B"
shows "F g (A \<union> B) * F g (A \<inter> B) = F g A * F g B"
- -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
+ -- \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close>
using assms proof (induct A)
case empty then show ?case by simp
next
@@ -140,7 +140,7 @@
assumes "A = B"
assumes g_h: "\<And>x. x \<in> B \<Longrightarrow> g x = h x"
shows "F g A = F h B"
- using g_h unfolding `A = B`
+ using g_h unfolding \<open>A = B\<close>
by (induct B rule: infinite_finite_induct) auto
lemma strong_cong [cong]:
@@ -207,9 +207,9 @@
assumes "finite T" and "S \<subseteq> T" and "\<forall>i \<in> T - S. h i = 1"
and "\<And>x. x \<in> S \<Longrightarrow> g x = h x" shows "F g S = F h T"
proof-
- have eq: "T = S \<union> (T - S)" using `S \<subseteq> T` by blast
- have d: "S \<inter> (T - S) = {}" using `S \<subseteq> T` by blast
- from `finite T` `S \<subseteq> T` have f: "finite S" "finite (T - S)"
+ have eq: "T = S \<union> (T - S)" using \<open>S \<subseteq> T\<close> by blast
+ have d: "S \<inter> (T - S) = {}" using \<open>S \<subseteq> T\<close> by blast
+ from \<open>finite T\<close> \<open>S \<subseteq> T\<close> have f: "finite S" "finite (T - S)"
by (auto intro: finite_subset)
show ?thesis using assms(4)
by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)])
@@ -268,7 +268,7 @@
also have "\<dots> = F g (T - T')"
using bij by (rule reindex_bij_betw)
also have "\<dots> = F g T"
- using nn `finite S` by (intro mono_neutral_cong_left) auto
+ using nn \<open>finite S\<close> by (intro mono_neutral_cong_left) auto
finally show ?thesis .
qed simp
qed
@@ -280,7 +280,7 @@
proof (subst reindex_bij_betw_not_neutral [symmetric])
show "bij_betw h (A - {x \<in> A. (g \<circ> h) x = 1}) (h ` A - h ` {x \<in> A. (g \<circ> h) x = 1})"
using nz by (auto intro!: inj_onI simp: bij_betw_def)
-qed (insert `finite A`, auto)
+qed (insert \<open>finite A\<close>, auto)
lemma reindex_bij_witness_not_neutral:
assumes fin: "finite S'" "finite T'"
@@ -369,7 +369,7 @@
have "\<forall>i\<in>A - A \<inter> B. (if i \<in> A \<inter> B then g i else 1) = 1"
by simp
moreover have "A \<inter> B \<subseteq> A" by blast
- ultimately have "F ?g (A \<inter> B) = F ?g A" using `finite A`
+ ultimately have "F ?g (A \<inter> B) = F ?g A" using \<open>finite A\<close>
by (intro mono_neutral_left) auto
then show ?thesis by simp
qed
@@ -395,7 +395,7 @@
and H: "F g (\<Union>B) = (F o F) g B" by auto
then have "F g (A \<union> \<Union>B) = F g A * F g (\<Union>B)"
by (simp add: union_inter_neutral)
- with `finite B` `A \<notin> B` show ?case
+ with \<open>finite B\<close> \<open>A \<notin> B\<close> show ?case
by (simp add: H)
qed
@@ -430,7 +430,7 @@
assumes trivial: "\<And>a. a \<in> C - A \<Longrightarrow> g a = 1" "\<And>b. b \<in> C - B \<Longrightarrow> h b = 1"
shows "F g A = F h B \<longleftrightarrow> F g C = F h C"
proof -
- from `finite C` subset have
+ from \<open>finite C\<close> subset have
"finite A" and "finite B" and "finite (C - A)" and "finite (C - B)"
by (auto elim: finite_subset)
from subset have [simp]: "A - (C - A) = A" by auto
@@ -438,12 +438,12 @@
from subset have "C = A \<union> (C - A)" by auto
then have "F g C = F g (A \<union> (C - A))" by simp
also have "\<dots> = F g (A - (C - A)) * F g (C - A - A) * F g (A \<inter> (C - A))"
- using `finite A` `finite (C - A)` by (simp only: union_diff2)
+ using \<open>finite A\<close> \<open>finite (C - A)\<close> by (simp only: union_diff2)
finally have P: "F g C = F g A" using trivial by simp
from subset have "C = B \<union> (C - B)" by auto
then have "F h C = F h (B \<union> (C - B))" by simp
also have "\<dots> = F h (B - (C - B)) * F h (C - B - B) * F h (B \<inter> (C - B))"
- using `finite B` `finite (C - B)` by (simp only: union_diff2)
+ using \<open>finite B\<close> \<open>finite (C - B)\<close> by (simp only: union_diff2)
finally have Q: "F h C = F h B" using trivial by simp
from P Q show ?thesis by simp
qed
@@ -462,7 +462,7 @@
notation Groups.one ("1")
-subsection {* Generalized summation over a set *}
+subsection \<open>Generalized summation over a set\<close>
context comm_monoid_add
begin
@@ -486,8 +486,8 @@
end
-text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
-written @{text"\<Sum>x\<in>A. e"}. *}
+text\<open>Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
+written @{text"\<Sum>x\<in>A. e"}.\<close>
syntax
"_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_./ _)" [0, 51, 10] 10)
@@ -496,12 +496,12 @@
syntax (HTML output)
"_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(2\<Sum>_\<in>_./ _)" [0, 51, 10] 10)
-translations -- {* Beware of argument permutation! *}
+translations -- \<open>Beware of argument permutation!\<close>
"SUM i:A. b" == "CONST setsum (%i. b) A"
"\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
-text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Sum>x|P. e"}. *}
+text\<open>Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Sum>x|P. e"}.\<close>
syntax
"_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
@@ -514,7 +514,7 @@
"SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
"\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
-print_translation {*
+print_translation \<open>
let
fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
if x <> y then raise Match
@@ -528,9 +528,9 @@
end
| setsum_tr' _ = raise Match;
in [(@{const_syntax setsum}, K setsum_tr')] end
-*}
+\<close>
-text {* TODO generalization candidates *}
+text \<open>TODO generalization candidates\<close>
lemma setsum_image_gen:
assumes fS: "finite S"
@@ -545,7 +545,7 @@
qed
-subsubsection {* Properties in more restricted classes of structures *}
+subsubsection \<open>Properties in more restricted classes of structures\<close>
lemma setsum_Un: "finite A ==> finite B ==>
(setsum f (A Un B) :: 'a :: ab_group_add) =
@@ -618,15 +618,15 @@
proof-
from assms(3) obtain a where a: "a:A" "f a < g a" by blast
have "setsum f A = setsum f ((A-{a}) \<union> {a})"
- by(simp add:insert_absorb[OF `a:A`])
+ by(simp add:insert_absorb[OF \<open>a:A\<close>])
also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
- using `finite A` by(subst setsum.union_disjoint) auto
+ using \<open>finite A\<close> by(subst setsum.union_disjoint) auto
also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
by(rule setsum_mono)(simp add: assms(2))
also have "setsum f {a} < setsum g {a}" using a by simp
also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
- using `finite A` by(subst setsum.union_disjoint[symmetric]) auto
- also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
+ using \<open>finite A\<close> by(subst setsum.union_disjoint[symmetric]) auto
+ also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF \<open>a:A\<close>])
finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono)
qed
@@ -880,7 +880,7 @@
lemma setsum_Un_nat: "finite A ==> finite B ==>
(setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
- -- {* For the natural numbers, we have subtraction. *}
+ -- \<open>For the natural numbers, we have subtraction.\<close>
by (subst setsum.union_inter [symmetric], auto simp add: algebra_simps)
lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
@@ -932,7 +932,7 @@
by (induct A rule: infinite_finite_induct) simp_all
-subsubsection {* Cardinality as special case of @{const setsum} *}
+subsubsection \<open>Cardinality as special case of @{const setsum}\<close>
lemma card_eq_setsum:
"card A = setsum (\<lambda>x. 1) A"
@@ -1008,7 +1008,7 @@
by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos)
-subsubsection {* Cardinality of products *}
+subsubsection \<open>Cardinality of products\<close>
lemma card_SigmaI [simp]:
"\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
@@ -1029,7 +1029,7 @@
by (simp add: card_cartesian_product)
-subsection {* Generalized product over a set *}
+subsection \<open>Generalized product over a set\<close>
context comm_monoid_mult
begin
@@ -1060,12 +1060,12 @@
syntax (HTML output)
"_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(2\<Prod>_\<in>_./ _)" [0, 51, 10] 10)
-translations -- {* Beware of argument permutation! *}
+translations -- \<open>Beware of argument permutation!\<close>
"PROD i:A. b" == "CONST setprod (%i. b) A"
"\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
-text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Prod>x|P. e"}. *}
+text\<open>Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Prod>x|P. e"}.\<close>
syntax
"_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(4PROD _ |/ _./ _)" [0,0,10] 10)
@@ -1102,7 +1102,7 @@
end
-subsubsection {* Properties in more restricted classes of structures *}
+subsubsection \<open>Properties in more restricted classes of structures\<close>
context comm_semiring_1
begin
@@ -1111,11 +1111,11 @@
assumes "finite A" and "a \<in> A" and "b = f a"
shows "b dvd setprod f A"
proof -
- from `finite A` have "setprod f (insert a (A - {a})) = f a * setprod f (A - {a})"
+ from \<open>finite A\<close> have "setprod f (insert a (A - {a})) = f a * setprod f (A - {a})"
by (intro setprod.insert) auto
- also from `a \<in> A` have "insert a (A - {a}) = A" by blast
+ also from \<open>a \<in> A\<close> have "insert a (A - {a}) = A" by blast
finally have "setprod f A = f a * setprod f (A - {a})" .
- with `b = f a` show ?thesis by simp
+ with \<open>b = f a\<close> show ?thesis by simp
qed
lemma dvd_setprodI [intro]:
@@ -1170,7 +1170,7 @@
next
case False with insert have "a \<in> B" by simp
def C \<equiv> "B - {a}"
- with `finite B` `a \<in> B`
+ with \<open>finite B\<close> \<open>a \<in> B\<close>
have *: "B = insert a C" "finite C" "a \<notin> C" by auto
with insert show ?thesis by (auto simp add: insert_commute ac_simps)
qed
--- a/src/HOL/Groups_List.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Groups_List.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,6 +1,6 @@
(* Author: Tobias Nipkow, TU Muenchen *)
-section {* Sum and product over lists *}
+section \<open>Sum and product over lists\<close>
theory Groups_List
imports List
@@ -56,7 +56,7 @@
notation Groups.one ("1")
-subsection {* List summation *}
+subsection \<open>List summation\<close>
context monoid_add
begin
@@ -101,7 +101,7 @@
end
-text {* Some syntactic sugar for summing a function over a list: *}
+text \<open>Some syntactic sugar for summing a function over a list:\<close>
syntax
"_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10)
@@ -110,11 +110,11 @@
syntax (HTML output)
"_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
-translations -- {* Beware of argument permutation! *}
+translations -- \<open>Beware of argument permutation!\<close>
"SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)"
"\<Sum>x\<leftarrow>xs. b" == "CONST listsum (CONST map (%x. b) xs)"
-text {* TODO duplicates *}
+text \<open>TODO duplicates\<close>
lemmas listsum_simps = listsum.Nil listsum.Cons
lemmas listsum_append = listsum.append
lemmas listsum_rev = listsum.rev
@@ -190,7 +190,7 @@
"(\<Sum>x\<leftarrow>xs. 0) = 0"
by (induct xs) (simp_all add: distrib_right)
-text{* For non-Abelian groups @{text xs} needs to be reversed on one side: *}
+text\<open>For non-Abelian groups @{text xs} needs to be reversed on one side:\<close>
lemma (in ab_group_add) uminus_listsum_map:
"- listsum (map f xs) = listsum (map (uminus \<circ> f) xs)"
by (induct xs) simp_all
@@ -232,7 +232,7 @@
"listsum (map f [k..l]) = setsum f (set [k..l])"
by (simp add: listsum_distinct_conv_setsum_set)
-text {* General equivalence between @{const listsum} and @{const setsum} *}
+text \<open>General equivalence between @{const listsum} and @{const setsum}\<close>
lemma (in monoid_add) listsum_setsum_nth:
"listsum xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
using interv_listsum_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
@@ -245,14 +245,14 @@
proof cases
assume "x \<in> set xs"
have "?l = f x + (\<Sum>x\<in>set xs. count_list xs x * f x)" by (simp add: Cons.IH)
- also have "set xs = insert x (set xs - {x})" using `x \<in> set xs`by blast
+ also have "set xs = insert x (set xs - {x})" using \<open>x \<in> set xs\<close>by blast
also have "f x + (\<Sum>x\<in>insert x (set xs - {x}). count_list xs x * f x) = ?r"
by (simp add: setsum.insert_remove eq_commute)
finally show ?thesis .
next
assume "x \<notin> set xs"
hence "\<And>xa. xa \<in> set xs \<Longrightarrow> x \<noteq> xa" by blast
- thus ?thesis by (simp add: Cons.IH `x \<notin> set xs`)
+ thus ?thesis by (simp add: Cons.IH \<open>x \<notin> set xs\<close>)
qed
qed simp
@@ -270,7 +270,7 @@
qed
-subsection {* Further facts about @{const List.n_lists} *}
+subsection \<open>Further facts about @{const List.n_lists}\<close>
lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n"
by (induct n) (auto simp add: comp_def length_concat listsum_triv)
@@ -298,7 +298,7 @@
qed
-subsection {* Tools setup *}
+subsection \<open>Tools setup\<close>
lemmas setsum_code = setsum.set_conv_list
@@ -325,7 +325,7 @@
end
-subsection {* List product *}
+subsection \<open>List product\<close>
context monoid_mult
begin
@@ -370,7 +370,7 @@
end
-text {* Some syntactic sugar: *}
+text \<open>Some syntactic sugar:\<close>
syntax
"_listprod" :: "pttrn => 'a list => 'b => 'b" ("(3PROD _<-_. _)" [0, 51, 10] 10)
@@ -379,7 +379,7 @@
syntax (HTML output)
"_listprod" :: "pttrn => 'a list => 'b => 'b" ("(3\<Prod>_\<leftarrow>_. _)" [0, 51, 10] 10)
-translations -- {* Beware of argument permutation! *}
+translations -- \<open>Beware of argument permutation!\<close>
"PROD x<-xs. b" == "CONST listprod (CONST map (%x. b) xs)"
"\<Prod>x\<leftarrow>xs. b" == "CONST listprod (CONST map (%x. b) xs)"
--- a/src/HOL/HOL.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/HOL.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson
*)
-section {* The basis of Higher-Order Logic *}
+section \<open>The basis of Higher-Order Logic\<close>
theory HOL
imports Pure "~~/src/Tools/Code_Generator"
@@ -53,13 +53,13 @@
\<close>
-subsection {* Primitive logic *}
+subsection \<open>Primitive logic\<close>
-subsubsection {* Core syntax *}
+subsubsection \<open>Core syntax\<close>
-setup {* Axclass.class_axiomatization (@{binding type}, []) *}
+setup \<open>Axclass.class_axiomatization (@{binding type}, [])\<close>
default_sort type
-setup {* Object_Logic.add_base_sort @{sort type} *}
+setup \<open>Object_Logic.add_base_sort @{sort type}\<close>
axiomatization where fun_arity: "OFCLASS('a \<Rightarrow> 'b, type_class)"
instance "fun" :: (type, type) type by (rule fun_arity)
@@ -90,7 +90,7 @@
Ex1 :: "('a => bool) => bool" (binder "EX! " 10)
-subsubsection {* Additional concrete syntax *}
+subsubsection \<open>Additional concrete syntax\<close>
notation (output)
eq (infix "=" 50)
@@ -127,11 +127,11 @@
syntax "_The" :: "[pttrn, bool] => 'a" ("(3THE _./ _)" [0, 10] 10)
translations "THE x. P" == "CONST The (%x. P)"
-print_translation {*
+print_translation \<open>
[(@{const_syntax The}, fn _ => fn [Abs abs] =>
let val (x, t) = Syntax_Trans.atomic_abs_tr' abs
in Syntax.const @{syntax_const "_The"} $ x $ t end)]
-*} -- {* To avoid eta-contraction of body *}
+\<close> -- \<open>To avoid eta-contraction of body\<close>
nonterminal letbinds and letbind
syntax
@@ -165,15 +165,15 @@
Ex1 (binder "?! " 10)
-subsubsection {* Axioms and basic definitions *}
+subsubsection \<open>Axioms and basic definitions\<close>
axiomatization where
refl: "t = (t::'a)" and
subst: "s = t \<Longrightarrow> P s \<Longrightarrow> P t" and
ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"
- -- {*Extensionality is built into the meta-logic, and this rule expresses
+ -- \<open>Extensionality is built into the meta-logic, and this rule expresses
a related property. It is an eta-expanded version of the traditional
- rule, and similar to the ABS rule of HOL*} and
+ rule, and similar to the ABS rule of HOL\<close> and
the_eq_trivial: "(THE x. x = a) = (a::'a)"
@@ -209,9 +209,9 @@
class default = fixes default :: 'a
-subsection {* Fundamental rules *}
+subsection \<open>Fundamental rules\<close>
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
lemma sym: "s = t ==> t = s"
by (erule subst) (rule refl)
@@ -230,7 +230,7 @@
shows "A = B"
by (unfold meq) (rule refl)
-text {* Useful with @{text erule} for proving equalities from known equalities. *}
+text \<open>Useful with @{text erule} for proving equalities from known equalities.\<close>
(* a = b
| |
c = d *)
@@ -241,7 +241,7 @@
apply assumption+
done
-text {* For calculational reasoning: *}
+text \<open>For calculational reasoning:\<close>
lemma forw_subst: "a = b ==> P b ==> P a"
by (rule ssubst)
@@ -250,15 +250,15 @@
by (rule subst)
-subsubsection {* Congruence rules for application *}
+subsubsection \<open>Congruence rules for application\<close>
-text {* Similar to @{text AP_THM} in Gordon's HOL. *}
+text \<open>Similar to @{text AP_THM} in Gordon's HOL.\<close>
lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)"
apply (erule subst)
apply (rule refl)
done
-text {* Similar to @{text AP_TERM} in Gordon's HOL and FOL's @{text subst_context}. *}
+text \<open>Similar to @{text AP_TERM} in Gordon's HOL and FOL's @{text subst_context}.\<close>
lemma arg_cong: "x=y ==> f(x)=f(y)"
apply (erule subst)
apply (rule refl)
@@ -274,10 +274,10 @@
apply (rule refl)
done
-ML {* fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong} *}
+ML \<open>fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\<close>
-subsubsection {* Equality of booleans -- iff *}
+subsubsection \<open>Equality of booleans -- iff\<close>
lemma iffI: assumes "P ==> Q" and "Q ==> P" shows "P=Q"
by (iprover intro: iff [THEN mp, THEN mp] impI assms)
@@ -301,7 +301,7 @@
by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1])
-subsubsection {*True*}
+subsubsection \<open>True\<close>
lemma TrueI: "True"
unfolding True_def by (rule refl)
@@ -313,7 +313,7 @@
by (erule iffD2) (rule TrueI)
-subsubsection {*Universal quantifier*}
+subsubsection \<open>Universal quantifier\<close>
lemma allI: assumes "!!x::'a. P(x)" shows "ALL x. P(x)"
unfolding All_def by (iprover intro: ext eqTrueI assms)
@@ -337,12 +337,12 @@
by (iprover intro: minor major major [THEN spec])
-subsubsection {* False *}
+subsubsection \<open>False\<close>
-text {*
+text \<open>
Depends upon @{text spec}; it is impossible to do propositional
logic before quantifiers!
-*}
+\<close>
lemma FalseE: "False ==> P"
apply (unfold False_def)
@@ -353,7 +353,7 @@
by (erule eqTrueE [THEN FalseE])
-subsubsection {* Negation *}
+subsubsection \<open>Negation\<close>
lemma notI:
assumes "P ==> False"
@@ -383,7 +383,7 @@
by (erule notE [THEN notI]) (erule meta_mp)
-subsubsection {*Implication*}
+subsubsection \<open>Implication\<close>
lemma impE:
assumes "P-->Q" "P" "Q ==> R"
@@ -414,7 +414,7 @@
by (erule subst, erule ssubst, assumption)
-subsubsection {*Existential quantifier*}
+subsubsection \<open>Existential quantifier\<close>
lemma exI: "P x ==> EX x::'a. P x"
apply (unfold Ex_def)
@@ -430,7 +430,7 @@
done
-subsubsection {*Conjunction*}
+subsubsection \<open>Conjunction\<close>
lemma conjI: "[| P; Q |] ==> P&Q"
apply (unfold and_def)
@@ -461,7 +461,7 @@
by (iprover intro: conjI assms)
-subsubsection {*Disjunction*}
+subsubsection \<open>Disjunction\<close>
lemma disjI1: "P ==> P|Q"
apply (unfold or_def)
@@ -482,7 +482,7 @@
major [unfolded or_def, THEN spec, THEN mp, THEN mp])
-subsubsection {*Classical logic*}
+subsubsection \<open>Classical logic\<close>
lemma classical:
assumes prem: "~P ==> P"
@@ -520,14 +520,14 @@
by (iprover intro: classical p1 p2 notE)
-subsubsection {*Unique existence*}
+subsubsection \<open>Unique existence\<close>
lemma ex1I:
assumes "P a" "!!x. P(x) ==> x=a"
shows "EX! x. P(x)"
by (unfold Ex1_def, iprover intro: assms exI conjI allI impI)
-text{*Sometimes easier to use: the premises have no shared variables. Safe!*}
+text\<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close>
lemma ex_ex1I:
assumes ex_prem: "EX x. P(x)"
and eq: "!!x y. [| P(x); P(y) |] ==> x=y"
@@ -550,7 +550,7 @@
done
-subsubsection {*Classical intro rules for disjunction and existential quantifiers*}
+subsubsection \<open>Classical intro rules for disjunction and existential quantifiers\<close>
lemma disjCI:
assumes "~Q ==> P" shows "P|Q"
@@ -561,10 +561,10 @@
lemma excluded_middle: "~P | P"
by (iprover intro: disjCI)
-text {*
+text \<open>
case distinction as a natural deduction rule.
Note that @{term "~P"} is the second case, not the first
-*}
+\<close>
lemma case_split [case_names True False]:
assumes prem1: "P ==> Q"
and prem2: "~P ==> Q"
@@ -611,7 +611,7 @@
done
-subsubsection {* Intuitionistic Reasoning *}
+subsubsection \<open>Intuitionistic Reasoning\<close>
lemma impE':
assumes 1: "P --> Q"
@@ -655,7 +655,7 @@
and [Pure.elim?] = iffD1 iffD2 impE
-subsubsection {* Atomizing meta-level connectives *}
+subsubsection \<open>Atomizing meta-level connectives\<close>
axiomatization where
eq_reflection: "x = y \<Longrightarrow> x \<equiv> y" (*admissible axiom*)
@@ -690,7 +690,7 @@
lemma atomize_eq [atomize, code]: "(x == y) == Trueprop (x = y)"
proof
assume "x == y"
- show "x = y" by (unfold `x == y`) (rule refl)
+ show "x = y" by (unfold \<open>x == y\<close>) (rule refl)
next
assume "x = y"
then show "x == y" by (rule eq_reflection)
@@ -717,7 +717,7 @@
and [symmetric, defn] = atomize_all atomize_imp atomize_eq
-subsubsection {* Atomizing elimination rules *}
+subsubsection \<open>Atomizing elimination rules\<close>
lemma atomize_exL[atomize_elim]: "(!!x. P x ==> Q) == ((EX x. P x) ==> Q)"
by rule iprover+
@@ -731,23 +731,23 @@
lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop A" ..
-subsection {* Package setup *}
+subsection \<open>Package setup\<close>
ML_file "Tools/hologic.ML"
-subsubsection {* Sledgehammer setup *}
+subsubsection \<open>Sledgehammer setup\<close>
-text {*
+text \<open>
Theorems blacklisted to Sledgehammer. These theorems typically produce clauses
that are prolific (match too many equality or membership literals) and relate to
seldom-used facts. Some duplicate other rules.
-*}
+\<close>
named_theorems no_atp "theorems that should be filtered out by Sledgehammer"
-subsubsection {* Classical Reasoner setup *}
+subsubsection \<open>Classical Reasoner setup\<close>
lemma imp_elim: "P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R"
by (rule classical) iprover
@@ -758,7 +758,7 @@
lemma thin_refl:
"\<And>X. \<lbrakk> x=x; PROP W \<rbrakk> \<Longrightarrow> PROP W" .
-ML {*
+ML \<open>
structure Hypsubst = Hypsubst
(
val dest_eq = HOLogic.dest_eq
@@ -786,9 +786,9 @@
structure Basic_Classical: BASIC_CLASSICAL = Classical;
open Basic_Classical;
-*}
+\<close>
-setup {*
+setup \<open>
(*prevent substitution on bool*)
let
fun non_bool_eq (@{const_name HOL.eq}, Type (_, [T, _])) = T <> @{typ bool}
@@ -801,7 +801,7 @@
in
Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac' ctxt ORELSE' tac)
end
-*}
+\<close>
declare iffI [intro!]
and notI [intro!]
@@ -824,7 +824,7 @@
declare exE [elim!]
allE [elim]
-ML {* val HOL_cs = claset_of @{context} *}
+ML \<open>val HOL_cs = claset_of @{context}\<close>
lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P"
apply (erule swap)
@@ -848,11 +848,11 @@
apply (rule prem)
apply assumption
apply (rule allI)+
-apply (tactic {* eresolve_tac @{context} [Classical.dup_elim @{context} @{thm allE}] 1 *})
+apply (tactic \<open>eresolve_tac @{context} [Classical.dup_elim @{context} @{thm allE}] 1\<close>)
apply iprover
done
-ML {*
+ML \<open>
structure Blast = Blast
(
structure Classical = Classical
@@ -864,10 +864,10 @@
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
);
val blast_tac = Blast.blast_tac;
-*}
+\<close>
-subsubsection {*THE: definite description operator*}
+subsubsection \<open>THE: definite description operator\<close>
lemma the_equality [intro]:
assumes "P a"
@@ -900,7 +900,7 @@
by blast
-subsubsection {* Simplifier *}
+subsubsection \<open>Simplifier\<close>
lemma eta_contract_eq: "(%s. f s) = f" ..
@@ -980,7 +980,7 @@
lemma imp_conjL: "((P&Q) -->R) = (P --> (Q --> R))" by iprover
lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover
-text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
+text \<open>These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}.\<close>
lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast
lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast
@@ -995,7 +995,7 @@
lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast
lemma not_iff: "(P~=Q) = (P = (~Q))" by blast
lemma disj_not1: "(~P | Q) = (P --> Q)" by blast
-lemma disj_not2: "(P | ~Q) = (Q --> P)" -- {* changes orientation :-( *}
+lemma disj_not2: "(P | ~Q) = (Q --> P)" -- \<open>changes orientation :-(\<close>
by blast
lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast
@@ -1003,8 +1003,8 @@
lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q"
- -- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
- -- {* cases boil down to the same thing. *}
+ -- \<open>Avoids duplication of subgoals after @{text split_if}, when the true and false\<close>
+ -- \<open>cases boil down to the same thing.\<close>
by blast
lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast
@@ -1018,9 +1018,9 @@
lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover
lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover
-text {*
+text \<open>
\medskip The @{text "&"} congruence rule: not included by default!
- May slow rewrite proofs down by as much as 50\% *}
+ May slow rewrite proofs down by as much as 50\%\<close>
lemma conj_cong:
"(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))"
@@ -1030,14 +1030,14 @@
"(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))"
by iprover
-text {* The @{text "|"} congruence rule: not included by default! *}
+text \<open>The @{text "|"} congruence rule: not included by default!\<close>
lemma disj_cong:
"(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))"
by blast
-text {* \medskip if-then-else rules *}
+text \<open>\medskip if-then-else rules\<close>
lemma if_True [code]: "(if True then x else y) = x"
by (unfold If_def) blast
@@ -1070,17 +1070,17 @@
lemma if_bool_eq_conj:
"(if P then Q else R) = ((P-->Q) & (~P-->R))"
- -- {* This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol. *}
+ -- \<open>This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol.\<close>
by (rule split_if)
lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))"
- -- {* And this form is useful for expanding @{text "if"}s on the LEFT. *}
+ -- \<open>And this form is useful for expanding @{text "if"}s on the LEFT.\<close>
by (simplesubst split_if) blast
lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover
lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover
-text {* \medskip let rules for simproc *}
+text \<open>\medskip let rules for simproc\<close>
lemma Let_folded: "f x \<equiv> g x \<Longrightarrow> Let x f \<equiv> Let x g"
by (unfold Let_def)
@@ -1088,11 +1088,11 @@
lemma Let_unfold: "f x \<equiv> g \<Longrightarrow> Let x f \<equiv> g"
by (unfold Let_def)
-text {*
+text \<open>
The following copy of the implication operator is useful for
fine-tuning congruence rules. It instructs the simplifier to simplify
its premise.
-*}
+\<close>
definition simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) where
"simp_implies \<equiv> op ==>"
@@ -1159,19 +1159,19 @@
by blast
ML_file "Tools/simpdata.ML"
-ML {* open Simpdata *}
+ML \<open>open Simpdata\<close>
-setup {*
+setup \<open>
map_theory_simpset (put_simpset HOL_basic_ss) #>
Simplifier.method_setup Splitter.split_modifiers
-*}
+\<close>
-simproc_setup defined_Ex ("EX x. P x") = {* fn _ => Quantifier1.rearrange_ex *}
-simproc_setup defined_All ("ALL x. P x") = {* fn _ => Quantifier1.rearrange_all *}
+simproc_setup defined_Ex ("EX x. P x") = \<open>fn _ => Quantifier1.rearrange_ex\<close>
+simproc_setup defined_All ("ALL x. P x") = \<open>fn _ => Quantifier1.rearrange_all\<close>
-text {* Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}: *}
+text \<open>Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}:\<close>
-simproc_setup neq ("x = y") = {* fn _ =>
+simproc_setup neq ("x = y") = \<open>fn _ =>
let
val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI};
fun is_neq eq lhs rhs thm =
@@ -1188,9 +1188,9 @@
| NONE => NONE)
| _ => NONE);
in proc end;
-*}
+\<close>
-simproc_setup let_simp ("Let x f") = {*
+simproc_setup let_simp ("Let x f") = \<open>
let
val (f_Let_unfold, x_Let_unfold) =
let val [(_ $ (f $ x) $ _)] = Thm.prems_of @{thm Let_unfold}
@@ -1253,7 +1253,7 @@
end
| _ => NONE)
end
-end *}
+end\<close>
lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
proof
@@ -1283,7 +1283,7 @@
"!!P Q. (EX x. P | Q x) = (P | (EX x. Q x))"
"!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)"
"!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))"
- -- {* Miniscoping: pushing in existential quantifiers. *}
+ -- \<open>Miniscoping: pushing in existential quantifiers.\<close>
by (iprover | blast)+
lemma all_simps:
@@ -1293,7 +1293,7 @@
"!!P Q. (ALL x. P | Q x) = (P | (ALL x. Q x))"
"!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)"
"!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))"
- -- {* Miniscoping: pushing in universal quantifiers. *}
+ -- \<open>Miniscoping: pushing in universal quantifiers.\<close>
by (iprover | blast)+
lemmas [simp] =
@@ -1330,9 +1330,9 @@
lemmas [cong] = imp_cong simp_implies_cong
lemmas [split] = split_if
-ML {* val HOL_ss = simpset_of @{context} *}
+ML \<open>val HOL_ss = simpset_of @{context}\<close>
-text {* Simplifies x assuming c and y assuming ~c *}
+text \<open>Simplifies x assuming c and y assuming ~c\<close>
lemma if_cong:
assumes "b = c"
and "c \<Longrightarrow> x = u"
@@ -1340,20 +1340,20 @@
shows "(if b then x else y) = (if c then u else v)"
using assms by simp
-text {* Prevents simplification of x and y:
- faster and allows the execution of functional programs. *}
+text \<open>Prevents simplification of x and y:
+ faster and allows the execution of functional programs.\<close>
lemma if_weak_cong [cong]:
assumes "b = c"
shows "(if b then x else y) = (if c then x else y)"
using assms by (rule arg_cong)
-text {* Prevents simplification of t: much faster *}
+text \<open>Prevents simplification of t: much faster\<close>
lemma let_weak_cong:
assumes "a = b"
shows "(let x = a in t x) = (let x = b in t x)"
using assms by (rule arg_cong)
-text {* To tidy up the result of a simproc. Only the RHS will be simplified. *}
+text \<open>To tidy up the result of a simproc. Only the RHS will be simplified.\<close>
lemma eq_cong2:
assumes "u = u'"
shows "(t \<equiv> u) \<equiv> (t \<equiv> u')"
@@ -1363,23 +1363,23 @@
"f (if c then x else y) = (if c then f x else f y)"
by simp
-text{*As a simplification rule, it replaces all function equalities by
- first-order equalities.*}
+text\<open>As a simplification rule, it replaces all function equalities by
+ first-order equalities.\<close>
lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
by auto
-subsubsection {* Generic cases and induction *}
+subsubsection \<open>Generic cases and induction\<close>
-text {* Rule projections: *}
-ML {*
+text \<open>Rule projections:\<close>
+ML \<open>
structure Project_Rule = Project_Rule
(
val conjunct1 = @{thm conjunct1}
val conjunct2 = @{thm conjunct2}
val mp = @{thm mp}
);
-*}
+\<close>
context
begin
@@ -1435,10 +1435,10 @@
lemma induct_trueI: "induct_true"
by (simp add: induct_true_def)
-text {* Method setup. *}
+text \<open>Method setup.\<close>
ML_file "~~/src/Tools/induct.ML"
-ML {*
+ML \<open>
structure Induct = Induct
(
val cases_default = @{thm case_split}
@@ -1450,11 +1450,11 @@
| dest_def _ = NONE
fun trivial_tac ctxt = match_tac ctxt @{thms induct_trueI}
)
-*}
+\<close>
ML_file "~~/src/Tools/induction.ML"
-declaration {*
+declaration \<open>
fn _ => Induct.map_simpset (fn ss => ss
addsimprocs
[Simplifier.simproc_global @{theory} "swap_induct_false"
@@ -1479,9 +1479,9 @@
|> Simplifier.set_mksimps (fn ctxt =>
Simpdata.mksimps Simpdata.mksimps_pairs ctxt #>
map (rewrite_rule ctxt (map Thm.symmetric @{thms induct_rulify_fallback}))))
-*}
+\<close>
-text {* Pre-simplification of induction and cases rules *}
+text \<open>Pre-simplification of induction and cases rules\<close>
lemma [induct_simp]: "(\<And>x. induct_equal x t \<Longrightarrow> PROP P x) \<equiv> PROP P t"
unfolding induct_equal_def
@@ -1538,10 +1538,10 @@
ML_file "~~/src/Tools/induct_tacs.ML"
-subsubsection {* Coherent logic *}
+subsubsection \<open>Coherent logic\<close>
ML_file "~~/src/Tools/coherent.ML"
-ML {*
+ML \<open>
structure Coherent = Coherent
(
val atomize_elimL = @{thm atomize_elimL};
@@ -1550,12 +1550,12 @@
val atomize_disjL = @{thm atomize_disjL};
val operator_names = [@{const_name HOL.disj}, @{const_name HOL.conj}, @{const_name Ex}];
);
-*}
+\<close>
-subsubsection {* Reorienting equalities *}
+subsubsection \<open>Reorienting equalities\<close>
-ML {*
+ML \<open>
signature REORIENT_PROC =
sig
val add : (term -> bool) -> theory -> theory
@@ -1584,10 +1584,10 @@
| _ => NONE
end;
end;
-*}
+\<close>
-subsection {* Other simple lemmas and lemma duplicates *}
+subsection \<open>Other simple lemmas and lemma duplicates\<close>
lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x"
by blast+
@@ -1616,9 +1616,9 @@
"(\<not> \<not>(P)) = P"
by blast+
-subsection {* Basic ML bindings *}
+subsection \<open>Basic ML bindings\<close>
-ML {*
+ML \<open>
val FalseE = @{thm FalseE}
val Let_def = @{thm Let_def}
val TrueI = @{thm TrueI}
@@ -1667,16 +1667,16 @@
val subst = @{thm subst}
val sym = @{thm sym}
val trans = @{thm trans}
-*}
+\<close>
ML_file "Tools/cnf.ML"
-section {* @{text NO_MATCH} simproc *}
+section \<open>@{text NO_MATCH} simproc\<close>
-text {*
+text \<open>
The simplification procedure can be used to avoid simplification of terms of a certain form
-*}
+\<close>
definition NO_MATCH :: "'a \<Rightarrow> 'b \<Rightarrow> bool" where "NO_MATCH pat val \<equiv> True"
@@ -1684,23 +1684,23 @@
declare [[coercion_args NO_MATCH - -]]
-simproc_setup NO_MATCH ("NO_MATCH pat val") = {* fn _ => fn ctxt => fn ct =>
+simproc_setup NO_MATCH ("NO_MATCH pat val") = \<open>fn _ => fn ctxt => fn ct =>
let
val thy = Proof_Context.theory_of ctxt
val dest_binop = Term.dest_comb #> apfst (Term.dest_comb #> snd)
val m = Pattern.matches thy (dest_binop (Thm.term_of ct))
in if m then NONE else SOME @{thm NO_MATCH_def} end
-*}
+\<close>
-text {*
+text \<open>
This setup ensures that a rewrite rule of the form @{term "NO_MATCH pat val \<Longrightarrow> t"}
is only applied, if the pattern @{term pat} does not match the value @{term val}.
-*}
+\<close>
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
-subsubsection {* Generic code generator preprocessor setup *}
+subsubsection \<open>Generic code generator preprocessor setup\<close>
lemma conj_left_cong:
"P \<longleftrightarrow> Q \<Longrightarrow> P \<and> R \<longleftrightarrow> Q \<and> R"
@@ -1710,16 +1710,16 @@
"P \<longleftrightarrow> Q \<Longrightarrow> P \<or> R \<longleftrightarrow> Q \<or> R"
by (fact arg_cong)
-setup {*
+setup \<open>
Code_Preproc.map_pre (put_simpset HOL_basic_ss) #>
Code_Preproc.map_post (put_simpset HOL_basic_ss) #>
Code_Simp.map_ss (put_simpset HOL_basic_ss #>
Simplifier.add_cong @{thm conj_left_cong} #>
Simplifier.add_cong @{thm disj_left_cong})
-*}
+\<close>
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
class equal =
fixes equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
@@ -1740,16 +1740,16 @@
declare eq_equal [symmetric, code_post]
declare eq_equal [code]
-setup {*
+setup \<open>
Code_Preproc.map_pre (fn ctxt =>
ctxt addsimprocs [Simplifier.simproc_global_i @{theory} "equal" [@{term HOL.eq}]
(fn _ => fn Const (_, Type ("fun", [Type _, _])) => SOME @{thm eq_equal} | _ => NONE)])
-*}
+\<close>
-subsubsection {* Generic code generator foundation *}
+subsubsection \<open>Generic code generator foundation\<close>
-text {* Datatype @{typ bool} *}
+text \<open>Datatype @{typ bool}\<close>
code_datatype True False
@@ -1771,7 +1771,7 @@
and "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P"
and "(P \<longrightarrow> True) \<longleftrightarrow> True" by simp_all
-text {* More about @{typ prop} *}
+text \<open>More about @{typ prop}\<close>
lemma [code nbe]:
shows "(True \<Longrightarrow> PROP Q) \<equiv> PROP Q"
@@ -1784,7 +1784,7 @@
declare Trueprop_code [symmetric, code_post]
-text {* Equality *}
+text \<open>Equality\<close>
declare simp_thms(6) [code nbe]
@@ -1803,42 +1803,42 @@
"equal TYPE('a) TYPE('a) \<longleftrightarrow> True"
by (simp add: equal)
-setup {* Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"}) *}
+setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})\<close>
lemma equal_alias_cert: "OFCLASS('a, equal_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> equal)" (is "?ofclass \<equiv> ?equal")
proof
assume "PROP ?ofclass"
show "PROP ?equal"
- by (tactic {* ALLGOALS (resolve_tac @{context} [Thm.unconstrainT @{thm eq_equal}]) *})
- (fact `PROP ?ofclass`)
+ by (tactic \<open>ALLGOALS (resolve_tac @{context} [Thm.unconstrainT @{thm eq_equal}])\<close>)
+ (fact \<open>PROP ?ofclass\<close>)
next
assume "PROP ?equal"
show "PROP ?ofclass" proof
- qed (simp add: `PROP ?equal`)
+ qed (simp add: \<open>PROP ?equal\<close>)
qed
-setup {* Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>equal \<Rightarrow> 'a \<Rightarrow> bool"}) *}
+setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>equal \<Rightarrow> 'a \<Rightarrow> bool"})\<close>
-setup {* Nbe.add_const_alias @{thm equal_alias_cert} *}
+setup \<open>Nbe.add_const_alias @{thm equal_alias_cert}\<close>
-text {* Cases *}
+text \<open>Cases\<close>
lemma Let_case_cert:
assumes "CASE \<equiv> (\<lambda>x. Let x f)"
shows "CASE x \<equiv> f x"
using assms by simp_all
-setup {*
+setup \<open>
Code.add_case @{thm Let_case_cert} #>
Code.add_undefined @{const_name undefined}
-*}
+\<close>
declare [[code abort: undefined]]
-subsubsection {* Generic code generator target languages *}
+subsubsection \<open>Generic code generator target languages\<close>
-text {* type @{typ bool} *}
+text \<open>type @{typ bool}\<close>
code_printing
type_constructor bool \<rightharpoonup>
@@ -1885,14 +1885,14 @@
code_module Pure \<rightharpoonup>
(SML) HOL and (OCaml) HOL and (Haskell) HOL and (Scala) HOL
-text {* using built-in Haskell equality *}
+text \<open>using built-in Haskell equality\<close>
code_printing
type_class equal \<rightharpoonup> (Haskell) "Eq"
| constant HOL.equal \<rightharpoonup> (Haskell) infix 4 "=="
| constant HOL.eq \<rightharpoonup> (Haskell) infix 4 "=="
-text {* undefined *}
+text \<open>undefined\<close>
code_printing
constant undefined \<rightharpoonup>
@@ -1902,9 +1902,9 @@
and (Scala) "!sys.error(\"undefined\")"
-subsubsection {* Evaluation and normalization by evaluation *}
+subsubsection \<open>Evaluation and normalization by evaluation\<close>
-method_setup eval = {*
+method_setup eval = \<open>
let
fun eval_tac ctxt =
let val conv = Code_Runtime.dynamic_holds_conv ctxt
@@ -1915,25 +1915,25 @@
in
Scan.succeed (SIMPLE_METHOD' o eval_tac)
end
-*} "solve goal by evaluation"
+\<close> "solve goal by evaluation"
-method_setup normalization = {*
+method_setup normalization = \<open>
Scan.succeed (fn ctxt =>
SIMPLE_METHOD'
(CHANGED_PROP o
(CONVERSION (Nbe.dynamic_conv ctxt)
THEN_ALL_NEW (TRY o resolve_tac ctxt [TrueI]))))
-*} "solve goal by normalization"
+\<close> "solve goal by normalization"
-subsection {* Counterexample Search Units *}
+subsection \<open>Counterexample Search Units\<close>
-subsubsection {* Quickcheck *}
+subsubsection \<open>Quickcheck\<close>
quickcheck_params [size = 5, iterations = 50]
-subsubsection {* Nitpick setup *}
+subsubsection \<open>Nitpick setup\<close>
named_theorems nitpick_unfold "alternative definitions of constants as needed by Nitpick"
and nitpick_simp "equational specification of constants as needed by Nitpick"
@@ -1944,16 +1944,16 @@
if_bool_eq_disj [no_atp]
-subsection {* Preprocessing for the predicate compiler *}
+subsection \<open>Preprocessing for the predicate compiler\<close>
named_theorems code_pred_def "alternative definitions of constants for the Predicate Compiler"
and code_pred_inline "inlining definitions for the Predicate Compiler"
and code_pred_simp "simplification rules for the optimisations in the Predicate Compiler"
-subsection {* Legacy tactics and ML bindings *}
+subsection \<open>Legacy tactics and ML bindings\<close>
-ML {*
+ML \<open>
(* combination of (spec RS spec RS ...(j times) ... spec RS mp) *)
local
fun wrong_prem (Const (@{const_name All}, _) $ Abs (_, _, t)) = wrong_prem t
@@ -1971,7 +1971,7 @@
in
fun nnf_conv ctxt = Simplifier.rewrite (put_simpset nnf_ss ctxt);
end
-*}
+\<close>
hide_const (open) eq equal
--- a/src/HOL/Hilbert_Choice.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Hilbert_Choice.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,14 +3,14 @@
Copyright 2001 University of Cambridge
*)
-section {* Hilbert's Epsilon-Operator and the Axiom of Choice *}
+section \<open>Hilbert's Epsilon-Operator and the Axiom of Choice\<close>
theory Hilbert_Choice
imports Nat Wellfounded
keywords "specification" :: thy_goal
begin
-subsection {* Hilbert's epsilon *}
+subsection \<open>Hilbert's epsilon\<close>
axiomatization Eps :: "('a => bool) => 'a" where
someI: "P x ==> P (Eps P)"
@@ -24,11 +24,11 @@
translations
"SOME x. P" == "CONST Eps (%x. P)"
-print_translation {*
+print_translation \<open>
[(@{const_syntax Eps}, fn _ => fn [Abs abs] =>
let val (x, t) = Syntax_Trans.atomic_abs_tr' abs
in Syntax.const @{syntax_const "_Eps"} $ x $ t end)]
-*} -- {* to avoid eta-contraction of body *}
+\<close> -- \<open>to avoid eta-contraction of body\<close>
definition inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
"inv_into A f == %x. SOME y. y : A & f y = x"
@@ -37,22 +37,22 @@
"inv == inv_into UNIV"
-subsection {*Hilbert's Epsilon-operator*}
+subsection \<open>Hilbert's Epsilon-operator\<close>
-text{*Easier to apply than @{text someI} if the witness comes from an
-existential formula*}
+text\<open>Easier to apply than @{text someI} if the witness comes from an
+existential formula\<close>
lemma someI_ex [elim?]: "\<exists>x. P x ==> P (SOME x. P x)"
apply (erule exE)
apply (erule someI)
done
-text{*Easier to apply than @{text someI} because the conclusion has only one
-occurrence of @{term P}.*}
+text\<open>Easier to apply than @{text someI} because the conclusion has only one
+occurrence of @{term P}.\<close>
lemma someI2: "[| P a; !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
by (blast intro: someI)
-text{*Easier to apply than @{text someI2} if the witness comes from an
-existential formula*}
+text\<open>Easier to apply than @{text someI2} if the witness comes from an
+existential formula\<close>
lemma someI2_ex: "[| \<exists>a. P a; !!x. P x ==> Q x |] ==> Q (SOME x. P x)"
by (blast intro: someI2)
@@ -81,7 +81,7 @@
done
-subsection{*Axiom of Choice, Proved Using the Description Operator*}
+subsection\<open>Axiom of Choice, Proved Using the Description Operator\<close>
lemma choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>f. \<forall>x. Q x (f x)"
by (fast elim: someI)
@@ -114,7 +114,7 @@
qed
-subsection {*Function Inverse*}
+subsection \<open>Function Inverse\<close>
lemma inv_def: "inv f = (%y. SOME x. f x = y)"
by(simp add: inv_into_def)
@@ -152,7 +152,7 @@
lemma inj_imp_inv_eq: "[| inj f; ALL x. f(g x) = x |] ==> inv f = g"
by (blast intro: inv_into_f_eq)
-text{*But is it useful?*}
+text\<open>But is it useful?\<close>
lemma inj_transfer:
assumes injf: "inj f" and minor: "!!y. y \<in> range(f) ==> P(inv f y)"
shows "P x"
@@ -288,7 +288,7 @@
ultimately show "finite (UNIV :: 'a set)" by simp
qed
-text {*
+text \<open>
Every infinite set contains a countable subset. More precisely we
show that a set @{text S} is infinite if and only if there exists an
injective function from the naturals into @{text S}.
@@ -298,12 +298,12 @@
infinite set @{text S}. The idea is to construct a sequence of
non-empty and infinite subsets of @{text S} obtained by successively
removing elements of @{text S}.
-*}
+\<close>
lemma infinite_countable_subset:
assumes inf: "\<not> finite (S::'a set)"
shows "\<exists>f. inj (f::nat \<Rightarrow> 'a) \<and> range f \<subseteq> S"
- -- {* Courtesy of Stephan Merz *}
+ -- \<open>Courtesy of Stephan Merz\<close>
proof -
def Sseq \<equiv> "rec_nat S (\<lambda>n T. T - {SOME e. e \<in> T})"
def pick \<equiv> "\<lambda>n. (SOME e. e \<in> Sseq n)"
@@ -321,7 +321,7 @@
qed
lemma infinite_iff_countable_subset: "\<not> finite S \<longleftrightarrow> (\<exists>f. inj (f::nat \<Rightarrow> 'a) \<and> range f \<subseteq> S)"
- -- {* Courtesy of Stephan Merz *}
+ -- \<open>Courtesy of Stephan Merz\<close>
using finite_imageD finite_subset infinite_UNIV_char_0 infinite_countable_subset by auto
lemma image_inv_into_cancel:
@@ -345,9 +345,9 @@
have 1: "bij_betw ?f' A' A" using assms
by (auto simp add: bij_betw_inv_into)
obtain a' where 2: "a' \<in> A'" and 3: "?f' a' = a"
- using 1 `a \<in> A` unfolding bij_betw_def by force
+ using 1 \<open>a \<in> A\<close> unfolding bij_betw_def by force
hence "?f'' a = a'"
- using `a \<in> A` 1 3 by (auto simp add: f_inv_into_f bij_betw_def)
+ using \<open>a \<in> A\<close> 1 3 by (auto simp add: f_inv_into_f bij_betw_def)
moreover have "f a = a'" using assms 2 3
by (auto simp add: bij_betw_def)
ultimately show "?f'' a = f a" by simp
@@ -432,7 +432,7 @@
using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff)
-subsection {* The Cantor-Bernstein Theorem *}
+subsection \<open>The Cantor-Bernstein Theorem\<close>
lemma Cantor_Bernstein_aux:
shows "\<exists>A' h. A' \<le> A \<and>
@@ -534,11 +534,11 @@
ultimately show ?thesis unfolding bij_betw_def by auto
qed
-subsection {*Other Consequences of Hilbert's Epsilon*}
+subsection \<open>Other Consequences of Hilbert's Epsilon\<close>
-text {*Hilbert's Epsilon and the @{term split} Operator*}
+text \<open>Hilbert's Epsilon and the @{term split} Operator\<close>
-text{*Looping simprule*}
+text\<open>Looping simprule\<close>
lemma split_paired_Eps: "(SOME x. P x) = (SOME (a,b). P(a,b))"
by simp
@@ -549,7 +549,7 @@
by blast
-text{*A relation is wellfounded iff it has no infinite descending chain*}
+text\<open>A relation is wellfounded iff it has no infinite descending chain\<close>
lemma wf_iff_no_infinite_down_chain:
"wf r = (~(\<exists>f. \<forall>i. (f(Suc i),f i) \<in> r))"
apply (simp only: wf_eq_minimal)
@@ -569,15 +569,15 @@
lemma wf_no_infinite_down_chainE:
assumes "wf r" obtains k where "(f (Suc k), f k) \<notin> r"
-using `wf r` wf_iff_no_infinite_down_chain[of r] by blast
+using \<open>wf r\<close> wf_iff_no_infinite_down_chain[of r] by blast
-text{*A dynamically-scoped fact for TFL *}
+text\<open>A dynamically-scoped fact for TFL\<close>
lemma tfl_some: "\<forall>P x. P x --> P (Eps P)"
by (blast intro: someI)
-subsection {* Least value operator *}
+subsection \<open>Least value operator\<close>
definition
LeastM :: "['a => 'b::ord, 'a => bool] => 'a" where
@@ -630,7 +630,7 @@
by (rule LeastM_nat_lemma [THEN conjunct2, THEN spec, THEN mp], assumption, assumption)
-subsection {* Greatest value operator *}
+subsection \<open>Greatest value operator\<close>
definition
GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" where
@@ -699,7 +699,7 @@
done
-text {* \medskip Specialization to @{text GREATEST}. *}
+text \<open>\medskip Specialization to @{text GREATEST}.\<close>
lemma GreatestI: "P (k::nat) ==> \<forall>y. P y --> y < b ==> P (GREATEST x. P x)"
apply (simp add: Greatest_def)
@@ -713,9 +713,9 @@
done
-subsection {* An aside: bounded accessible part *}
+subsection \<open>An aside: bounded accessible part\<close>
-text {* Finite monotone eventually stable sequences *}
+text \<open>Finite monotone eventually stable sequences\<close>
lemma finite_mono_remains_stable_implies_strict_prefix:
fixes f :: "nat \<Rightarrow> 'a::order"
@@ -728,7 +728,7 @@
assume "\<not> ?thesis"
then have "\<And>n. f n \<noteq> f (Suc n)" by auto
then have "\<And>n. f n < f (Suc n)"
- using `mono f` by (auto simp: le_less mono_iff_le_Suc)
+ using \<open>mono f\<close> by (auto simp: le_less mono_iff_le_Suc)
with lift_Suc_mono_less_iff[of f]
have *: "\<And>n m. n < m \<Longrightarrow> f n < f m" by auto
have "inj f"
@@ -737,7 +737,7 @@
assume "f x = f y"
then show "x = y" by (cases x y rule: linorder_cases) (auto dest: *)
qed
- with `finite (range f)` have "finite (UNIV::nat set)"
+ with \<open>finite (range f)\<close> have "finite (UNIV::nat set)"
by (rule finite_imageD)
then show False by simp
qed
@@ -754,7 +754,7 @@
using eq[rule_format, of "n - 1"] N
by (cases n) (auto simp add: le_Suc_eq)
qed simp
- from this[of n] `N \<le> n` show "f N = f n" by auto
+ from this[of n] \<open>N \<le> n\<close> show "f N = f n" by auto
next
fix n m :: nat assume "m < n" "n \<le> N"
then show "f m < f n"
@@ -763,7 +763,7 @@
then have "i < N" by simp
then have "f i \<noteq> f (Suc i)"
unfolding N_def by (rule not_less_Least)
- with `mono f` show ?case by (simp add: mono_iff_le_Suc less_le)
+ with \<open>mono f\<close> show ?case by (simp add: mono_iff_le_Suc less_le)
qed auto
qed
qed
@@ -799,7 +799,7 @@
qed
-subsection {* More on injections, bijections, and inverses *}
+subsection \<open>More on injections, bijections, and inverses\<close>
lemma infinite_imp_bij_betw:
assumes INF: "\<not> finite A"
@@ -913,7 +913,7 @@
by (auto intro: inj_on_inv_into)
-subsection {* Specification package -- Hilbertized version *}
+subsection \<open>Specification package -- Hilbertized version\<close>
lemma exE_some: "[| Ex P ; c == Eps P |] ==> P c"
by (simp only: someI_ex)
--- a/src/HOL/Inductive.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Inductive.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Markus Wenzel, TU Muenchen
*)
-section {* Knaster-Tarski Fixpoint Theorem and inductive definitions *}
+section \<open>Knaster-Tarski Fixpoint Theorem and inductive definitions\<close>
theory Inductive
imports Complete_Lattices Ctr_Sugar
@@ -14,24 +14,24 @@
"primrec" :: thy_decl
begin
-subsection {* Least and greatest fixed points *}
+subsection \<open>Least and greatest fixed points\<close>
context complete_lattice
begin
definition
lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
- "lfp f = Inf {u. f u \<le> u}" --{*least fixed point*}
+ "lfp f = Inf {u. f u \<le> u}" --\<open>least fixed point\<close>
definition
gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
- "gfp f = Sup {u. u \<le> f u}" --{*greatest fixed point*}
+ "gfp f = Sup {u. u \<le> f u}" --\<open>greatest fixed point\<close>
-subsection{* Proof of Knaster-Tarski Theorem using @{term lfp} *}
+subsection\<open>Proof of Knaster-Tarski Theorem using @{term lfp}\<close>
-text{*@{term "lfp f"} is the least upper bound of
- the set @{term "{u. f(u) \<le> u}"} *}
+text\<open>@{term "lfp f"} is the least upper bound of
+ the set @{term "{u. f(u) \<le> u}"}\<close>
lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A"
by (auto simp add: lfp_def intro: Inf_lower)
@@ -54,7 +54,7 @@
by (rule lfp_unfold) (simp add:mono_def)
-subsection {* General induction rules for least fixed points *}
+subsection \<open>General induction rules for least fixed points\<close>
lemma lfp_ordinal_induct[case_names mono step union]:
fixes f :: "'a\<Colon>complete_lattice \<Rightarrow> 'a"
@@ -101,8 +101,8 @@
using assms by (rule lfp_ordinal_induct)
-text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct},
- to control unfolding*}
+text\<open>Definition forms of @{text lfp_unfold} and @{text lfp_induct},
+ to control unfolding\<close>
lemma def_lfp_unfold: "[| h==lfp(f); mono(f) |] ==> h = f(h)"
by (auto intro!: lfp_unfold)
@@ -124,10 +124,10 @@
by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans)
-subsection {* Proof of Knaster-Tarski Theorem using @{term gfp} *}
+subsection \<open>Proof of Knaster-Tarski Theorem using @{term gfp}\<close>
-text{*@{term "gfp f"} is the greatest lower bound of
- the set @{term "{u. u \<le> f(u)}"} *}
+text\<open>@{term "gfp f"} is the greatest lower bound of
+ the set @{term "{u. u \<le> f(u)}"}\<close>
lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f"
by (auto simp add: gfp_def intro: Sup_upper)
@@ -145,9 +145,9 @@
by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3)
-subsection {* Coinduction rules for greatest fixed points *}
+subsection \<open>Coinduction rules for greatest fixed points\<close>
-text{*weak version*}
+text\<open>weak version\<close>
lemma weak_coinduct: "[| a: X; X \<subseteq> f(X) |] ==> a : gfp(f)"
by (rule gfp_upperbound [THEN subsetD]) auto
@@ -169,7 +169,7 @@
apply assumption
done
-text{*strong version, thanks to Coen and Frost*}
+text\<open>strong version, thanks to Coen and Frost\<close>
lemma coinduct_set: "[| mono(f); a: X; X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
by (rule weak_coinduct[rotated], rule coinduct_lemma) blast+
@@ -203,10 +203,10 @@
by (intro order_trans[OF ind _] monoD[OF mono]) auto
qed (auto intro: mono Inf_greatest)
-subsection {* Even Stronger Coinduction Rule, by Martin Coen *}
+subsection \<open>Even Stronger Coinduction Rule, by Martin Coen\<close>
-text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
- @{term lfp} and @{term gfp}*}
+text\<open>Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
+ @{term lfp} and @{term gfp}\<close>
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
by (iprover intro: subset_refl monoI Un_mono monoD)
@@ -230,8 +230,8 @@
apply (simp_all)
done
-text{*Definition forms of @{text gfp_unfold} and @{text coinduct},
- to control unfolding*}
+text\<open>Definition forms of @{text gfp_unfold} and @{text coinduct},
+ to control unfolding\<close>
lemma def_gfp_unfold: "[| A==gfp(f); mono(f) |] ==> A = f(A)"
by (auto intro!: gfp_unfold)
@@ -255,11 +255,11 @@
"[| A==gfp(f); mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
by (auto intro!: coinduct3)
-text{*Monotonicity of @{term gfp}!*}
+text\<open>Monotonicity of @{term gfp}!\<close>
lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g"
by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans)
-subsection {* Rules for fixed point calculus *}
+subsection \<open>Rules for fixed point calculus\<close>
lemma lfp_rolling:
@@ -339,9 +339,9 @@
qed
qed
-subsection {* Inductive predicates and sets *}
+subsection \<open>Inductive predicates and sets\<close>
-text {* Package setup. *}
+text \<open>Package setup.\<close>
theorems basic_monos =
subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj
@@ -356,9 +356,9 @@
induct_rulify_fallback
-subsection {* Inductive datatypes and primitive recursion *}
+subsection \<open>Inductive datatypes and primitive recursion\<close>
-text {* Package setup. *}
+text \<open>Package setup.\<close>
ML_file "Tools/Old_Datatype/old_datatype_aux.ML"
ML_file "Tools/Old_Datatype/old_datatype_prop.ML"
@@ -370,14 +370,14 @@
ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML"
ML_file "Tools/BNF/bnf_lfp_rec_sugar.ML"
-text{* Lambda-abstractions with pattern matching: *}
+text\<open>Lambda-abstractions with pattern matching:\<close>
syntax
"_lam_pats_syntax" :: "cases_syn => 'a => 'b" ("(%_)" 10)
syntax (xsymbols)
"_lam_pats_syntax" :: "cases_syn => 'a => 'b" ("(\<lambda>_)" 10)
-parse_translation {*
+parse_translation \<open>
let
fun fun_tr ctxt [cs] =
let
@@ -385,6 +385,6 @@
val ft = Case_Translation.case_tr true ctxt [x, cs];
in lambda x ft end
in [(@{syntax_const "_lam_pats_syntax"}, fun_tr)] end
-*}
+\<close>
end
--- a/src/HOL/Inequalities.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Inequalities.thy Sat Jul 18 22:58:50 2015 +0200
@@ -16,7 +16,7 @@
case (Suc i)
have 0: "i = nat((n-1) - m)" "m \<le> n-1" using Suc(2,3) by arith+
have "\<Sum> {m..n} = \<Sum> {m..1+(n-1)}" by simp
- also have "\<dots> = \<Sum> {m..n-1} + n" using `m \<le> n`
+ also have "\<dots> = \<Sum> {m..n-1} + n" using \<open>m \<le> n\<close>
by(subst atLeastAtMostPlus1_int_conv) simp_all
also have "\<dots> = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n"
by(simp add: Suc(1)[OF 0])
@@ -43,7 +43,7 @@
hence "{m..<n} = {m..n-1}" by auto
hence "\<Sum>{m..<n} = \<Sum>{m..n-1}" by simp
also have "\<dots> = (n*(n-1) - m*(m-1)) div 2"
- using assms `m < n` by (simp add: Setsum_Icc_nat mult.commute)
+ using assms \<open>m < n\<close> by (simp add: Setsum_Icc_nat mult.commute)
finally show ?thesis .
next
assume "\<not> m < n" with assms show ?thesis by simp
--- a/src/HOL/Int.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Int.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Tobias Nipkow, Florian Haftmann, TU Muenchen
*)
-section {* The Integers as Equivalence Classes over Pairs of Natural Numbers *}
+section \<open>The Integers as Equivalence Classes over Pairs of Natural Numbers\<close>
theory Int
imports Equiv_Relations Power Quotient Fun_Def
begin
-subsection {* Definition of integers as a quotient type *}
+subsection \<open>Definition of integers as a quotient type\<close>
definition intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" where
"intrel = (\<lambda>(x, y) (u, v). x + v = u + y)"
@@ -32,7 +32,7 @@
"(!!x y. z = Abs_Integ (x, y) ==> P) ==> P"
by (induct z) auto
-subsection {* Integers form a commutative ring *}
+subsection \<open>Integers form a commutative ring\<close>
instantiation int :: comm_ring_1
begin
@@ -85,7 +85,7 @@
obtains (diff) m n where "z = int m - int n"
by transfer clarsimp
-subsection {* Integers are totally ordered *}
+subsection \<open>Integers are totally ordered\<close>
instantiation int :: linorder
begin
@@ -118,7 +118,7 @@
end
-subsection {* Ordering properties of arithmetic operations *}
+subsection \<open>Ordering properties of arithmetic operations\<close>
instance int :: ordered_cancel_ab_semigroup_add
proof
@@ -127,9 +127,9 @@
by transfer clarsimp
qed
-text{*Strict Monotonicity of Multiplication*}
+text\<open>Strict Monotonicity of Multiplication\<close>
-text{*strict, in 1st argument; proof is by induction on k>0*}
+text\<open>strict, in 1st argument; proof is by induction on k>0\<close>
lemma zmult_zless_mono2_lemma:
"(i::int)<j ==> 0<k ==> int k * i < int k * j"
apply (induct k)
@@ -156,7 +156,7 @@
apply (auto simp add: zmult_zless_mono2_lemma)
done
-text{*The integers form an ordered integral domain*}
+text\<open>The integers form an ordered integral domain\<close>
instantiation int :: linordered_idom
begin
@@ -198,7 +198,7 @@
for z1 z2 w :: int
-subsection {* Embedding of the Integers into any @{text ring_1}: @{text of_int}*}
+subsection \<open>Embedding of the Integers into any @{text ring_1}: @{text of_int}\<close>
context ring_1
begin
@@ -225,7 +225,7 @@
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
by (transfer fixing: times) (clarsimp simp add: algebra_simps of_nat_mult)
-text{*Collapse nested embeddings*}
+text\<open>Collapse nested embeddings\<close>
lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n"
by (induct n) auto
@@ -249,7 +249,7 @@
by transfer (clarsimp simp add: algebra_simps
of_nat_add [symmetric] simp del: of_nat_add)
-text{*Special cases where either operand is zero*}
+text\<open>Special cases where either operand is zero\<close>
lemma of_int_eq_0_iff [simp]:
"of_int z = 0 \<longleftrightarrow> z = 0"
using of_int_eq_iff [of z 0] by simp
@@ -263,7 +263,7 @@
context linordered_idom
begin
-text{*Every @{text linordered_idom} has characteristic zero.*}
+text\<open>Every @{text linordered_idom} has characteristic zero.\<close>
subclass ring_char_0 ..
lemma of_int_le_iff [simp]:
@@ -316,7 +316,7 @@
apply simp
done
-subsection {* Magnitude of an Integer, as a Natural Number: @{text nat} *}
+subsection \<open>Magnitude of an Integer, as a Natural Number: @{text nat}\<close>
lift_definition nat :: "int \<Rightarrow> nat" is "\<lambda>(x, y). x - y"
by auto
@@ -336,7 +336,7 @@
lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
by transfer (clarsimp, arith)
-text{*An alternative condition is @{term "0 \<le> w"} *}
+text\<open>An alternative condition is @{term "0 \<le> w"}\<close>
corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)"
by (simp add: nat_le_eq_zle linorder_not_le [symmetric])
@@ -432,11 +432,11 @@
by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
-text {* For termination proofs: *}
+text \<open>For termination proofs:\<close>
lemma measure_function_int[measure_function]: "is_measure (nat o abs)" ..
-subsection{*Lemmas about the Function @{term of_nat} and Orderings*}
+subsection\<open>Lemmas about the Function @{term of_nat} and Orderings\<close>
lemma negative_zless_0: "- (int (Suc n)) < (0 \<Colon> int)"
by (simp add: order_less_le del: of_nat_Suc)
@@ -479,10 +479,10 @@
lemma int_Suc0_eq_1: "int (Suc 0) = 1"
by simp
-text{*This version is proved for all ordered rings, not just integers!
+text\<open>This version is proved for all ordered rings, not just integers!
It is proved here because attribute @{text arith_split} is not available
in theory @{text Rings}.
- But is it really better than just rewriting with @{text abs_if}?*}
+ But is it really better than just rewriting with @{text abs_if}?\<close>
lemma abs_split [arith_split, no_atp]:
"P(abs(a::'a::linordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
@@ -493,19 +493,19 @@
apply (rule_tac x="b - Suc a" in exI, arith)
done
-subsection {* Cases and induction *}
+subsection \<open>Cases and induction\<close>
-text{*Now we replace the case analysis rule by a more conventional one:
-whether an integer is negative or not.*}
+text\<open>Now we replace the case analysis rule by a more conventional one:
+whether an integer is negative or not.\<close>
-text{*This version is symmetric in the two subgoals.*}
+text\<open>This version is symmetric in the two subgoals.\<close>
theorem int_cases2 [case_names nonneg nonpos, cases type: int]:
"\<lbrakk>!! n. z = int n \<Longrightarrow> P; !! n. z = - (int n) \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
apply (cases "z < 0")
apply (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym])
done
-text{*This is the default, with a negative case.*}
+text\<open>This is the default, with a negative case.\<close>
theorem int_cases [case_names nonneg neg, cases type: int]:
"[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P"
apply (cases "z < 0")
@@ -524,14 +524,14 @@
using assms by (rule nonneg_eq_int)
lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
- -- {* Unfold all @{text let}s involving constants *}
- by (fact Let_numeral) -- {* FIXME drop *}
+ -- \<open>Unfold all @{text let}s involving constants\<close>
+ by (fact Let_numeral) -- \<open>FIXME drop\<close>
lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)"
- -- {* Unfold all @{text let}s involving constants *}
- by (fact Let_neg_numeral) -- {* FIXME drop *}
+ -- \<open>Unfold all @{text let}s involving constants\<close>
+ by (fact Let_neg_numeral) -- \<open>FIXME drop\<close>
-text {* Unfold @{text min} and @{text max} on numerals. *}
+text \<open>Unfold @{text min} and @{text max} on numerals.\<close>
lemmas max_number_of [simp] =
max_def [of "numeral u" "numeral v"]
@@ -546,9 +546,9 @@
min_def [of "- numeral u" "- numeral v"] for u v
-subsubsection {* Binary comparisons *}
+subsubsection \<open>Binary comparisons\<close>
-text {* Preliminaries *}
+text \<open>Preliminaries\<close>
lemma le_imp_0_less:
assumes le: "0 \<le> z"
@@ -572,7 +572,7 @@
add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric])
qed
-subsubsection {* Comparisons, for Ordered Rings *}
+subsubsection \<open>Comparisons, for Ordered Rings\<close>
lemmas double_eq_0_iff = double_zero
@@ -599,7 +599,7 @@
qed
-subsection {* The Set of Integers *}
+subsection \<open>The Set of Integers\<close>
context ring_1
begin
@@ -654,7 +654,7 @@
obtains (of_int) z where "q = of_int z"
unfolding Ints_def
proof -
- from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def .
+ from \<open>q \<in> \<int>\<close> have "q \<in> range of_int" unfolding Ints_def .
then obtain z where "q = of_int z" ..
then show thesis ..
qed
@@ -665,7 +665,7 @@
end
-text {* The premise involving @{term Ints} prevents @{term "a = 1/2"}. *}
+text \<open>The premise involving @{term Ints} prevents @{term "a = 1/2"}.\<close>
lemma Ints_double_eq_0_iff:
assumes in_Ints: "a \<in> Ints"
@@ -718,7 +718,7 @@
qed
-subsection {* @{term setsum} and @{term setprod} *}
+subsection \<open>@{term setsum} and @{term setprod}\<close>
lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
apply (cases "finite A")
@@ -744,27 +744,27 @@
lemmas int_setprod = of_nat_setprod [where 'a=int]
-text {* Legacy theorems *}
+text \<open>Legacy theorems\<close>
lemmas zle_int = of_nat_le_iff [where 'a=int]
lemmas int_int_eq = of_nat_eq_iff [where 'a=int]
lemmas numeral_1_eq_1 = numeral_One
-subsection {* Setting up simplification procedures *}
+subsection \<open>Setting up simplification procedures\<close>
lemmas of_int_simps =
of_int_0 of_int_1 of_int_add of_int_mult
ML_file "Tools/int_arith.ML"
-declaration {* K Int_Arith.setup *}
+declaration \<open>K Int_Arith.setup\<close>
simproc_setup fast_arith ("(m::'a::linordered_idom) < n" |
"(m::'a::linordered_idom) <= n" |
"(m::'a::linordered_idom) = n") =
- {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
+ \<open>fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct)\<close>
-subsection{*More Inequality Reasoning*}
+subsection\<open>More Inequality Reasoning\<close>
lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
by arith
@@ -782,15 +782,15 @@
by arith
-subsection{*The functions @{term nat} and @{term int}*}
+subsection\<open>The functions @{term nat} and @{term int}\<close>
-text{*Simplify the term @{term "w + - z"}*}
+text\<open>Simplify the term @{term "w + - z"}\<close>
lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
using zless_nat_conj [of 1 z] by auto
-text{*This simplifies expressions of the form @{term "int n = z"} where
- z is an integer literal.*}
+text\<open>This simplifies expressions of the form @{term "int n = z"} where
+ z is an integer literal.\<close>
lemmas int_eq_iff_numeral [simp] = int_eq_iff [of _ "numeral v"] for v
lemma split_nat [arith_split]:
@@ -895,7 +895,7 @@
subsection "Induction principles for int"
-text{*Well-founded segments of the integers*}
+text\<open>Well-founded segments of the integers\<close>
definition
int_ge_less_than :: "int => (int * int) set"
@@ -910,8 +910,8 @@
by (rule wf_subset [OF wf_measure])
qed
-text{*This variant looks odd, but is typical of the relations suggested
-by RankFinder.*}
+text\<open>This variant looks odd, but is typical of the relations suggested
+by RankFinder.\<close>
definition
int_ge_less_than2 :: "int => (int * int) set"
@@ -1021,7 +1021,7 @@
qed
qed
-subsection{*Intermediate value theorems*}
+subsection\<open>Intermediate value theorems\<close>
lemma int_val_lemma:
"(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->
@@ -1053,7 +1053,7 @@
done
-subsection{*Products and 1, by T. M. Rasmussen*}
+subsection\<open>Products and 1, by T. M. Rasmussen\<close>
lemma zabs_less_one_iff [simp]: "(\<bar>z\<bar> < 1) = (z = (0::int))"
by arith
@@ -1110,18 +1110,18 @@
qed
-subsection {* Further theorems on numerals *}
+subsection \<open>Further theorems on numerals\<close>
-subsubsection{*Special Simplification for Constants*}
+subsubsection\<open>Special Simplification for Constants\<close>
-text{*These distributive laws move literals inside sums and differences.*}
+text\<open>These distributive laws move literals inside sums and differences.\<close>
lemmas distrib_right_numeral [simp] = distrib_right [of _ _ "numeral v"] for v
lemmas distrib_left_numeral [simp] = distrib_left [of "numeral v"] for v
lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v
lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v
-text{*These are actually for fields, like real: but where else to put them?*}
+text\<open>These are actually for fields, like real: but where else to put them?\<close>
lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w
lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w
@@ -1129,8 +1129,8 @@
lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w
-text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}. It looks
- strange, but then other simprocs simplify the quotient.*}
+text \<open>Replaces @{text "inverse #nn"} by @{text "1/#nn"}. It looks
+ strange, but then other simprocs simplify the quotient.\<close>
lemmas inverse_eq_divide_numeral [simp] =
inverse_eq_divide [of "numeral w"] for w
@@ -1138,8 +1138,8 @@
lemmas inverse_eq_divide_neg_numeral [simp] =
inverse_eq_divide [of "- numeral w"] for w
-text {*These laws simplify inequalities, moving unary minus from a term
-into the literal.*}
+text \<open>These laws simplify inequalities, moving unary minus from a term
+into the literal.\<close>
lemmas equation_minus_iff_numeral [no_atp] =
equation_minus_iff [of "numeral v"] for v
@@ -1159,10 +1159,10 @@
lemmas minus_less_iff_numeral [no_atp] =
minus_less_iff [of _ "numeral v"] for v
--- {* FIXME maybe simproc *}
+-- \<open>FIXME maybe simproc\<close>
-text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
+text \<open>Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"})\<close>
lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v
lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v
@@ -1170,7 +1170,7 @@
lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v
-text {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
+text \<open>Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="})\<close>
lemmas le_divide_eq_numeral1 [simp] =
pos_le_divide_eq [of "numeral w", OF zero_less_numeral]
@@ -1197,9 +1197,9 @@
divide_eq_eq [of _ "- numeral w"] for w
-subsubsection{*Optional Simplification Rules Involving Constants*}
+subsubsection\<open>Optional Simplification Rules Involving Constants\<close>
-text{*Simplify quotients that are compared with a literal constant.*}
+text\<open>Simplify quotients that are compared with a literal constant.\<close>
lemmas le_divide_eq_numeral =
le_divide_eq [of "numeral w"]
@@ -1226,14 +1226,14 @@
divide_eq_eq [of _ _ "- numeral w"] for w
-text{*Not good as automatic simprules because they cause case splits.*}
+text\<open>Not good as automatic simprules because they cause case splits.\<close>
lemmas divide_const_simps =
le_divide_eq_numeral divide_le_eq_numeral less_divide_eq_numeral
divide_less_eq_numeral eq_divide_eq_numeral divide_eq_eq_numeral
le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
-subsection {* The divides relation *}
+subsection \<open>The divides relation\<close>
lemma zdvd_antisym_nonneg:
"0 <= m ==> 0 <= n ==> m dvd n ==> n dvd m ==> m = (n::int)"
@@ -1247,11 +1247,11 @@
assume "a = 0" with assms show ?thesis by simp
next
assume "a \<noteq> 0"
- from `a dvd b` obtain k where k:"b = a*k" unfolding dvd_def by blast
- from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast
+ from \<open>a dvd b\<close> obtain k where k:"b = a*k" unfolding dvd_def by blast
+ from \<open>b dvd a\<close> obtain k' where k':"a = b*k'" unfolding dvd_def by blast
from k k' have "a = a*k*k'" by simp
with mult_cancel_left1[where c="a" and b="k*k'"]
- have kk':"k*k' = 1" using `a\<noteq>0` by (simp add: mult.assoc)
+ have kk':"k*k' = 1" using \<open>a\<noteq>0\<close> by (simp add: mult.assoc)
hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff)
thus ?thesis using k k' by auto
qed
@@ -1267,11 +1267,11 @@
assumes "i \<noteq> 0" and "d dvd i"
shows "\<bar>d\<bar> \<le> \<bar>i\<bar>"
proof -
- from `d dvd i` obtain k where "i = d * k" ..
- with `i \<noteq> 0` have "k \<noteq> 0" by auto
+ from \<open>d dvd i\<close> obtain k where "i = d * k" ..
+ with \<open>i \<noteq> 0\<close> have "k \<noteq> 0" by auto
then have "1 \<le> \<bar>k\<bar>" and "0 \<le> \<bar>d\<bar>" by auto
then have "\<bar>d\<bar> * 1 \<le> \<bar>d\<bar> * \<bar>k\<bar>" by (rule mult_left_mono)
- with `i = d * k` show ?thesis by (simp add: abs_mult)
+ with \<open>i = d * k\<close> show ?thesis by (simp add: abs_mult)
qed
lemma zdvd_not_zless:
@@ -1281,10 +1281,10 @@
proof
from assms have "0 < n" by auto
assume "n dvd m" then obtain k where k: "m = n * k" ..
- with `0 < m` have "0 < n * k" by auto
- with `0 < n` have "0 < k" by (simp add: zero_less_mult_iff)
- with k `0 < n` `m < n` have "n * k < n * 1" by simp
- with `0 < n` `0 < k` show False unfolding mult_less_cancel_left by auto
+ with \<open>0 < m\<close> have "0 < n * k" by auto
+ with \<open>0 < n\<close> have "0 < k" by (simp add: zero_less_mult_iff)
+ with k \<open>0 < n\<close> \<open>m < n\<close> have "n * k < n * 1" by simp
+ with \<open>0 < n\<close> \<open>0 < k\<close> show False unfolding mult_less_cancel_left by auto
qed
lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)"
@@ -1381,17 +1381,17 @@
assume "a dvd (x + t)"
then obtain l where "x + t = a * l" by (rule dvdE)
then have "x = a * l - t" by simp
- with `d = a * k` show "a dvd x + c * d + t" by simp
+ with \<open>d = a * k\<close> show "a dvd x + c * d + t" by simp
next
assume "a dvd x + c * d + t"
then obtain l where "x + c * d + t = a * l" by (rule dvdE)
then have "x = a * l - c * d - t" by simp
- with `d = a * k` show "a dvd (x + t)" by simp
+ with \<open>d = a * k\<close> show "a dvd (x + t)" by simp
qed
qed
-subsection {* Finiteness of intervals *}
+subsection \<open>Finiteness of intervals\<close>
lemma finite_interval_int1 [iff]: "finite {i :: int. a <= i & i <= b}"
proof (cases "a <= b")
@@ -1421,9 +1421,9 @@
by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto
-subsection {* Configuration of the code generator *}
+subsection \<open>Configuration of the code generator\<close>
-text {* Constructors *}
+text \<open>Constructors\<close>
definition Pos :: "num \<Rightarrow> int" where
[simp, code_abbrev]: "Pos = numeral"
@@ -1434,7 +1434,7 @@
code_datatype "0::int" Pos Neg
-text {* Auxiliary operations *}
+text \<open>Auxiliary operations\<close>
definition dup :: "int \<Rightarrow> int" where
[simp]: "dup k = k + k"
@@ -1465,7 +1465,7 @@
apply (simp_all only: minus_add add.assoc left_minus)
done
-text {* Implementations *}
+text \<open>Implementations\<close>
lemma one_int_code [code, code_unfold]:
"1 = Pos Num.One"
@@ -1567,7 +1567,7 @@
by simp_all
-text {* Serializer setup *}
+text \<open>Serializer setup\<close>
code_identifier
code_module Int \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
@@ -1577,7 +1577,7 @@
hide_const (open) Pos Neg sub dup
-subsection {* Legacy theorems *}
+subsection \<open>Legacy theorems\<close>
lemmas inj_int = inj_of_nat [where 'a=int]
lemmas zadd_int = of_nat_add [where 'a=int, symmetric]
@@ -1609,12 +1609,12 @@
lemmas zpower_int = int_power [symmetric]
-text {* De-register @{text "int"} as a quotient type: *}
+text \<open>De-register @{text "int"} as a quotient type:\<close>
lifting_update int.lifting
lifting_forget int.lifting
-text{*Also the class for fields with characteristic zero.*}
+text\<open>Also the class for fields with characteristic zero.\<close>
class field_char_0 = field + ring_char_0
subclass (in linordered_field) field_char_0 ..
--- a/src/HOL/Lattices.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Lattices.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,19 +2,19 @@
Author: Tobias Nipkow
*)
-section {* Abstract lattices *}
+section \<open>Abstract lattices\<close>
theory Lattices
imports Groups
begin
-subsection {* Abstract semilattice *}
+subsection \<open>Abstract semilattice\<close>
-text {*
+text \<open>
These locales provide a basic structure for interpretation into
bigger structures; extensions require careful thinking, otherwise
undesired effects may occur due to interpretation.
-*}
+\<close>
no_notation times (infixl "*" 70)
no_notation Groups.one ("1")
@@ -73,7 +73,7 @@
by simp
then have "a = (a * b) * c"
by (simp add: assoc)
- with `a = a * b` [symmetric] have "a = a * c" by simp
+ with \<open>a = a * b\<close> [symmetric] have "a = a * c" by simp
then show "a \<preceq> c" by (rule orderI)
qed
@@ -153,7 +153,7 @@
notation Groups.one ("1")
-subsection {* Syntactic infimum and supremum operations *}
+subsection \<open>Syntactic infimum and supremum operations\<close>
class inf =
fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70)
@@ -162,7 +162,7 @@
fixes sup :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65)
-subsection {* Concrete lattices *}
+subsection \<open>Concrete lattices\<close>
notation
less_eq (infix "\<sqsubseteq>" 50) and
@@ -179,7 +179,7 @@
and sup_least: "y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<squnion> z \<sqsubseteq> x"
begin
-text {* Dual lattice *}
+text \<open>Dual lattice\<close>
lemma dual_semilattice:
"class.semilattice_inf sup greater_eq greater"
@@ -191,7 +191,7 @@
class lattice = semilattice_inf + semilattice_sup
-subsubsection {* Intro and elim rules*}
+subsubsection \<open>Intro and elim rules\<close>
context semilattice_inf
begin
@@ -266,7 +266,7 @@
end
-subsubsection {* Equational laws *}
+subsubsection \<open>Equational laws\<close>
context semilattice_inf
begin
@@ -373,7 +373,7 @@
lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2
-text{* Towards distributivity *}
+text\<open>Towards distributivity\<close>
lemma distrib_sup_le: "x \<squnion> (y \<sqinter> z) \<sqsubseteq> (x \<squnion> y) \<sqinter> (x \<squnion> z)"
by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2)
@@ -381,7 +381,7 @@
lemma distrib_inf_le: "(x \<sqinter> y) \<squnion> (x \<sqinter> z) \<sqsubseteq> x \<sqinter> (y \<squnion> z)"
by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2)
-text{* If you have one of them, you have them all. *}
+text\<open>If you have one of them, you have them all.\<close>
lemma distrib_imp1:
assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
@@ -411,7 +411,7 @@
end
-subsubsection {* Strict order *}
+subsubsection \<open>Strict order\<close>
context semilattice_inf
begin
@@ -442,7 +442,7 @@
end
-subsection {* Distributive lattices *}
+subsection \<open>Distributive lattices\<close>
class distrib_lattice = lattice +
assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
@@ -479,7 +479,7 @@
end
-subsection {* Bounded lattices and boolean algebras *}
+subsection \<open>Bounded lattices and boolean algebras\<close>
class bounded_semilattice_inf_top = semilattice_inf + order_top
begin
@@ -710,7 +710,7 @@
end
-subsection {* @{text "min/max"} as special case of lattice *}
+subsection \<open>@{text "min/max"} as special case of lattice\<close>
context linorder
begin
@@ -788,7 +788,7 @@
by (auto intro: antisym simp add: max_def fun_eq_iff)
-subsection {* Uniqueness of inf and sup *}
+subsection \<open>Uniqueness of inf and sup\<close>
lemma (in semilattice_inf) inf_unique:
fixes f (infixl "\<triangle>" 70)
@@ -815,7 +815,7 @@
qed
-subsection {* Lattice on @{typ bool} *}
+subsection \<open>Lattice on @{typ bool}\<close>
instantiation bool :: boolean_algebra
begin
@@ -850,7 +850,7 @@
by auto
-subsection {* Lattice on @{typ "_ \<Rightarrow> _"} *}
+subsection \<open>Lattice on @{typ "_ \<Rightarrow> _"}\<close>
instantiation "fun" :: (type, semilattice_sup) semilattice_sup
begin
@@ -921,7 +921,7 @@
qed (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+
-subsection {* Lattice on unary and binary predicates *}
+subsection \<open>Lattice on unary and binary predicates\<close>
lemma inf1I: "A x \<Longrightarrow> B x \<Longrightarrow> (A \<sqinter> B) x"
by (simp add: inf_fun_def)
@@ -965,10 +965,10 @@
lemma sup2E: "(A \<squnion> B) x y \<Longrightarrow> (A x y \<Longrightarrow> P) \<Longrightarrow> (B x y \<Longrightarrow> P) \<Longrightarrow> P"
by (simp add: sup_fun_def) iprover
-text {*
+text \<open>
\medskip Classical introduction rule: no commitment to @{text A} vs
@{text B}.
-*}
+\<close>
lemma sup1CI: "(\<not> B x \<Longrightarrow> A x) \<Longrightarrow> (A \<squnion> B) x"
by (auto simp add: sup_fun_def)
--- a/src/HOL/Lattices_Big.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Lattices_Big.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,19 +3,19 @@
with contributions by Jeremy Avigad
*)
-section {* Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets *}
+section \<open>Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets\<close>
theory Lattices_Big
imports Finite_Set Option
begin
-subsection {* Generic lattice operations over a set *}
+subsection \<open>Generic lattice operations over a set\<close>
no_notation times (infixl "*" 70)
no_notation Groups.one ("1")
-subsubsection {* Without neutral element *}
+subsubsection \<open>Without neutral element\<close>
locale semilattice_set = semilattice
begin
@@ -50,9 +50,9 @@
assumes "finite A" and "x \<notin> A" and "A \<noteq> {}"
shows "F (insert x A) = x * F A"
proof -
- from `A \<noteq> {}` obtain b where "b \<in> A" by blast
+ from \<open>A \<noteq> {}\<close> obtain b where "b \<in> A" by blast
then obtain B where *: "A = insert b B" "b \<notin> B" by (blast dest: mk_disjoint_insert)
- with `finite A` and `x \<notin> A`
+ with \<open>finite A\<close> and \<open>x \<notin> A\<close>
have "finite (insert x B)" and "b \<notin> insert x B" by auto
then have "F (insert b (insert x B)) = x * F (insert b B)"
by (simp add: eq_fold)
@@ -64,7 +64,7 @@
shows "x * F A = F A"
proof -
from assms have "A \<noteq> {}" by auto
- with `finite A` show ?thesis using `x \<in> A`
+ with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close>
by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem)
qed
@@ -102,7 +102,7 @@
lemma closed:
assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x * y \<in> {x, y}"
shows "F A \<in> A"
-using `finite A` `A \<noteq> {}` proof (induct rule: finite_ne_induct)
+using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct)
case singleton then show ?case by simp
next
case insert with elem show ?case by force
@@ -156,7 +156,7 @@
shows "F A \<preceq> a"
proof -
from assms have "A \<noteq> {}" by auto
- from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
+ from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis
proof (induct rule: finite_ne_induct)
case singleton thus ?case by (simp add: refl)
next
@@ -173,7 +173,7 @@
case True then show ?thesis by (simp add: refl)
next
case False
- have B: "B = A \<union> (B - A)" using `A \<subseteq> B` by blast
+ have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast
then have "F B = F (A \<union> (B - A))" by simp
also have "\<dots> = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset)
also have "\<dots> \<preceq> F A" by simp
@@ -183,7 +183,7 @@
end
-subsubsection {* With neutral element *}
+subsubsection \<open>With neutral element\<close>
locale semilattice_neutr_set = semilattice_neutr
begin
@@ -213,7 +213,7 @@
shows "x * F A = F A"
proof -
from assms have "A \<noteq> {}" by auto
- with `finite A` show ?thesis using `x \<in> A`
+ with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close>
by (induct A rule: finite_ne_induct) (auto simp add: ac_simps)
qed
@@ -246,7 +246,7 @@
lemma closed:
assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x * y \<in> {x, y}"
shows "F A \<in> A"
-using `finite A` `A \<noteq> {}` proof (induct rule: finite_ne_induct)
+using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct)
case singleton then show ?case by simp
next
case insert with elem show ?case by force
@@ -279,7 +279,7 @@
shows "F A \<preceq> a"
proof -
from assms have "A \<noteq> {}" by auto
- from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
+ from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis
proof (induct rule: finite_ne_induct)
case singleton thus ?case by (simp add: refl)
next
@@ -296,7 +296,7 @@
case True then show ?thesis by (simp add: refl)
next
case False
- have B: "B = A \<union> (B - A)" using `A \<subseteq> B` by blast
+ have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast
then have "F B = F (A \<union> (B - A))" by simp
also have "\<dots> = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset)
also have "\<dots> \<preceq> F A" by simp
@@ -309,7 +309,7 @@
notation Groups.one ("1")
-subsection {* Lattice operations on finite sets *}
+subsection \<open>Lattice operations on finite sets\<close>
context semilattice_inf
begin
@@ -348,7 +348,7 @@
end
-subsection {* Infimum and Supremum over non-empty sets *}
+subsection \<open>Infimum and Supremum over non-empty sets\<close>
context lattice
begin
@@ -357,9 +357,9 @@
assumes "finite A" and "A \<noteq> {}"
shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA"
proof -
- from `A \<noteq> {}` obtain a where "a \<in> A" by blast
- with `finite A` have "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> a" by (rule Inf_fin.coboundedI)
- moreover from `finite A` `a \<in> A` have "a \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI)
+ from \<open>A \<noteq> {}\<close> obtain a where "a \<in> A" by blast
+ with \<open>finite A\<close> have "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> a" by (rule Inf_fin.coboundedI)
+ moreover from \<open>finite A\<close> \<open>a \<in> A\<close> have "a \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI)
ultimately show ?thesis by (rule order_trans)
qed
@@ -477,7 +477,7 @@
end
-subsection {* Minimum and Maximum over non-empty sets *}
+subsection \<open>Minimum and Maximum over non-empty sets\<close>
context linorder
begin
@@ -506,7 +506,7 @@
end
-text {* An aside: @{const Min}/@{const Max} on linear orders as special case of @{const Inf_fin}/@{const Sup_fin} *}
+text \<open>An aside: @{const Min}/@{const Max} on linear orders as special case of @{const Inf_fin}/@{const Sup_fin}\<close>
lemma Inf_fin_Min:
"Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \<Rightarrow> 'a)"
@@ -566,9 +566,9 @@
proof (cases "A = {}")
case True then show ?thesis by simp
next
- case False with `finite A` have "Min (insert a A) = min a (Min A)"
+ case False with \<open>finite A\<close> have "Min (insert a A) = min a (Min A)"
by simp
- moreover from `finite A` `A \<noteq> {}` min have "a \<le> Min A" by simp
+ moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> min have "a \<le> Min A" by simp
ultimately show ?thesis by (simp add: min.absorb1)
qed
@@ -578,9 +578,9 @@
proof (cases "A = {}")
case True then show ?thesis by simp
next
- case False with `finite A` have "Max (insert a A) = max a (Max A)"
+ case False with \<open>finite A\<close> have "Max (insert a A) = max a (Max A)"
by simp
- moreover from `finite A` `A \<noteq> {}` max have "Max A \<le> a" by simp
+ moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> max have "Max A \<le> a" by simp
ultimately show ?thesis by (simp add: max.absorb1)
qed
@@ -600,7 +600,7 @@
and "x \<in> A"
shows "Min A = x"
proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
+ from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto
with assms show "Min A \<ge> x" by simp
next
from assms show "x \<ge> Min A" by simp
@@ -612,7 +612,7 @@
and "x \<in> A"
shows "Max A = x"
proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
+ from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto
with assms show "Max A \<le> x" by simp
next
from assms show "x \<le> Max A" by simp
@@ -687,14 +687,14 @@
assumes "finite A" and "A \<noteq> {}"
shows "f (Min A) = Min (f ` A)"
proof (rule linorder_class.Min_eqI [symmetric])
- from `finite A` show "finite (f ` A)" by simp
+ from \<open>finite A\<close> show "finite (f ` A)" by simp
from assms show "f (Min A) \<in> f ` A" by simp
fix x
assume "x \<in> f ` A"
then obtain y where "y \<in> A" and "x = f y" ..
with assms have "Min A \<le> y" by auto
- with `mono f` have "f (Min A) \<le> f y" by (rule monoE)
- with `x = f y` show "f (Min A) \<le> x" by simp
+ with \<open>mono f\<close> have "f (Min A) \<le> f y" by (rule monoE)
+ with \<open>x = f y\<close> show "f (Min A) \<le> x" by simp
qed
lemma mono_Max_commute:
@@ -702,14 +702,14 @@
assumes "finite A" and "A \<noteq> {}"
shows "f (Max A) = Max (f ` A)"
proof (rule linorder_class.Max_eqI [symmetric])
- from `finite A` show "finite (f ` A)" by simp
+ from \<open>finite A\<close> show "finite (f ` A)" by simp
from assms show "f (Max A) \<in> f ` A" by simp
fix x
assume "x \<in> f ` A"
then obtain y where "y \<in> A" and "x = f y" ..
with assms have "y \<le> Max A" by auto
- with `mono f` have "f y \<le> f (Max A)" by (rule monoE)
- with `x = f y` show "x \<le> f (Max A)" by simp
+ with \<open>mono f\<close> have "f y \<le> f (Max A)" by (rule monoE)
+ with \<open>x = f y\<close> show "x \<le> f (Max A)" by simp
qed
lemma finite_linorder_max_induct [consumes 1, case_names empty insert]:
@@ -727,18 +727,18 @@
show "P A"
proof (cases "A = {}")
assume "A = {}"
- then show "P A" using `P {}` by simp
+ then show "P A" using \<open>P {}\<close> by simp
next
let ?B = "A - {Max A}"
let ?A = "insert (Max A) ?B"
- have "finite ?B" using `finite A` by simp
+ have "finite ?B" using \<open>finite A\<close> by simp
assume "A \<noteq> {}"
- with `finite A` have "Max A : A" by auto
+ with \<open>finite A\<close> have "Max A : A" by auto
then have A: "?A = A" using insert_Diff_single insert_absorb by auto
- then have "P ?B" using `P {}` step IH [of ?B] by blast
+ then have "P ?B" using \<open>P {}\<close> step IH [of ?B] by blast
moreover
- have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastforce
- ultimately show "P A" using A insert_Diff_single step [OF `finite ?B`] by fastforce
+ have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF \<open>finite A\<close>] by fastforce
+ ultimately show "P A" using A insert_Diff_single step [OF \<open>finite ?B\<close>] by fastforce
qed
qed
@@ -770,7 +770,7 @@
shows "\<not> finite X"
proof
assume "finite X"
- with `X \<noteq> {}` have "Max X \<in> X" "\<forall>x\<in>X. x \<le> Max X"
+ with \<open>X \<noteq> {}\<close> have "Max X \<in> X" "\<forall>x\<in>X. x \<le> Max X"
by auto
with *[of "Max X"] show False
by auto
--- a/src/HOL/Lazy_Sequence.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Lazy_Sequence.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,12 +1,12 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* Lazy sequences *}
+section \<open>Lazy sequences\<close>
theory Lazy_Sequence
imports Predicate
begin
-subsection {* Type of lazy sequences *}
+subsection \<open>Type of lazy sequences\<close>
datatype (plugins only: code extraction) (dead 'a) lazy_sequence =
lazy_sequence_of_list "'a list"
@@ -171,12 +171,12 @@
| Some ((), xq) \<Rightarrow> empty)"
-subsection {* Code setup *}
+subsection \<open>Code setup\<close>
code_reflect Lazy_Sequence
datatypes lazy_sequence = Lazy_Sequence
-ML {*
+ML \<open>
signature LAZY_SEQUENCE =
sig
datatype 'a lazy_sequence = Lazy_Sequence of (unit -> ('a * 'a Lazy_Sequence.lazy_sequence) option)
@@ -197,19 +197,19 @@
fun yieldn k = Predicate.anamorph yield k;
end;
-*}
+\<close>
-subsection {* Generator Sequences *}
+subsection \<open>Generator Sequences\<close>
-subsubsection {* General lazy sequence operation *}
+subsubsection \<open>General lazy sequence operation\<close>
definition product :: "'a lazy_sequence \<Rightarrow> 'b lazy_sequence \<Rightarrow> ('a \<times> 'b) lazy_sequence"
where
"product s1 s2 = bind s1 (\<lambda>a. bind s2 (\<lambda>b. single (a, b)))"
-subsubsection {* Small lazy typeclasses *}
+subsubsection \<open>Small lazy typeclasses\<close>
class small_lazy =
fixes small_lazy :: "natural \<Rightarrow> 'a lazy_sequence"
@@ -226,8 +226,8 @@
instantiation int :: small_lazy
begin
-text {* maybe optimise this expression -> append (single x) xs == cons x xs
-Performance difference? *}
+text \<open>maybe optimise this expression -> append (single x) xs == cons x xs
+Performance difference?\<close>
function small_lazy' :: "int \<Rightarrow> int \<Rightarrow> int lazy_sequence"
where
@@ -268,8 +268,8 @@
end
-subsection {* With Hit Bound Value *}
-text {* assuming in negative context *}
+subsection \<open>With Hit Bound Value\<close>
+text \<open>assuming in negative context\<close>
type_synonym 'a hit_bound_lazy_sequence = "'a option lazy_sequence"
--- a/src/HOL/Lifting.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Lifting.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,7 +3,7 @@
Author: Cezary Kaliszyk and Christian Urban
*)
-section {* Lifting package *}
+section \<open>Lifting package\<close>
theory Lifting
imports Equiv_Relations Transfer
@@ -14,7 +14,7 @@
"setup_lifting" "lifting_forget" "lifting_update" :: thy_decl
begin
-subsection {* Function map *}
+subsection \<open>Function map\<close>
context
begin
@@ -24,7 +24,7 @@
"(id ---> id) = id"
by (simp add: fun_eq_iff)
-subsection {* Quotient Predicate *}
+subsection \<open>Quotient Predicate\<close>
definition
"Quotient R Abs Rep T \<longleftrightarrow>
@@ -55,7 +55,7 @@
by blast
lemma Quotient_rel:
- "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *}
+ "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- \<open>orientation does not loop on rewriting\<close>
using a unfolding Quotient_def
by blast
@@ -122,7 +122,7 @@
lemma identity_quotient: "Quotient (op =) id id (op =)"
unfolding Quotient_def by simp
-text {* TODO: Use one of these alternatives as the real definition. *}
+text \<open>TODO: Use one of these alternatives as the real definition.\<close>
lemma Quotient_alt_def:
"Quotient R Abs Rep T \<longleftrightarrow>
@@ -195,7 +195,7 @@
then show ?thesis using assms(2) by (auto intro: apply_rsp')
qed
-subsection {* Quotient composition *}
+subsection \<open>Quotient composition\<close>
lemma Quotient_compose:
assumes 1: "Quotient R1 Abs1 Rep1 T1"
@@ -207,7 +207,7 @@
"equivp R \<Longrightarrow> reflp R"
by (erule equivpE)
-subsection {* Respects predicate *}
+subsection \<open>Respects predicate\<close>
definition Respects :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set"
where "Respects R = {x. R x x}"
@@ -272,7 +272,7 @@
unfolding eq_onp_def by (drule Quotient_rep_reflp) blast
-text {* Generating transfer rules for quotients. *}
+text \<open>Generating transfer rules for quotients.\<close>
context
fixes R Abs Rep T
@@ -294,7 +294,7 @@
end
-text {* Generating transfer rules for total quotients. *}
+text \<open>Generating transfer rules for total quotients.\<close>
context
fixes R Abs Rep T
@@ -318,7 +318,7 @@
end
-text {* Generating transfer rules for a type defined with @{text "typedef"}. *}
+text \<open>Generating transfer rules for a type defined with @{text "typedef"}.\<close>
context
fixes Rep Abs A T
@@ -349,15 +349,15 @@
end
-text {* Generating the correspondence rule for a constant defined with
- @{text "lift_definition"}. *}
+text \<open>Generating the correspondence rule for a constant defined with
+ @{text "lift_definition"}.\<close>
lemma Quotient_to_transfer:
assumes "Quotient R Abs Rep T" and "R c c" and "c' \<equiv> Abs c"
shows "T c c'"
using assms by (auto dest: Quotient_cr_rel)
-text {* Proving reflexivity *}
+text \<open>Proving reflexivity\<close>
lemma Quotient_to_left_total:
assumes q: "Quotient R Abs Rep T"
@@ -383,7 +383,7 @@
lemma reflp_ge_eq:
"reflp R \<Longrightarrow> R \<ge> op=" unfolding reflp_def by blast
-text {* Proving a parametrized correspondence relation *}
+text \<open>Proving a parametrized correspondence relation\<close>
definition POS :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool" where
"POS A B \<equiv> A \<le> B"
@@ -443,7 +443,7 @@
shows "R' f g"
using assms unfolding POS_def by auto
-text {* Proving a parametrized correspondence relation *}
+text \<open>Proving a parametrized correspondence relation\<close>
lemma fun_mono:
assumes "A \<ge> C"
@@ -484,7 +484,7 @@
apply (intro choice)
by metis
-subsection {* Domains *}
+subsection \<open>Domains\<close>
lemma composed_equiv_rel_eq_onp:
assumes "left_unique R"
@@ -551,7 +551,7 @@
shows "(rel_set A) (Collect (Domainp A)) UNIV"
using assms unfolding right_total_def rel_set_def Domainp_iff by blast
-subsection {* ML setup *}
+subsection \<open>ML setup\<close>
ML_file "Tools/Lifting/lifting_util.ML"
--- a/src/HOL/Lifting_Set.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Lifting_Set.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Author: Brian Huffman and Ondrej Kuncar
*)
-section {* Setup for Lifting/Transfer for the set type *}
+section \<open>Setup for Lifting/Transfer for the set type\<close>
theory Lifting_Set
imports Lifting
begin
-subsection {* Relator and predicator properties *}
+subsection \<open>Relator and predicator properties\<close>
lemma rel_setD1: "\<lbrakk> rel_set R A B; x \<in> A \<rbrakk> \<Longrightarrow> \<exists>y \<in> B. R x y"
and rel_setD2: "\<lbrakk> rel_set R A B; y \<in> B \<rbrakk> \<Longrightarrow> \<exists>x \<in> A. R x y"
@@ -82,21 +82,21 @@
proof
def f \<equiv> "\<lambda>x. THE y. R x y"
{ fix x assume "x \<in> X"
- with `rel_set R X Y` `bi_unique R` have "R x (f x)"
+ with \<open>rel_set R X Y\<close> \<open>bi_unique R\<close> have "R x (f x)"
by (simp add: bi_unique_def rel_set_def f_def) (metis theI)
- with assms `x \<in> X`
+ with assms \<open>x \<in> X\<close>
have "R x (f x)" "\<forall>x'\<in>X. R x' (f x) \<longrightarrow> x = x'" "\<forall>y\<in>Y. R x y \<longrightarrow> y = f x" "f x \<in> Y"
by (fastforce simp add: bi_unique_def rel_set_def)+ }
note * = this
moreover
{ fix y assume "y \<in> Y"
- with `rel_set R X Y` *(3) `y \<in> Y` have "\<exists>x\<in>X. y = f x"
+ with \<open>rel_set R X Y\<close> *(3) \<open>y \<in> Y\<close> have "\<exists>x\<in>X. y = f x"
by (fastforce simp: rel_set_def) }
ultimately show "\<forall>x\<in>X. R x (f x)" "Y = image f X" "inj_on f X"
by (auto simp: inj_on_def image_iff)
qed
-subsection {* Quotient theorem for the Lifting package *}
+subsection \<open>Quotient theorem for the Lifting package\<close>
lemma Quotient_set[quot_map]:
assumes "Quotient R Abs Rep T"
@@ -108,9 +108,9 @@
done
-subsection {* Transfer rules for the Transfer package *}
+subsection \<open>Transfer rules for the Transfer package\<close>
-subsubsection {* Unconditional transfer rules *}
+subsubsection \<open>Unconditional transfer rules\<close>
context
begin
@@ -183,7 +183,7 @@
unfolding SUP_def [abs_def] by transfer_prover
-subsubsection {* Rules requiring bi-unique, bi-total or right-total relations *}
+subsubsection \<open>Rules requiring bi-unique, bi-total or right-total relations\<close>
lemma member_transfer [transfer_rule]:
assumes "bi_unique A"
@@ -281,7 +281,7 @@
proof (rule rel_funI)+
fix f :: "'b \<Rightarrow> 'a" and g S T
assume "rel_fun A (op =) f g" "rel_set A S T"
- with `bi_unique A` obtain i where "bij_betw i S T" "\<And>x. x \<in> S \<Longrightarrow> f x = g (i x)"
+ with \<open>bi_unique A\<close> obtain i where "bij_betw i S T" "\<And>x. x \<in> S \<Longrightarrow> f x = g (i x)"
by (auto elim: bi_unique_rel_set_lemma simp: rel_fun_def bij_betw_def)
then show "F f S = F g T"
by (simp add: reindex_bij_betw)
--- a/src/HOL/Limited_Sequence.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Limited_Sequence.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,13 +1,13 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* Depth-Limited Sequences with failure element *}
+section \<open>Depth-Limited Sequences with failure element\<close>
theory Limited_Sequence
imports Lazy_Sequence
begin
-subsection {* Depth-Limited Sequence *}
+subsection \<open>Depth-Limited Sequence\<close>
type_synonym 'a dseq = "natural \<Rightarrow> bool \<Rightarrow> 'a lazy_sequence option"
@@ -80,7 +80,7 @@
| Some xq \<Rightarrow> Some (Lazy_Sequence.map f xq))"
-subsection {* Positive Depth-Limited Sequence *}
+subsection \<open>Positive Depth-Limited Sequence\<close>
type_synonym 'a pos_dseq = "natural \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
@@ -121,7 +121,7 @@
"pos_map f xq = (\<lambda>i. Lazy_Sequence.map f (xq i))"
-subsection {* Negative Depth-Limited Sequence *}
+subsection \<open>Negative Depth-Limited Sequence\<close>
type_synonym 'a neg_dseq = "natural \<Rightarrow> 'a Lazy_Sequence.hit_bound_lazy_sequence"
@@ -162,7 +162,7 @@
"neg_map f xq = (\<lambda>i. Lazy_Sequence.hb_map f (xq i))"
-subsection {* Negation *}
+subsection \<open>Negation\<close>
definition pos_not_seq :: "unit neg_dseq \<Rightarrow> unit pos_dseq"
where
@@ -175,7 +175,7 @@
| Some ((), xq) => Lazy_Sequence.empty)"
-ML {*
+ML \<open>
signature LIMITED_SEQUENCE =
sig
type 'a dseq = Code_Numeral.natural -> bool -> 'a Lazy_Sequence.lazy_sequence option
@@ -198,7 +198,7 @@
| SOME s => let val (xs, s') = Lazy_Sequence.yieldn n s in (xs, fn _ => fn _ => SOME s') end);
end;
-*}
+\<close>
code_reserved Eval Limited_Sequence
--- a/src/HOL/Limits.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Limits.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,13 +5,13 @@
Author: Jeremy Avigad
*)
-section {* Limits on Real Vector Spaces *}
+section \<open>Limits on Real Vector Spaces\<close>
theory Limits
imports Real_Vector_Spaces
begin
-subsection {* Filter going to infinity norm *}
+subsection \<open>Filter going to infinity norm\<close>
definition at_infinity :: "'a::real_normed_vector filter" where
"at_infinity = (INF r. principal {x. r \<le> norm x})"
@@ -47,7 +47,7 @@
by (simp add: filterlim_at_top_imp_at_infinity filterlim_compose filterlim_real_sequentially)
-subsubsection {* Boundedness *}
+subsubsection \<open>Boundedness\<close>
definition Bfun :: "('a \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a filter \<Rightarrow> bool" where
Bfun_metric_def: "Bfun f F = (\<exists>y. \<exists>K>0. eventually (\<lambda>x. dist (f x) y \<le> K) F)"
@@ -72,7 +72,7 @@
by (intro always_eventually) (metis dist_commute dist_triangle)
with * have "eventually (\<lambda>x. dist (f x) 0 \<le> K + dist 0 y) F"
by eventually_elim auto
- with `0 < K` show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F"
+ with \<open>0 < K\<close> show "\<exists>K>0. eventually (\<lambda>x. dist (f x) 0 \<le> K) F"
by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto
qed auto
@@ -105,7 +105,7 @@
done
-subsubsection {* Bounded Sequences *}
+subsubsection \<open>Bounded Sequences\<close>
lemma BseqI': "(\<And>n. norm (X n) \<le> K) \<Longrightarrow> Bseq X"
by (intro BfunI) (auto simp: eventually_sequentially)
@@ -155,7 +155,7 @@
then show "\<exists>N. \<forall>n. norm (X n) \<le> real (Suc N)" ..
qed (force simp add: real_of_nat_Suc)
-text{* alternative definition for Bseq *}
+text\<open>alternative definition for Bseq\<close>
lemma Bseq_iff: "Bseq X = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
apply (simp add: Bseq_def)
apply (simp (no_asm) add: lemma_NBseq_def)
@@ -175,9 +175,9 @@
lemma Bseq_iff1a: "Bseq X = (\<exists>N. \<forall>n. norm (X n) < real(Suc N))"
by (simp add: Bseq_def lemma_NBseq_def2)
-subsubsection{*A Few More Equivalence Theorems for Boundedness*}
+subsubsection\<open>A Few More Equivalence Theorems for Boundedness\<close>
-text{*alternative formulation for boundedness*}
+text\<open>alternative formulation for boundedness\<close>
lemma Bseq_iff2: "Bseq X = (\<exists>k > 0. \<exists>x. \<forall>n. norm (X(n) + -x) \<le> k)"
apply (unfold Bseq_def, safe)
apply (rule_tac [2] x = "k + norm x" in exI)
@@ -189,7 +189,7 @@
apply simp
done
-text{*alternative formulation for boundedness*}
+text\<open>alternative formulation for boundedness\<close>
lemma Bseq_iff3:
"Bseq X \<longleftrightarrow> (\<exists>k>0. \<exists>N. \<forall>n. norm (X n + - X N) \<le> k)" (is "?P \<longleftrightarrow> ?Q")
proof
@@ -201,7 +201,7 @@
by (auto intro: order_trans norm_triangle_ineq4)
then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
by simp
- with `0 < K + norm (X 0)` show ?Q by blast
+ with \<open>0 < K + norm (X 0)\<close> show ?Q by blast
next
assume ?Q then show ?P by (auto simp add: Bseq_iff2)
qed
@@ -213,7 +213,7 @@
done
-subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
+subsubsection\<open>Upper Bounds and Lubs of Bounded Sequences\<close>
lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
by (simp add: Bseq_def)
@@ -231,9 +231,9 @@
lemma decseq_bounded: "decseq X \<Longrightarrow> \<forall>i. (B::real) \<le> X i \<Longrightarrow> Bseq X"
by (intro Bseq_eq_bounded[of X B "X 0"]) (auto simp: decseq_def)
-subsection {* Bounded Monotonic Sequences *}
+subsection \<open>Bounded Monotonic Sequences\<close>
-subsubsection{*A Bounded and Monotonic Sequence Converges*}
+subsubsection\<open>A Bounded and Monotonic Sequence Converges\<close>
(* TODO: delete *)
(* FIXME: one use in NSA/HSEQ.thy *)
@@ -244,7 +244,7 @@
apply blast
done
-subsection {* Convergence to Zero *}
+subsection \<open>Convergence to Zero\<close>
definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
where "Zfun f F = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) F)"
@@ -301,7 +301,7 @@
also have "norm (f x) * K \<le> norm (f x) * 0"
using K norm_ge_zero by (rule mult_left_mono)
finally show ?case
- using `0 < r` by simp
+ using \<open>0 < r\<close> by simp
qed
qed
qed
@@ -400,7 +400,7 @@
\<Longrightarrow> (g ---> 0) F"
by (simp add: Zfun_imp_Zfun tendsto_Zfun_iff)
-subsubsection {* Distance and norms *}
+subsubsection \<open>Distance and norms\<close>
lemma tendsto_dist [tendsto_intros]:
fixes l m :: "'a :: metric_space"
@@ -481,7 +481,7 @@
"((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) F \<longleftrightarrow> (f ---> 0) F"
by (fold real_norm_def, rule tendsto_norm_zero_iff)
-subsubsection {* Addition and subtraction *}
+subsubsection \<open>Addition and subtraction\<close>
lemma tendsto_add [tendsto_intros]:
fixes a b :: "'a::real_normed_vector"
@@ -564,7 +564,7 @@
lemmas real_tendsto_sandwich = tendsto_sandwich[where 'b=real]
-subsubsection {* Linear operators and multiplication *}
+subsubsection \<open>Linear operators and multiplication\<close>
lemma (in bounded_linear) tendsto:
"(g ---> a) F \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) F"
@@ -679,7 +679,7 @@
shows "(\<And>i. i \<in> S \<Longrightarrow> continuous_on s (f i)) \<Longrightarrow> continuous_on s (\<lambda>x. \<Prod>i\<in>S. f i x)"
unfolding continuous_on_def by (auto intro: tendsto_setprod)
-subsubsection {* Inverse and division *}
+subsubsection \<open>Inverse and division\<close>
lemma (in bounded_bilinear) Zfun_prod_Bfun:
assumes f: "Zfun f F"
@@ -890,14 +890,14 @@
qed
qed force
-subsection {* Relate @{const at}, @{const at_left} and @{const at_right} *}
+subsection \<open>Relate @{const at}, @{const at_left} and @{const at_right}\<close>
-text {*
+text \<open>
This lemmas are useful for conversion between @{term "at x"} to @{term "at_left x"} and
@{term "at_right x"} and also @{term "at_right 0"}.
-*}
+\<close>
lemmas filterlim_split_at_real = filterlim_split_at[where 'a=real]
@@ -991,10 +991,10 @@
show "\<exists>b. \<forall>x. b \<le> norm x \<longrightarrow> norm (inverse x :: 'a) < r"
proof (intro exI[of _ "inverse (r / 2)"] allI impI)
fix x :: 'a
- from `0 < r` have "0 < inverse (r / 2)" by simp
+ from \<open>0 < r\<close> have "0 < inverse (r / 2)" by simp
also assume *: "inverse (r / 2) \<le> norm x"
finally show "norm (inverse x) < r"
- using * `0 < r` by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
+ using * \<open>0 < r\<close> by (subst nonzero_norm_inverse) (simp_all add: inverse_eq_divide field_simps)
qed
qed
@@ -1106,12 +1106,12 @@
by (simp add: inverse_eq_divide lim_at_infinity_0 comp_def)
-text {*
+text \<open>
We only show rules for multiplication and addition when the functions are either against a real
value or against infinity. Further rules are easy to derive by using @{thm filterlim_uminus_at_top}.
-*}
+\<close>
lemma filterlim_tendsto_pos_mult_at_top:
assumes f: "(f ---> c) F" and c: "0 < c"
@@ -1120,7 +1120,7 @@
unfolding filterlim_at_top_gt[where c=0]
proof safe
fix Z :: real assume "0 < Z"
- from f `0 < c` have "eventually (\<lambda>x. c / 2 < f x) F"
+ from f \<open>0 < c\<close> have "eventually (\<lambda>x. c / 2 < f x) F"
by (auto dest!: tendstoD[where e="c / 2"] elim!: eventually_elim1
simp: dist_real_def abs_real_def split: split_if_asm)
moreover from g have "eventually (\<lambda>x. (Z / c * 2) \<le> g x) F"
@@ -1128,9 +1128,9 @@
ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
proof eventually_elim
fix x assume "c / 2 < f x" "Z / c * 2 \<le> g x"
- with `0 < Z` `0 < c` have "c / 2 * (Z / c * 2) \<le> f x * g x"
+ with \<open>0 < Z\<close> \<open>0 < c\<close> have "c / 2 * (Z / c * 2) \<le> f x * g x"
by (intro mult_mono) (auto simp: zero_le_divide_iff)
- with `0 < c` show "Z \<le> f x * g x"
+ with \<open>0 < c\<close> show "Z \<le> f x * g x"
by simp
qed
qed
@@ -1149,7 +1149,7 @@
ultimately show "eventually (\<lambda>x. Z \<le> f x * g x) F"
proof eventually_elim
fix x assume "1 \<le> f x" "Z \<le> g x"
- with `0 < Z` have "1 * Z \<le> f x * g x"
+ with \<open>0 < Z\<close> have "1 * Z \<le> f x * g x"
by (intro mult_mono) (auto simp: zero_le_divide_iff)
then show "Z \<le> f x * g x"
by simp
@@ -1172,7 +1172,7 @@
fixes f :: "real \<Rightarrow> real"
assumes "0 < n" and f: "LIM x F. f x :> at_top"
shows "LIM x F. (f x)^n :: real :> at_top"
-using `0 < n` proof (induct n)
+using \<open>0 < n\<close> proof (induct n)
case (Suc n) with f show ?case
by (cases "n = 0") (auto intro!: filterlim_at_top_mult_at_top)
qed simp
@@ -1264,7 +1264,7 @@
qed simp
-subsection {* Limits of Sequences *}
+subsection \<open>Limits of Sequences\<close>
lemma [trans]: "X=Y ==> Y ----> z ==> X ----> z"
by simp
@@ -1300,7 +1300,7 @@
shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
by (rule Bfun_inverse)
-text{* Transformation of limit. *}
+text\<open>Transformation of limit.\<close>
lemma eventually_at2:
"eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
@@ -1346,7 +1346,7 @@
using assms(1,2) by auto
qed
-text{* Common case assuming being away from some crucial point like 0. *}
+text\<open>Common case assuming being away from some crucial point like 0.\<close>
lemma Lim_transform_away_within:
fixes a b :: "'a::t1_space"
@@ -1369,7 +1369,7 @@
shows "(g ---> l) (at a)"
using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp
-text{* Alternatively, within an open set. *}
+text\<open>Alternatively, within an open set.\<close>
lemma Lim_transform_within_open:
assumes "open S" and "a \<in> S"
@@ -1383,7 +1383,7 @@
show "(f ---> l) (at a)" by fact
qed
-text{* A congruence rule allowing us to transform limits assuming not at point. *}
+text\<open>A congruence rule allowing us to transform limits assuming not at point.\<close>
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
@@ -1402,7 +1402,7 @@
shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
unfolding tendsto_def eventually_at_topological
using assms by simp
-text{*An unbounded sequence's inverse tends to 0*}
+text\<open>An unbounded sequence's inverse tends to 0\<close>
lemma LIMSEQ_inverse_zero:
"\<forall>r::real. \<exists>N. \<forall>n\<ge>N. r < X n \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> 0"
@@ -1411,14 +1411,14 @@
apply (metis abs_le_D1 linorder_le_cases linorder_not_le)
done
-text{*The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity*}
+text\<open>The sequence @{term "1/n"} tends to 0 as @{term n} tends to infinity\<close>
lemma LIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----> 0"
by (metis filterlim_compose tendsto_inverse_0 filterlim_mono order_refl filterlim_Suc
filterlim_compose[OF filterlim_real_sequentially] at_top_le_at_infinity)
-text{*The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
-infinity is now easily proved*}
+text\<open>The sequence @{term "r + 1/n"} tends to @{term r} as @{term n} tends to
+infinity is now easily proved\<close>
lemma LIMSEQ_inverse_real_of_nat_add:
"(%n. r + inverse(real(Suc n))) ----> r"
@@ -1434,7 +1434,7 @@
using tendsto_mult [OF tendsto_const LIMSEQ_inverse_real_of_nat_add_minus [of 1]]
by auto
-subsection {* Convergence on sequences *}
+subsection \<open>Convergence on sequences\<close>
lemma convergent_add:
fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
@@ -1471,7 +1471,7 @@
done
-text {* A monotone sequence converges to its least upper bound. *}
+text \<open>A monotone sequence converges to its least upper bound.\<close>
lemma LIMSEQ_incseq_SUP:
fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
@@ -1489,7 +1489,7 @@
by (rule order_tendstoI)
(auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
-text{*Main monotonicity theorem*}
+text\<open>Main monotonicity theorem\<close>
lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
@@ -1517,10 +1517,10 @@
assumes "incseq X" and "\<forall>i. X i \<le> B"
obtains L where "X ----> L" "\<forall>i. X i \<le> L"
proof atomize_elim
- from incseq_bounded[OF assms] `incseq X` Bseq_monoseq_convergent[of X]
+ from incseq_bounded[OF assms] \<open>incseq X\<close> Bseq_monoseq_convergent[of X]
obtain L where "X ----> L"
by (auto simp: convergent_def monoseq_def incseq_def)
- with `incseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
+ with \<open>incseq X\<close> show "\<exists>L. X ----> L \<and> (\<forall>i. X i \<le> L)"
by (auto intro!: exI[of _ L] incseq_le)
qed
@@ -1529,17 +1529,17 @@
assumes "decseq X" and "\<forall>i. B \<le> X i"
obtains L where "X ----> L" "\<forall>i. L \<le> X i"
proof atomize_elim
- from decseq_bounded[OF assms] `decseq X` Bseq_monoseq_convergent[of X]
+ from decseq_bounded[OF assms] \<open>decseq X\<close> Bseq_monoseq_convergent[of X]
obtain L where "X ----> L"
by (auto simp: convergent_def monoseq_def decseq_def)
- with `decseq X` show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
+ with \<open>decseq X\<close> show "\<exists>L. X ----> L \<and> (\<forall>i. L \<le> X i)"
by (auto intro!: exI[of _ L] decseq_le)
qed
-subsubsection {* Cauchy Sequences are Bounded *}
+subsubsection \<open>Cauchy Sequences are Bounded\<close>
-text{*A Cauchy sequence is bounded -- this is the standard
- proof mechanization rather than the nonstandard proof*}
+text\<open>A Cauchy sequence is bounded -- this is the standard
+ proof mechanization rather than the nonstandard proof\<close>
lemma lemmaCauchy: "\<forall>n \<ge> M. norm (X M - X n) < (1::real)
==> \<forall>n \<ge> M. norm (X n :: 'a::real_normed_vector) < 1 + norm (X M)"
@@ -1549,11 +1549,11 @@
apply simp
done
-subsection {* Power Sequences *}
+subsection \<open>Power Sequences\<close>
-text{*The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
+text\<open>The sequence @{term "x^n"} tends to 0 if @{term "0\<le>x"} and @{term
"x<1"}. Proof will use (NS) Cauchy equivalence for convergence and
- also fact that bounded and monotonic sequence converges.*}
+ also fact that bounded and monotonic sequence converges.\<close>
lemma Bseq_realpow: "[| 0 \<le> (x::real); x \<le> 1 |] ==> Bseq (%n. x ^ n)"
apply (simp add: Bseq_def)
@@ -1598,7 +1598,7 @@
lemma LIMSEQ_divide_realpow_zero: "1 < x \<Longrightarrow> (\<lambda>n. a / (x ^ n) :: real) ----> 0"
by (rule tendsto_divide_0 [OF tendsto_const filterlim_realpow_sequentially_gt1]) simp
-text{*Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}*}
+text\<open>Limit of @{term "c^n"} for @{term"\<bar>c\<bar> < 1"}\<close>
lemma LIMSEQ_rabs_realpow_zero: "\<bar>c\<bar> < 1 \<Longrightarrow> (\<lambda>n. \<bar>c\<bar> ^ n :: real) ----> 0"
by (rule LIMSEQ_realpow_zero [OF abs_ge_zero])
@@ -1607,7 +1607,7 @@
by (rule LIMSEQ_power_zero) simp
-subsection {* Limits of Functions *}
+subsection \<open>Limits of Functions\<close>
lemma LIM_eq:
fixes a :: "'a::real_normed_vector" and L :: "'b::real_normed_vector"
@@ -1702,7 +1702,7 @@
qed
-subsection {* Continuity *}
+subsection \<open>Continuity\<close>
lemma LIM_isCont_iff:
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::topological_space"
@@ -1777,7 +1777,7 @@
shows "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
by (auto intro: continuous_setsum)
-subsection {* Uniform Continuity *}
+subsection \<open>Uniform Continuity\<close>
definition
isUCont :: "['a::metric_space \<Rightarrow> 'b::metric_space] \<Rightarrow> bool" where
@@ -1823,13 +1823,13 @@
shows "0 \<le> f x"
proof (rule tendsto_le_const)
show "(f ---> f x) (at_left x)"
- using `isCont f x` by (simp add: filterlim_at_split isCont_def)
+ using \<open>isCont f x\<close> by (simp add: filterlim_at_split isCont_def)
show "eventually (\<lambda>x. 0 \<le> f x) (at_left x)"
using ev by (auto simp: eventually_at dist_real_def intro!: exI[of _ "x - b"])
qed simp
-subsection {* Nested Intervals and Bisection -- Needed for Compactness *}
+subsection \<open>Nested Intervals and Bisection -- Needed for Compactness\<close>
lemma nested_sequence_unique:
assumes "\<forall>n. f n \<le> f (Suc n)" "\<forall>n. g (Suc n) \<le> g n" "\<forall>n. f n \<le> g n" "(\<lambda>n. f n - g n) ----> 0"
@@ -1839,17 +1839,17 @@
have "decseq g" unfolding decseq_Suc_iff by fact
{ fix n
- from `decseq g` have "g n \<le> g 0" by (rule decseqD) simp
- with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f n \<le> g 0" by auto }
+ from \<open>decseq g\<close> have "g n \<le> g 0" by (rule decseqD) simp
+ with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f n \<le> g 0" by auto }
then obtain u where "f ----> u" "\<forall>i. f i \<le> u"
- using incseq_convergent[OF `incseq f`] by auto
+ using incseq_convergent[OF \<open>incseq f\<close>] by auto
moreover
{ fix n
- from `incseq f` have "f 0 \<le> f n" by (rule incseqD) simp
- with `\<forall>n. f n \<le> g n`[THEN spec, of n] have "f 0 \<le> g n" by simp }
+ from \<open>incseq f\<close> have "f 0 \<le> f n" by (rule incseqD) simp
+ with \<open>\<forall>n. f n \<le> g n\<close>[THEN spec, of n] have "f 0 \<le> g n" by simp }
then obtain l where "g ----> l" "\<forall>i. l \<le> g i"
- using decseq_convergent[OF `decseq g`] by auto
- moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]]
+ using decseq_convergent[OF \<open>decseq g\<close>] by auto
+ moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF \<open>f ----> u\<close> \<open>g ----> l\<close>]]
ultimately show ?thesis by auto
qed
@@ -1877,7 +1877,7 @@
qed fact
then obtain x where x: "\<And>n. l n \<le> x" "\<And>n. x \<le> u n" and "l ----> x" "u ----> x" by auto
obtain d where "0 < d" and d: "\<And>a b. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> b - a < d \<Longrightarrow> P a b"
- using `l 0 \<le> x` `x \<le> u 0` local[of x] by auto
+ using \<open>l 0 \<le> x\<close> \<open>x \<le> u 0\<close> local[of x] by auto
show "P a b"
proof (rule ccontr)
@@ -1885,12 +1885,12 @@
{ fix n have "\<not> P (l n) (u n)"
proof (induct n)
case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto
- qed (simp add: `\<not> P a b`) }
+ qed (simp add: \<open>\<not> P a b\<close>) }
moreover
{ have "eventually (\<lambda>n. x - d / 2 < l n) sequentially"
- using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto
+ using \<open>0 < d\<close> \<open>l ----> x\<close> by (intro order_tendstoD[of _ x]) auto
moreover have "eventually (\<lambda>n. u n < x + d / 2) sequentially"
- using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto
+ using \<open>0 < d\<close> \<open>u ----> x\<close> by (intro order_tendstoD[of _ x]) auto
ultimately have "eventually (\<lambda>n. P (l n) (u n)) sequentially"
proof eventually_elim
fix n assume "x - d / 2 < l n" "u n < x + d / 2"
@@ -1919,7 +1919,7 @@
with C(2) obtain c where "x \<in> c" "open c" "c \<in> C" by auto
then obtain e where "0 < e" "{x - e <..< x + e} \<subseteq> c"
by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff)
- with `c \<in> C` show ?case
+ with \<open>c \<in> C\<close> show ?case
by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto
qed
qed simp
@@ -1932,7 +1932,7 @@
shows "\<exists>c d. f`S = {c..d} \<and> c \<le> d"
proof -
have S: "compact S" "S \<noteq> {}"
- using `a \<le> b` by (auto simp: S_def)
+ using \<open>a \<le> b\<close> by (auto simp: S_def)
obtain c where "c \<in> S" "\<forall>d\<in>S. f d \<le> f c"
using continuous_attains_sup[OF S f] by auto
moreover obtain d where "d \<in> S" "\<forall>c\<in>S. f d \<le> f c"
@@ -1945,9 +1945,9 @@
by auto
qed
-subsection {* Boundedness of continuous functions *}
+subsection \<open>Boundedness of continuous functions\<close>
-text{*By bisection, function continuous on closed interval is bounded above*}
+text\<open>By bisection, function continuous on closed interval is bounded above\<close>
lemma isCont_eq_Ub:
fixes f :: "real \<Rightarrow> 'a::linorder_topology"
@@ -2006,7 +2006,7 @@
qed
-text{*Continuity of inverse function*}
+text\<open>Continuity of inverse function\<close>
lemma isCont_inverse_function:
fixes f g :: "real \<Rightarrow> real"
@@ -2056,7 +2056,7 @@
==> isCont g (f x)"
by (rule isCont_inverse_function)
-text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
+text\<open>Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110\<close>
lemma LIM_fun_gt_zero:
fixes f :: "real \<Rightarrow> real"
shows "f -- c --> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>r. 0 < r \<and> (\<forall>x. x \<noteq> c \<and> \<bar>c - x\<bar> < r \<longrightarrow> 0 < f x)"
--- a/src/HOL/List.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/List.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Tobias Nipkow
*)
-section {* The datatype of finite lists *}
+section \<open>The datatype of finite lists\<close>
theory List
imports Sledgehammer Code_Numeral Lifting_Set
@@ -20,29 +20,29 @@
datatype_compat list
lemma [case_names Nil Cons, cases type: list]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
"(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P"
by (rule list.exhaust)
lemma [case_names Nil Cons, induct type: list]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
"P [] \<Longrightarrow> (\<And>a list. P list \<Longrightarrow> P (a # list)) \<Longrightarrow> P list"
by (rule list.induct)
-text {* Compatibility: *}
-
-setup {* Sign.mandatory_path "list" *}
+text \<open>Compatibility:\<close>
+
+setup \<open>Sign.mandatory_path "list"\<close>
lemmas inducts = list.induct
lemmas recs = list.rec
lemmas cases = list.case
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
lemmas set_simps = list.set (* legacy *)
syntax
- -- {* list Enumeration *}
+ -- \<open>list Enumeration\<close>
"_list" :: "args => 'a list" ("[(_)]")
translations
@@ -50,7 +50,7 @@
"[x]" == "x#[]"
-subsection {* Basic list processing functions *}
+subsection \<open>Basic list processing functions\<close>
primrec (nonexhaustive) last :: "'a list \<Rightarrow> 'a" where
"last (x # xs) = (if xs = [] then x else last xs)"
@@ -78,7 +78,7 @@
"filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
syntax
- -- {* Special syntax for filter *}
+ -- \<open>Special syntax for filter\<close>
"_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])")
translations
@@ -108,19 +108,19 @@
primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
drop_Nil: "drop n [] = []" |
drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
- -- {*Warning: simpset does not contain this definition, but separate
- theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
+ -- \<open>Warning: simpset does not contain this definition, but separate
+ theorems for @{text "n = 0"} and @{text "n = Suc k"}\<close>
primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
take_Nil:"take n [] = []" |
take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
- -- {*Warning: simpset does not contain this definition, but separate
- theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
+ -- \<open>Warning: simpset does not contain this definition, but separate
+ theorems for @{text "n = 0"} and @{text "n = Suc k"}\<close>
primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where
nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
- -- {*Warning: simpset does not contain this definition, but separate
- theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
+ -- \<open>Warning: simpset does not contain this definition, but separate
+ theorems for @{text "n = 0"} and @{text "n = Suc k"}\<close>
primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
"list_update [] i v = []" |
@@ -151,8 +151,8 @@
"zip xs [] = []" |
zip_Cons: "zip xs (y # ys) =
(case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
- -- {*Warning: simpset does not contain this definition, but separate
- theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
+ -- \<open>Warning: simpset does not contain this definition, but separate
+ theorems for @{text "xs = []"} and @{text "xs = z # zs"}\<close>
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
"product [] _ = []" |
@@ -228,9 +228,9 @@
replicate_0: "replicate 0 x = []" |
replicate_Suc: "replicate (Suc n) x = x # replicate n x"
-text {*
+text \<open>
Function @{text size} is overloaded for all datatypes. Users may
- refer to the list version as @{text length}. *}
+ refer to the list version as @{text length}.\<close>
abbreviation length :: "'a list \<Rightarrow> nat" where
"length \<equiv> size"
@@ -263,7 +263,7 @@
"splice xs [] = xs" |
"splice (x#xs) (y#ys) = x # y # splice xs ys"
-text{*
+text\<open>
\begin{figure}[htbp]
\fbox{
\begin{tabular}{l}
@@ -322,10 +322,10 @@
\end{figure}
Figure~\ref{fig:Characteristic} shows characteristic examples
that should give an intuitive understanding of the above functions.
-*}
-
-text{* The following simple sort functions are intended for proofs,
-not for efficient implementations. *}
+\<close>
+
+text\<open>The following simple sort functions are intended for proofs,
+not for efficient implementations.\<close>
context linorder
begin
@@ -368,9 +368,9 @@
end
-subsubsection {* List comprehension *}
-
-text{* Input syntax for Haskell-like list comprehension notation.
+subsubsection \<open>List comprehension\<close>
+
+text\<open>Input syntax for Haskell-like list comprehension notation.
Typical example: @{text"[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]"},
the list of all pairs of distinct elements from @{text xs} and @{text ys}.
The syntax is as in Haskell, except that @{text"|"} becomes a dot
@@ -393,7 +393,7 @@
It is easy to write short list comprehensions which stand for complex
expressions. During proofs, they may become unreadable (and
mangled). In such cases it can be advisable to introduce separate
-definitions for the list comprehensions in question. *}
+definitions for the list comprehensions in question.\<close>
nonterminal lc_qual and lc_quals
@@ -424,7 +424,7 @@
syntax (HTML output)
"_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
-parse_translation {*
+parse_translation \<open>
let
val NilC = Syntax.const @{const_syntax Nil};
val ConsC = Syntax.const @{const_syntax Cons};
@@ -480,9 +480,9 @@
in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
in [(@{syntax_const "_listcompr"}, lc_tr)] end
-*}
-
-ML_val {*
+\<close>
+
+ML_val \<open>
let
val read = Syntax.read_term @{context} o Syntax.implode_input;
fun check s1 s2 =
@@ -517,14 +517,14 @@
check \<open>[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]\<close>
\<open>concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)\<close>
end;
-*}
+\<close>
(*
term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
*)
-ML {*
+ML \<open>
(* Simproc for rewriting list comprehensions applied to List.set to set
comprehension. *)
@@ -744,16 +744,16 @@
end
end
-*}
+\<close>
simproc_setup list_to_set_comprehension ("set xs") =
- {* K List_to_Set_Comprehension.simproc *}
+ \<open>K List_to_Set_Comprehension.simproc\<close>
code_datatype set coset
hide_const (open) coset
-subsubsection {* @{const Nil} and @{const Cons} *}
+subsubsection \<open>@{const Nil} and @{const Cons}\<close>
lemma not_Cons_self [simp]:
"xs \<noteq> x # xs"
@@ -780,7 +780,7 @@
assumes single: "\<And>x. P [x]"
assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)"
shows "P xs"
-using `xs \<noteq> []` proof (induct xs)
+using \<open>xs \<noteq> []\<close> proof (induct xs)
case Nil then show ?case by simp
next
case (Cons x xs)
@@ -802,12 +802,12 @@
by (auto intro!: inj_onI)
-subsubsection {* @{const length} *}
-
-text {*
+subsubsection \<open>@{const length}\<close>
+
+text \<open>
Needs to come before @{text "@"} because of theorem @{text
append_eq_append_conv}.
-*}
+\<close>
lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
by (induct xs) auto
@@ -890,7 +890,7 @@
lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
by (rule Eq_FalseI) auto
-simproc_setup list_neq ("(xs::'a list) = ys") = {*
+simproc_setup list_neq ("(xs::'a list) = ys") = \<open>
(*
Reduces xs=ys to False if xs and ys cannot be of the same length.
This is the case if the atomic sublists of one are a submultiset
@@ -927,10 +927,10 @@
then prove_neq() else NONE
end;
in K list_neq end;
-*}
-
-
-subsubsection {* @{text "@"} -- append *}
+\<close>
+
+
+subsubsection \<open>@{text "@"} -- append\<close>
lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
by (induct xs) auto
@@ -1007,7 +1007,7 @@
by(cases ys) auto
-text {* Trivial rules for solving @{text "@"}-equations automatically. *}
+text \<open>Trivial rules for solving @{text "@"}-equations automatically.\<close>
lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
by simp
@@ -1021,14 +1021,14 @@
by (drule sym) simp
-text {*
+text \<open>
Simplification procedure for all list equalities.
Currently only tries to rearrange @{text "@"} to see if
- both lists end in a singleton list,
- or both lists end in the same list.
-*}
-
-simproc_setup list_eq ("(xs::'a list) = ys") = {*
+\<close>
+
+simproc_setup list_eq ("(xs::'a list) = ys") = \<open>
let
fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
(case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
@@ -1067,10 +1067,10 @@
else NONE
end;
in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end;
-*}
-
-
-subsubsection {* @{const map} *}
+\<close>
+
+
+subsubsection \<open>@{const map}\<close>
lemma hd_map: "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
by (cases xs) simp_all
@@ -1199,7 +1199,7 @@
declare map.id [simp]
-subsubsection {* @{const rev} *}
+subsubsection \<open>@{const rev}\<close>
lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
by (induct xs) auto
@@ -1247,7 +1247,7 @@
and single: "\<And>x. P [x]"
and snoc': "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (xs@[x])"
shows "P xs"
-using `xs \<noteq> []` proof (induct xs rule: rev_induct)
+using \<open>xs \<noteq> []\<close> proof (induct xs rule: rev_induct)
case (snoc x xs) then show ?case
proof (cases xs)
case Nil thus ?thesis by (simp add: single)
@@ -1260,7 +1260,7 @@
by(rule rev_cases[of xs]) auto
-subsubsection {* @{const set} *}
+subsubsection \<open>@{const set}\<close>
declare list.set[code_post] --"pretty output"
@@ -1370,7 +1370,7 @@
next
assume "\<not> P x"
hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
- thus ?thesis using `\<not> P x` Cons(1) by (metis append_Cons set_ConsD)
+ thus ?thesis using \<open>\<not> P x\<close> Cons(1) by (metis append_Cons set_ConsD)
qed
qed
@@ -1397,7 +1397,7 @@
next
assume "\<not> P x"
hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
- thus ?thesis using `\<not> P x` snoc(1) by fastforce
+ thus ?thesis using \<open>\<not> P x\<close> snoc(1) by fastforce
qed
qed
@@ -1423,7 +1423,7 @@
by (induct xs) auto
-subsubsection {* @{const filter} *}
+subsubsection \<open>@{const filter}\<close>
lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
by (induct xs) auto
@@ -1492,7 +1492,7 @@
hence eq: "?S' = insert 0 (Suc ` ?S)"
by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
have "length (filter p (x # xs)) = Suc(card ?S)"
- using Cons `p x` by simp
+ using Cons \<open>p x\<close> by simp
also have "\<dots> = Suc(card(Suc ` ?S))" using fin
by (simp add: card_image)
also have "\<dots> = card ?S'" using eq fin
@@ -1503,7 +1503,7 @@
hence eq: "?S' = Suc ` ?S"
by(auto simp add: image_def split:nat.split elim:lessE)
have "length (filter p (x # xs)) = card ?S"
- using Cons `\<not> p x` by simp
+ using Cons \<open>\<not> p x\<close> by simp
also have "\<dots> = card(Suc ` ?S)" using fin
by (simp add: card_image)
also have "\<dots> = card ?S'" using eq fin
@@ -1562,7 +1562,7 @@
by (induct ys) simp_all
-subsubsection {* List partitioning *}
+subsubsection \<open>List partitioning\<close>
primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
"partition P [] = ([], [])" |
@@ -1602,7 +1602,7 @@
declare partition.simps[simp del]
-subsubsection {* @{const concat} *}
+subsubsection \<open>@{const concat}\<close>
lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
by (induct xs) auto
@@ -1638,7 +1638,7 @@
by (simp add: concat_eq_concat_iff)
-subsubsection {* @{const nth} *}
+subsubsection \<open>@{const nth}\<close>
lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
by auto
@@ -1707,10 +1707,10 @@
proof (rule ccontr)
assume "n \<noteq> 0"
then have "n > 0" by simp
- with `?lhs` have "xs ! (n - 1) = x" by simp
- moreover from `n > 0` `n \<le> length xs` have "n - 1 < length xs" by simp
+ with \<open>?lhs\<close> have "xs ! (n - 1) = x" by simp
+ moreover from \<open>n > 0\<close> \<open>n \<le> length xs\<close> have "n - 1 < length xs" by simp
ultimately have "\<exists>i<length xs. xs ! i = x" by auto
- with `x \<notin> set xs` in_set_conv_nth [of x xs] show False by simp
+ with \<open>x \<notin> set xs\<close> in_set_conv_nth [of x xs] show False by simp
qed
next
assume ?rhs then show ?lhs by simp
@@ -1721,7 +1721,7 @@
shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume "?lhs" with assms have "n > 0" by (cases n) simp_all
- with `?lhs` show ?rhs by simp
+ with \<open>?lhs\<close> show ?rhs by simp
next
assume "?rhs" then show "?lhs" by simp
qed
@@ -1779,7 +1779,7 @@
qed
-subsubsection {* @{const list_update} *}
+subsubsection \<open>@{const list_update}\<close>
lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
by (induct xs arbitrary: i) (auto split: nat.split)
@@ -1867,7 +1867,7 @@
by simp_all
-subsubsection {* @{const last} and @{const butlast} *}
+subsubsection \<open>@{const last} and @{const butlast}\<close>
lemma last_snoc [simp]: "last (xs @ [x]) = x"
by (induct xs) auto
@@ -1961,7 +1961,7 @@
by fastforce
-subsubsection {* @{const take} and @{const drop} *}
+subsubsection \<open>@{const take} and @{const drop}\<close>
lemma take_0 [simp]: "take 0 xs = []"
by (induct xs) auto
@@ -2208,7 +2208,7 @@
by(auto simp: set_conv_nth image_def) (metis Suc_le_eq nth_take order_trans)
-subsubsection {* @{const takeWhile} and @{const dropWhile} *}
+subsubsection \<open>@{const takeWhile} and @{const dropWhile}\<close>
lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
by (induct xs) auto
@@ -2344,7 +2344,7 @@
proof (rule classical)
assume "\<not> ?thesis"
hence "length (takeWhile P xs) < length xs" using assms by simp
- thus ?thesis using all `\<not> ?thesis` nth_length_takeWhile[of P xs] by auto
+ thus ?thesis using all \<open>\<not> ?thesis\<close> nth_length_takeWhile[of P xs] by auto
qed
lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
@@ -2387,7 +2387,7 @@
by (induct xs) auto
-subsubsection {* @{const zip} *}
+subsubsection \<open>@{const zip}\<close>
lemma zip_Nil [simp]: "zip [] ys = []"
by (induct ys) auto
@@ -2474,7 +2474,7 @@
"map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
by (auto simp: zip_map2)
-text{* Courtesy of Andreas Lochbihler: *}
+text\<open>Courtesy of Andreas Lochbihler:\<close>
lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
by(induct xs) auto
@@ -2560,7 +2560,7 @@
qed
-subsubsection {* @{const list_all2} *}
+subsubsection \<open>@{const list_all2}\<close>
lemma list_all2_lengthD [intro?]:
"list_all2 P xs ys ==> length xs = length ys"
@@ -2719,7 +2719,7 @@
lemma list_all2_same: "list_all2 P xs xs \<longleftrightarrow> (\<forall>x\<in>set xs. P x x)"
by(auto simp add: list_all2_conv_all_nth set_conv_nth)
-subsubsection {* @{const List.product} and @{const product_lists} *}
+subsubsection \<open>@{const List.product} and @{const product_lists}\<close>
lemma set_product[simp]: "set (List.product xs ys) = set xs \<times> set ys"
by (induct xs) auto
@@ -2749,16 +2749,16 @@
proof (intro equalityI subsetI, unfold mem_Collect_eq)
fix xs assume "xs \<in> ?L"
then have "length xs = length xss" by (rule in_set_product_lists_length)
- from this `xs \<in> ?L` show "?R xs" by (induct xs xss rule: list_induct2) auto
+ from this \<open>xs \<in> ?L\<close> show "?R xs" by (induct xs xss rule: list_induct2) auto
next
fix xs assume "?R xs"
then show "xs \<in> ?L" by induct auto
qed
-subsubsection {* @{const fold} with natural argument order *}
-
-lemma fold_simps [code]: -- {* eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala *}
+subsubsection \<open>@{const fold} with natural argument order\<close>
+
+lemma fold_simps [code]: -- \<open>eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\<close>
"fold f [] s = s"
"fold f (x # xs) s = fold f xs (f x s)"
by simp_all
@@ -2818,7 +2818,7 @@
lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))"
by (induct xss) simp_all
-text {* @{const Finite_Set.fold} and @{const fold} *}
+text \<open>@{const Finite_Set.fold} and @{const fold}\<close>
lemma (in comp_fun_commute) fold_set_fold_remdups:
"Finite_Set.fold f y (set xs) = fold f (remdups xs) y"
@@ -2900,9 +2900,9 @@
declare SUP_set_fold [code]
-subsubsection {* Fold variants: @{const foldr} and @{const foldl} *}
-
-text {* Correspondence *}
+subsubsection \<open>Fold variants: @{const foldr} and @{const foldl}\<close>
+
+text \<open>Correspondence\<close>
lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)"
by (induct xs) simp_all
@@ -2910,7 +2910,7 @@
lemma foldl_conv_fold: "foldl f s xs = fold (\<lambda>x s. f s x) xs s"
by (induct xs arbitrary: s) simp_all
-lemma foldr_conv_foldl: -- {* The ``Third Duality Theorem'' in Bird \& Wadler: *}
+lemma foldr_conv_foldl: -- \<open>The ``Third Duality Theorem'' in Bird \& Wadler:\<close>
"foldr f xs a = foldl (\<lambda>x y. f y x) a (rev xs)"
by (simp add: foldr_conv_fold foldl_conv_fold)
@@ -2953,10 +2953,10 @@
by (simp add: fold_append_concat_rev foldr_conv_fold)
-subsubsection {* @{const upt} *}
+subsubsection \<open>@{const upt}\<close>
lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])"
--- {* simp does not terminate! *}
+-- \<open>simp does not terminate!\<close>
by (induct j) auto
lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n
@@ -2976,14 +2976,14 @@
done
lemma upt_Suc_append: "i <= j ==> [i..<(Suc j)] = [i..<j]@[j]"
--- {* Only needed if @{text upt_Suc} is deleted from the simpset. *}
+-- \<open>Only needed if @{text upt_Suc} is deleted from the simpset.\<close>
by simp
lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
by (simp add: upt_rec)
lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
--- {* LOOPS as a simprule, since @{text "j <= j"}. *}
+-- \<open>LOOPS as a simprule, since @{text "j <= j"}.\<close>
by (induct k) auto
lemma length_upt [simp]: "length [i..<j] = j - i"
@@ -3030,12 +3030,12 @@
(!!i. i < k --> xs!i = ys!i) ==> 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 {* Both lists must be non-empty *}
+txt \<open>Both lists must be non-empty\<close>
apply (case_tac xs, simp)
apply (case_tac ys, clarify)
apply (simp (no_asm_use))
apply clarify
-txt {* prenexing's needed, not miniscoping *}
+txt \<open>prenexing's needed, not miniscoping\<close>
apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
apply blast
done
@@ -3056,7 +3056,7 @@
done
lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
--- {* The famous take-lemma. *}
+-- \<open>The famous take-lemma.\<close>
apply (drule_tac x = "max (length xs) (length ys)" in spec)
apply (simp add: le_max_iff_disj)
done
@@ -3086,7 +3086,7 @@
by (simp add: nth_Cons')
-subsubsection {* @{text upto}: interval-list on @{typ int} *}
+subsubsection \<open>@{text upto}: interval-list on @{typ int}\<close>
function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
"upto i j = (if i \<le> j then i # [i+1..j] else [])"
@@ -3124,7 +3124,7 @@
unfolding upto.simps[of i j] by auto
qed
-text{* Tail recursive version for code generation: *}
+text\<open>Tail recursive version for code generation:\<close>
definition upto_aux :: "int \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
"upto_aux i j js = [i..j] @ js"
@@ -3137,7 +3137,7 @@
by(simp add: upto_aux_def)
-subsubsection {* @{const distinct} and @{const remdups} and @{const remdups_adj} *}
+subsubsection \<open>@{const distinct} and @{const remdups} and @{const remdups_adj}\<close>
lemma distinct_tl: "distinct xs \<Longrightarrow> distinct (tl xs)"
by (cases xs) simp_all
@@ -3242,8 +3242,8 @@
\<rbrakk> \<Longrightarrow> distinct (concat xs)"
by (induct xs) auto
-text {* It is best to avoid this indexed version of distinct, but
-sometimes it is useful. *}
+text \<open>It is best to avoid this indexed version of distinct, but
+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)"
@@ -3375,8 +3375,8 @@
shows "distinct (butlast xs)"
proof (cases "xs = []")
case False
- from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
- with `distinct xs` show ?thesis by simp
+ from \<open>xs \<noteq> []\<close> obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
+ with \<open>distinct xs\<close> show ?thesis by simp
qed (auto)
lemma remdups_map_remdups:
@@ -3391,7 +3391,7 @@
assume "length xs' = length ys'"
assume "xs' = take n xs"
with assms have "distinct xs'" by simp
- with `length xs' = length ys'` show "distinct (zip xs' ys')"
+ with \<open>length xs' = length ys'\<close> show "distinct (zip xs' ys')"
by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
qed
@@ -3403,7 +3403,7 @@
assume "length xs' = length ys'"
assume "ys' = take n ys"
with assms have "distinct ys'" by simp
- with `length xs' = length ys'` show "distinct (zip xs' ys')"
+ with \<open>length xs' = length ys'\<close> show "distinct (zip xs' ys')"
by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
qed
@@ -3450,7 +3450,7 @@
let ?x1 = "if ?cond then id else Cons x1"
let ?f = "\<lambda> i. if i = 0 then 0 else ?Succ (f (i - 1))"
have ys: "ys = ?x1 zs" unfolding ys by (cases ?cond, auto)
- have mono: "mono ?f" using `mono f` unfolding mono_def by auto
+ have mono: "mono ?f" using \<open>mono f\<close> unfolding mono_def by auto
show ?case unfolding ys
proof (intro exI[of _ ?f] conjI allI impI)
show "mono ?f" by fact
@@ -3652,7 +3652,7 @@
qed
-subsubsection {* @{const insert} *}
+subsubsection \<open>@{const insert}\<close>
lemma in_set_insert [simp]:
"x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
@@ -3676,9 +3676,9 @@
by (simp add: List.insert_def)
-subsubsection {* @{const List.union} *}
-
-text{* This is all one should need to know about union: *}
+subsubsection \<open>@{const List.union}\<close>
+
+text\<open>This is all one should need to know about union:\<close>
lemma set_union[simp]: "set (List.union xs ys) = set xs \<union> set ys"
unfolding List.union_def
by(induct xs arbitrary: ys) simp_all
@@ -3688,7 +3688,7 @@
by(induct xs arbitrary: ys) simp_all
-subsubsection {* @{const List.find} *}
+subsubsection \<open>@{const List.find}\<close>
lemma find_None_iff: "List.find P xs = None \<longleftrightarrow> \<not> (\<exists>x. x \<in> set xs \<and> P x)"
proof (induction xs)
@@ -3727,7 +3727,7 @@
by (induct xs) simp_all
-subsubsection {* @{const count_list} *}
+subsubsection \<open>@{const count_list}\<close>
lemma count_notin[simp]: "x \<notin> set xs \<Longrightarrow> count_list xs x = 0"
by (induction xs) auto
@@ -3743,7 +3743,7 @@
by (metis Diff_eq setsum.remove)
-subsubsection {* @{const List.extract} *}
+subsubsection \<open>@{const List.extract}\<close>
lemma extract_None_iff: "List.extract P xs = None \<longleftrightarrow> \<not> (\<exists> x\<in>set xs. P x)"
by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits)
@@ -3771,7 +3771,7 @@
(metis dropWhile_eq_Nil_conv list.distinct(1))
-subsubsection {* @{const remove1} *}
+subsubsection \<open>@{const remove1}\<close>
lemma remove1_append:
"remove1 x (xs @ ys) =
@@ -3827,7 +3827,7 @@
by (induct xs) simp_all
-subsubsection {* @{const removeAll} *}
+subsubsection \<open>@{const removeAll}\<close>
lemma removeAll_filter_not_eq:
"removeAll x = filter (\<lambda>y. x \<noteq> y)"
@@ -3873,7 +3873,7 @@
by (rule map_removeAll_inj_on, erule subset_inj_on, rule subset_UNIV)
-subsubsection {* @{const replicate} *}
+subsubsection \<open>@{const replicate}\<close>
lemma length_replicate [simp]: "length (replicate n x) = n"
by (induct n) auto
@@ -3907,14 +3907,14 @@
lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
by (induct n) auto
-text{* Courtesy of Matthias Daum: *}
+text\<open>Courtesy of Matthias Daum:\<close>
lemma append_replicate_commute:
"replicate n x @ replicate k x = replicate k x @ replicate n x"
apply (simp add: replicate_add [symmetric])
apply (simp add: add.commute)
done
-text{* Courtesy of Andreas Lochbihler: *}
+text\<open>Courtesy of Andreas Lochbihler:\<close>
lemma filter_replicate:
"filter P (replicate n x) = (if P x then replicate n x else [])"
by(induct n) auto
@@ -3931,7 +3931,7 @@
lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
by (induct n arbitrary: i)(auto simp: nth_Cons split: nat.split)
-text{* Courtesy of Matthias Daum (2 lemmas): *}
+text\<open>Courtesy of Matthias Daum (2 lemmas):\<close>
lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x"
apply (case_tac "k \<le> i")
apply (simp add: min_def)
@@ -4028,19 +4028,19 @@
case True
then have "concat (replicate 1 xs') = xs'"
and "concat (replicate 1 xs') = ys'"
- using `ys' = xs' @ ws` by auto
+ using \<open>ys' = xs' @ ws\<close> by auto
then show ?thesis by blast
next
case False
- from `ys' = xs' @ ws` and `xs' @ ys' = ys' @ xs'`
+ from \<open>ys' = xs' @ ws\<close> and \<open>xs' @ ys' = ys' @ xs'\<close>
have "xs' @ ws = ws @ xs'" by simp
then have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ws"
- using False and `xs' \<noteq> []` and `ys' = xs' @ ws` and len
+ using False and \<open>xs' \<noteq> []\<close> and \<open>ys' = xs' @ ws\<close> and len
by (intro less.hyps) auto
then obtain m n zs where *: "concat (replicate m zs) = xs'"
and "concat (replicate n zs) = ws" by blast
then have "concat (replicate (m + n) zs) = ys'"
- using `ys' = xs' @ ws`
+ using \<open>ys' = xs' @ ws\<close>
by (simp add: replicate_add)
with * show ?thesis by blast
qed
@@ -4059,7 +4059,7 @@
and "concat (replicate n zs) = ys"
using comm_append_are_replicate[of xs ys, OF assms] by blast
then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
- using `xs \<noteq> []` and `ys \<noteq> []`
+ using \<open>xs \<noteq> []\<close> and \<open>ys \<noteq> []\<close>
by (auto simp: replicate_add)
then show ?thesis by blast
qed
@@ -4081,7 +4081,7 @@
by (subst foldr_fold [symmetric]) simp_all
-subsubsection {* @{const enumerate} *}
+subsubsection \<open>@{const enumerate}\<close>
lemma enumerate_simps [simp, code]:
"enumerate n [] = []"
@@ -4148,7 +4148,7 @@
by (cases "n \<le> m") (simp_all add: zip_map2 zip_same_conv_map enumerate_eq_zip)
-subsubsection {* @{const rotate1} and @{const rotate} *}
+subsubsection \<open>@{const rotate1} and @{const rotate}\<close>
lemma rotate0[simp]: "rotate 0 = id"
by(simp add:rotate_def)
@@ -4242,7 +4242,7 @@
using mod_less_divisor[of "length xs" n] by arith
-subsubsection {* @{const sublist} --- a generalization of @{const nth} to sets *}
+subsubsection \<open>@{const sublist} --- a generalization of @{const nth} to sets\<close>
lemma sublist_empty [simp]: "sublist xs {} = []"
by (auto simp add: sublist_def)
@@ -4325,7 +4325,7 @@
qed
-subsubsection {* @{const sublists} and @{const List.n_lists} *}
+subsubsection \<open>@{const sublists} and @{const List.n_lists}\<close>
lemma length_sublists:
"length (sublists xs) = 2 ^ length xs"
@@ -4376,7 +4376,7 @@
qed
-subsubsection {* @{const splice} *}
+subsubsection \<open>@{const splice}\<close>
lemma splice_Nil2 [simp, code]: "splice xs [] = xs"
by (cases xs) simp_all
@@ -4388,7 +4388,7 @@
by (induct xs ys rule: splice.induct) auto
-subsubsection {* Transpose *}
+subsubsection \<open>Transpose\<close>
function transpose where
"transpose [] = []" |
@@ -4481,7 +4481,7 @@
using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
show ?thesis
- unfolding transpose.simps `i = Suc j` nth_Cons_Suc "3.hyps"[OF j_less]
+ unfolding transpose.simps \<open>i = Suc j\<close> nth_Cons_Suc "3.hyps"[OF j_less]
apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
apply (rule list.exhaust)
by auto
@@ -4501,7 +4501,7 @@
qed
-subsubsection {* (In)finiteness *}
+subsubsection \<open>(In)finiteness\<close>
lemma finite_maxlen:
"finite (M::'a list set) ==> EX n. ALL s:M. size s < n"
@@ -4523,7 +4523,7 @@
assumes "finite A"
shows finite_lists_length_eq: "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
and card_lists_length_eq: "card {xs. set xs \<subseteq> A \<and> length xs = n} = (card A)^n"
- using `finite A`
+ using \<open>finite A\<close>
by (induct n)
(auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong)
@@ -4532,14 +4532,14 @@
(is "finite ?S")
proof-
have "?S = (\<Union>n\<in>{0..n}. {xs. set xs \<subseteq> A \<and> length xs = n})" by auto
- thus ?thesis by (auto intro!: finite_lists_length_eq[OF `finite A`] simp only:)
+ thus ?thesis by (auto intro!: finite_lists_length_eq[OF \<open>finite A\<close>] simp only:)
qed
lemma card_lists_length_le:
assumes "finite A" shows "card {xs. set xs \<subseteq> A \<and> length xs \<le> n} = (\<Sum>i\<le>n. card A^i)"
proof -
have "(\<Sum>i\<le>n. card A^i) = card (\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i})"
- using `finite A`
+ using \<open>finite A\<close>
by (subst card_UN_disjoint)
(auto simp add: card_lists_length_eq finite_lists_length_eq)
also have "(\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i}) = {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
@@ -4563,7 +4563,7 @@
from Suc have "k < card A" by simp
moreover have "finite A" using assms by (simp add: card_ge_0_finite)
moreover have "finite {xs. ?k_list k xs}"
- using finite_lists_length_eq[OF `finite A`, of k]
+ using finite_lists_length_eq[OF \<open>finite A\<close>, of k]
by - (rule finite_subset, auto)
moreover have "\<And>i j. i \<noteq> j \<longrightarrow> {i} \<times> (A - set i) \<inter> {j} \<times> (A - set j) = {}"
by auto
@@ -4588,13 +4588,13 @@
by simp
-subsection {* Sorting *}
-
-text{* Currently it is not shown that @{const sort} returns a
+subsection \<open>Sorting\<close>
+
+text\<open>Currently it is not shown that @{const sort} returns a
permutation of its input because the nicest proof is via multisets,
which are not yet available. Alternatively one could define a function
that counts the number of occurrences of an element in a list and use
-that instead of multisets to state the correctness property. *}
+that instead of multisets to state the correctness property.\<close>
context linorder
begin
@@ -4728,8 +4728,8 @@
assumes "xs \<noteq> []" and "sorted xs"
shows "sorted (butlast xs)"
proof -
- from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
- with `sorted xs` show ?thesis by (simp add: sorted_append)
+ from \<open>xs \<noteq> []\<close> obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
+ with \<open>sorted xs\<close> show ?thesis by (simp add: sorted_append)
qed
lemma insort_not_Nil [simp]:
@@ -4759,7 +4759,7 @@
case False
then have "f x \<noteq> f a" using Cons.prems by auto
then have "f x < f a" using Cons.prems by (auto simp: sorted_Cons)
- with `f x \<noteq> f a` show ?thesis using Cons by (auto simp: sorted_Cons insort_is_Cons)
+ with \<open>f x \<noteq> f a\<close> show ?thesis using Cons by (auto simp: sorted_Cons insort_is_Cons)
qed (auto simp: sorted_Cons insort_is_Cons)
qed simp
@@ -4767,9 +4767,9 @@
assumes "a \<in> set xs" and "sorted xs"
shows "insort a (remove1 a xs) = xs"
proof (rule insort_key_remove1)
- from `a \<in> set xs` show "a \<in> set xs" .
- from `sorted xs` show "sorted (map (\<lambda>x. x) xs)" by simp
- from `a \<in> set xs` have "a \<in> set (filter (op = a) xs)" by auto
+ from \<open>a \<in> set xs\<close> show "a \<in> set xs" .
+ from \<open>sorted xs\<close> show "sorted (map (\<lambda>x. x) xs)" by simp
+ from \<open>a \<in> set xs\<close> have "a \<in> set (filter (op = a) xs)" by auto
then have "set (filter (op = a) xs) \<noteq> {}" by auto
then have "filter (op = a) xs \<noteq> []" by (auto simp only: set_empty)
then have "length (filter (op = a) xs) > 0" by simp
@@ -4811,7 +4811,7 @@
proof -
from assms have "map f xs = map f ys"
by (simp add: sorted_distinct_set_unique)
- with `inj_on f (set xs \<union> set ys)` show "xs = ys"
+ with \<open>inj_on f (set xs \<union> set ys)\<close> show "xs = ys"
by (blast intro: map_inj_on)
qed
@@ -4982,7 +4982,7 @@
by (simp add: enumerate_eq_zip)
-subsubsection {* @{const transpose} on sorted lists *}
+subsubsection \<open>@{const transpose} on sorted lists\<close>
lemma sorted_transpose[simp]:
shows "sorted (rev (map length (transpose xs)))"
@@ -5034,10 +5034,10 @@
assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
shows "length (filter (\<lambda>ys. i < length ys) (transpose xs)) = length (xs ! i)"
proof -
- have "xs \<noteq> []" using `i < length xs` by auto
+ have "xs \<noteq> []" using \<open>i < length xs\<close> by auto
note filter_equals_takeWhile_sorted_rev[OF sorted, simp]
{ fix j assume "j \<le> i"
- note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this `i < length xs`]
+ note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this \<open>i < length xs\<close>]
} note sortedE = this[consumes 1]
have "{j. j < length (transpose xs) \<and> i < length (transpose xs ! j)}
@@ -5052,13 +5052,13 @@
next
fix j assume "j < length (xs ! i)"
thus "j < length (transpose xs)"
- using foldr_max_sorted[OF sorted] `xs \<noteq> []` sortedE[OF le0]
+ using foldr_max_sorted[OF sorted] \<open>xs \<noteq> []\<close> sortedE[OF le0]
by (auto simp: length_transpose comp_def foldr_map)
have "Suc i \<le> length (takeWhile (\<lambda>ys. j < length ys) xs)"
- using `i < length xs` `j < length (xs ! i)` less_Suc_eq_le
+ using \<open>i < length xs\<close> \<open>j < length (xs ! i)\<close> less_Suc_eq_le
by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE)
- with nth_transpose[OF `j < length (transpose xs)`]
+ with nth_transpose[OF \<open>j < length (transpose xs)\<close>]
show "i < length (transpose xs ! j)" by simp
qed
thus ?thesis by (simp add: length_filter_conv_card)
@@ -5078,10 +5078,10 @@
from j have j_less: "j < length (xs ! i)" using length by simp
have i_less_tW: "Suc i \<le> length (takeWhile (\<lambda>ys. Suc j \<le> length ys) xs)"
proof (rule length_takeWhile_less_P_nth)
- show "Suc i \<le> length xs" using `i < length xs` by simp
+ show "Suc i \<le> length xs" using \<open>i < length xs\<close> by simp
fix k assume "k < Suc i"
hence "k \<le> i" by auto
- with sorted_rev_nth_mono[OF sorted this] `i < length xs`
+ with sorted_rev_nth_mono[OF sorted this] \<open>i < length xs\<close>
have "length (xs ! i) \<le> length (xs ! k)" by simp
thus "Suc j \<le> length (xs ! k)" using j_less by simp
qed
@@ -5131,19 +5131,19 @@
qed
-subsubsection {* @{text sorted_list_of_set} *}
-
-text{* This function maps (finite) linearly ordered sets to sorted
+subsubsection \<open>@{text sorted_list_of_set}\<close>
+
+text\<open>This function maps (finite) linearly ordered sets to sorted
lists. Warning: in most cases it is not a good idea to convert from
sets to lists but one should convert in the other direction (via
-@{const set}). *}
-
-subsubsection {* @{text sorted_list_of_set} *}
-
-text{* This function maps (finite) linearly ordered sets to sorted
+@{const set}).\<close>
+
+subsubsection \<open>@{text sorted_list_of_set}\<close>
+
+text\<open>This function maps (finite) linearly ordered sets to sorted
lists. Warning: in most cases it is not a good idea to convert from
sets to lists but one should convert in the other direction (via
-@{const set}). *}
+@{const set}).\<close>
context linorder
begin
@@ -5206,7 +5206,7 @@
by (rule sorted_distinct_set_unique) simp_all
-subsubsection {* @{text lists}: the list-forming operator over sets *}
+subsubsection \<open>@{text lists}: the list-forming operator over sets\<close>
inductive_set
lists :: "'a set => 'a list set"
@@ -5253,7 +5253,7 @@
lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set]
lemma in_listsp_conv_set: "(listsp A xs) = (\<forall>x \<in> set xs. A x)"
--- {* eliminate @{text listsp} in favour of @{text set} *}
+-- \<open>eliminate @{text listsp} in favour of @{text set}\<close>
by (induct xs) auto
lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
@@ -5284,7 +5284,7 @@
then show ?thesis by auto
qed
-subsubsection {* Inductive definition for membership *}
+subsubsection \<open>Inductive definition for membership\<close>
inductive ListMem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
where
@@ -5300,10 +5300,10 @@
done
-subsubsection {* Lists as Cartesian products *}
-
-text{*@{text"set_Cons A Xs"}: the set of lists with head drawn from
-@{term A} and tail drawn from @{term Xs}.*}
+subsubsection \<open>Lists as Cartesian products\<close>
+
+text\<open>@{text"set_Cons A Xs"}: the set of lists with head drawn from
+@{term A} and tail drawn from @{term Xs}.\<close>
definition set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" where
"set_Cons A XS = {z. \<exists>x xs. z = x # xs \<and> x \<in> A \<and> xs \<in> XS}"
@@ -5311,22 +5311,22 @@
lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A"
by (auto simp add: set_Cons_def)
-text{*Yields the set of lists, all of the same length as the argument and
-with elements drawn from the corresponding element of the argument.*}
+text\<open>Yields the set of lists, all of the same length as the argument and
+with elements drawn from the corresponding element of the argument.\<close>
primrec listset :: "'a set list \<Rightarrow> 'a list set" where
"listset [] = {[]}" |
"listset (A # As) = set_Cons A (listset As)"
-subsection {* Relations on Lists *}
-
-subsubsection {* Length Lexicographic Ordering *}
-
-text{*These orderings preserve well-foundedness: shorter lists
- precede longer lists. These ordering are not used in dictionaries.*}
+subsection \<open>Relations on Lists\<close>
+
+subsubsection \<open>Length Lexicographic Ordering\<close>
+
+text\<open>These orderings preserve well-foundedness: shorter lists
+ precede longer lists. These ordering are not used in dictionaries.\<close>
-primrec -- {*The lexicographic ordering for lists of the specified length*}
+primrec -- \<open>The lexicographic ordering for lists of the specified length\<close>
lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
"lexn r 0 = {}" |
"lexn r (Suc n) =
@@ -5334,11 +5334,11 @@
{(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
definition lex :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
-"lex r = (\<Union>n. lexn r n)" -- {*Holds only between lists of the same length*}
+"lex r = (\<Union>n. lexn r n)" -- \<open>Holds only between lists of the same length\<close>
definition lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
"lenlex r = inv_image (less_than <*lex*> lex r) (\<lambda>xs. (length xs, xs))"
- -- {*Compares lists by their length and then lexicographically*}
+ -- \<open>Compares lists by their length and then lexicographically\<close>
lemma wf_lexn: "wf r ==> wf (lexn r n)"
apply (induct n, simp, simp)
@@ -5403,11 +5403,11 @@
done
-subsubsection {* Lexicographic Ordering *}
-
-text {* Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
+subsubsection \<open>Lexicographic Ordering\<close>
+
+text \<open>Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
This ordering does \emph{not} preserve well-foundedness.
- Author: N. Voelker, March 2005. *}
+ Author: N. Voelker, March 2005.\<close>
definition lexord :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
"lexord r = {(x,y). \<exists> a v. y = x @ a # v \<or>
@@ -5462,7 +5462,7 @@
apply (rule_tac x="drop (Suc i) y" in exI)
by (simp add: Cons_nth_drop_Suc)
--- {* lexord is extension of partial ordering List.lex *}
+-- \<open>lexord is extension of partial ordering List.lex\<close>
lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
apply (rule_tac x = y in spec)
apply (induct_tac x, clarsimp)
@@ -5471,7 +5471,7 @@
lemma lexord_irreflexive: "ALL x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
by (induct xs) auto
-text{* By Ren\'e Thiemann: *}
+text\<open>By Ren\'e Thiemann:\<close>
lemma lexord_partial_trans:
"(\<And>x y z. x \<in> set xs \<Longrightarrow> (x,y) \<in> r \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> (x,z) \<in> r)
\<Longrightarrow> (xs,ys) \<in> lexord r \<Longrightarrow> (ys,zs) \<in> lexord r \<Longrightarrow> (xs,zs) \<in> lexord r"
@@ -5545,15 +5545,15 @@
assumes hyp: "(a, b) \<in> lexord R"
shows "(b, a) \<notin> lexord R"
proof -
- from `asym R` have "asym (lexord R)" by (rule lexord_asym)
+ from \<open>asym R\<close> have "asym (lexord R)" by (rule lexord_asym)
then show ?thesis by (rule asym.cases) (auto simp add: hyp)
qed
-text {*
+text \<open>
Predicate version of lexicographic order integrated with Isabelle's order type classes.
Author: Andreas Lochbihler
-*}
+\<close>
context ord begin
@@ -5720,9 +5720,9 @@
end
-subsubsection {* Lexicographic combination of measure functions *}
-
-text {* These are useful for termination proofs *}
+subsubsection \<open>Lexicographic combination of measure functions\<close>
+
+text \<open>These are useful for termination proofs\<close>
definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
@@ -5744,7 +5744,7 @@
by auto
-subsubsection {* Lifting Relations to Lists: one element *}
+subsubsection \<open>Lifting Relations to Lists: one element\<close>
definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
"listrel1 r = {(xs,ys).
@@ -5843,7 +5843,7 @@
qed
-text{* Accessible part and wellfoundedness: *}
+text\<open>Accessible part and wellfoundedness:\<close>
lemma Cons_acc_listrel1I [intro!]:
"x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
@@ -5873,7 +5873,7 @@
by (auto simp: wf_acc_iff
intro: lists_accD lists_accI[THEN Cons_in_lists_iff[THEN iffD1, THEN conjunct1]])
-subsubsection {* Lifting Relations to Lists: all elements *}
+subsubsection \<open>Lifting Relations to Lists: all elements\<close>
inductive_set
listrel :: "('a \<times> 'b) set \<Rightarrow> ('a list \<times> 'b list) set"
@@ -5958,7 +5958,7 @@
"listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
by (auto simp add: set_Cons_def intro: listrel.intros)
-text {* Relating @{term listrel1}, @{term listrel} and closures: *}
+text \<open>Relating @{term listrel1}, @{term listrel} and closures:\<close>
lemma listrel1_rtrancl_subset_rtrancl_listrel1:
"listrel1 (r^*) \<subseteq> (listrel1 r)^*"
@@ -6027,7 +6027,7 @@
by(fast intro:rtrancl_listrel1_if_listrel)
-subsection {* Size function *}
+subsection \<open>Size function\<close>
lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (size_list f)"
by (rule is_measure_trivial)
@@ -6054,7 +6054,7 @@
by (induct xs) force+
-subsection {* Monad operation *}
+subsection \<open>Monad operation\<close>
definition bind :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
"bind xs f = concat (map f xs)"
@@ -6067,7 +6067,7 @@
by (simp_all add: bind_def)
-subsection {* Transfer *}
+subsection \<open>Transfer\<close>
definition embed_list :: "nat list \<Rightarrow> int list" where
"embed_list l = map int l"
@@ -6100,10 +6100,10 @@
*)
-subsection {* Code generation *}
-
-text{* Optional tail recursive version of @{const map}. Can avoid
-stack overflow in some target languages. *}
+subsection \<open>Code generation\<close>
+
+text\<open>Optional tail recursive version of @{const map}. Can avoid
+stack overflow in some target languages.\<close>
fun map_tailrec_rev :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
"map_tailrec_rev f [] bs = bs" |
@@ -6116,20 +6116,20 @@
definition map_tailrec :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
"map_tailrec f as = rev (map_tailrec_rev f as [])"
-text{* Code equation: *}
+text\<open>Code equation:\<close>
lemma map_eq_map_tailrec: "map = map_tailrec"
by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
-subsubsection {* Counterparts for set-related operations *}
+subsubsection \<open>Counterparts for set-related operations\<close>
definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
[code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"
-text {*
+text \<open>
Use @{text member} only for generating executable code. Otherwise use
@{prop "x \<in> set xs"} instead --- it is much easier to reason about.
-*}
+\<close>
lemma member_rec [code]:
"member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
@@ -6151,11 +6151,11 @@
definition list_ex1 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
list_ex1_iff [code_abbrev]: "list_ex1 P xs \<longleftrightarrow> (\<exists>! x. x \<in> set xs \<and> P x)"
-text {*
+text \<open>
Usually you should prefer @{text "\<forall>x\<in>set xs"}, @{text "\<exists>x\<in>set xs"}
and @{text "\<exists>!x. x\<in>set xs \<and> _"} over @{const list_all}, @{const list_ex}
and @{const list_ex1} in specifications.
-*}
+\<close>
lemma list_all_simps [code]:
"list_all P (x # xs) \<longleftrightarrow> P x \<and> list_all P xs"
@@ -6220,7 +6220,7 @@
by (simp add: list_ex1_iff can_select_def)
-text {* Executable checks for relations on sets *}
+text \<open>Executable checks for relations on sets\<close>
definition listrel1p :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"listrel1p r xs ys = ((xs, ys) \<in> listrel1 {(x, y). r x y})"
@@ -6250,7 +6250,7 @@
"lexordp r (x # xs) (y # ys) = (r x y | (x = y & lexordp r xs ys))"
unfolding lexordp_def by auto
-text {* Bounded quantification and summation over nats. *}
+text \<open>Bounded quantification and summation over nats.\<close>
lemma atMost_upto [code_unfold]:
"{..n} = set [0..<Suc n]"
@@ -6290,7 +6290,7 @@
"(\<exists>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
by auto
-text{* Bounded @{text LEAST} operator: *}
+text\<open>Bounded @{text LEAST} operator:\<close>
definition "Bleast S P = (LEAST x. x \<in> S \<and> P x)"
@@ -6321,7 +6321,7 @@
declare Bleast_def[symmetric, code_unfold]
-text {* Summation over ints. *}
+text \<open>Summation over ints.\<close>
lemma greaterThanLessThan_upto [code_unfold]:
"{i<..<j::int} = set [i+1..j - 1]"
@@ -6338,14 +6338,14 @@
lemmas atLeastAtMost_upto [code_unfold] = set_upto [symmetric]
-subsubsection {* Optimizing by rewriting *}
+subsubsection \<open>Optimizing by rewriting\<close>
definition null :: "'a list \<Rightarrow> bool" where
[code_abbrev]: "null xs \<longleftrightarrow> xs = []"
-text {*
+text \<open>
Efficient emptyness check is implemented by @{const null}.
-*}
+\<close>
lemma null_rec [code]:
"null (x # xs) \<longleftrightarrow> False"
@@ -6367,10 +6367,10 @@
definition map_filter :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
[code_post]: "map_filter f xs = map (the \<circ> f) (filter (\<lambda>x. f x \<noteq> None) xs)"
-text {*
+text \<open>
Operations @{const maps} and @{const map_filter} avoid
intermediate lists on execution -- do not use for proving.
-*}
+\<close>
lemma maps_simps [code]:
"maps f (x # xs) = f x @ maps f xs"
@@ -6390,8 +6390,8 @@
"map f (filter P xs) = map_filter (\<lambda>x. if P x then Some (f x) else None) xs"
by (simp add: map_filter_def)
-text {* Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
-and similiarly for @{text"\<exists>"}. *}
+text \<open>Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
+and similiarly for @{text"\<exists>"}.\<close>
definition all_interval_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"all_interval_nat P i j \<longleftrightarrow> (\<forall>n \<in> {i..<j}. P n)"
@@ -6439,7 +6439,7 @@
"list_ex P [i..j] \<longleftrightarrow> \<not> (all_interval_int (Not \<circ> P) i j)"
by (simp add: list_ex_iff all_interval_int_def)
-text {* optimized code (tail-recursive) for @{term length} *}
+text \<open>optimized code (tail-recursive) for @{term length}\<close>
definition gen_length :: "nat \<Rightarrow> 'a list \<Rightarrow> nat"
where "gen_length n xs = n + length xs"
@@ -6457,9 +6457,9 @@
hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length
-subsubsection {* Pretty lists *}
-
-ML {*
+subsubsection \<open>Pretty lists\<close>
+
+ML \<open>
(* Code generation for list literals. *)
signature LIST_CODE =
@@ -6507,7 +6507,7 @@
end
end;
-*}
+\<close>
code_printing
type_constructor list \<rightharpoonup>
@@ -6525,7 +6525,7 @@
| constant "HOL.equal :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool" \<rightharpoonup>
(Haskell) infix 4 "=="
-setup {* fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"] *}
+setup \<open>fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"]\<close>
code_reserved SML
list
@@ -6534,7 +6534,7 @@
list
-subsubsection {* Use convenient predefined operations *}
+subsubsection \<open>Use convenient predefined operations\<close>
code_printing
constant "op @" \<rightharpoonup>
@@ -6566,7 +6566,7 @@
(Haskell) "any"
-subsubsection {* Implementation of sets by lists *}
+subsubsection \<open>Implementation of sets by lists\<close>
lemma is_empty_set [code]:
"Set.is_empty (set xs) \<longleftrightarrow> List.null xs"
@@ -6617,7 +6617,7 @@
"List.coset [] \<le> set [] \<longleftrightarrow> False"
by auto
-text {* A frequent case -- avoid intermediate sets *}
+text \<open>A frequent case -- avoid intermediate sets\<close>
lemma [code_unfold]:
"set xs \<subseteq> set ys \<longleftrightarrow> list_all (\<lambda>x. x \<in> set ys) xs"
by (auto simp: list_all_iff)
@@ -6657,7 +6657,7 @@
hide_const (open) map_project
-text {* Operations on relations *}
+text \<open>Operations on relations\<close>
lemma product_code [code]:
"Product_Type.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
@@ -6684,9 +6684,9 @@
by (simp add: wf_iff_acyclic_if_finite)
-subsection {* Setup for Lifting/Transfer *}
-
-subsubsection {* Transfer rules for the Transfer package *}
+subsection \<open>Setup for Lifting/Transfer\<close>
+
+subsubsection \<open>Transfer rules for the Transfer package\<close>
context
begin
--- a/src/HOL/MacLaurin.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/MacLaurin.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,16 +4,16 @@
Conversion of Mac Laurin to Isar by Lukas Bulwahn and Bernhard Häupler, 2005
*)
-section{*MacLaurin Series*}
+section\<open>MacLaurin Series\<close>
theory MacLaurin
imports Transcendental
begin
-subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*}
+subsection\<open>Maclaurin's Theorem with Lagrange Form of Remainder\<close>
-text{*This is a very long, messy proof even now that it's been broken down
-into lemmas.*}
+text\<open>This is a very long, messy proof even now that it's been broken down
+into lemmas.\<close>
lemma Maclaurin_lemma:
"0 < h ==>
@@ -61,7 +61,7 @@
moreover
have "(n - m) * t ^ (n - Suc m) * B / (fact (n - m)) =
B * (t ^ (n - Suc m) / (fact (n - Suc m)))"
- using `0 < n - m`
+ using \<open>0 < n - m\<close>
by (simp add: divide_simps fact_reduce)
ultimately show "DERIV (difg m) t :> difg (Suc m) t"
unfolding difg_def by simp
@@ -120,7 +120,7 @@
have "m < n" using m by simp
have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m) t :> 0"
- using `m < n`
+ using \<open>m < n\<close>
proof (induct m)
case 0
show ?case
@@ -140,26 +140,26 @@
proof (rule Rolle)
show "0 < t" by fact
show "difg (Suc m') 0 = difg (Suc m') t"
- using t `Suc m' < n` by (simp add: difg_Suc_eq_0 difg_eq_0)
+ using t \<open>Suc m' < n\<close> by (simp add: difg_Suc_eq_0 difg_eq_0)
show "\<forall>x. 0 \<le> x \<and> x \<le> t \<longrightarrow> isCont (difg (Suc m')) x"
- using `t < h` `Suc m' < n` by (simp add: isCont_difg)
+ using \<open>t < h\<close> \<open>Suc m' < n\<close> by (simp add: isCont_difg)
show "\<forall>x. 0 < x \<and> x < t \<longrightarrow> difg (Suc m') differentiable (at x)"
- using `t < h` `Suc m' < n` by (simp add: differentiable_difg)
+ using \<open>t < h\<close> \<open>Suc m' < n\<close> by (simp add: differentiable_difg)
qed
thus ?case
- using `t < h` by auto
+ using \<open>t < h\<close> by auto
qed
then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0" by fast
hence "difg (Suc m) t = 0"
- using `m < n` by (simp add: difg_Suc_eq_0)
+ using \<open>m < n\<close> by (simp add: difg_Suc_eq_0)
show ?thesis
proof (intro exI conjI)
show "0 < t" by fact
show "t < h" by fact
show "f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
- using `difg (Suc m) t = 0`
+ using \<open>difg (Suc m) t = 0\<close>
by (simp add: m f_h difg_def)
qed
qed
@@ -248,7 +248,7 @@
by (blast intro: Maclaurin_minus)
-subsection{*More Convenient "Bidirectional" Version.*}
+subsection\<open>More Convenient "Bidirectional" Version.\<close>
(* not good for PVS sin_approx, cos_approx *)
@@ -266,27 +266,27 @@
(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) +
diff n t / (fact n) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
proof cases
- assume "n = 0" with `diff 0 = f` show ?thesis by force
+ assume "n = 0" with \<open>diff 0 = f\<close> show ?thesis by force
next
assume "n \<noteq> 0"
show ?thesis
proof (cases rule: linorder_cases)
- assume "x = 0" with `n \<noteq> 0` `diff 0 = f` DERIV
+ assume "x = 0" with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
have "\<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by (auto simp add: Maclaurin_bi_le_lemma)
thus ?thesis ..
next
assume "x < 0"
- with `n \<noteq> 0` DERIV
+ with \<open>n \<noteq> 0\<close> DERIV
have "\<exists>t>x. t < 0 \<and> diff 0 x = ?f x t" by (intro Maclaurin_minus) auto
then guess t ..
- with `x < 0` `diff 0 = f` have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
+ with \<open>x < 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
thus ?thesis ..
next
assume "x > 0"
- with `n \<noteq> 0` `diff 0 = f` DERIV
+ with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV
have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t" by (intro Maclaurin) auto
then guess t ..
- with `x > 0` `diff 0 = f` have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
+ with \<open>x > 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
thus ?thesis ..
qed
qed
@@ -304,13 +304,13 @@
assume "x < 0"
with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t" by (intro Maclaurin_minus) auto
then guess t ..
- with `x < 0` have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
+ with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
thus ?thesis ..
next
assume "x > 0"
with assms have "\<exists>t>0. t < x \<and> f x = ?f x t " by (intro Maclaurin) auto
then guess t ..
- with `x > 0` have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
+ with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" by simp
thus ?thesis ..
qed
@@ -347,13 +347,13 @@
show ?thesis
proof cases
assume "x = 0"
- with `n \<noteq> 0` have "(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) = diff 0 0"
+ with \<open>n \<noteq> 0\<close> have "(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) = diff 0 0"
by (intro Maclaurin_zero) auto
- with INIT `x = 0` `n \<noteq> 0` have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by force
+ with INIT \<open>x = 0\<close> \<open>n \<noteq> 0\<close> have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by force
thus ?thesis ..
next
assume "x \<noteq> 0"
- with INIT `n \<noteq> 0` DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
+ with INIT \<open>n \<noteq> 0\<close> DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
by (intro Maclaurin_all_lt) auto
then guess t ..
hence "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t" by simp
@@ -370,7 +370,7 @@
by (blast intro: Maclaurin_all_le)
-subsection{*Version for Exponential Function*}
+subsection\<open>Version for Exponential Function\<close>
lemma Maclaurin_exp_lt:
fixes x::real
@@ -396,7 +396,7 @@
by (auto simp: numeral_3_eq_3 power2_eq_square power_Suc)
-subsection{*Version for Sine Function*}
+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::nat)"
@@ -415,7 +415,7 @@
by (induct "n", auto)
-text{*It is unclear why so many variant results are needed.*}
+text\<open>It is unclear why so many variant results are needed.\<close>
lemma sin_expansion_lemma:
"sin (x + real (Suc m) * pi / 2) =
@@ -485,7 +485,7 @@
done
-subsection{*Maclaurin Expansion for Cosine Function*}
+subsection\<open>Maclaurin Expansion for Cosine Function\<close>
lemma sumr_cos_zero_one [simp]:
"(\<Sum>m<(Suc n). cos_coeff m * 0 ^ m) = 1"
--- a/src/HOL/Main.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Main.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,15 +1,15 @@
-section {* Main HOL *}
+section \<open>Main HOL\<close>
theory Main
imports Predicate_Compile Quickcheck_Narrowing Extraction Coinduction Nitpick BNF_Greatest_Fixpoint Filter
begin
-text {*
+text \<open>
Classical Higher-order Logic -- only ``Main'', excluding real and
complex numbers etc.
-*}
+\<close>
-text {* See further @{cite "Nipkow-et-al:2002:tutorial"} *}
+text \<open>See further @{cite "Nipkow-et-al:2002:tutorial"}\<close>
no_notation
bot ("\<bottom>") and
--- a/src/HOL/Map.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Map.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,7 +5,7 @@
The datatype of `maps' (written ~=>); strongly resembles maps in VDM.
*)
-section {* Maps *}
+section \<open>Maps\<close>
theory Map
imports List
@@ -89,13 +89,13 @@
by simp_all
-subsection {* @{term [source] empty} *}
+subsection \<open>@{term [source] empty}\<close>
lemma empty_upd_none [simp]: "empty(x := None) = empty"
by (rule ext) simp
-subsection {* @{term [source] map_upd} *}
+subsection \<open>@{term [source] map_upd}\<close>
lemma map_upd_triv: "t k = Some x ==> t(k|->x) = t"
by (rule ext) simp
@@ -131,7 +131,7 @@
done
-subsection {* @{term [source] map_of} *}
+subsection \<open>@{term [source] map_of}\<close>
lemma map_of_eq_None_iff:
"(map_of xys x = None) = (x \<notin> fst ` (set xys))"
@@ -204,12 +204,12 @@
case Nil show ?case by simp
next
case (Cons y ys x xs z zs)
- from `map_of (zip (x#xs) (y#ys)) = map_of (zip (x#xs) (z#zs))`
+ from \<open>map_of (zip (x#xs) (y#ys)) = map_of (zip (x#xs) (z#zs))\<close>
have map_of: "map_of (zip xs ys)(x \<mapsto> y) = map_of (zip xs zs)(x \<mapsto> z)" by simp
from Cons have "length ys = length xs" and "length zs = length xs"
and "x \<notin> set xs" by simp_all
then have "map_of (zip xs ys) = map_of (zip xs zs)" using map_of by (rule map_of_zip_upd)
- with Cons.hyps `distinct (x # xs)` have "ys = zs" by simp
+ with Cons.hyps \<open>distinct (x # xs)\<close> have "ys = zs" by simp
moreover from map_of have "y = z" by (rule map_upd_eqD1)
ultimately show ?case by simp
qed
@@ -254,7 +254,7 @@
using dom_map_option [of "\<lambda>_. g" m] by (simp add: comp_def)
-subsection {* @{const map_option} related *}
+subsection \<open>@{const map_option} related\<close>
lemma map_option_o_empty [simp]: "map_option f o empty = empty"
by (rule ext) simp
@@ -264,7 +264,7 @@
by (rule ext) simp
-subsection {* @{term [source] map_comp} related *}
+subsection \<open>@{term [source] map_comp} related\<close>
lemma map_comp_empty [simp]:
"m \<circ>\<^sub>m empty = empty"
@@ -285,7 +285,7 @@
by (auto simp add: map_comp_def split: option.splits)
-subsection {* @{text "++"} *}
+subsection \<open>@{text "++"}\<close>
lemma map_add_empty[simp]: "m ++ empty = m"
by(simp add: map_add_def)
@@ -352,7 +352,7 @@
by (induct ps) (auto simp add: fun_eq_iff map_add_def)
-subsection {* @{term [source] restrict_map} *}
+subsection \<open>@{term [source] restrict_map}\<close>
lemma restrict_map_to_empty [simp]: "m|`{} = empty"
by (simp add: restrict_map_def)
@@ -405,7 +405,7 @@
by (simp add: restrict_map_def fun_eq_iff)
-subsection {* @{term [source] map_upds} *}
+subsection \<open>@{term [source] map_upds}\<close>
lemma map_upds_Nil1 [simp]: "m([] [|->] bs) = m"
by (simp add: map_upds_def)
@@ -485,7 +485,7 @@
done
-subsection {* @{term [source] dom} *}
+subsection \<open>@{term [source] dom}\<close>
lemma dom_eq_empty_conv [simp]: "dom f = {} \<longleftrightarrow> f = empty"
by (auto simp: dom_def)
@@ -605,14 +605,14 @@
proof
fix m assume "m \<in> ?S'"
hence 1: "dom m = A" by force
- hence 2: "ran m \<subseteq> B" using `m \<in> ?S'` by(auto simp: dom_def ran_def)
+ hence 2: "ran m \<subseteq> B" using \<open>m \<in> ?S'\<close> by(auto simp: dom_def ran_def)
from 1 2 show "m \<in> ?S" by blast
qed
qed
with assms show ?thesis by(simp add: finite_set_of_finite_funs)
qed
-subsection {* @{term [source] ran} *}
+subsection \<open>@{term [source] ran}\<close>
lemma ranI: "m a = Some b ==> b : ran m"
by(auto simp: ran_def)
@@ -644,7 +644,7 @@
lemma ran_map_option: "ran (\<lambda>x. map_option f (m x)) = f ` ran m"
by(auto simp add: ran_def)
-subsection {* @{text "map_le"} *}
+subsection \<open>@{text "map_le"}\<close>
lemma map_le_empty [simp]: "empty \<subseteq>\<^sub>m g"
by (simp add: map_le_def)
@@ -702,14 +702,14 @@
assume "dom f = {x}"
then obtain v where "f x = Some v" by auto
hence "[x \<mapsto> v] \<subseteq>\<^sub>m f" by(auto simp add: map_le_def)
- moreover have "f \<subseteq>\<^sub>m [x \<mapsto> v]" using `dom f = {x}` `f x = Some v`
+ moreover have "f \<subseteq>\<^sub>m [x \<mapsto> v]" using \<open>dom f = {x}\<close> \<open>f x = Some v\<close>
by(auto simp add: map_le_def)
ultimately have "f = [x \<mapsto> v]" by-(rule map_le_antisym)
thus "\<exists>v. f = [x \<mapsto> v]" by blast
qed
-subsection {* Various *}
+subsection \<open>Various\<close>
lemma set_map_of_compr:
assumes distinct: "distinct (map fst xs)"
@@ -725,7 +725,7 @@
{(k', v'). (map_of xs(k \<mapsto> v)) k' = Some v'}"
by (auto split: if_splits)
from Cons have "set xs = {(k, v). map_of xs k = Some v}" by simp
- with * `x = (k, v)` show ?case by simp
+ with * \<open>x = (k, v)\<close> show ?case by simp
qed
lemma map_of_inject_set:
@@ -733,9 +733,9 @@
shows "map_of xs = map_of ys \<longleftrightarrow> set xs = set ys" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
- moreover from `distinct (map fst xs)` have "set xs = {(k, v). map_of xs k = Some v}"
+ moreover from \<open>distinct (map fst xs)\<close> have "set xs = {(k, v). map_of xs k = Some v}"
by (rule set_map_of_compr)
- moreover from `distinct (map fst ys)` have "set ys = {(k, v). map_of ys k = Some v}"
+ moreover from \<open>distinct (map fst ys)\<close> have "set ys = {(k, v). map_of ys k = Some v}"
by (rule set_map_of_compr)
ultimately show ?rhs by simp
next
@@ -744,12 +744,12 @@
fix k
show "map_of xs k = map_of ys k" proof (cases "map_of xs k")
case None
- with `?rhs` have "map_of ys k = None"
+ with \<open>?rhs\<close> have "map_of ys k = None"
by (simp add: map_of_eq_None_iff)
with None show ?thesis by simp
next
case (Some v)
- with distinct `?rhs` have "map_of ys k = Some v"
+ with distinct \<open>?rhs\<close> have "map_of ys k = Some v"
by simp
with Some show ?thesis by simp
qed
--- a/src/HOL/Meson.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Meson.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,15 +5,15 @@
Copyright 2001 University of Cambridge
*)
-section {* MESON Proof Method *}
+section \<open>MESON Proof Method\<close>
theory Meson
imports Nat
begin
-subsection {* Negation Normal Form *}
+subsection \<open>Negation Normal Form\<close>
-text {* de Morgan laws *}
+text \<open>de Morgan laws\<close>
lemma not_conjD: "~(P&Q) ==> ~P | ~Q"
and not_disjD: "~(P|Q) ==> ~P & ~Q"
@@ -22,32 +22,32 @@
and not_exD: "!!P. ~(\<exists>x. P(x)) ==> \<forall>x. ~P(x)"
by fast+
-text {* Removal of @{text "-->"} and @{text "<->"} (positive and
-negative occurrences) *}
+text \<open>Removal of @{text "-->"} and @{text "<->"} (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)"
- -- {* Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF *}
+ -- \<open>Much more efficient than @{prop "(P & ~Q) | (Q & ~P)"} for computing CNF\<close>
and not_refl_disj_D: "x ~= x | P ==> P"
by fast+
-subsection {* Pulling out the existential quantifiers *}
+subsection \<open>Pulling out the existential quantifiers\<close>
-text {* Conjunction *}
+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)"
by fast+
-text {* Disjunction *}
+text \<open>Disjunction\<close>
lemma disj_exD: "!!P Q. (\<exists>x. P(x)) | (\<exists>x. Q(x)) ==> \<exists>x. P(x) | Q(x)"
- -- {* DO NOT USE with forall-Skolemization: makes fewer schematic variables!! *}
- -- {* With ex-Skolemization, makes fewer Skolem constants *}
+ -- \<open>DO NOT USE with forall-Skolemization: makes fewer schematic variables!!\<close>
+ -- \<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)"
by fast+
@@ -59,31 +59,31 @@
by fast+
-text{* Generation of contrapositives *}
+text\<open>Generation of contrapositives\<close>
-text{*Inserts negated disjunct after removing the negation; P is a literal.
+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.*}
+ hence the negated disjuncts.\<close>
lemma make_neg_rule: "~P|Q ==> ((~P==>P) ==> Q)"
by blast
-text{*Version for Plaisted's "Postive refinement" of the Meson procedure*}
+text\<open>Version for Plaisted's "Postive refinement" of the Meson procedure\<close>
lemma make_refined_neg_rule: "~P|Q ==> (P ==> Q)"
by blast
-text{*@{term P} should be a literal*}
+text\<open>@{term P} should be a literal\<close>
lemma make_pos_rule: "P|Q ==> ((P==>~P) ==> Q)"
by blast
-text{*Versions of @{text make_neg_rule} and @{text make_pos_rule} that don't
-insert new assumptions, for ordinary resolution.*}
+text\<open>Versions of @{text make_neg_rule} and @{text make_pos_rule} that don't
+insert new assumptions, for ordinary resolution.\<close>
lemmas make_neg_rule' = make_refined_neg_rule
lemma make_pos_rule': "[|P|Q; ~P|] ==> Q"
by blast
-text{* Generation of a goal clause -- put away the final literal *}
+text\<open>Generation of a goal clause -- put away the final literal\<close>
lemma make_neg_goal: "~P ==> ((~P==>P) ==> False)"
by blast
@@ -92,9 +92,9 @@
by blast
-subsection {* Lemmas for Forward Proof *}
+subsection \<open>Lemmas for Forward Proof\<close>
-text{*There is a similarity to congruence rules*}
+text\<open>There is a similarity to congruence rules\<close>
(*NOTE: could handle conjunctions (faster?) by
nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *)
@@ -117,7 +117,7 @@
by blast
-subsection {* Clausification helper *}
+subsection \<open>Clausification helper\<close>
lemma TruepropI: "P \<equiv> Q \<Longrightarrow> Trueprop P \<equiv> Trueprop Q"
by simp
@@ -129,7 +129,7 @@
by auto
-text{* Combinator translation helpers *}
+text\<open>Combinator translation helpers\<close>
definition COMBI :: "'a \<Rightarrow> 'a" where
"COMBI P = P"
@@ -177,7 +177,7 @@
done
-subsection {* Skolemization helpers *}
+subsection \<open>Skolemization helpers\<close>
definition skolem :: "'a \<Rightarrow> 'a" where
"skolem = (\<lambda>x. x)"
@@ -189,7 +189,7 @@
lemmas skolem_COMBK_D = iffD2 [OF skolem_COMBK_iff]
-subsection {* Meson package *}
+subsection \<open>Meson package\<close>
ML_file "Tools/Meson/meson.ML"
ML_file "Tools/Meson/meson_clausify.ML"
--- a/src/HOL/Metis.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Metis.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* Metis Proof Method *}
+section \<open>Metis Proof Method\<close>
theory Metis
imports ATP
@@ -15,7 +15,7 @@
declare [[ML_print_depth = 10]]
-subsection {* Literal selection and lambda-lifting helpers *}
+subsection \<open>Literal selection and lambda-lifting helpers\<close>
definition select :: "'a \<Rightarrow> 'a" where
"select = (\<lambda>x. x)"
@@ -38,7 +38,7 @@
unfolding lambda_def by assumption
-subsection {* Metis package *}
+subsection \<open>Metis package\<close>
ML_file "Tools/Metis/metis_generate.ML"
ML_file "Tools/Metis/metis_reconstruct.ML"
--- a/src/HOL/Nat.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Nat.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,7 +5,7 @@
and * (for div and mod, see theory Divides).
*)
-section {* Natural numbers *}
+section \<open>Natural numbers\<close>
theory Nat
imports Inductive Typedef Fun Fields
@@ -18,18 +18,18 @@
ML_file "~~/src/Provers/Arith/fast_lin_arith.ML"
-subsection {* Type @{text ind} *}
+subsection \<open>Type @{text ind}\<close>
typedecl ind
axiomatization Zero_Rep :: ind and Suc_Rep :: "ind => ind" where
- -- {* the axiom of infinity in 2 parts *}
+ -- \<open>the axiom of infinity in 2 parts\<close>
Suc_Rep_inject: "Suc_Rep x = Suc_Rep y ==> x = y" and
Suc_Rep_not_Zero_Rep: "Suc_Rep x \<noteq> Zero_Rep"
-subsection {* Type nat *}
-
-text {* Type definition *}
+subsection \<open>Type nat\<close>
+
+text \<open>Type definition\<close>
inductive Nat :: "ind \<Rightarrow> bool" where
Zero_RepI: "Nat Zero_Rep"
@@ -79,7 +79,7 @@
shows "P n"
using assms
apply (unfold Zero_nat_def Suc_def)
-apply (rule Rep_Nat_inverse [THEN subst]) -- {* types force good instantiation *}
+apply (rule Rep_Nat_inverse [THEN subst]) -- \<open>types force good instantiation\<close>
apply (erule Nat_Rep_Nat [THEN Nat.induct])
apply (iprover elim: Nat_Abs_Nat_inverse [THEN subst])
done
@@ -96,8 +96,8 @@
apply (simp only: Suc_not_Zero)
done
--- {* Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}. *}
-setup {* Sign.mandatory_path "old" *}
+-- \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
+setup \<open>Sign.mandatory_path "old"\<close>
old_rep_datatype "0 \<Colon> nat" Suc
apply (erule nat_induct0, assumption)
@@ -105,10 +105,10 @@
apply (rule nat.distinct(1))
done
-setup {* Sign.parent_path *}
-
--- {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
-setup {* Sign.mandatory_path "nat" *}
+setup \<open>Sign.parent_path\<close>
+
+-- \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>
+setup \<open>Sign.mandatory_path "nat"\<close>
declare
old.nat.inject[iff del]
@@ -119,14 +119,14 @@
lemmas rec = old.nat.rec
lemmas simps = nat.inject nat.distinct nat.case nat.rec
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
abbreviation rec_nat :: "'a \<Rightarrow> (nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a" where
"rec_nat \<equiv> old.rec_nat"
declare nat.sel[code del]
-hide_const (open) Nat.pred -- {* hide everything related to the selector *}
+hide_const (open) Nat.pred -- \<open>hide everything related to the selector\<close>
hide_fact
nat.case_eq_if
nat.collapse
@@ -137,12 +137,12 @@
nat.split_sel_asm
lemma nat_exhaust [case_names 0 Suc, cases type: nat]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
"(y = 0 \<Longrightarrow> P) \<Longrightarrow> (\<And>nat. y = Suc nat \<Longrightarrow> P) \<Longrightarrow> P"
by (rule old.nat.exhaust)
lemma nat_induct [case_names 0 Suc, induct type: nat]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
fixes n
assumes "P 0" and "\<And>n. P n \<Longrightarrow> P (Suc n)"
shows "P n"
@@ -152,7 +152,7 @@
nat_exhaust
nat_induct0
-ML {*
+ML \<open>
val nat_basic_lfp_sugar =
let
val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global @{theory} @{type_name nat});
@@ -162,9 +162,9 @@
{T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]],
ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}}
end;
-*}
-
-setup {*
+\<close>
+
+setup \<open>
let
fun basic_lfp_sugars_of _ [@{typ nat}] _ _ ctxt =
([], [0], [nat_basic_lfp_sugar], [], [], TrueI (*dummy*), [], false, ctxt)
@@ -175,9 +175,9 @@
{nested_simps = [], is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of,
rewrite_nested_rec_call = NONE}
end
-*}
-
-text {* Injectiveness and distinctness lemmas *}
+\<close>
+
+text \<open>Injectiveness and distinctness lemmas\<close>
lemma inj_Suc[simp]: "inj_on Suc N"
by (simp add: inj_on_def)
@@ -197,8 +197,8 @@
lemma Suc_n_not_n: "Suc n \<noteq> n"
by (rule not_sym, rule n_not_Suc_n)
-text {* A special form of induction for reasoning
- about @{term "m < n"} and @{term "m - n"} *}
+text \<open>A special form of induction for reasoning
+ about @{term "m < n"} and @{term "m - n"}\<close>
lemma diff_induct: "(!!x. P x 0) ==> (!!y. P 0 (Suc y)) ==>
(!!x y. P x y ==> P (Suc x) (Suc y)) ==> P m n"
@@ -210,7 +210,7 @@
done
-subsection {* Arithmetic operators *}
+subsection \<open>Arithmetic operators\<close>
instantiation nat :: comm_monoid_diff
begin
@@ -290,7 +290,7 @@
end
-text {* Difference distributes over multiplication *}
+text \<open>Difference distributes over multiplication\<close>
lemma diff_mult_distrib:
"((m::nat) - n) * k = (m * k) - (n * k)"
@@ -301,7 +301,7 @@
by (fact right_diff_distrib')
-subsubsection {* Addition *}
+subsubsection \<open>Addition\<close>
lemma nat_add_left_cancel:
fixes k m n :: nat
@@ -313,7 +313,7 @@
shows "m + k = n + k \<longleftrightarrow> m = n"
by (fact add_right_cancel)
-text {* Reasoning about @{text "m + 0 = 0"}, etc. *}
+text \<open>Reasoning about @{text "m + 0 = 0"}, etc.\<close>
lemma add_is_0 [iff]:
fixes m n :: nat
@@ -347,7 +347,7 @@
unfolding One_nat_def by simp
-subsubsection {* Difference *}
+subsubsection \<open>Difference\<close>
lemma diff_self_eq_0 [simp]: "(m\<Colon>nat) - m = 0"
by (fact diff_cancel)
@@ -379,7 +379,7 @@
lemma diff_Suc_1 [simp]: "Suc n - 1 = n"
unfolding One_nat_def by simp
-subsubsection {* Multiplication *}
+subsubsection \<open>Multiplication\<close>
lemma add_mult_distrib2: "k * (m + n) = (k * m) + ((k * n)::nat)"
by (fact distrib_left)
@@ -427,9 +427,9 @@
by (subst mult_cancel1) simp
-subsection {* Orders on @{typ nat} *}
-
-subsubsection {* Operation definition *}
+subsection \<open>Orders on @{typ nat}\<close>
+
+subsubsection \<open>Operation definition\<close>
instantiation nat :: linorder
begin
@@ -532,7 +532,7 @@
by default (auto intro: less_Suc_eq_le [THEN iffD2])
-subsubsection {* Introduction properties *}
+subsubsection \<open>Introduction properties\<close>
lemma lessI [iff]: "n < Suc n"
by (simp add: less_Suc_eq_le)
@@ -541,7 +541,7 @@
by (simp add: less_Suc_eq_le)
-subsubsection {* Elimination properties *}
+subsubsection \<open>Elimination properties\<close>
lemma less_not_refl: "~ n < (n::nat)"
by (rule order_less_irrefl)
@@ -570,7 +570,7 @@
lemma Suc_mono: "m < n ==> Suc m < Suc n"
by simp
-text {* "Less than" is antisymmetric, sort of *}
+text \<open>"Less than" is antisymmetric, sort of\<close>
lemma less_antisym: "\<lbrakk> \<not> n < m; n < Suc m \<rbrakk> \<Longrightarrow> m = n"
unfolding not_less less_Suc_eq_le by (rule antisym)
@@ -588,7 +588,7 @@
done
-subsubsection {* Inductive (?) properties *}
+subsubsection \<open>Inductive (?) properties\<close>
lemma Suc_lessI: "m < n ==> Suc m \<noteq> n ==> Suc m < n"
unfolding less_eq_Suc_le [of m] le_less by simp
@@ -630,14 +630,14 @@
apply (blast dest: Suc_lessD)
done
-text {* Can be used with @{text less_Suc_eq} to get @{term "n = m | n < m"} *}
+text \<open>Can be used with @{text less_Suc_eq} to get @{term "n = m | n < m"}\<close>
lemma not_less_eq: "\<not> m < n \<longleftrightarrow> n < Suc m"
unfolding not_less less_Suc_eq_le ..
lemma not_less_eq_eq: "\<not> m \<le> n \<longleftrightarrow> Suc n \<le> m"
unfolding not_le Suc_le_eq ..
-text {* Properties of "less than or equal" *}
+text \<open>Properties of "less than or equal"\<close>
lemma le_imp_less_Suc: "m \<le> n ==> m < Suc n"
unfolding less_Suc_eq_le .
@@ -654,18 +654,18 @@
lemma Suc_leI: "m < n ==> Suc(m) \<le> n"
unfolding Suc_le_eq .
-text {* Stronger version of @{text Suc_leD} *}
+text \<open>Stronger version of @{text Suc_leD}\<close>
lemma Suc_le_lessD: "Suc m \<le> n ==> m < n"
unfolding Suc_le_eq .
lemma less_imp_le_nat: "m < n ==> m \<le> (n::nat)"
unfolding less_eq_Suc_le by (rule Suc_leD)
-text {* For instance, @{text "(Suc m < Suc n) = (Suc m \<le> n) = (m < n)"} *}
+text \<open>For instance, @{text "(Suc m < Suc n) = (Suc m \<le> n) = (m < n)"}\<close>
lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq
-text {* Equivalence of @{term "m \<le> n"} and @{term "m < n | m = n"} *}
+text \<open>Equivalence of @{term "m \<le> n"} and @{term "m < n | m = n"}\<close>
lemma less_or_eq_imp_le: "m < n | m = n ==> m \<le> (n::nat)"
unfolding le_less .
@@ -673,7 +673,7 @@
lemma le_eq_less_or_eq: "(m \<le> (n::nat)) = (m < n | m=n)"
by (rule le_less)
-text {* Useful with @{text blast}. *}
+text \<open>Useful with @{text blast}.\<close>
lemma eq_imp_le: "(m::nat) = n ==> m \<le> n"
by auto
@@ -717,7 +717,7 @@
lemma neq0_conv[iff]: fixes n :: nat shows "(n \<noteq> 0) = (0 < n)"
by (cases n) simp_all
-text {* This theorem is useful with @{text blast} *}
+text \<open>This theorem is useful with @{text blast}\<close>
lemma gr0I: "((n::nat) = 0 ==> False) ==> 0 < n"
by (rule neq0_conv[THEN iffD1], iprover)
@@ -730,12 +730,12 @@
lemma Suc_le_D: "(Suc n \<le> m') ==> (? m. m' = Suc m)"
by (induct m') simp_all
-text {* Useful in certain inductive arguments *}
+text \<open>Useful in certain inductive arguments\<close>
lemma less_Suc_eq_0_disj: "(m < Suc n) = (m = 0 | (\<exists>j. m = Suc j & j < n))"
by (cases m) simp_all
-subsubsection {* Monotonicity of Addition *}
+subsubsection \<open>Monotonicity of Addition\<close>
lemma Suc_pred [simp]: "n>0 ==> Suc (n - Suc 0) = n"
by (simp add: diff_Suc split: nat.split)
@@ -752,17 +752,17 @@
lemma add_gr_0 [iff]: "!!m::nat. (m + n > 0) = (m>0 | n>0)"
by(auto dest:gr0_implies_Suc)
-text {* strict, in 1st argument *}
+text \<open>strict, in 1st argument\<close>
lemma add_less_mono1: "i < j ==> i + k < j + (k::nat)"
by (induct k) simp_all
-text {* strict, in both arguments *}
+text \<open>strict, in both arguments\<close>
lemma add_less_mono: "[|i < j; k < l|] ==> i + k < j + (l::nat)"
apply (rule add_less_mono1 [THEN less_trans], assumption+)
apply (induct j, simp_all)
done
-text {* Deleted @{text less_natE}; use @{text "less_imp_Suc_add RS exE"} *}
+text \<open>Deleted @{text less_natE}; use @{text "less_imp_Suc_add RS exE"}\<close>
lemma less_imp_Suc_add: "m < n ==> (\<exists>k. n = Suc (m + k))"
apply (induct n)
apply (simp_all add: order_le_less)
@@ -773,20 +773,20 @@
lemma le_Suc_ex: "(k::nat) \<le> l \<Longrightarrow> (\<exists>n. l = k + n)"
by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add)
-text {* strict, in 1st argument; proof is by induction on @{text "k > 0"} *}
+text \<open>strict, in 1st argument; proof is by induction on @{text "k > 0"}\<close>
lemma mult_less_mono2: "(i::nat) < j ==> 0<k ==> k * i < k * j"
apply(auto simp: gr0_conv_Suc)
apply (induct_tac m)
apply (simp_all add: add_less_mono)
done
-text {* Addition is the inverse of subtraction:
- if @{term "n \<le> m"} then @{term "n + (m - n) = m"}. *}
+text \<open>Addition is the inverse of subtraction:
+ if @{term "n \<le> m"} then @{term "n + (m - n) = m"}.\<close>
lemma add_diff_inverse_nat: "~ m < n ==> n + (m - n) = (m::nat)"
by (induct m n rule: diff_induct) simp_all
-text{*The naturals form an ordered @{text semidom}*}
+text\<open>The naturals form an ordered @{text semidom}\<close>
instance nat :: linordered_semidom
proof
show "0 < (1::nat)" by simp
@@ -798,7 +798,7 @@
qed
-subsubsection {* @{term min} and @{term max} *}
+subsubsection \<open>@{term min} and @{term max}\<close>
lemma mono_Suc: "mono Suc"
by (rule monoI) simp
@@ -868,9 +868,9 @@
by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans)
-subsubsection {* Additional theorems about @{term "op \<le>"} *}
-
-text {* Complete induction, aka course-of-values induction *}
+subsubsection \<open>Additional theorems about @{term "op \<le>"}\<close>
+
+text \<open>Complete induction, aka course-of-values induction\<close>
instance nat :: wellorder proof
fix P and n :: nat
@@ -947,7 +947,7 @@
shows "P a"
by (induct m\<equiv>"f a" arbitrary: a rule: less_induct) (auto intro: step)
-text {* old style induction rules: *}
+text \<open>old style induction rules:\<close>
lemma measure_induct:
fixes f :: "'a \<Rightarrow> nat"
shows "(\<And>x. \<forall>y. f y < f x \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a"
@@ -958,7 +958,7 @@
shows "P n"
by (rule less_induct) (auto intro: step simp:le_simps)
-text{*An induction rule for estabilishing binary relations*}
+text\<open>An induction rule for estabilishing binary relations\<close>
lemma less_Suc_induct:
assumes less: "i < j"
and step: "!!i. P i (Suc i)"
@@ -980,16 +980,16 @@
thus "P i j" by (simp add: j)
qed
-text {* The method of infinite descent, frequently used in number theory.
+text \<open>The method of infinite descent, frequently used in number theory.
Provided by Roelof Oosterhuis.
$P(n)$ is true for all $n\in\mathbb{N}$ if
\begin{itemize}
\item case ``0'': given $n=0$ prove $P(n)$,
\item case ``smaller'': given $n>0$ and $\neg P(n)$ prove there exists
a smaller integer $m$ such that $\neg P(m)$.
-\end{itemize} *}
-
-text{* A compact version without explicit base case: *}
+\end{itemize}\<close>
+
+text\<open>A compact version without explicit base case:\<close>
lemma infinite_descent:
"\<lbrakk> !!n::nat. \<not> P n \<Longrightarrow> \<exists>m<n. \<not> P m \<rbrakk> \<Longrightarrow> P n"
by (induct n rule: less_induct) auto
@@ -998,14 +998,14 @@
"\<lbrakk> P 0; !!n. n>0 \<Longrightarrow> \<not> P n \<Longrightarrow> (\<exists>m::nat. m < n \<and> \<not>P m) \<rbrakk> \<Longrightarrow> P n"
by (rule infinite_descent) (case_tac "n>0", auto)
-text {*
+text \<open>
Infinite descent using a mapping to $\mathbb{N}$:
$P(x)$ is true for all $x\in D$ if there exists a $V: D \to \mathbb{N}$ and
\begin{itemize}
\item case ``0'': given $V(x)=0$ prove $P(x)$,
\item case ``smaller'': given $V(x)>0$ and $\neg P(x)$ prove there exists a $y \in D$ such that $V(y)<V(x)$ and $~\neg P(y)$.
\end{itemize}
-NB: the proof also shows how to use the previous lemma. *}
+NB: the proof also shows how to use the previous lemma.\<close>
corollary infinite_descent0_measure [case_names 0 smaller]:
assumes A0: "!!x. V x = (0::nat) \<Longrightarrow> P x"
@@ -1027,7 +1027,7 @@
ultimately show "P x" by auto
qed
-text{* Again, without explicit base case: *}
+text\<open>Again, without explicit base case:\<close>
lemma infinite_descent_measure:
assumes "!!x. \<not> P x \<Longrightarrow> \<exists>y. (V::'a\<Rightarrow>nat) y < V x \<and> \<not> P y" shows "P x"
proof -
@@ -1040,18 +1040,18 @@
ultimately show "P x" by auto
qed
-text {* A [clumsy] way of lifting @{text "<"}
- monotonicity to @{text "\<le>"} monotonicity *}
+text \<open>A [clumsy] way of lifting @{text "<"}
+ monotonicity to @{text "\<le>"} monotonicity\<close>
lemma less_mono_imp_le_mono:
"\<lbrakk> !!i j::nat. i < j \<Longrightarrow> f i < f j; i \<le> j \<rbrakk> \<Longrightarrow> f i \<le> ((f j)::nat)"
by (simp add: order_le_less) (blast)
-text {* non-strict, in 1st argument *}
+text \<open>non-strict, in 1st argument\<close>
lemma add_le_mono1: "i \<le> j ==> i + k \<le> j + (k::nat)"
by (rule add_right_mono)
-text {* non-strict, in both arguments *}
+text \<open>non-strict, in both arguments\<close>
lemma add_le_mono: "[| i \<le> j; k \<le> l |] ==> i + k \<le> j + (l::nat)"
by (rule add_mono)
@@ -1109,13 +1109,13 @@
lemma add_leE: "(m::nat) + k \<le> n ==> (m \<le> n ==> k \<le> n ==> R) ==> R"
by (blast dest: add_leD1 add_leD2)
-text {* needs @{text "!!k"} for @{text ac_simps} to work *}
+text \<open>needs @{text "!!k"} for @{text ac_simps} to work\<close>
lemma less_add_eq_less: "!!k::nat. k < l ==> m + l = k + n ==> m < n"
by (force simp del: add_Suc_right
simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps)
-subsubsection {* More results about difference *}
+subsubsection \<open>More results about difference\<close>
lemma Suc_diff_le: "n \<le> m ==> Suc m - n = Suc (m - n)"
by (induct m n rule: diff_induct) simp_all
@@ -1169,7 +1169,7 @@
by (simp add: order_less_imp_le)
qed
-text {* a nice rewrite for bounded subtraction *}
+text \<open>a nice rewrite for bounded subtraction\<close>
lemma nat_minus_add_max:
fixes n m :: nat
shows "n - m + m = max n m"
@@ -1177,14 +1177,14 @@
lemma nat_diff_split:
"P(a - b::nat) = ((a<b --> P 0) & (ALL d. a = b + d --> P d))"
- -- {* elimination of @{text -} on @{text nat} *}
+ -- \<open>elimination of @{text -} on @{text nat}\<close>
by (cases "a < b")
(auto simp add: diff_is_0_eq [THEN iffD2] diff_add_inverse
not_less le_less dest!: add_eq_self_zero add_eq_self_zero[OF sym])
lemma nat_diff_split_asm:
"P(a - b::nat) = (~ (a < b & ~ P 0 | (EX d. a = b + d & ~ P d)))"
- -- {* elimination of @{text -} on @{text nat} in assumptions *}
+ -- \<open>elimination of @{text -} on @{text nat} in assumptions\<close>
by (auto split: nat_diff_split)
lemma Suc_pred': "0 < n ==> n = Suc(n - 1)"
@@ -1206,7 +1206,7 @@
by (fact Let_def)
-subsubsection {* Monotonicity of multiplication *}
+subsubsection \<open>Monotonicity of multiplication\<close>
lemma mult_le_mono1: "i \<le> (j::nat) ==> i * k \<le> j * k"
by (simp add: mult_right_mono)
@@ -1214,15 +1214,15 @@
lemma mult_le_mono2: "i \<le> (j::nat) ==> k * i \<le> k * j"
by (simp add: mult_left_mono)
-text {* @{text "\<le>"} monotonicity, BOTH arguments *}
+text \<open>@{text "\<le>"} monotonicity, BOTH arguments\<close>
lemma mult_le_mono: "i \<le> (j::nat) ==> k \<le> l ==> i * k \<le> j * l"
by (simp add: mult_mono)
lemma mult_less_mono1: "(i::nat) < j ==> 0 < k ==> i * k < j * k"
by (simp add: mult_strict_right_mono)
-text{*Differs from the standard @{text zero_less_mult_iff} in that
- there are no negative numbers.*}
+text\<open>Differs from the standard @{text zero_less_mult_iff} in that
+ there are no negative numbers.\<close>
lemma nat_0_less_mult_iff [simp]: "(0 < (m::nat) * n) = (0 < m & 0 < n)"
apply (induct m)
apply simp
@@ -1265,7 +1265,7 @@
lemma le_cube: "(m::nat) \<le> m * (m * m)"
by (cases m) (auto intro: le_add1)
-text {* Lemma for @{text gcd} *}
+text \<open>Lemma for @{text gcd}\<close>
lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = 1 | m = 0"
apply (drule sym)
apply (rule disjCI)
@@ -1284,7 +1284,7 @@
with assms show "n * m \<le> n * q" by simp
qed
-text {* the lattice order on @{typ nat} *}
+text \<open>the lattice order on @{typ nat}\<close>
instantiation nat :: distrib_lattice
begin
@@ -1302,12 +1302,12 @@
end
-subsection {* Natural operation of natural numbers on functions *}
-
-text {*
+subsection \<open>Natural operation of natural numbers on functions\<close>
+
+text \<open>
We use the same logical constant for the power operations on
functions and relations, in order to share the same syntax.
-*}
+\<close>
consts compow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -1320,7 +1320,7 @@
notation (HTML output)
compower ("(_\<^bsup>_\<^esup>)" [1000] 1000)
-text {* @{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f} *}
+text \<open>@{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f}\<close>
overloading
funpow == "compow :: nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)"
@@ -1345,7 +1345,7 @@
lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right
-text {* for code generation *}
+text \<open>for code generation\<close>
definition funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
funpow_code_def [code_abbrev]: "funpow = compow"
@@ -1392,7 +1392,7 @@
by (induct n arbitrary: A B)
(auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def)
-subsection {* Kleene iteration *}
+subsection \<open>Kleene iteration\<close>
lemma Kleene_iter_lpfp:
assumes "mono f" and "f p \<le> p" shows "(f^^k) (bot::'a::order_bot) \<le> p"
@@ -1453,7 +1453,7 @@
by (intro gfp_upperbound) (simp del: funpow.simps)
qed
-subsection {* Embedding of the naturals into any @{text semiring_1}: @{term of_nat} *}
+subsection \<open>Embedding of the naturals into any @{text semiring_1}: @{term of_nat}\<close>
context semiring_1
begin
@@ -1477,7 +1477,7 @@
primrec of_nat_aux :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
"of_nat_aux inc 0 i = i"
-| "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" -- {* tail recursive *}
+| "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" -- \<open>tail recursive\<close>
lemma of_nat_code:
"of_nat n = of_nat_aux (\<lambda>i. i + 1) n 0"
@@ -1496,8 +1496,8 @@
declare of_nat_code [code]
-text{*Class for unital semirings with characteristic zero.
- Includes non-ordered rings like the complex numbers.*}
+text\<open>Class for unital semirings with characteristic zero.
+ Includes non-ordered rings like the complex numbers.\<close>
class semiring_char_0 = semiring_1 +
assumes inj_of_nat: "inj of_nat"
@@ -1506,7 +1506,7 @@
lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \<longleftrightarrow> m = n"
by (auto intro: inj_of_nat injD)
-text{*Special cases where either operand is zero*}
+text\<open>Special cases where either operand is zero\<close>
lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0])
@@ -1545,12 +1545,12 @@
lemma of_nat_less_imp_less: "of_nat m < of_nat n \<Longrightarrow> m < n"
by simp
-text{*Every @{text linordered_semidom} has characteristic zero.*}
+text\<open>Every @{text linordered_semidom} has characteristic zero.\<close>
subclass semiring_char_0 proof
qed (auto intro!: injI simp add: eq_iff)
-text{*Special cases where either operand is zero*}
+text\<open>Special cases where either operand is zero\<close>
lemma of_nat_le_0_iff [simp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
by (rule of_nat_le_iff [of _ 0, simplified])
@@ -1583,7 +1583,7 @@
by (auto simp add: fun_eq_iff)
-subsection {* The set of natural numbers *}
+subsection \<open>The set of natural numbers\<close>
context semiring_1
begin
@@ -1626,7 +1626,7 @@
obtains (of_nat) n where "x = of_nat n"
unfolding Nats_def
proof -
- from `x \<in> \<nat>` have "x \<in> range of_nat" unfolding Nats_def .
+ from \<open>x \<in> \<nat>\<close> have "x \<in> range of_nat" unfolding Nats_def .
then obtain n where "x = of_nat n" ..
then show thesis ..
qed
@@ -1638,7 +1638,7 @@
end
-subsection {* Further arithmetic facts concerning the natural numbers *}
+subsection \<open>Further arithmetic facts concerning the natural numbers\<close>
lemma subst_equals:
assumes 1: "t = s" and 2: "u = t"
@@ -1649,26 +1649,26 @@
simproc_setup nateq_cancel_sums
("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") =
- {* fn phi => try o Nat_Arith.cancel_eq_conv *}
+ \<open>fn phi => try o Nat_Arith.cancel_eq_conv\<close>
simproc_setup natless_cancel_sums
("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") =
- {* fn phi => try o Nat_Arith.cancel_less_conv *}
+ \<open>fn phi => try o Nat_Arith.cancel_less_conv\<close>
simproc_setup natle_cancel_sums
("(l::nat) + m \<le> n" | "(l::nat) \<le> m + n" | "Suc m \<le> n" | "m \<le> Suc n") =
- {* fn phi => try o Nat_Arith.cancel_le_conv *}
+ \<open>fn phi => try o Nat_Arith.cancel_le_conv\<close>
simproc_setup natdiff_cancel_sums
("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") =
- {* fn phi => try o Nat_Arith.cancel_diff_conv *}
+ \<open>fn phi => try o Nat_Arith.cancel_diff_conv\<close>
ML_file "Tools/lin_arith.ML"
-setup {* Lin_Arith.global_setup *}
-declaration {* K Lin_Arith.setup *}
+setup \<open>Lin_Arith.global_setup\<close>
+declaration \<open>K Lin_Arith.setup\<close>
simproc_setup fast_arith_nat ("(m::nat) < n" | "(m::nat) <= n" | "(m::nat) = n") =
- {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
+ \<open>fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct)\<close>
(* Because of this simproc, the arithmetic solver is really only
useful to detect inconsistencies among the premises for subgoals which are
*not* themselves (in)equalities, because the latter activate
@@ -1688,7 +1688,7 @@
case True
then show ?thesis
by (induct n n' rule: less_Suc_induct [consumes 1]) (auto intro: mono)
-qed (insert `n \<le> n'`, auto) -- {* trivial for @{prop "n = n'"} *}
+qed (insert \<open>n \<le> n'\<close>, auto) -- \<open>trivial for @{prop "n = n'"}\<close>
lemma lift_Suc_antimono_le:
assumes mono: "\<And>n. f n \<ge> f (Suc n)" and "n \<le> n'"
@@ -1697,12 +1697,12 @@
case True
then show ?thesis
by (induct n n' rule: less_Suc_induct [consumes 1]) (auto intro: mono)
-qed (insert `n \<le> n'`, auto) -- {* trivial for @{prop "n = n'"} *}
+qed (insert \<open>n \<le> n'\<close>, auto) -- \<open>trivial for @{prop "n = n'"}\<close>
lemma lift_Suc_mono_less:
assumes mono: "\<And>n. f n < f (Suc n)" and "n < n'"
shows "f n < f n'"
-using `n < n'`
+using \<open>n < n'\<close>
by (induct n n' rule: less_Suc_induct [consumes 1]) (auto intro: mono)
lemma lift_Suc_mono_less_iff:
@@ -1735,7 +1735,7 @@
qed
-text{*Subtraction laws, mostly by Clemens Ballarin*}
+text\<open>Subtraction laws, mostly by Clemens Ballarin\<close>
lemma diff_less_mono: "[| a < (b::nat); c \<le> a |] ==> a-c < b-c"
by arith
@@ -1753,20 +1753,20 @@
by arith
lemma le_diff_conv2: "k \<le> j ==> (i \<le> j-k) = (i+k \<le> (j::nat))"
- by (fact le_diff_conv2) -- {* FIXME delete *}
+ by (fact le_diff_conv2) -- \<open>FIXME delete\<close>
lemma diff_diff_cancel [simp]: "i \<le> (n::nat) ==> n - (n - i) = i"
by arith
lemma le_add_diff: "k \<le> (n::nat) ==> m \<le> n + m - k"
- by (fact le_add_diff) -- {* FIXME delete *}
+ by (fact le_add_diff) -- \<open>FIXME delete\<close>
(*Replaces the previous diff_less and le_diff_less, which had the stronger
second premise n\<le>m*)
lemma diff_less[simp]: "!!m::nat. [| 0<n; 0<m |] ==> m - n < m"
by arith
-text {* Simplification of relational expressions involving subtraction *}
+text \<open>Simplification of relational expressions involving subtraction\<close>
lemma diff_diff_eq: "[| k \<le> m; k \<le> (n::nat) |] ==> ((m-k) - (n-k)) = (m-n)"
by (simp split add: nat_diff_split)
@@ -1782,7 +1782,7 @@
lemma le_diff_iff: "[| k \<le> m; k \<le> (n::nat) |] ==> (m-k \<le> n-k) = (m\<le>n)"
by (auto split add: nat_diff_split)
-text{*(Anti)Monotonicity of subtraction -- by Stephan Merz*}
+text\<open>(Anti)Monotonicity of subtraction -- by Stephan Merz\<close>
(* Monotonicity of subtraction in first argument *)
lemma diff_le_mono: "m \<le> (n::nat) ==> (m-l) \<le> (n-l)"
@@ -1810,7 +1810,7 @@
with a k_le_n show "x = y" by auto
qed
-text{*Rewriting to pull differences out*}
+text\<open>Rewriting to pull differences out\<close>
lemma diff_diff_right [simp]: "k\<le>j --> i - (j - k) = i + (k::nat) - j"
by arith
@@ -1832,10 +1832,10 @@
lemmas add_diff_assoc2 = diff_add_assoc2[symmetric]
declare diff_diff_left [simp] add_diff_assoc [simp] add_diff_assoc2[simp]
-text{*At present we prove no analogue of @{text not_less_Least} or @{text
-Least_Suc}, since there appears to be no need.*}
-
-text{*Lemmas for ex/Factorization*}
+text\<open>At present we prove no analogue of @{text not_less_Least} or @{text
+Least_Suc}, since there appears to be no need.\<close>
+
+text\<open>Lemmas for ex/Factorization\<close>
lemma one_less_mult: "[| Suc 0 < n; Suc 0 < m |] ==> Suc 0 < m*n"
by (cases m) auto
@@ -1846,7 +1846,7 @@
lemma n_less_n_mult_m: "[| Suc 0 < n; Suc 0 < m |] ==> n<n*m"
by (cases m) auto
-text {* Specialized induction principles that work "backwards": *}
+text \<open>Specialized induction principles that work "backwards":\<close>
lemma inc_induct[consumes 1, case_names base step]:
assumes less: "i \<le> j"
@@ -1873,7 +1873,7 @@
using less
proof (induct d=="j - i - 1" arbitrary: i)
case (0 i)
- with `i < j` have "j = Suc i" by simp
+ with \<open>i < j\<close> have "j = Suc i" by simp
with base show ?case by simp
next
case (Suc d i)
@@ -1888,7 +1888,7 @@
lemma zero_induct: "P k ==> (!!n. P (Suc n) ==> P n) ==> P 0"
using inc_induct[of 0 k P] by blast
-text {* Further induction rule similar to @{thm inc_induct} *}
+text \<open>Further induction rule similar to @{thm inc_induct}\<close>
lemma dec_induct[consumes 1, case_names base step]:
"i \<le> j \<Longrightarrow> P i \<Longrightarrow> (\<And>n. i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n \<Longrightarrow> P (Suc n)) \<Longrightarrow> P j"
@@ -1920,7 +1920,7 @@
shows "mono Q \<Longrightarrow> antimono (\<lambda>i. (Q ^^ i) \<top>)"
by (auto intro!: funpow_increasing simp: antimono_def)
-subsection {* The divides relation on @{typ nat} *}
+subsection \<open>The divides relation on @{typ nat}\<close>
lemma dvd_1_left [iff]: "Suc 0 dvd k"
unfolding dvd_def by simp
@@ -1935,7 +1935,7 @@
unfolding dvd_def
by (force dest: mult_eq_self_implies_10 simp add: mult.assoc)
-text {* @{term "op dvd"} is a partial order *}
+text \<open>@{term "op dvd"} is a partial order\<close>
interpretation dvd: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
proof qed (auto intro: dvd_refl dvd_trans dvd_antisym)
@@ -2006,7 +2006,7 @@
qed
-subsection {* Aliases *}
+subsection \<open>Aliases\<close>
lemma nat_mult_1: "(1::nat) * n = n"
by (fact mult_1_left)
@@ -2015,10 +2015,10 @@
by (fact mult_1_right)
-subsection {* Size of a datatype value *}
+subsection \<open>Size of a datatype value\<close>
class size =
- fixes size :: "'a \<Rightarrow> nat" -- {* see further theory @{text Wellfounded} *}
+ fixes size :: "'a \<Rightarrow> nat" -- \<open>see further theory @{text Wellfounded}\<close>
instantiation nat :: size
begin
@@ -2031,7 +2031,7 @@
end
-subsection {* Code module namespace *}
+subsection \<open>Code module namespace\<close>
code_identifier
code_module Nat \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
--- a/src/HOL/Nat_Transfer.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Nat_Transfer.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,13 +1,13 @@
(* Authors: Jeremy Avigad and Amine Chaieb *)
-section {* Generic transfer machinery; specific transfer from nats to ints and back. *}
+section \<open>Generic transfer machinery; specific transfer from nats to ints and back.\<close>
theory Nat_Transfer
imports Int
begin
-subsection {* Generic transfer machinery *}
+subsection \<open>Generic transfer machinery\<close>
definition transfer_morphism:: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
where "transfer_morphism f A \<longleftrightarrow> True"
@@ -18,9 +18,9 @@
ML_file "Tools/legacy_transfer.ML"
-subsection {* Set up transfer from nat to int *}
+subsection \<open>Set up transfer from nat to int\<close>
-text {* set up transfer direction *}
+text \<open>set up transfer direction\<close>
lemma transfer_morphism_nat_int: "transfer_morphism nat (op <= (0::int))" ..
@@ -30,7 +30,7 @@
labels: nat_int
]
-text {* basic functions and relations *}
+text \<open>basic functions and relations\<close>
lemma transfer_nat_int_numerals [transfer key: transfer_morphism_nat_int]:
"(0::nat) = nat 0"
@@ -79,7 +79,7 @@
by (auto simp add: zdvd_int)
-text {* first-order quantifiers *}
+text \<open>first-order quantifiers\<close>
lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"
by (simp split add: split_nat)
@@ -113,14 +113,14 @@
cong: all_cong ex_cong]
-text {* if *}
+text \<open>if\<close>
lemma nat_if_cong [transfer key: transfer_morphism_nat_int]:
"(if P then (nat x) else (nat y)) = nat (if P then x else y)"
by auto
-text {* operations with sets *}
+text \<open>operations with sets\<close>
definition
nat_set :: "int set \<Rightarrow> bool"
@@ -187,7 +187,7 @@
]
-text {* setsum and setprod *}
+text \<open>setsum and setprod\<close>
(* this handles the case where the *domain* of f is nat *)
lemma transfer_nat_int_sum_prod:
@@ -256,9 +256,9 @@
cong: transfer_nat_int_sum_prod_cong]
-subsection {* Set up transfer from int to nat *}
+subsection \<open>Set up transfer from int to nat\<close>
-text {* set up transfer direction *}
+text \<open>set up transfer direction\<close>
lemma transfer_morphism_int_nat: "transfer_morphism int (\<lambda>n. True)" ..
@@ -269,7 +269,7 @@
]
-text {* basic functions and relations *}
+text \<open>basic functions and relations\<close>
definition
is_nat :: "int \<Rightarrow> bool"
@@ -317,7 +317,7 @@
]
-text {* first-order quantifiers *}
+text \<open>first-order quantifiers\<close>
lemma transfer_int_nat_quantifiers:
"(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))"
@@ -332,7 +332,7 @@
return: transfer_int_nat_quantifiers]
-text {* if *}
+text \<open>if\<close>
lemma int_if_cong: "(if P then (int x) else (int y)) =
int (if P then x else y)"
@@ -342,7 +342,7 @@
-text {* operations with sets *}
+text \<open>operations with sets\<close>
lemma transfer_int_nat_set_functions:
"nat_set A \<Longrightarrow> card A = card (nat ` A)"
@@ -392,7 +392,7 @@
]
-text {* setsum and setprod *}
+text \<open>setsum and setprod\<close>
(* this handles the case where the *domain* of f is int *)
lemma transfer_int_nat_sum_prod:
--- a/src/HOL/Nitpick.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Nitpick.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,7 +5,7 @@
Nitpick: Yet another counterexample generator for Isabelle/HOL.
*)
-section {* Nitpick: Yet Another Counterexample Generator for Isabelle/HOL *}
+section \<open>Nitpick: Yet Another Counterexample Generator for Isabelle/HOL\<close>
theory Nitpick
imports Record
@@ -30,9 +30,9 @@
Quot :: "'a \<Rightarrow> 'b"
safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
-text {*
+text \<open>
Alternative definitions.
-*}
+\<close>
lemma Ex1_unfold[nitpick_unfold]: "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
apply (rule eq_reflection)
@@ -79,10 +79,10 @@
"fold_graph' f z {} z" |
"\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
-text {*
+text \<open>
The following lemmas are not strictly necessary but they help the
\textit{specialize} optimization.
-*}
+\<close>
lemma The_psimp[nitpick_psimp]: "P = (op =) x \<Longrightarrow> The P = x"
by auto
@@ -114,10 +114,10 @@
"size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
by (cases xs) auto
-text {*
+text \<open>
Auxiliary definitions used to provide an alternative representation for
@{text rat} and @{text real}.
-*}
+\<close>
function nat_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))"
@@ -220,7 +220,7 @@
ML_file "Tools/Nitpick/nitpick_commands.ML"
ML_file "Tools/Nitpick/nitpick_tests.ML"
-setup {*
+setup \<open>
Nitpick_HOL.register_ersatz_global
[(@{const_name card}, @{const_name card'}),
(@{const_name setsum}, @{const_name setsum'}),
@@ -228,7 +228,7 @@
(@{const_name wf}, @{const_name wf'}),
(@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
(@{const_name wfrec}, @{const_name wfrec'})]
-*}
+\<close>
hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The FunBox PairBox Word prod
refl' wf' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac
--- a/src/HOL/NthRoot.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/NthRoot.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Conversion to Isar and new proofs by Lawrence C Paulson, 2004
*)
-section {* Nth Roots of Real Numbers *}
+section \<open>Nth Roots of Real Numbers\<close>
theory NthRoot
imports Deriv Binomial
@@ -21,9 +21,9 @@
shows "0 < n \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a ^ n = b ^ n \<longleftrightarrow> a = b"
using power_eq_imp_eq_base[of a n b] by auto
-subsection {* Existence of Nth Root *}
+subsection \<open>Existence of Nth Root\<close>
-text {* Existence follows from the Intermediate Value Theorem *}
+text \<open>Existence follows from the Intermediate Value Theorem\<close>
lemma realpow_pos_nth:
assumes n: "0 < n"
@@ -52,21 +52,21 @@
lemma realpow_pos_nth2: "(0::real) < a \<Longrightarrow> \<exists>r>0. r ^ Suc n = a"
by (blast intro: realpow_pos_nth)
-text {* Uniqueness of nth positive root *}
+text \<open>Uniqueness of nth positive root\<close>
lemma realpow_pos_nth_unique: "\<lbrakk>0 < n; 0 < a\<rbrakk> \<Longrightarrow> \<exists>!r. 0 < r \<and> r ^ n = (a::real)"
by (auto intro!: realpow_pos_nth simp: power_eq_iff_eq_base)
-subsection {* Nth Root *}
+subsection \<open>Nth Root\<close>
-text {* We define roots of negative reals such that
+text \<open>We define roots of negative reals such that
@{term "root n (- x) = - root n x"}. This allows
- us to omit side conditions from many theorems. *}
+ us to omit side conditions from many theorems.\<close>
lemma inj_sgn_power: assumes "0 < n" shows "inj (\<lambda>y. sgn y * \<bar>y\<bar>^n :: real)" (is "inj ?f")
proof (rule injI)
have x: "\<And>a b :: real. (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b) \<Longrightarrow> a \<noteq> b" by auto
- fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\<bar>x\<bar>" "\<bar>y\<bar>"] `0<n` show "x = y"
+ fix x y assume "?f x = ?f y" with power_eq_iff_eq_base[of n "\<bar>x\<bar>" "\<bar>y\<bar>"] \<open>0<n\<close> show "x = y"
by (cases rule: linorder_cases[of 0 x, case_product linorder_cases[of 0 y]])
(simp_all add: x)
qed
@@ -87,13 +87,13 @@
assumes "0 < n" shows "sgn (root n x) * \<bar>(root n x)\<bar>^n = x" (is "?f (root n x) = x")
proof cases
assume "x \<noteq> 0"
- with realpow_pos_nth[OF `0 < n`, of "\<bar>x\<bar>"] obtain r where "0 < r" "r ^ n = \<bar>x\<bar>" by auto
- with `x \<noteq> 0` have S: "x \<in> range ?f"
+ with realpow_pos_nth[OF \<open>0 < n\<close>, of "\<bar>x\<bar>"] obtain r where "0 < r" "r ^ n = \<bar>x\<bar>" by auto
+ with \<open>x \<noteq> 0\<close> have S: "x \<in> range ?f"
by (intro image_eqI[of _ _ "sgn x * r"])
(auto simp: abs_mult sgn_mult power_mult_distrib abs_sgn_eq mult_sgn_abs)
- from `0 < n` f_the_inv_into_f[OF inj_sgn_power[OF `0 < n`] this] show ?thesis
+ from \<open>0 < n\<close> f_the_inv_into_f[OF inj_sgn_power[OF \<open>0 < n\<close>] this] show ?thesis
by (simp add: root_def)
-qed (insert `0 < n` root_sgn_power[of n 0], simp)
+qed (insert \<open>0 < n\<close> root_sgn_power[of n 0], simp)
lemma split_root: "P (root n x) \<longleftrightarrow> (n = 0 \<longrightarrow> P 0) \<and> (0 < n \<longrightarrow> (\<forall>y. sgn y * \<bar>y\<bar>^n = x \<longrightarrow> P y))"
apply (cases "n = 0")
@@ -151,7 +151,7 @@
lemma real_root_one [simp]: "0 < n \<Longrightarrow> root n 1 = 1"
by (simp add: real_root_pos_unique)
-text {* Root function is strictly monotonic, hence injective *}
+text \<open>Root function is strictly monotonic, hence injective\<close>
lemma real_root_le_mono: "\<lbrakk>0 < n; x \<le> y\<rbrakk> \<Longrightarrow> root n x \<le> root n y"
by (auto simp add: order_le_less real_root_less_mono)
@@ -195,7 +195,7 @@
lemma real_root_eq_1_iff [simp]: "0 < n \<Longrightarrow> (root n x = 1) = (x = 1)"
by (insert real_root_eq_iff [where y=1], simp)
-text {* Roots of multiplication and division *}
+text \<open>Roots of multiplication and division\<close>
lemma real_root_mult: "root n (x * y) = root n x * root n y"
by (auto split: split_root elim!: sgn_power_injE simp: sgn_mult abs_mult power_mult_distrib)
@@ -212,7 +212,7 @@
lemma real_root_power: "0 < n \<Longrightarrow> root n (x ^ k) = root n x ^ k"
by (induct k) (simp_all add: real_root_mult)
-text {* Roots of roots *}
+text \<open>Roots of roots\<close>
lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x"
by (simp add: odd_real_root_unique)
@@ -224,7 +224,7 @@
lemma real_root_commute: "root m (root n x) = root n (root m x)"
by (simp add: real_root_mult_exp [symmetric] mult.commute)
-text {* Monotonicity in first argument *}
+text \<open>Monotonicity in first argument\<close>
lemma real_root_strict_decreasing:
"\<lbrakk>0 < n; n < N; 1 < x\<rbrakk> \<Longrightarrow> root N x < root n x"
@@ -248,7 +248,7 @@
"\<lbrakk>0 < n; n < N; 0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> root n x \<le> root N x"
by (auto simp add: order_le_less real_root_strict_increasing)
-text {* Continuity and derivatives *}
+text \<open>Continuity and derivatives\<close>
lemma isCont_real_root: "isCont (root n) x"
proof cases
@@ -320,8 +320,8 @@
show "\<forall>y. x - 1 < y \<and> y < 0 \<longrightarrow> - (root n y ^ n) = y"
proof (rule allI, rule impI, erule conjE)
fix y assume "x - 1 < y" and "y < 0"
- hence "root n (-y) ^ n = -y" using `0 < n` by simp
- with real_root_minus and `even n`
+ hence "root n (-y) ^ n = -y" using \<open>0 < n\<close> by simp
+ with real_root_minus and \<open>even n\<close>
show "- (root n y ^ n) = y" by simp
qed
next
@@ -342,7 +342,7 @@
DERIV_odd_real_root[THEN DERIV_cong]
DERIV_even_real_root[THEN DERIV_cong])
-subsection {* Square Root *}
+subsection \<open>Square Root\<close>
definition sqrt :: "real \<Rightarrow> real" where
"sqrt = root 2"
@@ -546,7 +546,7 @@
by (rule filterlim_at_top_at_top[where Q="\<lambda>x. True" and P="\<lambda>x. 0 < x" and g="power2"])
(auto intro: eventually_gt_at_top)
-subsection {* Square Root of Sum of Squares *}
+subsection \<open>Square Root of Sum of Squares\<close>
lemma sum_squares_bound:
fixes x:: "'a::linordered_field"
@@ -635,8 +635,8 @@
by (metis not_less not_less_iff_gr_or_eq numeral_less_iff real_sqrt_four real_sqrt_le_iff semiring_norm(75) semiring_norm(78) semiring_norm(85))
-text{*Needed for the infinitely close relation over the nonstandard
- complex numbers*}
+text\<open>Needed for the infinitely close relation over the nonstandard
+ complex numbers\<close>
lemma lemma_sqrt_hcomplex_capprox:
"[| 0 < u; x < u/2; y < u/2; 0 \<le> x; 0 \<le> y |] ==> sqrt (x\<^sup>2 + y\<^sup>2) < u"
apply (rule real_sqrt_sum_squares_less)
@@ -659,20 +659,20 @@
(simp_all add: at_infinity_eq_at_top_bot)
{ fix n :: nat assume "2 < n"
have "1 + (real (n - 1) * n) / 2 * x n^2 = 1 + of_nat (n choose 2) * x n^2"
- using `2 < n` unfolding gbinomial_def binomial_gbinomial
+ using \<open>2 < n\<close> unfolding gbinomial_def binomial_gbinomial
by (simp add: atLeast0AtMost atMost_Suc field_simps real_of_nat_diff numeral_2_eq_2 real_eq_of_nat[symmetric])
also have "\<dots> \<le> (\<Sum>k\<in>{0, 2}. of_nat (n choose k) * x n^k)"
by (simp add: x_def)
also have "\<dots> \<le> (\<Sum>k=0..n. of_nat (n choose k) * x n^k)"
- using `2 < n` by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
+ using \<open>2 < n\<close> by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
also have "\<dots> = (x n + 1) ^ n"
by (simp add: binomial_ring)
also have "\<dots> = n"
- using `2 < n` by (simp add: x_def)
+ using \<open>2 < n\<close> by (simp add: x_def)
finally have "real (n - 1) * (real n / 2 * (x n)\<^sup>2) \<le> real (n - 1) * 1"
by simp
then have "(x n)\<^sup>2 \<le> 2 / real n"
- using `2 < n` unfolding mult_le_cancel_left by (simp add: field_simps)
+ using \<open>2 < n\<close> unfolding mult_le_cancel_left by (simp add: field_simps)
from real_sqrt_le_mono[OF this] have "x n \<le> sqrt (2 / real n)"
by simp }
then show "eventually (\<lambda>n. x n \<le> sqrt (2 / real n)) sequentially"
@@ -697,21 +697,21 @@
(simp_all add: at_infinity_eq_at_top_bot)
{ fix n :: nat assume "1 < n"
have "1 + x n * n = 1 + of_nat (n choose 1) * x n^1"
- using `1 < n` unfolding gbinomial_def binomial_gbinomial by (simp add: real_eq_of_nat[symmetric])
+ using \<open>1 < n\<close> unfolding gbinomial_def binomial_gbinomial by (simp add: real_eq_of_nat[symmetric])
also have "\<dots> \<le> (\<Sum>k\<in>{0, 1}. of_nat (n choose k) * x n^k)"
by (simp add: x_def)
also have "\<dots> \<le> (\<Sum>k=0..n. of_nat (n choose k) * x n^k)"
- using `1 < n` `1 \<le> c` by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
+ using \<open>1 < n\<close> \<open>1 \<le> c\<close> by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
also have "\<dots> = (x n + 1) ^ n"
by (simp add: binomial_ring)
also have "\<dots> = c"
- using `1 < n` `1 \<le> c` by (simp add: x_def)
+ using \<open>1 < n\<close> \<open>1 \<le> c\<close> by (simp add: x_def)
finally have "x n \<le> c / n"
- using `1 \<le> c` `1 < n` by (simp add: field_simps) }
+ using \<open>1 \<le> c\<close> \<open>1 < n\<close> by (simp add: field_simps) }
then show "eventually (\<lambda>n. x n \<le> c / n) sequentially"
by (auto intro!: exI[of _ 3] simp: eventually_sequentially)
show "eventually (\<lambda>n. 0 \<le> x n) sequentially"
- using `1 \<le> c` by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def)
+ using \<open>1 \<le> c\<close> by (auto intro!: exI[of _ 1] simp: eventually_sequentially le_diff_eq x_def)
qed
from tendsto_add[OF this tendsto_const[of 1]] have "(\<lambda>n. root n c) ----> 1"
by (simp add: x_def) }
@@ -722,10 +722,10 @@
assume "1 \<le> c" with ge_1 show ?thesis by blast
next
assume "\<not> 1 \<le> c"
- with `0 < c` have "1 \<le> 1 / c"
+ with \<open>0 < c\<close> have "1 \<le> 1 / c"
by simp
then have "(\<lambda>n. 1 / root n (1 / c)) ----> 1 / 1"
- by (intro tendsto_divide tendsto_const ge_1 `1 \<le> 1 / c` one_neq_zero)
+ by (intro tendsto_divide tendsto_const ge_1 \<open>1 \<le> 1 / c\<close> one_neq_zero)
then show ?thesis
by (rule filterlim_cong[THEN iffD1, rotated 3])
(auto intro!: exI[of _ 1] simp: eventually_sequentially real_root_divide)
--- a/src/HOL/Num.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Num.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,24 +3,24 @@
Author: Brian Huffman
*)
-section {* Binary Numerals *}
+section \<open>Binary Numerals\<close>
theory Num
imports BNF_Least_Fixpoint
begin
-subsection {* The @{text num} type *}
+subsection \<open>The @{text num} type\<close>
datatype num = One | Bit0 num | Bit1 num
-text {* Increment function for type @{typ num} *}
+text \<open>Increment function for type @{typ num}\<close>
primrec inc :: "num \<Rightarrow> num" where
"inc One = Bit0 One" |
"inc (Bit0 x) = Bit1 x" |
"inc (Bit1 x) = Bit0 (inc x)"
-text {* Converting between type @{typ num} and type @{typ nat} *}
+text \<open>Converting between type @{typ num} and type @{typ nat}\<close>
primrec nat_of_num :: "num \<Rightarrow> nat" where
"nat_of_num One = Suc 0" |
@@ -44,10 +44,10 @@
"0 < n \<Longrightarrow> num_of_nat (n + n) = Bit0 (num_of_nat n)"
by (induct n) simp_all
-text {*
+text \<open>
Type @{typ num} is isomorphic to the strictly positive
natural numbers.
-*}
+\<close>
lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x"
by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos)
@@ -81,14 +81,14 @@
by (simp add: nat_of_num_inverse)
qed
-text {*
+text \<open>
From now on, there are two possible models for @{typ num}:
as positive naturals (rule @{text "num_induct"})
and as digit representation (rules @{text "num.induct"}, @{text "num.cases"}).
-*}
+\<close>
-subsection {* Numeral operations *}
+subsection \<open>Numeral operations\<close>
instantiation num :: "{plus,times,linorder}"
begin
@@ -174,7 +174,7 @@
using nat_of_num_pos [of n] nat_of_num_pos [of m]
by (auto simp add: less_eq_num_def less_num_def)
-text {* Rules using @{text One} and @{text inc} as constructors *}
+text \<open>Rules using @{text One} and @{text inc} as constructors\<close>
lemma add_One: "x + One = inc x"
by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc)
@@ -188,7 +188,7 @@
lemma mult_inc: "x * inc y = x * y + x"
by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc)
-text {* The @{const num_of_nat} conversion *}
+text \<open>The @{const num_of_nat} conversion\<close>
lemma num_of_nat_One:
"n \<le> 1 \<Longrightarrow> num_of_nat n = One"
@@ -198,7 +198,7 @@
"0 < m \<Longrightarrow> 0 < n \<Longrightarrow> num_of_nat (m + n) = num_of_nat m + num_of_nat n"
by (induct n) (auto simp add: add_One add_One_commute add_inc)
-text {* A double-and-decrement function *}
+text \<open>A double-and-decrement function\<close>
primrec BitM :: "num \<Rightarrow> num" where
"BitM One = One" |
@@ -211,7 +211,7 @@
lemma one_plus_BitM: "One + BitM n = Bit0 n"
unfolding add_One_commute BitM_plus_one ..
-text {* Squaring and exponentiation *}
+text \<open>Squaring and exponentiation\<close>
primrec sqr :: "num \<Rightarrow> num" where
"sqr One = One" |
@@ -230,12 +230,12 @@
by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult)
-subsection {* Binary numerals *}
+subsection \<open>Binary numerals\<close>
-text {*
+text \<open>
We embed binary representations into a generic algebraic
structure using @{text numeral}.
-*}
+\<close>
class numeral = one + semigroup_add
begin
@@ -275,14 +275,14 @@
end
-text {* Numeral syntax. *}
+text \<open>Numeral syntax.\<close>
syntax
"_Numeral" :: "num_const \<Rightarrow> 'a" ("_")
ML_file "Tools/numeral.ML"
-parse_translation {*
+parse_translation \<open>
let
fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
c $ numeral_tr [t] $ u
@@ -290,9 +290,9 @@
(Numeral.mk_number_syntax o #value o Lexicon.read_num) num
| numeral_tr ts = raise TERM ("numeral_tr", ts);
in [(@{syntax_const "_Numeral"}, K numeral_tr)] end
-*}
+\<close>
-typed_print_translation {*
+typed_print_translation \<open>
let
fun dest_num (Const (@{const_syntax Bit0}, _) $ n) = 2 * dest_num n
| dest_num (Const (@{const_syntax Bit1}, _) $ n) = 2 * dest_num n + 1
@@ -315,16 +315,16 @@
in
[(@{const_syntax numeral}, num_tr')]
end
-*}
+\<close>
-subsection {* Class-specific numeral rules *}
+subsection \<open>Class-specific numeral rules\<close>
-text {*
+text \<open>
@{const numeral} is a morphism.
-*}
+\<close>
-subsubsection {* Structures with addition: class @{text numeral} *}
+subsubsection \<open>Structures with addition: class @{text numeral}\<close>
context numeral
begin
@@ -350,9 +350,9 @@
end
-subsubsection {*
+subsubsection \<open>
Structures with negation: class @{text neg_numeral}
-*}
+\<close>
class neg_numeral = numeral + group_add
begin
@@ -361,7 +361,7 @@
"- Numeral1 = - 1"
by (simp add: numeral_One)
-text {* Numerals form an abelian subgroup. *}
+text \<open>Numerals form an abelian subgroup.\<close>
inductive is_num :: "'a \<Rightarrow> bool" where
"is_num 1" |
@@ -488,9 +488,9 @@
end
-subsubsection {*
+subsubsection \<open>
Structures with multiplication: class @{text semiring_numeral}
-*}
+\<close>
class semiring_numeral = semiring + monoid_mult
begin
@@ -514,9 +514,9 @@
end
-subsubsection {*
+subsubsection \<open>
Structures with a zero: class @{text semiring_1}
-*}
+\<close>
context semiring_1
begin
@@ -544,9 +544,9 @@
"nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))"
by (simp_all add: Let_def)
-subsubsection {*
+subsubsection \<open>
Equality: class @{text semiring_char_0}
-*}
+\<close>
context semiring_char_0
begin
@@ -577,11 +577,11 @@
end
-subsubsection {*
+subsubsection \<open>
Comparisons: class @{text linordered_semidom}
-*}
+\<close>
-text {* Could be perhaps more general than here. *}
+text \<open>Could be perhaps more general than here.\<close>
context linordered_semidom
begin
@@ -648,9 +648,9 @@
end
-subsubsection {*
+subsubsection \<open>
Multiplication and negation: class @{text ring_1}
-*}
+\<close>
context ring_1
begin
@@ -672,9 +672,9 @@
end
-subsubsection {*
+subsubsection \<open>
Equality using @{text iszero} for rings with non-zero characteristic
-*}
+\<close>
context ring_1
begin
@@ -705,7 +705,7 @@
lemma eq_iff_iszero_diff: "x = y \<longleftrightarrow> iszero (x - y)"
unfolding iszero_def by (rule eq_iff_diff_eq_0)
-text {* The @{text "eq_numeral_iff_iszero"} lemmas are not declared
+text \<open>The @{text "eq_numeral_iff_iszero"} lemmas are not declared
@{text "[simp]"} by default, because for rings of characteristic zero,
better simp rules are possible. For a type like integers mod @{text
"n"}, type-instantiated versions of these rules should be added to the
@@ -715,7 +715,7 @@
bh: Maybe it would not be so bad to just declare these as simp
rules anyway? I should test whether these rules take precedence over
the @{text "ring_char_0"} rules in the simplifier.
-*}
+\<close>
lemma eq_numeral_iff_iszero:
"numeral x = numeral y \<longleftrightarrow> iszero (sub x y)"
@@ -735,9 +735,9 @@
end
-subsubsection {*
+subsubsection \<open>
Equality and negation: class @{text ring_char_0}
-*}
+\<close>
class ring_char_0 = ring_1 + semiring_char_0
begin
@@ -811,9 +811,9 @@
end
-subsubsection {*
+subsubsection \<open>
Structures with negation and order: class @{text linordered_idom}
-*}
+\<close>
context linordered_idom
begin
@@ -944,9 +944,9 @@
end
-subsubsection {*
+subsubsection \<open>
Natural numbers
-*}
+\<close>
lemma Suc_1 [simp]: "Suc 1 = 2"
unfolding Suc_eq_plus1 by (rule one_add_one)
@@ -989,7 +989,7 @@
(*Maps #n to n for n = 1, 2*)
lemmas numerals = numeral_One [where 'a=nat] numeral_2_eq_2
-text {* Comparisons involving @{term Suc}. *}
+text \<open>Comparisons involving @{term Suc}.\<close>
lemma eq_numeral_Suc [simp]: "numeral k = Suc n \<longleftrightarrow> pred_numeral k = n"
by (simp add: numeral_eq_Suc)
@@ -1031,7 +1031,7 @@
"min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)"
by (simp add: numeral_eq_Suc)
-text {* For @{term case_nat} and @{term rec_nat}. *}
+text \<open>For @{term case_nat} and @{term rec_nat}.\<close>
lemma case_nat_numeral [simp]:
"case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)"
@@ -1051,13 +1051,13 @@
(let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))"
by (simp add: numeral_eq_Suc Let_def)
-text {* Case analysis on @{term "n < 2"} *}
+text \<open>Case analysis on @{term "n < 2"}\<close>
lemma less_2_cases: "n < 2 \<Longrightarrow> n = 0 \<or> n = Suc 0"
by (auto simp add: numeral_2_eq_2)
-text {* Removal of Small Numerals: 0, 1 and (in additive positions) 2 *}
-text {* bh: Are these rules really a good idea? *}
+text \<open>Removal of Small Numerals: 0, 1 and (in additive positions) 2\<close>
+text \<open>bh: Are these rules really a good idea?\<close>
lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
by simp
@@ -1065,7 +1065,7 @@
lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
by simp
-text {* Can be used to eliminate long strings of Sucs, but not by default. *}
+text \<open>Can be used to eliminate long strings of Sucs, but not by default.\<close>
lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
by simp
@@ -1073,7 +1073,7 @@
lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *)
-subsection {* Particular lemmas concerning @{term 2} *}
+subsection \<open>Particular lemmas concerning @{term 2}\<close>
context linordered_field
begin
@@ -1089,7 +1089,7 @@
end
-subsection {* Numeral equations as default simplification rules *}
+subsection \<open>Numeral equations as default simplification rules\<close>
declare (in numeral) numeral_One [simp]
declare (in numeral) numeral_plus_numeral [simp]
@@ -1101,7 +1101,7 @@
declare (in semiring_numeral) numeral_times_numeral [simp]
declare (in ring_1) mult_neg_numeral_simps [simp]
-subsection {* Setting up simprocs *}
+subsection \<open>Setting up simprocs\<close>
lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)"
by simp
@@ -1116,8 +1116,8 @@
"inverse Numeral1 = (Numeral1::'a::division_ring)"
by simp
-text{*Theorem lists for the cancellation simprocs. The use of a binary
-numeral for 1 reduces the number of special cases.*}
+text\<open>Theorem lists for the cancellation simprocs. The use of a binary
+numeral for 1 reduces the number of special cases.\<close>
lemma mult_1s:
fixes a :: "'a::semiring_numeral"
@@ -1128,18 +1128,18 @@
"b * - Numeral1 = - b"
by simp_all
-setup {*
+setup \<open>
Reorient_Proc.add
(fn Const (@{const_name numeral}, _) $ _ => true
| Const (@{const_name uminus}, _) $ (Const (@{const_name numeral}, _) $ _) => true
| _ => false)
-*}
+\<close>
simproc_setup reorient_numeral
("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc
-subsubsection {* Simplification of arithmetic operations on integer constants. *}
+subsubsection \<open>Simplification of arithmetic operations on integer constants.\<close>
lemmas arith_special = (* already declared simp above *)
add_numeral_special add_neg_numeral_special
@@ -1154,10 +1154,10 @@
mult_zero_left mult_zero_right
abs_numeral abs_neg_numeral
-text {*
+text \<open>
For making a minimal simpset, one must include these default simprules.
Also include @{text simp_thms}.
-*}
+\<close>
lemmas arith_simps =
add_num_simps mult_num_simps sub_num_simps
@@ -1174,7 +1174,7 @@
lemmas of_nat_simps =
of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult
-text {* Simplification of relational operations *}
+text \<open>Simplification of relational operations\<close>
lemmas eq_numeral_extra =
zero_neq_one one_neq_zero
@@ -1186,14 +1186,14 @@
eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra
lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
- -- {* Unfold all @{text let}s involving constants *}
+ -- \<open>Unfold all @{text let}s involving constants\<close>
unfolding Let_def ..
lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)"
- -- {* Unfold all @{text let}s involving constants *}
+ -- \<open>Unfold all @{text let}s involving constants\<close>
unfolding Let_def ..
-declaration {*
+declaration \<open>
let
fun number_of ctxt T n =
if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort numeral}))
@@ -1215,10 +1215,10 @@
@{thm of_nat_numeral}]
#> Lin_Arith.set_number_of number_of)
end
-*}
+\<close>
-subsubsection {* Simplification of arithmetic when nested to the right. *}
+subsubsection \<open>Simplification of arithmetic when nested to the right.\<close>
lemma add_numeral_left [simp]:
"numeral v + (numeral w + z) = (numeral(v + w) + z)"
@@ -1240,7 +1240,7 @@
hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec
-subsection {* code module namespace *}
+subsection \<open>code module namespace\<close>
code_identifier
code_module Num \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
--- a/src/HOL/Numeral_Simprocs.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Numeral_Simprocs.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,6 +1,6 @@
(* Author: Various *)
-section {* Combination and Cancellation Simprocs for Numeral Expressions *}
+section \<open>Combination and Cancellation Simprocs for Numeral Expressions\<close>
theory Numeral_Simprocs
imports Divides
@@ -23,12 +23,12 @@
declare split_div [of _ _ "numeral k", arith_split] for k
declare split_mod [of _ _ "numeral k", arith_split] for k
-text {* For @{text combine_numerals} *}
+text \<open>For @{text combine_numerals}\<close>
lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)"
by (simp add: add_mult_distrib)
-text {* For @{text cancel_numerals} *}
+text \<open>For @{text cancel_numerals}\<close>
lemma nat_diff_add_eq1:
"j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)"
@@ -62,7 +62,7 @@
"i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)"
by (auto split add: nat_diff_split simp add: add_mult_distrib)
-text {* For @{text cancel_numeral_factors} *}
+text \<open>For @{text cancel_numeral_factors}\<close>
lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)"
by auto
@@ -83,7 +83,7 @@
lemma nat_mult_dvd_cancel1: "0 < k \<Longrightarrow> (k*m) dvd (k*n::nat) = (m dvd n)"
by(auto)
-text {* For @{text cancel_factor} *}
+text \<open>For @{text cancel_factor}\<close>
lemmas nat_mult_le_cancel_disj = mult_le_cancel1
@@ -107,17 +107,17 @@
simproc_setup semiring_assoc_fold
("(a::'a::comm_semiring_1_cancel) * b") =
- {* fn phi => Numeral_Simprocs.assoc_fold *}
+ \<open>fn phi => Numeral_Simprocs.assoc_fold\<close>
(* TODO: see whether the type class can be generalized further *)
simproc_setup int_combine_numerals
("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") =
- {* fn phi => Numeral_Simprocs.combine_numerals *}
+ \<open>fn phi => Numeral_Simprocs.combine_numerals\<close>
simproc_setup field_combine_numerals
("(i::'a::{field,ring_char_0}) + j"
|"(i::'a::{field,ring_char_0}) - j") =
- {* fn phi => Numeral_Simprocs.field_combine_numerals *}
+ \<open>fn phi => Numeral_Simprocs.field_combine_numerals\<close>
simproc_setup inteq_cancel_numerals
("(l::'a::comm_ring_1) + m = n"
@@ -128,7 +128,7 @@
|"(l::'a::comm_ring_1) = m * n"
|"- (l::'a::comm_ring_1) = m"
|"(l::'a::comm_ring_1) = - m") =
- {* fn phi => Numeral_Simprocs.eq_cancel_numerals *}
+ \<open>fn phi => Numeral_Simprocs.eq_cancel_numerals\<close>
simproc_setup intless_cancel_numerals
("(l::'a::linordered_idom) + m < n"
@@ -139,7 +139,7 @@
|"(l::'a::linordered_idom) < m * n"
|"- (l::'a::linordered_idom) < m"
|"(l::'a::linordered_idom) < - m") =
- {* fn phi => Numeral_Simprocs.less_cancel_numerals *}
+ \<open>fn phi => Numeral_Simprocs.less_cancel_numerals\<close>
simproc_setup intle_cancel_numerals
("(l::'a::linordered_idom) + m \<le> n"
@@ -150,140 +150,140 @@
|"(l::'a::linordered_idom) \<le> m * n"
|"- (l::'a::linordered_idom) \<le> m"
|"(l::'a::linordered_idom) \<le> - m") =
- {* fn phi => Numeral_Simprocs.le_cancel_numerals *}
+ \<open>fn phi => Numeral_Simprocs.le_cancel_numerals\<close>
simproc_setup ring_eq_cancel_numeral_factor
("(l::'a::{idom,ring_char_0}) * m = n"
|"(l::'a::{idom,ring_char_0}) = m * n") =
- {* fn phi => Numeral_Simprocs.eq_cancel_numeral_factor *}
+ \<open>fn phi => Numeral_Simprocs.eq_cancel_numeral_factor\<close>
simproc_setup ring_less_cancel_numeral_factor
("(l::'a::linordered_idom) * m < n"
|"(l::'a::linordered_idom) < m * n") =
- {* fn phi => Numeral_Simprocs.less_cancel_numeral_factor *}
+ \<open>fn phi => Numeral_Simprocs.less_cancel_numeral_factor\<close>
simproc_setup ring_le_cancel_numeral_factor
("(l::'a::linordered_idom) * m <= n"
|"(l::'a::linordered_idom) <= m * n") =
- {* fn phi => Numeral_Simprocs.le_cancel_numeral_factor *}
+ \<open>fn phi => Numeral_Simprocs.le_cancel_numeral_factor\<close>
(* TODO: remove comm_ring_1 constraint if possible *)
simproc_setup int_div_cancel_numeral_factors
("((l::'a::{semiring_div,comm_ring_1,ring_char_0}) * m) div n"
|"(l::'a::{semiring_div,comm_ring_1,ring_char_0}) div (m * n)") =
- {* fn phi => Numeral_Simprocs.div_cancel_numeral_factor *}
+ \<open>fn phi => Numeral_Simprocs.div_cancel_numeral_factor\<close>
simproc_setup divide_cancel_numeral_factor
("((l::'a::{field,ring_char_0}) * m) / n"
|"(l::'a::{field,ring_char_0}) / (m * n)"
|"((numeral v)::'a::{field,ring_char_0}) / (numeral w)") =
- {* fn phi => Numeral_Simprocs.divide_cancel_numeral_factor *}
+ \<open>fn phi => Numeral_Simprocs.divide_cancel_numeral_factor\<close>
simproc_setup ring_eq_cancel_factor
("(l::'a::idom) * m = n" | "(l::'a::idom) = m * n") =
- {* fn phi => Numeral_Simprocs.eq_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.eq_cancel_factor\<close>
simproc_setup linordered_ring_le_cancel_factor
("(l::'a::linordered_idom) * m <= n"
|"(l::'a::linordered_idom) <= m * n") =
- {* fn phi => Numeral_Simprocs.le_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.le_cancel_factor\<close>
simproc_setup linordered_ring_less_cancel_factor
("(l::'a::linordered_idom) * m < n"
|"(l::'a::linordered_idom) < m * n") =
- {* fn phi => Numeral_Simprocs.less_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.less_cancel_factor\<close>
simproc_setup int_div_cancel_factor
("((l::'a::semiring_div) * m) div n"
|"(l::'a::semiring_div) div (m * n)") =
- {* fn phi => Numeral_Simprocs.div_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.div_cancel_factor\<close>
simproc_setup int_mod_cancel_factor
("((l::'a::semiring_div) * m) mod n"
|"(l::'a::semiring_div) mod (m * n)") =
- {* fn phi => Numeral_Simprocs.mod_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.mod_cancel_factor\<close>
simproc_setup dvd_cancel_factor
("((l::'a::idom) * m) dvd n"
|"(l::'a::idom) dvd (m * n)") =
- {* fn phi => Numeral_Simprocs.dvd_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.dvd_cancel_factor\<close>
simproc_setup divide_cancel_factor
("((l::'a::field) * m) / n"
|"(l::'a::field) / (m * n)") =
- {* fn phi => Numeral_Simprocs.divide_cancel_factor *}
+ \<open>fn phi => Numeral_Simprocs.divide_cancel_factor\<close>
ML_file "Tools/nat_numeral_simprocs.ML"
simproc_setup nat_combine_numerals
("(i::nat) + j" | "Suc (i + j)") =
- {* fn phi => Nat_Numeral_Simprocs.combine_numerals *}
+ \<open>fn phi => Nat_Numeral_Simprocs.combine_numerals\<close>
simproc_setup nateq_cancel_numerals
("(l::nat) + m = n" | "(l::nat) = m + n" |
"(l::nat) * m = n" | "(l::nat) = m * n" |
"Suc m = n" | "m = Suc n") =
- {* fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals *}
+ \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals\<close>
simproc_setup natless_cancel_numerals
("(l::nat) + m < n" | "(l::nat) < m + n" |
"(l::nat) * m < n" | "(l::nat) < m * n" |
"Suc m < n" | "m < Suc n") =
- {* fn phi => Nat_Numeral_Simprocs.less_cancel_numerals *}
+ \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numerals\<close>
simproc_setup natle_cancel_numerals
("(l::nat) + m \<le> n" | "(l::nat) \<le> m + n" |
"(l::nat) * m \<le> n" | "(l::nat) \<le> m * n" |
"Suc m \<le> n" | "m \<le> Suc n") =
- {* fn phi => Nat_Numeral_Simprocs.le_cancel_numerals *}
+ \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numerals\<close>
simproc_setup natdiff_cancel_numerals
("((l::nat) + m) - n" | "(l::nat) - (m + n)" |
"(l::nat) * m - n" | "(l::nat) - m * n" |
"Suc m - n" | "m - Suc n") =
- {* fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals *}
+ \<open>fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals\<close>
simproc_setup nat_eq_cancel_numeral_factor
("(l::nat) * m = n" | "(l::nat) = m * n") =
- {* fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor\<close>
simproc_setup nat_less_cancel_numeral_factor
("(l::nat) * m < n" | "(l::nat) < m * n") =
- {* fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor\<close>
simproc_setup nat_le_cancel_numeral_factor
("(l::nat) * m <= n" | "(l::nat) <= m * n") =
- {* fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor\<close>
simproc_setup nat_div_cancel_numeral_factor
("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
- {* fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor\<close>
simproc_setup nat_dvd_cancel_numeral_factor
("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
- {* fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\<close>
simproc_setup nat_eq_cancel_factor
("(l::nat) * m = n" | "(l::nat) = m * n") =
- {* fn phi => Nat_Numeral_Simprocs.eq_cancel_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_factor\<close>
simproc_setup nat_less_cancel_factor
("(l::nat) * m < n" | "(l::nat) < m * n") =
- {* fn phi => Nat_Numeral_Simprocs.less_cancel_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_factor\<close>
simproc_setup nat_le_cancel_factor
("(l::nat) * m <= n" | "(l::nat) <= m * n") =
- {* fn phi => Nat_Numeral_Simprocs.le_cancel_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_factor\<close>
simproc_setup nat_div_cancel_factor
("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
- {* fn phi => Nat_Numeral_Simprocs.div_cancel_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_factor\<close>
simproc_setup nat_dvd_cancel_factor
("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
- {* fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor *}
+ \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor\<close>
-declaration {*
+declaration \<open>
K (Lin_Arith.add_simprocs
[@{simproc semiring_assoc_fold},
@{simproc int_combine_numerals},
@@ -297,6 +297,6 @@
@{simproc natless_cancel_numerals},
@{simproc natle_cancel_numerals},
@{simproc natdiff_cancel_numerals}])
-*}
+\<close>
end
--- a/src/HOL/Option.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Option.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Folklore
*)
-section {* Datatype option *}
+section \<open>Datatype option\<close>
theory Option
imports Lifting Finite_Set
@@ -15,23 +15,23 @@
datatype_compat option
lemma [case_names None Some, cases type: option]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
"(y = None \<Longrightarrow> P) \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> P) \<Longrightarrow> P"
by (rule option.exhaust)
lemma [case_names None Some, induct type: option]:
- -- {* for backward compatibility -- names of variables differ *}
+ -- \<open>for backward compatibility -- names of variables differ\<close>
"P None \<Longrightarrow> (\<And>option. P (Some option)) \<Longrightarrow> P option"
by (rule option.induct)
-text {* Compatibility: *}
+text \<open>Compatibility:\<close>
-setup {* Sign.mandatory_path "option" *}
+setup \<open>Sign.mandatory_path "option"\<close>
lemmas inducts = option.induct
lemmas cases = option.case
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)"
by (induct x) auto
@@ -39,9 +39,9 @@
lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)"
by (induct x) auto
-text{*Although it may appear that both of these equalities are helpful
+text\<open>Although it may appear that both of these equalities are helpful
only when applied to assumptions, in practice it seems better to give
-them the uniform iff attribute. *}
+them the uniform iff attribute.\<close>
lemma inj_Some [simp]: "inj_on Some A"
by (rule inj_onI) simp
@@ -78,12 +78,12 @@
by(cases y) auto
-subsubsection {* Operations *}
+subsubsection \<open>Operations\<close>
lemma ospec [dest]: "(ALL x:set_option A. P x) ==> A = Some x ==> P x"
by simp
-setup {* map_theory_claset (fn ctxt => ctxt addSD2 ("ospec", @{thm ospec})) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSD2 ("ospec", @{thm ospec}))\<close>
lemma elem_set [iff]: "(x : set_option xo) = (xo = Some x)"
by (cases xo) auto
@@ -251,7 +251,7 @@
hide_fact (open) bind_cong
-subsection {* Transfer rules for the Transfer package *}
+subsection \<open>Transfer rules for the Transfer package\<close>
context
begin
@@ -269,7 +269,7 @@
end
-subsubsection {* Interaction with finite sets *}
+subsubsection \<open>Interaction with finite sets\<close>
lemma finite_option_UNIV [simp]:
"finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"
@@ -279,7 +279,7 @@
by default (simp add: UNIV_option_conv)
-subsubsection {* Code generator setup *}
+subsubsection \<open>Code generator setup\<close>
lemma equal_None_code_unfold [code_unfold]:
"HOL.equal x None \<longleftrightarrow> is_none x"
--- a/src/HOL/Order_Relation.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Order_Relation.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Andrei Popescu, TU Muenchen
*)
-section {* Orders as Relations *}
+section \<open>Orders as Relations\<close>
theory Order_Relation
imports Wfrec
begin
-subsection{* Orders on a set *}
+subsection\<open>Orders on a set\<close>
definition "preorder_on A r \<equiv> refl_on A r \<and> trans r"
@@ -56,7 +56,7 @@
by(simp add: order_on_defs trans_diff_Id)
-subsection{* Orders on the field *}
+subsection\<open>Orders on the field\<close>
abbreviation "Refl r \<equiv> refl_on (Field r) r"
@@ -115,7 +115,7 @@
qed
-subsection{* Orders on a type *}
+subsection\<open>Orders on a type\<close>
abbreviation "strict_linear_order \<equiv> strict_linear_order_on UNIV"
@@ -124,14 +124,14 @@
abbreviation "well_order \<equiv> well_order_on UNIV"
-subsection {* Order-like relations *}
+subsection \<open>Order-like relations\<close>
-text{* In this subsection, we develop basic concepts and results pertaining
+text\<open>In this subsection, we develop basic concepts and results pertaining
to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or
-total relations. We also further define upper and lower bounds operators. *}
+total relations. We also further define upper and lower bounds operators.\<close>
-subsubsection {* Auxiliaries *}
+subsubsection \<open>Auxiliaries\<close>
lemma refl_on_domain:
"\<lbrakk>refl_on A r; (a,b) : r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
@@ -177,14 +177,14 @@
using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force
-subsubsection {* The upper and lower bounds operators *}
+subsubsection \<open>The upper and lower bounds operators\<close>
-text{* Here we define upper (``above") and lower (``below") bounds operators.
+text\<open>Here we define upper (``above") and lower (``below") bounds operators.
We think of @{text "r"} as a {\em non-strict} relation. The suffix ``S"
at the names of some operators indicates that the bounds are strict -- e.g.,
@{text "underS a"} is the set of all strict lower bounds of @{text "a"} (w.r.t. @{text "r"}).
Capitalization of the first letter in the name reminds that the operator acts on sets, rather
-than on individual elements. *}
+than on individual elements.\<close>
definition under::"'a rel \<Rightarrow> 'a \<Rightarrow> 'a set"
where "under r a \<equiv> {b. (b,a) \<in> r}"
@@ -213,9 +213,9 @@
definition ofilter :: "'a rel \<Rightarrow> 'a set \<Rightarrow> bool"
where "ofilter r A \<equiv> (A \<le> Field r) \<and> (\<forall>a \<in> A. under r a \<le> A)"
-text{* Note: In the definitions of @{text "Above[S]"} and @{text "Under[S]"},
+text\<open>Note: In the definitions of @{text "Above[S]"} and @{text "Under[S]"},
we bounded comprehension by @{text "Field r"} in order to properly cover
- the case of @{text "A"} being empty. *}
+ the case of @{text "A"} being empty.\<close>
lemma underS_subset_under: "underS r a \<le> under r a"
by(auto simp add: underS_def under_def)
@@ -307,22 +307,22 @@
qed
-subsection {* Variations on Well-Founded Relations *}
+subsection \<open>Variations on Well-Founded Relations\<close>
-text {*
+text \<open>
This subsection contains some variations of the results from @{theory Wellfounded}:
\begin{itemize}
\item means for slightly more direct definitions by well-founded recursion;
\item variations of well-founded induction;
\item means for proving a linear order to be a well-order.
\end{itemize}
-*}
+\<close>
-subsubsection {* Characterizations of well-foundedness *}
+subsubsection \<open>Characterizations of well-foundedness\<close>
-text {* A transitive relation is well-founded iff it is ``locally'' well-founded,
-i.e., iff its restriction to the lower bounds of of any element is well-founded. *}
+text \<open>A transitive relation is well-founded iff it is ``locally'' well-founded,
+i.e., iff its restriction to the lower bounds of of any element is well-founded.\<close>
lemma trans_wf_iff:
assumes "trans r"
@@ -362,8 +362,8 @@
ultimately show ?thesis using R_def by blast
qed
-text {* The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded,
-allowing one to assume the set included in the field. *}
+text \<open>The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded,
+allowing one to assume the set included in the field.\<close>
lemma wf_eq_minimal2:
"wf r = (\<forall>A. A <= Field r \<and> A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. \<not> (a',a) \<in> r))"
@@ -408,11 +408,11 @@
qed
-subsubsection {* Characterizations of well-foundedness *}
+subsubsection \<open>Characterizations of well-foundedness\<close>
-text {* The next lemma and its corollary enable one to prove that
+text \<open>The next lemma and its corollary enable one to prove that
a linear order is a well-order in a way which is more standard than
-via well-foundedness of the strict version of the relation. *}
+via well-foundedness of the strict version of the relation.\<close>
lemma Linear_order_wf_diff_Id:
assumes LI: "Linear_order r"
--- a/src/HOL/Orderings.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Orderings.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson
*)
-section {* Abstract orderings *}
+section \<open>Abstract orderings\<close>
theory Orderings
imports HOL
@@ -12,13 +12,13 @@
ML_file "~~/src/Provers/order.ML"
ML_file "~~/src/Provers/quasi.ML" (* FIXME unused? *)
-subsection {* Abstract ordering *}
+subsection \<open>Abstract ordering\<close>
locale ordering =
fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50)
and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50)
assumes strict_iff_order: "a \<prec> b \<longleftrightarrow> a \<preceq> b \<and> a \<noteq> b"
- assumes refl: "a \<preceq> a" -- {* not @{text iff}: makes problems due to multiple (dual) interpretations *}
+ assumes refl: "a \<preceq> a" -- \<open>not @{text iff}: makes problems due to multiple (dual) interpretations\<close>
and antisym: "a \<preceq> b \<Longrightarrow> b \<preceq> a \<Longrightarrow> a = b"
and trans: "a \<preceq> b \<Longrightarrow> b \<preceq> c \<Longrightarrow> a \<preceq> c"
begin
@@ -39,7 +39,7 @@
"a \<preceq> b \<longleftrightarrow> a \<prec> b \<or> a = b"
by (auto simp add: strict_iff_order refl)
-lemma irrefl: -- {* not @{text iff}: makes problems due to multiple (dual) interpretations *}
+lemma irrefl: -- \<open>not @{text iff}: makes problems due to multiple (dual) interpretations\<close>
"\<not> a \<prec> a"
by (simp add: strict_iff_order)
@@ -85,7 +85,7 @@
end
-subsection {* Syntactic orders *}
+subsection \<open>Syntactic orders\<close>
class ord =
fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
@@ -120,7 +120,7 @@
end
-subsection {* Quasi orders *}
+subsection \<open>Quasi orders\<close>
class preorder = ord +
assumes less_le_not_le: "x < y \<longleftrightarrow> x \<le> y \<and> \<not> (y \<le> x)"
@@ -128,10 +128,10 @@
and order_trans: "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
begin
-text {* Reflexivity. *}
+text \<open>Reflexivity.\<close>
lemma eq_refl: "x = y \<Longrightarrow> x \<le> y"
- -- {* This form is useful with the classical reasoner. *}
+ -- \<open>This form is useful with the classical reasoner.\<close>
by (erule ssubst) (rule order_refl)
lemma less_irrefl [iff]: "\<not> x < x"
@@ -141,7 +141,7 @@
unfolding less_le_not_le by blast
-text {* Asymmetry. *}
+text \<open>Asymmetry.\<close>
lemma less_not_sym: "x < y \<Longrightarrow> \<not> (y < x)"
by (simp add: less_le_not_le)
@@ -150,7 +150,7 @@
by (drule less_not_sym, erule contrapos_np) simp
-text {* Transitivity. *}
+text \<open>Transitivity.\<close>
lemma less_trans: "x < y \<Longrightarrow> y < z \<Longrightarrow> x < z"
by (auto simp add: less_le_not_le intro: order_trans)
@@ -162,7 +162,7 @@
by (auto simp add: less_le_not_le intro: order_trans)
-text {* Useful for simplification, but too risky to include by default. *}
+text \<open>Useful for simplification, but too risky to include by default.\<close>
lemma less_imp_not_less: "x < y \<Longrightarrow> (\<not> y < x) \<longleftrightarrow> True"
by (blast elim: less_asym)
@@ -171,13 +171,13 @@
by (blast elim: less_asym)
-text {* Transitivity rules for calculational reasoning *}
+text \<open>Transitivity rules for calculational reasoning\<close>
lemma less_asym': "a < b \<Longrightarrow> b < a \<Longrightarrow> P"
by (rule less_asym)
-text {* Dual order *}
+text \<open>Dual order\<close>
lemma dual_preorder:
"class.preorder (op \<ge>) (op >)"
@@ -186,7 +186,7 @@
end
-subsection {* Partial orders *}
+subsection \<open>Partial orders\<close>
class order = preorder +
assumes antisym: "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
@@ -199,17 +199,17 @@
by default (auto intro: antisym order_trans simp add: less_le)
-text {* Reflexivity. *}
+text \<open>Reflexivity.\<close>
lemma le_less: "x \<le> y \<longleftrightarrow> x < y \<or> x = y"
- -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
+ -- \<open>NOT suitable for iff, since it can cause PROOF FAILED.\<close>
by (fact order.order_iff_strict)
lemma le_imp_less_or_eq: "x \<le> y \<Longrightarrow> x < y \<or> x = y"
unfolding less_le by blast
-text {* Useful for simplification, but too risky to include by default. *}
+text \<open>Useful for simplification, but too risky to include by default.\<close>
lemma less_imp_not_eq: "x < y \<Longrightarrow> (x = y) \<longleftrightarrow> False"
by auto
@@ -218,7 +218,7 @@
by auto
-text {* Transitivity rules for calculational reasoning *}
+text \<open>Transitivity rules for calculational reasoning\<close>
lemma neq_le_trans: "a \<noteq> b \<Longrightarrow> a \<le> b \<Longrightarrow> a < b"
by (fact order.not_eq_order_implies_strict)
@@ -227,7 +227,7 @@
by (rule order.not_eq_order_implies_strict)
-text {* Asymmetry. *}
+text \<open>Asymmetry.\<close>
lemma eq_iff: "x = y \<longleftrightarrow> x \<le> y \<and> y \<le> x"
by (blast intro: antisym)
@@ -239,7 +239,7 @@
by (fact order.strict_implies_not_eq)
-text {* Least value operator *}
+text \<open>Least value operator\<close>
definition (in ord)
Least :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" (binder "LEAST " 10) where
@@ -261,7 +261,7 @@
(blast intro: assms antisym)+
-text {* Dual order *}
+text \<open>Dual order\<close>
lemma dual_order:
"class.order (op \<ge>) (op >)"
@@ -270,7 +270,7 @@
end
-text {* Alternative introduction rule with bias towards strict order *}
+text \<open>Alternative introduction rule with bias towards strict order\<close>
lemma order_strictI:
fixes less (infix "\<sqsubset>" 50)
@@ -299,7 +299,7 @@
qed
-subsection {* Linear (total) orders *}
+subsection \<open>Linear (total) orders\<close>
class linorder = order +
assumes linear: "x \<le> y \<or> y \<le> x"
@@ -364,7 +364,7 @@
lemma not_leE: "\<not> y \<le> x \<Longrightarrow> x < y"
unfolding not_le .
-text {* Dual order *}
+text \<open>Dual order\<close>
lemma dual_linorder:
"class.linorder (op \<ge>) (op >)"
@@ -373,7 +373,7 @@
end
-text {* Alternative introduction rule with bias towards strict order *}
+text \<open>Alternative introduction rule with bias towards strict order\<close>
lemma linorder_strictI:
fixes less (infix "\<sqsubset>" 50)
@@ -383,7 +383,7 @@
shows "class.linorder less_eq less"
proof -
interpret order less_eq less
- by (fact `class.order less_eq less`)
+ by (fact \<open>class.order less_eq less\<close>)
show ?thesis
proof
fix a b
@@ -393,9 +393,9 @@
qed
-subsection {* Reasoning tools setup *}
+subsection \<open>Reasoning tools setup\<close>
-ML {*
+ML \<open>
signature ORDERS =
sig
val print_structures: Proof.context -> unit
@@ -492,22 +492,22 @@
(fn _ => Data.map (AList.delete struct_eq s));
end;
-*}
+\<close>
-attribute_setup order = {*
+attribute_setup order = \<open>
Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) || Args.del >> K NONE) --|
Args.colon (* FIXME || Scan.succeed true *) ) -- Scan.lift Args.name --
Scan.repeat Args.term
>> (fn ((SOME tag, n), ts) => Orders.add_struct (n, ts) tag
| ((NONE, n), ts) => Orders.del_struct (n, ts))
-*} "theorems controlling transitivity reasoner"
+\<close> "theorems controlling transitivity reasoner"
-method_setup order = {*
+method_setup order = \<open>
Scan.succeed (fn ctxt => SIMPLE_METHOD' (Orders.order_tac ctxt []))
-*} "transitivity reasoner"
+\<close> "transitivity reasoner"
-text {* Declarations to set up transitivity reasoner of partial and linear orders. *}
+text \<open>Declarations to set up transitivity reasoner of partial and linear orders.\<close>
context order
begin
@@ -592,16 +592,16 @@
end
-setup {*
+setup \<open>
map_theory_simpset (fn ctxt0 => ctxt0 addSolver
mk_solver "Transitivity" (fn ctxt => Orders.order_tac ctxt (Simplifier.prems_of ctxt)))
(*Adding the transitivity reasoners also as safe solvers showed a slight
speed up, but the reasoning strength appears to be not higher (at least
no breaking of additional proofs in the entire HOL distribution, as
of 5 March 2004, was observed).*)
-*}
+\<close>
-ML {*
+ML \<open>
local
fun prp t thm = Thm.prop_of thm = t; (* FIXME proper aconv!? *)
in
@@ -645,13 +645,13 @@
| _ => NONE);
end;
-*}
+\<close>
simproc_setup antisym_le ("(x::'a::order) \<le> y") = "K antisym_le_simproc"
simproc_setup antisym_less ("\<not> (x::'a::linorder) < y") = "K antisym_less_simproc"
-subsection {* Bounded quantifiers *}
+subsection \<open>Bounded quantifiers\<close>
syntax
"_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10)
@@ -702,7 +702,7 @@
"ALL x>=y. P" => "ALL x. x >= y \<longrightarrow> P"
"EX x>=y. P" => "EX x. x >= y \<and> P"
-print_translation {*
+print_translation \<open>
let
val All_binder = Mixfix.binder_name @{const_syntax All};
val Ex_binder = Mixfix.binder_name @{const_syntax Ex};
@@ -739,10 +739,10 @@
else raise Match)
| _ => raise Match));
in [tr' All_binder, tr' Ex_binder] end
-*}
+\<close>
-subsection {* Transitivity reasoning *}
+subsection \<open>Transitivity reasoning\<close>
context ord
begin
@@ -869,9 +869,9 @@
finally (ord_eq_less_trans) show ?thesis .
qed
-text {*
+text \<open>
Note that this list of rules is in reverse order of priorities.
-*}
+\<close>
lemmas [trans] =
order_less_subst2
@@ -945,8 +945,8 @@
ord_eq_less_trans
trans
-text {* These support proving chains of decreasing inequalities
- a >= b >= c ... in Isar proofs. *}
+text \<open>These support proving chains of decreasing inequalities
+ a >= b >= c ... in Isar proofs.\<close>
lemma xt1 [no_atp]:
"a = b ==> b > c ==> a > c"
@@ -1029,7 +1029,7 @@
*)
-subsection {* Monotonicity *}
+subsection \<open>Monotonicity\<close>
context order
begin
@@ -1100,7 +1100,7 @@
proof (cases "x = y")
case True then show ?thesis by simp
next
- case False with `x \<le> y` have "x < y" by simp
+ case False with \<open>x \<le> y\<close> have "x < y" by simp
with assms strict_monoD have "f x < f y" by auto
then show ?thesis by simp
qed
@@ -1121,8 +1121,8 @@
proof (rule ccontr)
assume "\<not> x \<le> y"
then have "y \<le> x" by simp
- with `mono f` obtain "f y \<le> f x" by (rule monoE)
- with `f x < f y` show False by simp
+ with \<open>mono f\<close> obtain "f y \<le> f x" by (rule monoE)
+ with \<open>f x < f y\<close> show False by simp
qed
qed
@@ -1133,12 +1133,12 @@
assume "f x = f y"
show "x = y" proof (cases x y rule: linorder_cases)
case less with assms strict_monoD have "f x < f y" by auto
- with `f x = f y` show ?thesis by simp
+ with \<open>f x = f y\<close> show ?thesis by simp
next
case equal then show ?thesis .
next
case greater with assms strict_monoD have "f y < f x" by auto
- with `f x = f y` show ?thesis by simp
+ with \<open>f x = f y\<close> show ?thesis by simp
qed
qed simp
@@ -1153,7 +1153,7 @@
show "x \<le> y" proof (rule ccontr)
assume "\<not> x \<le> y" then have "y < x" by simp
with assms strict_monoD have "f y < f x" by auto
- with `f x \<le> f y` show False by simp
+ with \<open>f x \<le> f y\<close> show False by simp
qed
qed
@@ -1166,7 +1166,7 @@
end
-subsection {* min and max -- fundamental *}
+subsection \<open>min and max -- fundamental\<close>
definition (in ord) min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"min a b = (if a \<le> b then a else b)"
@@ -1187,7 +1187,7 @@
by (simp add: max_def)
-subsection {* (Unique) top and bottom elements *}
+subsection \<open>(Unique) top and bottom elements\<close>
class bot =
fixes bot :: 'a ("\<bottom>")
@@ -1246,7 +1246,7 @@
end
-subsection {* Dense orders *}
+subsection \<open>Dense orders\<close>
class dense_order = order +
assumes dense: "x < y \<Longrightarrow> (\<exists>z. x < z \<and> z < y)"
@@ -1263,7 +1263,7 @@
hence "z < y" by simp
from dense[OF this]
obtain x where "x < y" and "z < x" by safe
- moreover have "x \<le> z" using assms[OF `x < y`] .
+ moreover have "x \<le> z" using assms[OF \<open>x < y\<close>] .
ultimately show False by auto
qed
@@ -1274,16 +1274,16 @@
shows "y \<le> z"
proof (rule dense_le)
fix w assume "w < y"
- from dense[OF `x < y`] obtain u where "x < u" "u < y" by safe
+ from dense[OF \<open>x < y\<close>] obtain u where "x < u" "u < y" by safe
from linear[of u w]
show "w \<le> z"
proof (rule disjE)
assume "u \<le> w"
- from less_le_trans[OF `x < u` `u \<le> w`] `w < y`
+ from less_le_trans[OF \<open>x < u\<close> \<open>u \<le> w\<close>] \<open>w < y\<close>
show "w \<le> z" by (rule *)
next
assume "w \<le> u"
- from `w \<le> u` *[OF `x < u` `u < y`]
+ from \<open>w \<le> u\<close> *[OF \<open>x < u\<close> \<open>u < y\<close>]
show "w \<le> z" by (rule order_trans)
qed
qed
@@ -1297,7 +1297,7 @@
hence "z < y" by simp
from dense[OF this]
obtain x where "x < y" and "z < x" by safe
- moreover have "y \<le> x" using assms[OF `z < x`] .
+ moreover have "y \<le> x" using assms[OF \<open>z < x\<close>] .
ultimately show False by auto
qed
@@ -1308,16 +1308,16 @@
shows "y \<le> z"
proof (rule dense_ge)
fix w assume "z < w"
- from dense[OF `z < x`] obtain u where "z < u" "u < x" by safe
+ from dense[OF \<open>z < x\<close>] obtain u where "z < u" "u < x" by safe
from linear[of u w]
show "y \<le> w"
proof (rule disjE)
assume "w \<le> u"
- from `z < w` le_less_trans[OF `w \<le> u` `u < x`]
+ from \<open>z < w\<close> le_less_trans[OF \<open>w \<le> u\<close> \<open>u < x\<close>]
show "y \<le> w" by (rule *)
next
assume "u \<le> w"
- from *[OF `z < u` `u < x`] `u \<le> w`
+ from *[OF \<open>z < u\<close> \<open>u < x\<close>] \<open>u \<le> w\<close>
show "y \<le> w" by (rule order_trans)
qed
qed
@@ -1333,7 +1333,7 @@
class unbounded_dense_linorder = dense_linorder + no_top + no_bot
-subsection {* Wellorders *}
+subsection \<open>Wellorders\<close>
class wellorder = linorder +
assumes less_induct [case_names less]: "(\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P a"
@@ -1359,9 +1359,9 @@
by auto
then show "x \<le> y" by auto
qed
- with `P x` have Least: "(LEAST a. P a) = x"
+ with \<open>P x\<close> have Least: "(LEAST a. P a) = x"
by (rule Least_equality)
- with `P x` show ?thesis by simp
+ with \<open>P x\<close> show ?thesis by simp
qed
qed
then show "P (LEAST x. P x)" and "(LEAST x. P x) \<le> k" by auto
@@ -1384,7 +1384,7 @@
and "\<And>a. \<lbrakk> P a; \<forall>b. P b \<longrightarrow> a \<le> b \<rbrakk> \<Longrightarrow> Q a"
shows "Q (Least P)"
proof (rule LeastI2_order)
- show "P (Least P)" using `P a` by (rule LeastI)
+ show "P (Least P)" using \<open>P a\<close> by (rule LeastI)
next
fix y assume "P y" thus "Least P \<le> y" by (rule Least_le)
next
@@ -1400,7 +1400,7 @@
end
-subsection {* Order on @{typ bool} *}
+subsection \<open>Order on @{typ bool}\<close>
instantiation bool :: "{order_bot, order_top, linorder}"
begin
@@ -1448,7 +1448,7 @@
by simp_all
-subsection {* Order on @{typ "_ \<Rightarrow> _"} *}
+subsection \<open>Order on @{typ "_ \<Rightarrow> _"}\<close>
instantiation "fun" :: (type, ord) ord
begin
@@ -1527,7 +1527,7 @@
unfolding mono_def le_fun_def by auto
-subsection {* Order on unary and binary predicates *}
+subsection \<open>Order on unary and binary predicates\<close>
lemma predicate1I:
assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
@@ -1582,7 +1582,7 @@
by (simp add: top_fun_def)
-subsection {* Name duplicates *}
+subsection \<open>Name duplicates\<close>
lemmas order_eq_refl = preorder_class.eq_refl
lemmas order_less_irrefl = preorder_class.less_irrefl
--- a/src/HOL/Parity.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Parity.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Jacques D. Fleuriot
*)
-section {* Parity in rings and semirings *}
+section \<open>Parity in rings and semirings\<close>
theory Parity
imports Nat_Transfer
begin
-subsection {* Ring structures with parity and @{text even}/@{text odd} predicates *}
+subsection \<open>Ring structures with parity and @{text even}/@{text odd} predicates\<close>
class semiring_parity = comm_semiring_1_cancel + numeral +
assumes odd_one [simp]: "\<not> 2 dvd 1"
@@ -115,7 +115,7 @@
end
-subsection {* Instances for @{typ nat} and @{typ int} *}
+subsection \<open>Instances for @{typ nat} and @{typ int}\<close>
lemma even_Suc_Suc_iff [simp]:
"2 dvd Suc (Suc n) \<longleftrightarrow> 2 dvd n"
@@ -227,7 +227,7 @@
by (simp add: even_int_iff [symmetric])
-subsection {* Parity and powers *}
+subsection \<open>Parity and powers\<close>
context comm_ring_1
begin
@@ -291,34 +291,34 @@
shows "a ^ n \<le> b ^ n"
proof -
have "0 \<le> \<bar>a\<bar>" by auto
- with `\<bar>a\<bar> \<le> \<bar>b\<bar>`
+ with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close>
have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" by (rule power_mono)
- with `even n` show ?thesis by (simp add: power_even_abs)
+ with \<open>even n\<close> show ?thesis by (simp add: power_even_abs)
qed
lemma power_mono_odd:
assumes "odd n" and "a \<le> b"
shows "a ^ n \<le> b ^ n"
proof (cases "b < 0")
- case True with `a \<le> b` have "- b \<le> - a" and "0 \<le> - b" by auto
+ case True with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto
hence "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono)
- with `odd n` show ?thesis by simp
+ with \<open>odd n\<close> show ?thesis by simp
next
case False then have "0 \<le> b" by auto
show ?thesis
proof (cases "a < 0")
- case True then have "n \<noteq> 0" and "a \<le> 0" using `odd n` [THEN odd_pos] by auto
- then have "a ^ n \<le> 0" unfolding power_le_zero_eq using `odd n` by auto
+ case True then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto
+ then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto
moreover
- from `0 \<le> b` have "0 \<le> b ^ n" by auto
+ from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto
ultimately show ?thesis by auto
next
case False then have "0 \<le> a" by auto
- with `a \<le> b` show ?thesis using power_mono by auto
+ with \<open>a \<le> b\<close> show ?thesis using power_mono by auto
qed
qed
-text {* Simplify, when the exponent is a numeral *}
+text \<open>Simplify, when the exponent is a numeral\<close>
lemma zero_le_power_eq_numeral [simp]:
"0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a"
@@ -345,7 +345,7 @@
end
-subsubsection {* Tools setup *}
+subsubsection \<open>Tools setup\<close>
declare transfer_morphism_int_nat [transfer add return:
even_int_iff
--- a/src/HOL/Partial_Function.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Partial_Function.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Alexander Krauss, TU Muenchen
*)
-section {* Partial Function Definitions *}
+section \<open>Partial Function Definitions\<close>
theory Partial_Function
imports Complete_Partial_Order Fun_Def_Base Option
@@ -20,7 +20,7 @@
case empty thus ?case by simp
next
case (insert x A)
- note chain = `Complete_Partial_Order.chain op \<le> (insert x A)`
+ note chain = \<open>Complete_Partial_Order.chain op \<le> (insert x A)\<close>
show ?case
proof(cases "A = {}")
case True thus ?thesis by simp
@@ -36,10 +36,10 @@
by(rule ccpo_Sup_least)(auto simp add: True intro: ccpo_Sup_upper[OF chain'])
hence "\<Squnion>insert x A = \<Squnion>A"
by(rule antisym)(blast intro: ccpo_Sup_upper[OF chain] ccpo_Sup_least[OF chain'])
- with `\<Squnion>A \<in> A` show ?thesis by simp
+ with \<open>\<Squnion>A \<in> A\<close> show ?thesis by simp
next
case False
- with chainD[OF chain, of x "\<Squnion>A"] `\<Squnion>A \<in> A`
+ with chainD[OF chain, of x "\<Squnion>A"] \<open>\<Squnion>A \<in> A\<close>
have "\<Squnion>insert x A = x"
by(auto intro: antisym ccpo_Sup_least[OF chain] order_trans[OF ccpo_Sup_upper[OF chain']] ccpo_Sup_upper[OF chain])
thus ?thesis by simp
@@ -47,10 +47,10 @@
qed
qed
-subsection {* Axiomatic setup *}
+subsection \<open>Axiomatic setup\<close>
-text {* This techical locale constains the requirements for function
- definitions with ccpo fixed points. *}
+text \<open>This techical locale constains the requirements for function
+ definitions with ccpo fixed points.\<close>
definition "fun_ord ord f g \<longleftrightarrow> (\<forall>x. ord (f x) (g x))"
definition "fun_lub L A = (\<lambda>x. L {y. \<exists>f\<in>A. y = f x})"
@@ -165,8 +165,8 @@
abbreviation "mono_body \<equiv> monotone le_fun leq"
abbreviation "admissible \<equiv> ccpo.admissible lub_fun le_fun"
-text {* Interpret manually, to avoid flooding everything with facts about
- orders *}
+text \<open>Interpret manually, to avoid flooding everything with facts about
+ orders\<close>
lemma ccpo: "class.ccpo lub_fun le_fun (mk_less le_fun)"
apply (rule ccpo)
@@ -174,13 +174,13 @@
apply (rule partial_function_definitions_axioms)
done
-text {* The crucial fixed-point theorem *}
+text \<open>The crucial fixed-point theorem\<close>
lemma mono_body_fixp:
"(\<And>x. mono_body (\<lambda>f. F f x)) \<Longrightarrow> fixp_fun F = F (fixp_fun F)"
by (rule ccpo.fixp_unfold[OF ccpo]) (auto simp: monotone_def fun_ord_def)
-text {* Version with curry/uncurry combinators, to be used by package *}
+text \<open>Version with curry/uncurry combinators, to be used by package\<close>
lemma fixp_rule_uc:
fixes F :: "'c \<Rightarrow> 'c" and
@@ -199,7 +199,7 @@
finally show "f = F f" .
qed
-text {* Fixpoint induction rule *}
+text \<open>Fixpoint induction rule\<close>
lemma fixp_induct_uc:
fixes F :: "'c \<Rightarrow> 'c"
@@ -221,7 +221,7 @@
done
-text {* Rules for @{term mono_body}: *}
+text \<open>Rules for @{term mono_body}:\<close>
lemma const_mono[partial_function_mono]: "monotone ord leq (\<lambda>f. c)"
by (rule monotoneI) (rule leq_refl)
@@ -229,7 +229,7 @@
end
-subsection {* Flat interpretation: tailrec and option *}
+subsection \<open>Flat interpretation: tailrec and option\<close>
definition
"flat_ord b x y \<longleftrightarrow> x = b \<or> x = y"
@@ -349,7 +349,7 @@
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])
- qed(simp add: `A \<noteq> {}`)
+ qed(simp add: \<open>A \<noteq> {}\<close>)
thus "P (img_lub f g lub A)" unfolding img_lub_def by simp
qed
@@ -403,11 +403,11 @@
case False
then obtain c where "c \<in> A" and "c \<noteq> b" by auto
{ fix z assume "z \<in> A"
- from chainD[OF ch `c \<in> A` this] have "z = c \<or> z = b"
- unfolding flat_ord_def using `c \<noteq> b` by auto }
+ from chainD[OF ch \<open>c \<in> A\<close> this] have "z = c \<or> z = b"
+ unfolding flat_ord_def using \<open>c \<noteq> b\<close> by auto }
with False have "A - {b} = {c}" by auto
with False have "flat_lub b A = c" by (auto simp: flat_lub_def)
- with `c \<in> A` lub show ?thesis by simp
+ with \<open>c \<in> A\<close> lub show ?thesis by simp
qed
lemma option_admissible: "option.admissible (%(f::'a \<Rightarrow> 'b option).
@@ -441,13 +441,13 @@
using step defined option.fixp_induct_uc[of U F C, OF mono eq inverse2 option_admissible]
unfolding fun_lub_def flat_lub_def by(auto 9 2)
-declaration {* Partial_Function.init "tailrec" @{term tailrec.fixp_fun}
+declaration \<open>Partial_Function.init "tailrec" @{term tailrec.fixp_fun}
@{term tailrec.mono_body} @{thm tailrec.fixp_rule_uc} @{thm tailrec.fixp_induct_uc}
- (SOME @{thm fixp_induct_tailrec[where c=undefined]}) *}
+ (SOME @{thm fixp_induct_tailrec[where c=undefined]})\<close>
-declaration {* Partial_Function.init "option" @{term option.fixp_fun}
+declaration \<open>Partial_Function.init "option" @{term option.fixp_fun}
@{term option.mono_body} @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc}
- (SOME @{thm fixp_induct_option}) *}
+ (SOME @{thm fixp_induct_option})\<close>
hide_const (open) chain
--- a/src/HOL/Power.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Power.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1997 University of Cambridge
*)
-section {* Exponentiation *}
+section \<open>Exponentiation\<close>
theory Power
imports Num Equiv_Relations
begin
-subsection {* Powers for Arbitrary Monoids *}
+subsection \<open>Powers for Arbitrary Monoids\<close>
class power = one + times
begin
@@ -24,7 +24,7 @@
notation (HTML output)
power ("(_\<^bsup>_\<^esup>)" [1000] 1000)
-text {* Special syntax for squares. *}
+text \<open>Special syntax for squares.\<close>
abbreviation (xsymbols)
power2 :: "'a \<Rightarrow> 'a" ("(_\<^sup>2)" [1000] 999) where
@@ -131,7 +131,7 @@
end
-text{*Extract constant factors from powers*}
+text\<open>Extract constant factors from powers\<close>
declare power_mult_distrib [where a = "numeral w" for w, simp]
declare power_mult_distrib [where b = "numeral w" for w, simp]
@@ -191,13 +191,13 @@
context comm_semiring_1
begin
-text {* The divides relation *}
+text \<open>The divides relation\<close>
lemma le_imp_power_dvd:
assumes "m \<le> n" shows "a ^ m dvd a ^ n"
proof
have "a ^ n = a ^ (m + (n - m))"
- using `m \<le> n` by simp
+ using \<open>m \<le> n\<close> by simp
also have "\<dots> = a ^ m * a ^ (n - m)"
by (rule power_add)
finally show "a ^ n = a ^ m * a ^ (n - m)" .
@@ -322,7 +322,7 @@
context division_ring
begin
-text {* FIXME reorient or rename to @{text nonzero_inverse_power} *}
+text \<open>FIXME reorient or rename to @{text nonzero_inverse_power}\<close>
lemma nonzero_power_inverse:
"a \<noteq> 0 \<Longrightarrow> inverse (a ^ n) = (inverse a) ^ n"
by (induct n)
@@ -342,7 +342,7 @@
end
-subsection {* Exponentiation on ordered types *}
+subsection \<open>Exponentiation on ordered types\<close>
context linordered_ring (* TODO: move *)
begin
@@ -386,7 +386,7 @@
by (fact order_trans [OF zero_le_one less_imp_le])
have "1 * 1 < a * 1" using gt1 by simp
also have "\<dots> \<le> a * a ^ n" using gt1
- by (simp only: mult_mono `0 \<le> a` one_le_power order_less_imp_le
+ by (simp only: mult_mono \<open>0 \<le> a\<close> one_le_power order_less_imp_le
zero_le_one order_refl)
finally show ?thesis by simp
qed
@@ -422,13 +422,13 @@
qed
qed
-text{*Surely we can strengthen this? It holds for @{text "0<a<1"} too.*}
+text\<open>Surely we can strengthen this? It holds for @{text "0<a<1"} too.\<close>
lemma power_inject_exp [simp]:
"1 < a \<Longrightarrow> a ^ m = a ^ n \<longleftrightarrow> m = n"
by (force simp add: order_antisym power_le_imp_le_exp)
-text{*Can relax the first premise to @{term "0<a"} in the case of the
-natural numbers.*}
+text\<open>Can relax the first premise to @{term "0<a"} in the case of the
+natural numbers.\<close>
lemma power_less_imp_less_exp:
"1 < a \<Longrightarrow> a ^ m < a ^ n \<Longrightarrow> m < n"
by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"]
@@ -439,7 +439,7 @@
by (induct n)
(auto simp add: mult_strict_mono le_less_trans [of 0 a b])
-text{*Lemma for @{text power_strict_decreasing}*}
+text\<open>Lemma for @{text power_strict_decreasing}\<close>
lemma power_Suc_less:
"0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a * a ^ n < a ^ n"
by (induct n)
@@ -458,7 +458,7 @@
done
qed
-text{*Proof resembles that of @{text power_strict_decreasing}*}
+text\<open>Proof resembles that of @{text power_strict_decreasing}\<close>
lemma power_decreasing [rule_format]:
"n \<le> N \<Longrightarrow> 0 \<le> a \<Longrightarrow> a \<le> 1 \<longrightarrow> a ^ N \<le> a ^ n"
proof (induct N)
@@ -475,7 +475,7 @@
"0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a ^ Suc n < 1"
using power_strict_decreasing [of 0 "Suc n" a] by simp
-text{*Proof again resembles that of @{text power_strict_decreasing}*}
+text\<open>Proof again resembles that of @{text power_strict_decreasing}\<close>
lemma power_increasing [rule_format]:
"n \<le> N \<Longrightarrow> 1 \<le> a \<Longrightarrow> a ^ n \<le> a ^ N"
proof (induct N)
@@ -488,7 +488,7 @@
done
qed
-text{*Lemma for @{text power_strict_increasing}*}
+text\<open>Lemma for @{text power_strict_increasing}\<close>
lemma power_less_power_Suc:
"1 < a \<Longrightarrow> a ^ n < a * a ^ n"
by (induct n) (auto simp add: mult_strict_left_mono less_trans [OF zero_less_one])
@@ -700,7 +700,7 @@
end
-subsection {* Miscellaneous rules *}
+subsection \<open>Miscellaneous rules\<close>
lemma self_le_power:
fixes x::"'a::linordered_semidom"
@@ -722,7 +722,7 @@
"(0::'a::{power, semiring_0}) ^ Suc n = 0"
by simp
-text{*It looks plausible as a simprule, but its effect can be strange.*}
+text\<open>It looks plausible as a simprule, but its effect can be strange.\<close>
lemma power_0_left:
"0 ^ n = (if n = 0 then 1 else (0::'a::{power, semiring_0}))"
by (induct n) simp_all
@@ -732,7 +732,7 @@
shows "n \<le> m \<Longrightarrow> a ^ (m - n) = a ^ m / a ^ n"
by (induct m n rule: diff_induct) (simp_all add: nz field_power_not_zero)
-text{*Perhaps these should be simprules.*}
+text\<open>Perhaps these should be simprules.\<close>
lemma power_inverse:
fixes a :: "'a::division_ring"
shows "inverse (a ^ n) = inverse a ^ n"
@@ -753,7 +753,7 @@
apply assumption
done
-text {* Simprules for comparisons where common factors can be cancelled. *}
+text \<open>Simprules for comparisons where common factors can be cancelled.\<close>
lemmas zero_compare_simps =
add_strict_increasing add_strict_increasing2 add_increasing
@@ -764,7 +764,7 @@
zero_le_power2 power2_less_0
-subsection {* Exponentiation for the Natural Numbers *}
+subsection \<open>Exponentiation for the Natural Numbers\<close>
lemma nat_one_le_power [simp]:
"Suc 0 \<le> i \<Longrightarrow> Suc 0 \<le> i ^ n"
@@ -782,9 +782,9 @@
"Suc 0 ^ n = Suc 0"
by simp
-text{*Valid for the naturals, but what if @{text"0<i<1"}?
+text\<open>Valid for the naturals, but what if @{text"0<i<1"}?
Premises cannot be weakened: consider the case where @{term "i=0"},
-@{term "m=1"} and @{term "n=0"}.*}
+@{term "m=1"} and @{term "n=0"}.\<close>
lemma nat_power_less_imp_less:
assumes nonneg: "0 < (i\<Colon>nat)"
assumes less: "i ^ m < i ^ n"
@@ -824,7 +824,7 @@
qed
qed
-subsubsection {* Cardinality of the Powerset *}
+subsubsection \<open>Cardinality of the Powerset\<close>
lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2"
unfolding UNIV_bool by simp
@@ -846,7 +846,7 @@
qed
-subsubsection {* Generalized sum over a set *}
+subsubsection \<open>Generalized sum over a set\<close>
lemma setsum_zero_power [simp]:
fixes c :: "nat \<Rightarrow> 'a::division_ring"
@@ -861,7 +861,7 @@
by auto
-subsubsection {* Generalized product over a set *}
+subsubsection \<open>Generalized product over a set\<close>
lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
apply (erule finite_induct)
@@ -910,7 +910,7 @@
ultimately show ?thesis by blast
qed
-subsection {* Code generator tweak *}
+subsection \<open>Code generator tweak\<close>
lemma power_power_power [code]:
"power = power.power (1::'a::{power}) (op *)"
--- a/src/HOL/Predicate.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Predicate.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Author: Lukas Bulwahn and Florian Haftmann, TU Muenchen
*)
-section {* Predicates as enumerations *}
+section \<open>Predicates as enumerations\<close>
theory Predicate
imports String
begin
-subsection {* The type of predicate enumerations (a monad) *}
+subsection \<open>The type of predicate enumerations (a monad)\<close>
datatype (plugins only: code extraction) (dead 'a) pred = Pred "'a \<Rightarrow> bool"
@@ -197,7 +197,7 @@
using assms by (cases A) (auto simp add: bot_pred_def)
-subsection {* Emptiness check and definite choice *}
+subsection \<open>Emptiness check and definite choice\<close>
definition is_empty :: "'a pred \<Rightarrow> bool" where
"is_empty A \<longleftrightarrow> A = \<bottom>"
@@ -321,7 +321,7 @@
using singleton_sup_aux [of default A B] by (simp only: singleton_sup_single_single)
-subsection {* Derived operations *}
+subsection \<open>Derived operations\<close>
definition if_pred :: "bool \<Rightarrow> unit pred" where
if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"
@@ -411,7 +411,7 @@
by (rule ext, rule pred_eqI, auto)+
-subsection {* Implementation *}
+subsection \<open>Implementation\<close>
datatype (plugins only: code extraction) (dead 'a) seq =
Empty
@@ -610,7 +610,7 @@
code_reflect Predicate
datatypes pred = Seq and seq = Empty | Insert | Join
-ML {*
+ML \<open>
signature PREDICATE =
sig
val anamorph: ('a -> ('b * 'a) option) -> int -> 'a -> 'b list * 'a
@@ -648,9 +648,9 @@
fun yieldn k = anamorph yield k;
end;
-*}
+\<close>
-text {* Conversion from and to sets *}
+text \<open>Conversion from and to sets\<close>
definition pred_of_set :: "'a set \<Rightarrow> 'a pred" where
"pred_of_set = Pred \<circ> (\<lambda>A x. x \<in> A)"
@@ -686,7 +686,7 @@
"set_of_seq (Predicate.Join P xq) = set_of_pred P \<union> set_of_seq xq"
by auto
-text {* Lazy Evaluation of an indexed function *}
+text \<open>Lazy Evaluation of an indexed function\<close>
function iterate_upto :: "(natural \<Rightarrow> 'a) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> 'a Predicate.pred"
where
@@ -698,7 +698,7 @@
termination by (relation "measure (%(f, n, m). nat_of_natural (m + 1 - n))")
(auto simp add: less_natural_def)
-text {* Misc *}
+text \<open>Misc\<close>
declare Inf_set_fold [where 'a = "'a Predicate.pred", code]
declare Sup_set_fold [where 'a = "'a Predicate.pred", code]
@@ -711,7 +711,7 @@
proof (rule sym)
interpret comp_fun_idem "sup :: 'a Predicate.pred \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a Predicate.pred"
by (fact comp_fun_idem_sup)
- from `finite A` show "?rhs = ?lhs" by (induct A) (auto intro!: pred_eqI)
+ from \<open>finite A\<close> show "?rhs = ?lhs" by (induct A) (auto intro!: pred_eqI)
qed
lemma pred_of_set_set_fold_sup:
--- a/src/HOL/Predicate_Compile.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Predicate_Compile.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Stefan Berghofer, Lukas Bulwahn, Florian Haftmann, TU Muenchen
*)
-section {* A compiler for predicates defined by introduction rules *}
+section \<open>A compiler for predicates defined by introduction rules\<close>
theory Predicate_Compile
imports Random_Sequence Quickcheck_Exhaustive
@@ -22,12 +22,12 @@
ML_file "Tools/Predicate_Compile/predicate_compile.ML"
-subsection {* Set membership as a generator predicate *}
+subsection \<open>Set membership as a generator predicate\<close>
-text {*
+text \<open>
Introduce a new constant for membership to allow
fine-grained control in code equations.
-*}
+\<close>
definition contains :: "'a set => 'a => bool"
where "contains A x \<longleftrightarrow> x : A"
@@ -68,7 +68,7 @@
"\<not> contains A x ==> Predicate.eval (Predicate.not_pred (contains_pred A x)) ()"
by(simp add: contains_pred_def contains_def not_pred_eq)
-setup {*
+setup \<open>
let
val Fun = Predicate_Compile_Aux.Fun
val Input = Predicate_Compile_Aux.Input
@@ -97,7 +97,7 @@
})],
needs_random = []}))
end
-*}
+\<close>
hide_const (open) contains contains_pred
hide_fact (open) pred_of_setE pred_of_setI pred_of_set_eq
--- a/src/HOL/Presburger.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Presburger.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Amine Chaieb, TU Muenchen
*)
-section {* Decision Procedure for Presburger Arithmetic *}
+section \<open>Decision Procedure for Presburger Arithmetic\<close>
theory Presburger
imports Groebner_Basis Set_Interval
@@ -12,7 +12,7 @@
ML_file "Tools/Qelim/qelim.ML"
ML_file "Tools/Qelim/cooper_procedure.ML"
-subsection{* The @{text "-\<infinity>"} and @{text "+\<infinity>"} Properties *}
+subsection\<open>The @{text "-\<infinity>"} and @{text "+\<infinity>"} Properties\<close>
lemma minf:
"\<lbrakk>\<exists>(z ::'a::linorder).\<forall>x<z. P x = P' x; \<exists>z.\<forall>x<z. Q x = Q' x\<rbrakk>
@@ -59,7 +59,7 @@
unfolding dvd_def mult.commute [of d]
by auto
-subsection{* The A and B sets *}
+subsection\<open>The A and B sets\<close>
lemma bset:
"\<lbrakk>\<forall>x.(\<forall>j \<in> {1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> P x \<longrightarrow> P(x - D) ;
\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> Q x \<longrightarrow> Q(x - D)\<rbrakk> \<Longrightarrow>
@@ -176,9 +176,9 @@
thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x + D) + t)" by auto
qed blast
-subsection{* Cooper's Theorem @{text "-\<infinity>"} and @{text "+\<infinity>"} Version *}
+subsection\<open>Cooper's Theorem @{text "-\<infinity>"} and @{text "+\<infinity>"} Version\<close>
-subsubsection{* First some trivial facts about periodic sets or predicates *}
+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)"
@@ -209,7 +209,7 @@
qed
qed auto
-subsubsection{* The @{text "-\<infinity>"} Version*}
+subsubsection\<open>The @{text "-\<infinity>"} Version\<close>
lemma decr_lemma: "0 < (d::int) \<Longrightarrow> x - (abs(x-z)+1) * d < z"
by(induct rule: int_gr_induct,simp_all add:int_distrib)
@@ -277,7 +277,7 @@
ultimately show ?thesis by blast
qed
-subsubsection {* The @{text "+\<infinity>"} Version*}
+subsubsection \<open>The @{text "+\<infinity>"} Version\<close>
lemma plusinfinity:
assumes dpos: "(0::int) < d" and
@@ -372,15 +372,15 @@
shows "d dvd t \<equiv> - d dvd t" and "d dvd t \<equiv> d dvd - t"
by simp_all
-text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*}
+text \<open>\bigskip Theorems for transforming predicates on nat to predicates on @{text int}\<close>
lemma zdiff_int_split: "P (int (x - y)) =
((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))"
by (cases "y \<le> x") (simp_all add: zdiff_int)
-text {*
+text \<open>
\medskip Specific instances of congruence rules, to prevent
- simplifier from looping. *}
+ simplifier from looping.\<close>
theorem imp_le_cong:
"\<lbrakk>x = x'; 0 \<le> x' \<Longrightarrow> P = P'\<rbrakk> \<Longrightarrow> (0 \<le> (x::int) \<longrightarrow> P) = (0 \<le> x' \<longrightarrow> P')"
@@ -392,7 +392,7 @@
ML_file "Tools/Qelim/cooper.ML"
-method_setup presburger = {*
+method_setup presburger = \<open>
let
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
@@ -408,7 +408,7 @@
(fn ((elim, add_ths), del_ths) => fn ctxt =>
SIMPLE_METHOD' (Cooper.tac elim add_ths del_ths ctxt))
end
-*} "Cooper's algorithm for Presburger arithmetic"
+\<close> "Cooper's algorithm for Presburger arithmetic"
declare dvd_eq_mod_eq_0 [symmetric, presburger]
declare mod_1 [presburger]
@@ -490,7 +490,7 @@
using even_int_iff [of n] by simp
-subsection {* Nice facts about division by @{term 4} *}
+subsection \<open>Nice facts about division by @{term 4}\<close>
lemma even_even_mod_4_iff:
"even (n::nat) \<longleftrightarrow> even (n mod 4)"
@@ -505,7 +505,7 @@
by presburger
-subsection {* Try0 *}
+subsection \<open>Try0\<close>
ML_file "Tools/try0.ML"
--- a/src/HOL/Product_Type.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Product_Type.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,36 +3,36 @@
Copyright 1992 University of Cambridge
*)
-section {* Cartesian products *}
+section \<open>Cartesian products\<close>
theory Product_Type
imports Typedef Inductive Fun
keywords "inductive_set" "coinductive_set" :: thy_decl
begin
-subsection {* @{typ bool} is a datatype *}
+subsection \<open>@{typ bool} is a datatype\<close>
free_constructors case_bool for True | False
by auto
-text {* Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}. *}
+text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
-setup {* Sign.mandatory_path "old" *}
+setup \<open>Sign.mandatory_path "old"\<close>
old_rep_datatype True False by (auto intro: bool_induct)
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
-text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
+text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>
-setup {* Sign.mandatory_path "bool" *}
+setup \<open>Sign.mandatory_path "bool"\<close>
lemmas induct = old.bool.induct
lemmas inducts = old.bool.inducts
lemmas rec = old.bool.rec
lemmas simps = bool.distinct bool.case bool.rec
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
declare case_split [cases type: bool]
-- "prefer plain propositional version"
@@ -50,14 +50,14 @@
shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)"
using assms by simp_all
-setup {* Code.add_case @{thm If_case_cert} *}
+setup \<open>Code.add_case @{thm If_case_cert}\<close>
code_printing
constant "HOL.equal :: bool \<Rightarrow> bool \<Rightarrow> bool" \<rightharpoonup> (Haskell) infix 4 "=="
| class_instance "bool" :: "equal" \<rightharpoonup> (Haskell) -
-subsection {* The @{text unit} type *}
+subsection \<open>The @{text unit} type\<close>
typedef unit = "{True}"
by auto
@@ -68,38 +68,38 @@
lemma unit_eq [no_atp]: "u = ()"
by (induct u) (simp add: Unity_def)
-text {*
+text \<open>
Simplification procedure for @{thm [source] unit_eq}. Cannot use
this rule directly --- it loops!
-*}
+\<close>
-simproc_setup unit_eq ("x::unit") = {*
+simproc_setup unit_eq ("x::unit") = \<open>
fn _ => fn _ => fn ct =>
if HOLogic.is_unit (Thm.term_of ct) then NONE
else SOME (mk_meta_eq @{thm unit_eq})
-*}
+\<close>
free_constructors case_unit for "()"
by auto
-text {* Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}. *}
+text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
-setup {* Sign.mandatory_path "old" *}
+setup \<open>Sign.mandatory_path "old"\<close>
old_rep_datatype "()" by simp
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
-text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
+text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>
-setup {* Sign.mandatory_path "unit" *}
+setup \<open>Sign.mandatory_path "unit"\<close>
lemmas induct = old.unit.induct
lemmas inducts = old.unit.inducts
lemmas rec = old.unit.rec
lemmas simps = unit.case unit.rec
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()"
by simp
@@ -107,11 +107,11 @@
lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P"
by (rule triv_forall_equality)
-text {*
+text \<open>
This rewrite counters the effect of simproc @{text unit_eq} on @{term
[source] "%u::unit. f u"}, replacing it by @{term [source]
f} rather than by @{term [source] "%u. f ()"}.
-*}
+\<close>
lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f"
by (rule ext) simp
@@ -217,9 +217,9 @@
Unit
-subsection {* The product type *}
+subsection \<open>The product type\<close>
-subsubsection {* Type definition *}
+subsubsection \<open>Type definition\<close>
definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
"Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
@@ -255,18 +255,18 @@
by (simp add: Pair_def Abs_prod_inject)
qed
-text {* Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}. *}
+text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
-setup {* Sign.mandatory_path "old" *}
+setup \<open>Sign.mandatory_path "old"\<close>
old_rep_datatype Pair
by (erule prod_cases) (rule prod.inject)
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
-text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
+text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>
-setup {* Sign.mandatory_path "prod" *}
+setup \<open>Sign.mandatory_path "prod"\<close>
declare
old.prod.inject[iff del]
@@ -276,21 +276,21 @@
lemmas rec = old.prod.rec
lemmas simps = prod.inject prod.case prod.rec
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
declare prod.case [nitpick_simp del]
declare prod.case_cong_weak [cong del]
-subsubsection {* Tuple syntax *}
+subsubsection \<open>Tuple syntax\<close>
abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
"split \<equiv> case_prod"
-text {*
+text \<open>
Patterns -- extends pre-defined type @{typ pttrn} used in
abstractions.
-*}
+\<close>
nonterminal tuple_args and patterns
@@ -310,12 +310,12 @@
"%(x, y, zs). b" == "CONST case_prod (%x (y, zs). b)"
"%(x, y). b" == "CONST case_prod (%x y. b)"
"_abs (CONST Pair x y) t" => "%(x, y). t"
- -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
- The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *}
+ -- \<open>The last rule accommodates tuples in `case C ... (x,y) ... => ...'
+ The (x,y) is parsed as `Pair x y' because it is logic, not pttrn\<close>
(*reconstruct pattern from (nested) splits, avoiding eta-contraction of body;
works best with enclosing "let", if "let" does not avoid eta-contraction*)
-print_translation {*
+print_translation \<open>
let
fun split_tr' [Abs (x, T, t as (Abs abs))] =
(* split (%x y. t) => %(x,y) t *)
@@ -348,10 +348,10 @@
end
| split_tr' _ = raise Match;
in [(@{const_syntax case_prod}, K split_tr')] end
-*}
+\<close>
(* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *)
-typed_print_translation {*
+typed_print_translation \<open>
let
fun split_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
| split_guess_names_tr' T [Abs (x, xT, t)] =
@@ -381,10 +381,10 @@
end)
| split_guess_names_tr' _ _ = raise Match;
in [(@{const_syntax case_prod}, K split_guess_names_tr')] end
-*}
+\<close>
-subsubsection {* Code generator setup *}
+subsubsection \<open>Code generator setup\<close>
code_printing
type_constructor prod \<rightharpoonup>
@@ -403,7 +403,7 @@
(Haskell) infix 4 "=="
-subsubsection {* Fundamental operations and properties *}
+subsubsection \<open>Fundamental operations and properties\<close>
lemma Pair_inject:
assumes "(a, b) = (a', b')"
@@ -448,7 +448,7 @@
by (simp add: fun_eq_iff split: prod.split)
lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f"
- -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
+ -- \<open>Subsumes the old @{text split_Pair} when @{term f} is the identity function.\<close>
by (simp add: fun_eq_iff split: prod.split)
lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)"
@@ -461,7 +461,7 @@
by (simp add: case_prod_unfold)
lemmas split_weak_cong = prod.case_cong_weak
- -- {* Prevents simplification of @{term c}: much faster *}
+ -- \<open>Prevents simplification of @{term c}: much faster\<close>
lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g"
by (simp add: split_eta)
@@ -474,22 +474,22 @@
next
fix x
assume "!!a b. PROP P (a, b)"
- from `PROP P (fst x, snd x)` show "PROP P x" by simp
+ from \<open>PROP P (fst x, snd x)\<close> show "PROP P x" by simp
qed
lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
by (cases x) simp
-text {*
+text \<open>
The rule @{thm [source] split_paired_all} does not work with the
Simplifier because it also affects premises in congrence rules,
where this can lead to premises of the form @{text "!!a b. ... =
?P(a, b)"} which cannot be solved by reflexivity.
-*}
+\<close>
lemmas split_tupled_all = split_paired_all unit_all_eq2
-ML {*
+ML \<open>
(* replace parameters of product type by individual component parameters *)
local (* filtering with exists_paired_all is an essential optimization *)
fun exists_paired_all (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) =
@@ -513,30 +513,30 @@
if exists_paired_all (Thm.prop_of th)
then full_simplify (put_simpset ss ctxt) th else th;
end;
-*}
+\<close>
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))"
- -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *}
+ -- \<open>@{text "[iff]"} is not a good idea because it makes @{text blast} loop\<close>
by fast
lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))"
by fast
lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))"
- -- {* Can't be added to simpset: loops! *}
+ -- \<open>Can't be added to simpset: loops!\<close>
by (simp add: split_eta)
-text {*
+text \<open>
Simplification procedure for @{thm [source] cond_split_eta}. Using
@{thm [source] split_eta} as a rewrite rule is not general enough,
and using @{thm [source] cond_split_eta} directly would render some
existing proofs very inefficient; similarly for @{text
split_beta}.
-*}
+\<close>
-ML {*
+ML \<open>
local
val cond_split_eta_ss =
simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms cond_split_eta});
@@ -578,11 +578,11 @@
| NONE => NONE)
| eta_proc _ _ = NONE;
end;
-*}
+\<close>
simproc_setup split_beta ("split f z") =
- {* fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct) *}
+ \<open>fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\<close>
simproc_setup split_eta ("split f") =
- {* fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct) *}
+ \<open>fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\<close>
lemmas split_beta [mono] = prod.case_eq_if
@@ -590,22 +590,22 @@
by (auto simp: fun_eq_iff)
lemmas split_split [no_atp] = prod.split
- -- {* For use with @{text split} and the Simplifier. *}
+ -- \<open>For use with @{text split} and the Simplifier.\<close>
-text {*
+text \<open>
@{thm [source] split_split} could be declared as @{text "[split]"}
done after the Splitter has been speeded up significantly;
precompute the constants involved and don't do anything unless the
current goal contains one of those constants.
-*}
+\<close>
lemmas split_split_asm [no_atp] = prod.split_asm
-text {*
+text \<open>
\medskip @{term split} used as a logical connective or set former.
\medskip These rules are for use with @{text blast}; could instead
- call @{text simp} using @{thm [source] prod.split} as rewrite. *}
+ call @{text simp} using @{thm [source] prod.split} as rewrite.\<close>
lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p"
apply (simp only: split_tupled_all)
@@ -651,7 +651,7 @@
declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!]
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]
-ML {*
+ML \<open>
local (* filtering with exists_p_split is an essential optimization *)
fun exists_p_split (Const (@{const_name case_prod},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
| exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
@@ -663,11 +663,11 @@
then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_conv}) i
else no_tac);
end;
-*}
+\<close>
(* This prevents applications of splitE for already splitted arguments leading
to quite time-consuming computations (in particular for nested tuples) *)
-setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac)) *}
+setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\<close>
lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
by (rule ext) fast
@@ -676,7 +676,7 @@
by (rule ext) fast
lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"
- -- {* Allows simplifications of nested splits in case of independent predicates. *}
+ -- \<open>Allows simplifications of nested splits in case of independent predicates.\<close>
by (rule ext) blast
(* Do NOT make this a simp rule as it
@@ -709,9 +709,9 @@
qed "The_split_eq";
*)
-text {*
+text \<open>
Setup of internal @{text split_rule}.
-*}
+\<close>
lemmas case_prodI = prod.case [THEN iffD2]
@@ -788,7 +788,7 @@
hide_const internal_split
-subsubsection {* Derived operations *}
+subsubsection \<open>Derived operations\<close>
definition curry :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c" where
"curry = (\<lambda>c x y. c (x, y))"
@@ -814,9 +814,9 @@
lemma curry_K: "curry (\<lambda>x. c) = (\<lambda>x y. c)"
by(simp add: fun_eq_iff)
-text {*
+text \<open>
The composition-uncurry combinator.
-*}
+\<close>
notation fcomp (infixl "\<circ>>" 60)
@@ -850,10 +850,10 @@
no_notation fcomp (infixl "\<circ>>" 60)
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
-text {*
+text \<open>
@{term map_prod} --- action of the product functor upon
functions.
-*}
+\<close>
definition map_prod :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
"map_prod f g = (\<lambda>(x, y). (f x, g y))"
@@ -988,7 +988,7 @@
context
begin
-local_setup {* Local_Theory.map_background_naming (Name_Space.mandatory_path "prod") *}
+local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\<close>
definition swap :: "'a \<times> 'b \<Rightarrow> 'b \<times> 'a"
where
@@ -1032,9 +1032,9 @@
"(case prod.swap p of (y, x) \<Rightarrow> f x y) = (case p of (x, y) \<Rightarrow> f x y)"
by (cases p) simp
-text {*
+text \<open>
Disjoint union of a family of sets -- Sigma.
-*}
+\<close>
definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set" where
Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
@@ -1064,13 +1064,13 @@
"[| c: Sigma A B;
!!x y.[| x:A; y:B(x); c=(x,y) |] ==> P
|] ==> P"
- -- {* The general elimination rule. *}
+ -- \<open>The general elimination rule.\<close>
by (unfold Sigma_def) blast
-text {*
+text \<open>
Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
eigenvariables.
-*}
+\<close>
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A"
by blast
@@ -1128,7 +1128,7 @@
lemma UN_Times_distrib:
"(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)"
- -- {* Suggested by Pierre Chartier *}
+ -- \<open>Suggested by Pierre Chartier\<close>
by blast
lemma split_paired_Ball_Sigma [simp, no_atp]:
@@ -1160,10 +1160,10 @@
lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)"
by blast
-text {*
+text \<open>
Non-dependent versions are needed to avoid the need for higher-order
matching, especially when the rules are re-oriented.
-*}
+\<close>
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)"
by (fact Sigma_Un_distrib1)
@@ -1245,7 +1245,7 @@
"x \<in> Product_Type.product A B \<longleftrightarrow> x \<in> A \<times> B"
by (simp add: product_def)
-text {* The following @{const map_prod} lemmas are due to Joachim Breitner: *}
+text \<open>The following @{const map_prod} lemmas are due to Joachim Breitner:\<close>
lemma map_prod_inj_on:
assumes "inj_on f A" and "inj_on g B"
@@ -1257,12 +1257,12 @@
assume "map_prod f g x = map_prod f g y"
hence "fst (map_prod f g x) = fst (map_prod f g y)" by (auto)
hence "f (fst x) = f (fst y)" by (cases x,cases y,auto)
- with `inj_on f A` and `fst x \<in> A` and `fst y \<in> A`
+ with \<open>inj_on f A\<close> and \<open>fst x \<in> A\<close> and \<open>fst y \<in> A\<close>
have "fst x = fst y" by (auto dest:dest:inj_onD)
- moreover from `map_prod f g x = map_prod f g y`
+ moreover from \<open>map_prod f g x = map_prod f g y\<close>
have "snd (map_prod f g x) = snd (map_prod f g y)" by (auto)
hence "g (snd x) = g (snd y)" by (cases x,cases y,auto)
- with `inj_on g B` and `snd x \<in> B` and `snd y \<in> B`
+ with \<open>inj_on g B\<close> and \<open>snd x \<in> B\<close> and \<open>snd y \<in> B\<close>
have "snd x = snd y" by (auto dest:dest:inj_onD)
ultimately show "x = y" by(rule prod_eqI)
qed
@@ -1274,9 +1274,9 @@
unfolding surj_def
proof
fix y :: "'b \<times> 'd"
- from `surj f` obtain a where "fst y = f a" by (auto elim:surjE)
+ from \<open>surj f\<close> obtain a where "fst y = f a" by (auto elim:surjE)
moreover
- from `surj g` obtain b where "snd y = g b" by (auto elim:surjE)
+ from \<open>surj g\<close> obtain b where "snd y = g b" by (auto elim:surjE)
ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto
thus "\<exists>x. y = map_prod f g x" by auto
qed
@@ -1289,39 +1289,39 @@
fix x :: "'a \<times> 'c"
assume "x \<in> {y\<Colon>'a \<times> 'c. \<exists>x\<Colon>'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
then obtain y where "y \<in> A \<times> B" and "x = map_prod f g y" by blast
- from `image f A = A'` and `y \<in> A \<times> B` have "f (fst y) \<in> A'" by auto
- moreover from `image g B = B'` and `y \<in> A \<times> B` have "g (snd y) \<in> B'" by auto
+ from \<open>image f A = A'\<close> and \<open>y \<in> A \<times> B\<close> have "f (fst y) \<in> A'" by auto
+ moreover from \<open>image g B = B'\<close> and \<open>y \<in> A \<times> B\<close> have "g (snd y) \<in> B'" by auto
ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')" by auto
- with `x = map_prod f g y` show "x \<in> A' \<times> B'" by (cases y, auto)
+ with \<open>x = map_prod f g y\<close> show "x \<in> A' \<times> B'" by (cases y, auto)
next
fix x :: "'a \<times> 'c"
assume "x \<in> A' \<times> B'" hence "fst x \<in> A'" and "snd x \<in> B'" by auto
- from `image f A = A'` and `fst x \<in> A'` have "fst x \<in> image f A" by auto
+ from \<open>image f A = A'\<close> and \<open>fst x \<in> A'\<close> have "fst x \<in> image f A" by auto
then obtain a where "a \<in> A" and "fst x = f a" by (rule imageE)
- moreover from `image g B = B'` and `snd x \<in> B'`
+ moreover from \<open>image g B = B'\<close> and \<open>snd x \<in> B'\<close>
obtain b where "b \<in> B" and "snd x = g b" by auto
ultimately have "(fst x, snd x) = map_prod f g (a,b)" by auto
- moreover from `a \<in> A` and `b \<in> B` have "(a , b) \<in> A \<times> B" by auto
+ moreover from \<open>a \<in> A\<close> and \<open>b \<in> B\<close> have "(a , b) \<in> A \<times> B" by auto
ultimately have "\<exists>y \<in> A \<times> B. x = map_prod f g y" by auto
thus "x \<in> {x. \<exists>y \<in> A \<times> B. x = map_prod f g y}" by auto
qed
-subsection {* Simproc for rewriting a set comprehension into a pointfree expression *}
+subsection \<open>Simproc for rewriting a set comprehension into a pointfree expression\<close>
ML_file "Tools/set_comprehension_pointfree.ML"
-setup {*
+setup \<open>
Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs
[Raw_Simplifier.make_simproc {name = "set comprehension", lhss = [@{cpat "Collect ?P"}],
proc = K Set_Comprehension_Pointfree.code_simproc, identifier = []}])
-*}
+\<close>
-subsection {* Inductively defined sets *}
+subsection \<open>Inductively defined sets\<close>
(* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
-simproc_setup Collect_mem ("Collect t") = {*
+simproc_setup Collect_mem ("Collect t") = \<open>
fn _ => fn ctxt => fn ct =>
(case Thm.term_of ct of
S as Const (@{const_name Collect}, Type (@{type_name fun}, [_, T])) $ t =>
@@ -1350,12 +1350,12 @@
| _ => NONE)
end
| _ => NONE)
-*}
+\<close>
ML_file "Tools/inductive_set.ML"
-subsection {* Legacy theorem bindings and duplicates *}
+subsection \<open>Legacy theorem bindings and duplicates\<close>
lemma PairE:
obtains x y where "p = (x, y)"
--- a/src/HOL/Quickcheck_Exhaustive.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Quickcheck_Exhaustive.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,19 +1,19 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* A simple counterexample generator performing exhaustive testing *}
+section \<open>A simple counterexample generator performing exhaustive testing\<close>
theory Quickcheck_Exhaustive
imports Quickcheck_Random
keywords "quickcheck_generator" :: thy_decl
begin
-subsection {* basic operations for exhaustive generators *}
+subsection \<open>basic operations for exhaustive generators\<close>
definition orelse :: "'a option => 'a option => 'a option" (infixr "orelse" 55)
where
[code_unfold]: "x orelse y = (case x of Some x' => Some x' | None => y)"
-subsection {* exhaustive generator type classes *}
+subsection \<open>exhaustive generator type classes\<close>
class exhaustive = term_of +
fixes exhaustive :: "('a \<Rightarrow> (bool * term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option"
@@ -229,7 +229,7 @@
end
-subsubsection {* A smarter enumeration scheme for functions over finite datatypes *}
+subsubsection \<open>A smarter enumeration scheme for functions over finite datatypes\<close>
class check_all = enum + term_of +
fixes check_all :: "('a * (unit \<Rightarrow> term) \<Rightarrow> (bool * term list) option) \<Rightarrow> (bool * term list) option"
@@ -505,12 +505,12 @@
end
-subsection {* Bounded universal quantifiers *}
+subsection \<open>Bounded universal quantifiers\<close>
class bounded_forall =
fixes bounded_forall :: "('a \<Rightarrow> bool) \<Rightarrow> natural \<Rightarrow> bool"
-subsection {* Fast exhaustive combinators *}
+subsection \<open>Fast exhaustive combinators\<close>
class fast_exhaustive = term_of +
fixes fast_exhaustive :: "('a \<Rightarrow> unit) \<Rightarrow> natural \<Rightarrow> unit"
@@ -524,7 +524,7 @@
| constant catch_Counterexample \<rightharpoonup>
(Quickcheck) "(((_); NONE) handle Exhaustive'_Generators.Counterexample ts => SOME ts)"
-subsection {* Continuation passing style functions as plus monad *}
+subsection \<open>Continuation passing style functions as plus monad\<close>
type_synonym 'a cps = "('a => term list option) => term list option"
@@ -610,7 +610,7 @@
where
"pos_bound_cps_not n = (%c i. case n (%u. Value []) i of No_value => c () | Value _ => None | Unknown_value => None)"
-subsection {* Defining generators for any first-order data type *}
+subsection \<open>Defining generators for any first-order data type\<close>
axiomatization unknown :: 'a
@@ -620,7 +620,7 @@
declare [[quickcheck_batch_tester = exhaustive]]
-subsection {* Defining generators for abstract types *}
+subsection \<open>Defining generators for abstract types\<close>
ML_file "Tools/Quickcheck/abstract_generators.ML"
--- a/src/HOL/Quickcheck_Narrowing.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Quickcheck_Narrowing.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,22 +1,22 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* Counterexample generator performing narrowing-based testing *}
+section \<open>Counterexample generator performing narrowing-based testing\<close>
theory Quickcheck_Narrowing
imports Quickcheck_Random
keywords "find_unused_assms" :: diag
begin
-subsection {* Counterexample generator *}
+subsection \<open>Counterexample generator\<close>
-subsubsection {* Code generation setup *}
+subsubsection \<open>Code generation setup\<close>
-setup {* Code_Target.add_derived_target ("Haskell_Quickcheck", [(Code_Haskell.target, I)]) *}
+setup \<open>Code_Target.add_derived_target ("Haskell_Quickcheck", [(Code_Haskell.target, I)])\<close>
code_printing
- code_module Typerep \<rightharpoonup> (Haskell_Quickcheck) {*
+ code_module Typerep \<rightharpoonup> (Haskell_Quickcheck) \<open>
data Typerep = Typerep String [Typerep]
-*}
+\<close>
| type_constructor typerep \<rightharpoonup> (Haskell_Quickcheck) "Typerep.Typerep"
| constant Typerep.Typerep \<rightharpoonup> (Haskell_Quickcheck) "Typerep.Typerep"
| type_constructor integer \<rightharpoonup> (Haskell_Quickcheck) "Prelude.Int"
@@ -27,7 +27,7 @@
constant "0::integer" \<rightharpoonup>
(Haskell_Quickcheck) "!(0/ ::/ Prelude.Int)"
-setup {*
+setup \<open>
let
val target = "Haskell_Quickcheck";
fun print _ = Code_Haskell.print_numeral "Prelude.Int";
@@ -35,10 +35,10 @@
Numeral.add_code @{const_name Code_Numeral.Pos} I print target
#> Numeral.add_code @{const_name Code_Numeral.Neg} (op ~) print target
end
-*}
+\<close>
-subsubsection {* Narrowing's deep representation of types and terms *}
+subsubsection \<open>Narrowing's deep representation of types and terms\<close>
datatype (plugins only: code extraction) narrowing_type =
Narrowing_sum_of_products "narrowing_type list list"
@@ -54,7 +54,7 @@
where
"map_cons f (Narrowing_cons ty cs) = Narrowing_cons ty (map (\<lambda>c. f o c) cs)"
-subsubsection {* From narrowing's deep representation of terms to @{theory Code_Evaluation}'s terms *}
+subsubsection \<open>From narrowing's deep representation of terms to @{theory Code_Evaluation}'s terms\<close>
class partial_term_of = typerep +
fixes partial_term_of :: "'a itself => narrowing_term => Code_Evaluation.term"
@@ -62,7 +62,7 @@
lemma partial_term_of_anything: "partial_term_of x nt \<equiv> t"
by (rule eq_reflection) (cases "partial_term_of x nt", cases t, simp)
-subsubsection {* Auxilary functions for Narrowing *}
+subsubsection \<open>Auxilary functions for Narrowing\<close>
consts nth :: "'a list => integer => 'a"
@@ -80,7 +80,7 @@
code_printing constant marker \<rightharpoonup> (Haskell_Quickcheck) "''\\0'"
-subsubsection {* Narrowing's basic operations *}
+subsubsection \<open>Narrowing's basic operations\<close>
type_synonym 'a narrowing = "integer => 'a narrowing_cons"
@@ -136,7 +136,7 @@
split: narrowing_cons.split narrowing_type.split)
qed
-subsubsection {* Narrowing generator type class *}
+subsubsection \<open>Narrowing generator type class\<close>
class narrowing =
fixes narrowing :: "integer => 'a narrowing_cons"
@@ -155,9 +155,9 @@
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))))"
-subsubsection {* class @{text is_testable} *}
+subsubsection \<open>class @{text is_testable}\<close>
-text {* The class @{text is_testable} ensures that all necessary type instances are generated. *}
+text \<open>The class @{text is_testable} ensures that all necessary type instances are generated.\<close>
class is_testable
@@ -170,7 +170,7 @@
"ensure_testable f = f"
-subsubsection {* Defining a simple datatype to represent functions in an incomplete and redundant way *}
+subsubsection \<open>Defining a simple datatype to represent functions in an incomplete and redundant way\<close>
datatype (plugins only: code quickcheck_narrowing extraction) (dead 'a, dead 'b) ffun =
Constant 'b
@@ -193,7 +193,7 @@
hide_type (open) cfun
hide_const (open) Constant eval_cfun Abs_cfun Rep_cfun
-subsubsection {* Setting up the counterexample generator *}
+subsubsection \<open>Setting up the counterexample generator\<close>
ML_file "Tools/Quickcheck/narrowing_generators.ML"
@@ -213,7 +213,7 @@
z = (conv :: _ => _ => unit) in f)"
unfolding Let_def ensure_testable_def ..
-subsection {* Narrowing for sets *}
+subsection \<open>Narrowing for sets\<close>
instantiation set :: (narrowing) narrowing
begin
@@ -224,7 +224,7 @@
end
-subsection {* Narrowing for integers *}
+subsection \<open>Narrowing for integers\<close>
definition drawn_from :: "'a list \<Rightarrow> 'a narrowing_cons"
@@ -319,11 +319,11 @@
(Generated'_Code.Const \"Groups.uminus'_class.uminus\" (mkFunT t t))
(mkNumber (- i)); } in mkNumber)"
-subsection {* The @{text find_unused_assms} command *}
+subsection \<open>The @{text find_unused_assms} command\<close>
ML_file "Tools/Quickcheck/find_unused_assms.ML"
-subsection {* Closing up *}
+subsection \<open>Closing up\<close>
hide_type narrowing_type narrowing_term narrowing_cons property
hide_const map_cons nth error toEnum marker empty Narrowing_cons conv non_empty ensure_testable all exists drawn_from around_zero
--- a/src/HOL/Quickcheck_Random.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Quickcheck_Random.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,6 +1,6 @@
(* Author: Florian Haftmann & Lukas Bulwahn, TU Muenchen *)
-section {* A simple counterexample generator performing random testing *}
+section \<open>A simple counterexample generator performing random testing\<close>
theory Quickcheck_Random
imports Random Code_Evaluation Enum
@@ -9,22 +9,22 @@
notation fcomp (infixl "\<circ>>" 60)
notation scomp (infixl "\<circ>\<rightarrow>" 60)
-setup {* Code_Target.add_derived_target ("Quickcheck", [(Code_Runtime.target, I)]) *}
+setup \<open>Code_Target.add_derived_target ("Quickcheck", [(Code_Runtime.target, I)])\<close>
-subsection {* Catching Match exceptions *}
+subsection \<open>Catching Match exceptions\<close>
axiomatization catch_match :: "'a => 'a => 'a"
code_printing
constant catch_match \<rightharpoonup> (Quickcheck) "((_) handle Match => _)"
-subsection {* The @{text random} class *}
+subsection \<open>The @{text random} class\<close>
class random = typerep +
fixes random :: "natural \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
-subsection {* Fundamental and numeric types*}
+subsection \<open>Fundamental and numeric types\<close>
instantiation bool :: random
begin
@@ -121,9 +121,9 @@
end
-subsection {* Complex generators *}
+subsection \<open>Complex generators\<close>
-text {* Towards @{typ "'a \<Rightarrow> 'b"} *}
+text \<open>Towards @{typ "'a \<Rightarrow> 'b"}\<close>
axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
\<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
@@ -147,7 +147,7 @@
end
-text {* Towards type copies and datatypes *}
+text \<open>Towards type copies and datatypes\<close>
definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a"
where "collapse f = (f \<circ>\<rightarrow> id)"
@@ -203,19 +203,19 @@
shows "random_aux k = rhs k"
using assms by (rule natural.induct)
-subsection {* Deriving random generators for datatypes *}
+subsection \<open>Deriving random generators for datatypes\<close>
ML_file "Tools/Quickcheck/quickcheck_common.ML"
ML_file "Tools/Quickcheck/random_generators.ML"
-subsection {* Code setup *}
+subsection \<open>Code setup\<close>
code_printing
constant random_fun_aux \<rightharpoonup> (Quickcheck) "Random'_Generators.random'_fun"
- -- {* With enough criminal energy this can be abused to derive @{prop False};
+ -- \<open>With enough criminal energy this can be abused to derive @{prop False};
for this reason we use a distinguished target @{text Quickcheck}
- not spoiling the regular trusted code generation *}
+ not spoiling the regular trusted code generation\<close>
code_reserved Quickcheck Random_Generators
--- a/src/HOL/Quotient.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Quotient.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Cezary Kaliszyk and Christian Urban
*)
-section {* Definition of Quotient Types *}
+section \<open>Definition of Quotient Types\<close>
theory Quotient
imports Lifting
@@ -12,12 +12,12 @@
"quotient_definition" :: thy_goal
begin
-text {*
+text \<open>
Basic definition for equivalence relations
that are represented by predicates.
-*}
+\<close>
-text {* Composition of Relations *}
+text \<open>Composition of Relations\<close>
abbreviation
rel_conj :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" (infixr "OOO" 75)
@@ -32,7 +32,7 @@
begin
interpretation lifting_syntax .
-subsection {* Quotient Predicate *}
+subsection \<open>Quotient Predicate\<close>
definition
"Quotient3 R Abs Rep \<longleftrightarrow>
@@ -64,7 +64,7 @@
by blast
lemma Quotient3_rel:
- "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- {* orientation does not loop on rewriting *}
+ "R r r \<and> R s s \<and> Abs r = Abs s \<longleftrightarrow> R r s" -- \<open>orientation does not loop on rewriting\<close>
using a
unfolding Quotient3_def
by blast
@@ -192,12 +192,12 @@
using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]
by simp
-text{*
+text\<open>
In the following theorem R1 can be instantiated with anything,
but we know some of the types of the Rep and Abs functions;
so by solving Quotient assumptions we can get a unique R1 that
will be provable; which is why we need to use @{text apply_rsp} and
- not the primed version *}
+ not the primed version\<close>
lemma apply_rspQ3:
fixes f g::"'a \<Rightarrow> 'c"
@@ -215,7 +215,7 @@
then show ?thesis using assms(2) by (auto intro: apply_rsp')
qed
-subsection {* lemmas for regularisation of ball and bex *}
+subsection \<open>lemmas for regularisation of ball and bex\<close>
lemma ball_reg_eqv:
fixes P :: "'a \<Rightarrow> bool"
@@ -315,7 +315,7 @@
shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)"
using assms by auto
-subsection {* Bounded abstraction *}
+subsection \<open>Bounded abstraction\<close>
definition
Babs :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
@@ -392,7 +392,7 @@
using a unfolding Quotient3_def Bex_def in_respects id_apply comp_def map_fun_def
by metis
-subsection {* @{text Bex1_rel} quantifier *}
+subsection \<open>@{text Bex1_rel} quantifier\<close>
definition
Bex1_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
@@ -506,7 +506,7 @@
apply simp
done
-subsection {* Various respects and preserve lemmas *}
+subsection \<open>Various respects and preserve lemmas\<close>
lemma quot_rel_rsp:
assumes a: "Quotient3 R Abs Rep"
@@ -605,7 +605,7 @@
have "R (SOME x. x \<in> Rep a) x" using r rep some_collect by metis
then have "R x (SOME x. x \<in> Rep a)" using part_equivp_symp[OF equivp] by fast
then show "R (SOME x. x \<in> Rep a) (SOME x. x \<in> Rep a)"
- using part_equivp_transp[OF equivp] by (metis `R (SOME x. x \<in> Rep a) x`)
+ using part_equivp_transp[OF equivp] by (metis \<open>R (SOME x. x \<in> Rep a) x\<close>)
qed
have "Collect (R (SOME x. x \<in> Rep a)) = (Rep a)" by (metis some_collect rep_prop)
then show "Abs (Collect (R (SOME x. x \<in> Rep a))) = a" using rep_inverse by auto
@@ -624,7 +624,7 @@
end
-subsection {* Quotient composition *}
+subsection \<open>Quotient composition\<close>
lemma OOO_quotient3:
fixes R1 :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
@@ -718,7 +718,7 @@
using assms
by (rule OOO_quotient3) auto
-subsection {* Quotient3 to Quotient *}
+subsection \<open>Quotient3 to Quotient\<close>
lemma Quotient3_to_Quotient:
assumes "Quotient3 R Abs Rep"
@@ -744,9 +744,9 @@
show "T = (\<lambda>x y. R x x \<and> Abs x = y)" using T_def equivp_reflp[OF eR] by simp
qed
-subsection {* ML setup *}
+subsection \<open>ML setup\<close>
-text {* Auxiliary data for the quotient package *}
+text \<open>Auxiliary data for the quotient package\<close>
named_theorems quot_equiv "equivalence relation theorems"
and quot_respect "respectfulness theorems"
@@ -763,7 +763,7 @@
lemmas [quot_equiv] = identity_equivp
-text {* Lemmas about simplifying id's. *}
+text \<open>Lemmas about simplifying id's.\<close>
lemmas [id_simps] =
id_def[symmetric]
map_fun_id
@@ -773,22 +773,22 @@
eq_comp_r
vimage_id
-text {* Translation functions for the lifting process. *}
+text \<open>Translation functions for the lifting process.\<close>
ML_file "Tools/Quotient/quotient_term.ML"
-text {* Definitions of the quotient types. *}
+text \<open>Definitions of the quotient types.\<close>
ML_file "Tools/Quotient/quotient_type.ML"
-text {* Definitions for quotient constants. *}
+text \<open>Definitions for quotient constants.\<close>
ML_file "Tools/Quotient/quotient_def.ML"
-text {*
+text \<open>
An auxiliary constant for recording some information
about the lifted theorem in a tactic.
-*}
+\<close>
definition
Quot_True :: "'a \<Rightarrow> bool"
where
@@ -809,55 +809,55 @@
begin
interpretation lifting_syntax .
-text {* Tactics for proving the lifted theorems *}
+text \<open>Tactics for proving the lifted theorems\<close>
ML_file "Tools/Quotient/quotient_tacs.ML"
end
-subsection {* Methods / Interface *}
+subsection \<open>Methods / Interface\<close>
method_setup lifting =
- {* Attrib.thms >> (fn thms => fn ctxt =>
- SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms)) *}
- {* lift theorems to quotient types *}
+ \<open>Attrib.thms >> (fn thms => fn ctxt =>
+ SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms))\<close>
+ \<open>lift theorems to quotient types\<close>
method_setup lifting_setup =
- {* Attrib.thm >> (fn thm => fn ctxt =>
- SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm)) *}
- {* set up the three goals for the quotient lifting procedure *}
+ \<open>Attrib.thm >> (fn thm => fn ctxt =>
+ SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm))\<close>
+ \<open>set up the three goals for the quotient lifting procedure\<close>
method_setup descending =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt [])) *}
- {* decend theorems to the raw level *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt []))\<close>
+ \<open>decend theorems to the raw level\<close>
method_setup descending_setup =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt [])) *}
- {* set up the three goals for the decending theorems *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt []))\<close>
+ \<open>set up the three goals for the decending theorems\<close>
method_setup partiality_descending =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt [])) *}
- {* decend theorems to the raw level *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt []))\<close>
+ \<open>decend theorems to the raw level\<close>
method_setup partiality_descending_setup =
- {* Scan.succeed (fn ctxt =>
- SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt [])) *}
- {* set up the three goals for the decending theorems *}
+ \<open>Scan.succeed (fn ctxt =>
+ SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt []))\<close>
+ \<open>set up the three goals for the decending theorems\<close>
method_setup regularize =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt)) *}
- {* prove the regularization goals from the quotient lifting procedure *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt))\<close>
+ \<open>prove the regularization goals from the quotient lifting procedure\<close>
method_setup injection =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt)) *}
- {* prove the rep/abs injection goals from the quotient lifting procedure *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt))\<close>
+ \<open>prove the rep/abs injection goals from the quotient lifting procedure\<close>
method_setup cleaning =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt)) *}
- {* prove the cleaning goals from the quotient lifting procedure *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt))\<close>
+ \<open>prove the cleaning goals from the quotient lifting procedure\<close>
attribute_setup quot_lifted =
- {* Scan.succeed Quotient_Tacs.lifted_attrib *}
- {* lift theorems to quotient types *}
+ \<open>Scan.succeed Quotient_Tacs.lifted_attrib\<close>
+ \<open>lift theorems to quotient types\<close>
no_notation
rel_conj (infixr "OOO" 75)
--- a/src/HOL/Random.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Random.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* A HOL random engine *}
+section \<open>A HOL random engine\<close>
theory Random
imports List Groups_List
@@ -11,7 +11,7 @@
notation scomp (infixl "\<circ>\<rightarrow>" 60)
-subsection {* Auxiliary functions *}
+subsection \<open>Auxiliary functions\<close>
fun log :: "natural \<Rightarrow> natural \<Rightarrow> natural" where
"log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
@@ -23,7 +23,7 @@
"minus_shift r k l = (if k < l then r + k - l else k - l)"
-subsection {* Random seeds *}
+subsection \<open>Random seeds\<close>
type_synonym seed = "natural \<times> natural"
@@ -45,7 +45,7 @@
in ((v'', w'), (v', w'')))"
-subsection {* Base selectors *}
+subsection \<open>Base selectors\<close>
fun iterate :: "natural \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
"iterate k f x = (if k = 0 then Pair x else f x \<circ>\<rightarrow> iterate (k - 1) f)"
@@ -136,12 +136,12 @@
qed
-subsection {* @{text ML} interface *}
+subsection \<open>@{text ML} interface\<close>
code_reflect Random_Engine
functions range select select_weight
-ML {*
+ML \<open>
structure Random_Engine =
struct
@@ -177,7 +177,7 @@
end;
end;
-*}
+\<close>
hide_type (open) seed
hide_const (open) inc_shift minus_shift log "next" split_seed
--- a/src/HOL/Random_Pred.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Random_Pred.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,7 +1,7 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* The Random-Predicate Monad *}
+section \<open>The Random-Predicate Monad\<close>
theory Random_Pred
imports Quickcheck_Random
--- a/src/HOL/Random_Sequence.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Random_Sequence.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,7 +1,7 @@
(* Author: Lukas Bulwahn, TU Muenchen *)
-section {* Various kind of sequences inside the random monad *}
+section \<open>Various kind of sequences inside the random monad\<close>
theory Random_Sequence
imports Random_Pred
--- a/src/HOL/Rat.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Rat.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,15 +2,15 @@
Author: Markus Wenzel, TU Muenchen
*)
-section {* Rational numbers *}
+section \<open>Rational numbers\<close>
theory Rat
imports GCD Archimedean_Field
begin
-subsection {* Rational numbers as quotient *}
+subsection \<open>Rational numbers as quotient\<close>
-subsubsection {* Construction of the type of rational numbers *}
+subsubsection \<open>Construction of the type of rational numbers\<close>
definition
ratrel :: "(int \<times> int) \<Rightarrow> (int \<times> int) \<Rightarrow> bool" where
@@ -52,7 +52,7 @@
lemma Domainp_cr_rat [transfer_domain_rule]: "Domainp pcr_rat = (\<lambda>x. snd x \<noteq> 0)"
by (simp add: rat.domain_eq)
-subsubsection {* Representation and basic operations *}
+subsubsection \<open>Representation and basic operations\<close>
lift_definition Fract :: "int \<Rightarrow> int \<Rightarrow> rat"
is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)"
@@ -72,12 +72,12 @@
by transfer simp
let ?a = "a div gcd a b"
let ?b = "b div gcd a b"
- from `b \<noteq> 0` have "?b * gcd a b = b"
+ from \<open>b \<noteq> 0\<close> have "?b * gcd a b = b"
by simp
- with `b \<noteq> 0` have "?b \<noteq> 0" by fastforce
- from `q = Fract a b` `b \<noteq> 0` `?b \<noteq> 0` have q: "q = Fract ?a ?b"
+ with \<open>b \<noteq> 0\<close> have "?b \<noteq> 0" by fastforce
+ from \<open>q = Fract a b\<close> \<open>b \<noteq> 0\<close> \<open>?b \<noteq> 0\<close> have q: "q = Fract ?a ?b"
by (simp add: eq_rat dvd_div_mult mult.commute [of a])
- from `b \<noteq> 0` have coprime: "coprime ?a ?b"
+ from \<open>b \<noteq> 0\<close> have coprime: "coprime ?a ?b"
by (auto intro: div_gcd_coprime_int)
show C proof (cases "b > 0")
case True
@@ -90,7 +90,7 @@
case False
note assms
moreover have "q = Fract (- ?a) (- ?b)" unfolding q by transfer simp
- moreover from False `b \<noteq> 0` have "- ?b > 0" by (simp add: pos_imp_zdiv_neg_iff)
+ moreover from False \<open>b \<noteq> 0\<close> have "- ?b > 0" by (simp add: pos_imp_zdiv_neg_iff)
moreover from coprime have "coprime (- ?a) (- ?b)" by simp
ultimately show C .
qed
@@ -240,11 +240,11 @@
case False
then obtain a b where "q = Fract a b" and "b > 0" and "coprime a b" by (cases q) auto
with False have "0 \<noteq> Fract a b" by simp
- with `b > 0` have "a \<noteq> 0" by (simp add: Zero_rat_def eq_rat)
- with Fract `q = Fract a b` `b > 0` `coprime a b` show C by blast
+ with \<open>b > 0\<close> have "a \<noteq> 0" by (simp add: Zero_rat_def eq_rat)
+ with Fract \<open>q = Fract a b\<close> \<open>b > 0\<close> \<open>coprime a b\<close> show C by blast
qed
-subsubsection {* Function @{text normalize} *}
+subsubsection \<open>Function @{text normalize}\<close>
lemma Fract_coprime: "Fract (a div gcd a b) (b div gcd a b) = Fract a b"
proof (cases "b = 0")
@@ -302,9 +302,9 @@
"q < 0 \<Longrightarrow> normalize (p, q) = normalize (- p, - q)"
by (simp add: normalize_def Let_def dvd_div_neg dvd_neg_div)
-text{*
+text\<open>
Decompose a fraction into normalized, i.e. coprime numerator and denominator:
-*}
+\<close>
definition quotient_of :: "rat \<Rightarrow> int \<times> int" where
"quotient_of x = (THE pair. x = Fract (fst pair) (snd pair) &
@@ -327,11 +327,11 @@
case False
with Fract Fract' have *: "c * b = a * d" and "c \<noteq> 0" by (auto simp add: eq_rat)
then have "c * b > 0 \<longleftrightarrow> a * d > 0" by auto
- with `b > 0` `d > 0` have "a > 0 \<longleftrightarrow> c > 0" by (simp add: zero_less_mult_iff)
- with `a \<noteq> 0` `c \<noteq> 0` have sgn: "sgn a = sgn c" by (auto simp add: not_less)
- from `coprime a b` `coprime c d` have "\<bar>a\<bar> * \<bar>d\<bar> = \<bar>c\<bar> * \<bar>b\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> \<bar>d\<bar> = \<bar>b\<bar>"
+ with \<open>b > 0\<close> \<open>d > 0\<close> have "a > 0 \<longleftrightarrow> c > 0" by (simp add: zero_less_mult_iff)
+ with \<open>a \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have sgn: "sgn a = sgn c" by (auto simp add: not_less)
+ from \<open>coprime a b\<close> \<open>coprime c d\<close> have "\<bar>a\<bar> * \<bar>d\<bar> = \<bar>c\<bar> * \<bar>b\<bar> \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> \<bar>d\<bar> = \<bar>b\<bar>"
by (simp add: coprime_crossproduct_int)
- with `b > 0` `d > 0` have "\<bar>a\<bar> * d = \<bar>c\<bar> * b \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> d = b" by simp
+ with \<open>b > 0\<close> \<open>d > 0\<close> have "\<bar>a\<bar> * d = \<bar>c\<bar> * b \<longleftrightarrow> \<bar>a\<bar> = \<bar>c\<bar> \<and> d = b" by simp
then have "a * sgn a * d = c * sgn c * b \<longleftrightarrow> a * sgn a = c * sgn c \<and> d = b" by (simp add: abs_sgn)
with sgn * show ?thesis by (auto simp add: sgn_0_0)
qed
@@ -387,7 +387,7 @@
by (auto simp add: quotient_of_inject)
-subsubsection {* Various *}
+subsubsection \<open>Various\<close>
lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l"
by (simp add: Fract_of_int_eq [symmetric])
@@ -404,7 +404,7 @@
thus ?thesis using Fract_of_int_quotient by simp
qed
-subsubsection {* The ordered field of rational numbers *}
+subsubsection \<open>The ordered field of rational numbers\<close>
lift_definition positive :: "rat \<Rightarrow> bool"
is "\<lambda>x. 0 < fst x * snd x"
@@ -418,7 +418,7 @@
hence "0 < a * b * d\<^sup>2 \<longleftrightarrow> 0 < c * d * b\<^sup>2"
by simp
thus "0 < a * b \<longleftrightarrow> 0 < c * d"
- using `b \<noteq> 0` and `d \<noteq> 0`
+ using \<open>b \<noteq> 0\<close> and \<open>d \<noteq> 0\<close>
by (simp add: zero_less_mult_iff)
qed
@@ -570,7 +570,7 @@
by (simp add: One_rat_def mult_le_cancel_right)
-subsubsection {* Rationals are an Archimedean field *}
+subsubsection \<open>Rationals are an Archimedean field\<close>
lemma rat_floor_lemma:
shows "of_int (a div b) \<le> Fract a b \<and> Fract a b < of_int (a div b + 1)"
@@ -613,9 +613,9 @@
by (simp add: Fract_of_int_quotient floor_divide_of_int_eq)
-subsection {* Linear arithmetic setup *}
+subsection \<open>Linear arithmetic setup\<close>
-declaration {*
+declaration \<open>
K (Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2]
(* not needed because x < (y::nat) can be rewritten as Suc x <= y: of_nat_less_iff RS iffD2 *)
#> Lin_Arith.add_inj_thms [@{thm of_int_le_iff} RS iffD2, @{thm of_int_eq_iff} RS iffD2]
@@ -632,10 +632,10 @@
#> Lin_Arith.add_simprocs Numeral_Simprocs.field_divide_cancel_numeral_factor
#> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat => rat"})
#> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int => rat"}))
-*}
+\<close>
-subsection {* Embedding from Rationals to other Fields *}
+subsection \<open>Embedding from Rationals to other Fields\<close>
context field_char_0
begin
@@ -735,7 +735,7 @@
using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq)
show "(of_rat (Fract a b) :: 'a::linordered_field) < of_rat (Fract c d)
\<longleftrightarrow> Fract a b < Fract c d"
- using not_zero `b * d > 0`
+ using not_zero \<open>b * d > 0\<close>
by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult)
qed
@@ -775,7 +775,7 @@
(simp add: of_rat_rat Fract_of_int_eq [symmetric])
qed
-text{*Collapse nested embeddings*}
+text\<open>Collapse nested embeddings\<close>
lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n"
by (induct n) (simp_all add: of_rat_add)
@@ -803,7 +803,7 @@
where
"rat_of_int \<equiv> of_int"
-subsection {* The Set of Rational Numbers *}
+subsection \<open>The Set of Rational Numbers\<close>
context field_char_0
begin
@@ -909,7 +909,7 @@
assumes "q \<in> \<rat>"
obtains (of_rat) r where "q = of_rat r"
proof -
- from `q \<in> \<rat>` have "q \<in> range of_rat" unfolding Rats_def .
+ from \<open>q \<in> \<rat>\<close> have "q \<in> range of_rat" unfolding Rats_def .
then obtain r where "q = of_rat r" ..
then show thesis ..
qed
@@ -921,9 +921,9 @@
lemma Rats_infinite: "\<not> finite \<rat>"
by (auto dest!: finite_imageD simp: inj_on_def infinite_UNIV_char_0 Rats_def)
-subsection {* Implementation of rational numbers as pairs of integers *}
+subsection \<open>Implementation of rational numbers as pairs of integers\<close>
-text {* Formal constructor *}
+text \<open>Formal constructor\<close>
definition Frct :: "int \<times> int \<Rightarrow> rat" where
[simp]: "Frct p = Fract (fst p) (snd p)"
@@ -933,7 +933,7 @@
by (cases q) (auto intro: quotient_of_eq)
-text {* Numerals *}
+text \<open>Numerals\<close>
declare quotient_of_Fract [code abstract]
@@ -967,7 +967,7 @@
by (simp_all add: Fract_of_int_quotient)
-text {* Operations *}
+text \<open>Operations\<close>
lemma rat_zero_code [code abstract]:
"quotient_of 0 = (0, 1)"
@@ -1052,7 +1052,7 @@
by (cases p) (simp add: quotient_of_Fract of_rat_rat)
-text {* Quickcheck *}
+text \<open>Quickcheck\<close>
definition (in term_syntax)
valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
@@ -1126,9 +1126,9 @@
end
-subsection {* Setup for Nitpick *}
+subsection \<open>Setup for Nitpick\<close>
-declaration {*
+declaration \<open>
Nitpick_HOL.register_frac_type @{type_name rat}
[(@{const_name zero_rat_inst.zero_rat}, @{const_name Nitpick.zero_frac}),
(@{const_name one_rat_inst.one_rat}, @{const_name Nitpick.one_frac}),
@@ -1139,7 +1139,7 @@
(@{const_name ord_rat_inst.less_rat}, @{const_name Nitpick.less_frac}),
(@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}),
(@{const_name field_char_0_class.of_rat}, @{const_name Nitpick.of_frac})]
-*}
+\<close>
lemmas [nitpick_unfold] = inverse_rat_inst.inverse_rat
one_rat_inst.one_rat ord_rat_inst.less_rat
@@ -1147,11 +1147,11 @@
uminus_rat_inst.uminus_rat zero_rat_inst.zero_rat
-subsection {* Float syntax *}
+subsection \<open>Float syntax\<close>
syntax "_Float" :: "float_const \<Rightarrow> 'a" ("_")
-parse_translation {*
+parse_translation \<open>
let
fun mk_frac str =
let
@@ -1165,14 +1165,14 @@
| float_tr [t as Const (str, _)] = mk_frac str
| float_tr ts = raise TERM ("float_tr", ts);
in [(@{syntax_const "_Float"}, K float_tr)] end
-*}
+\<close>
-text{* Test: *}
+text\<open>Test:\<close>
lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::rat)"
by simp
-subsection {* Hiding implementation details *}
+subsection \<open>Hiding implementation details\<close>
hide_const (open) normalize positive
--- a/src/HOL/Real.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Real.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,20 +7,20 @@
Construction of Cauchy Reals by Brian Huffman, 2010
*)
-section {* Development of the Reals using Cauchy Sequences *}
+section \<open>Development of the Reals using Cauchy Sequences\<close>
theory Real
imports Rat Conditionally_Complete_Lattices
begin
-text {*
+text \<open>
This theory contains a formalization of the real numbers as
equivalence classes of Cauchy sequences of rationals. See
@{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative
construction using Dedekind cuts.
-*}
+\<close>
-subsection {* Preliminary lemmas *}
+subsection \<open>Preliminary lemmas\<close>
lemma add_diff_add:
fixes a b c d :: "'a::ab_group_add"
@@ -52,7 +52,7 @@
show "r = r/2 + r/2" by simp
qed
-subsection {* Sequences that converge to zero *}
+subsection \<open>Sequences that converge to zero\<close>
definition
vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
@@ -120,7 +120,7 @@
thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
qed
-subsection {* Cauchy sequences *}
+subsection \<open>Cauchy sequences\<close>
definition
cauchy :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
@@ -196,7 +196,7 @@
also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
by (rule abs_triangle_ineq)
also have "\<dots> < Max (abs ` X ` {..k}) + 1"
- by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
+ by (rule add_le_less_mono, simp, simp add: k \<open>k \<le> n\<close>)
finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
qed
qed
@@ -218,7 +218,7 @@
show "0 < v/b" using v b(1) by simp
show "0 < u/a" using u a(1) by simp
show "r = a * (u/a) + (v/b) * b"
- using a(1) b(1) `r = u + v` by simp
+ using a(1) b(1) \<open>r = u + v\<close> by simp
qed
obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
using cauchyD [OF X s] ..
@@ -248,17 +248,17 @@
obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
using nz unfolding vanishes_def by (auto simp add: not_less)
obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
- using `0 < r` by (rule obtain_pos_sum)
+ using \<open>0 < r\<close> by (rule obtain_pos_sum)
obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
using cauchyD [OF X s] ..
obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
using r by fast
have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
- using i `i \<le> k` by auto
+ using i \<open>i \<le> k\<close> by auto
have "X k \<le> - r \<or> r \<le> X k"
- using `r \<le> \<bar>X k\<bar>` by auto
+ using \<open>r \<le> \<bar>X k\<bar>\<close> by auto
hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
- unfolding `r = s + t` using k by auto
+ unfolding \<open>r = s + t\<close> using k by auto
hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
using t by auto
@@ -282,7 +282,7 @@
from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
proof
- show "0 < b * r * b" by (simp add: `0 < r` b)
+ show "0 < b * r * b" by (simp add: \<open>0 < r\<close> b)
show "r = inverse b * (b * r * b) * inverse b"
using b by simp
qed
@@ -334,13 +334,13 @@
apply (intro mult_strict_mono' less_imp_inverse_less)
apply (simp_all add: a b i j k n)
done
- also note `inverse a * s * inverse b = r`
+ also note \<open>inverse a * s * inverse b = r\<close>
finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
qed
thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
qed
-subsection {* Equivalence relation on Cauchy sequences *}
+subsection \<open>Equivalence relation on Cauchy sequences\<close>
definition realrel :: "(nat \<Rightarrow> rat) \<Rightarrow> (nat \<Rightarrow> rat) \<Rightarrow> bool"
where "realrel = (\<lambda>X Y. cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n))"
@@ -368,7 +368,7 @@
by (fast intro: part_equivpI symp_realrel transp_realrel
realrel_refl cauchy_const)
-subsection {* The field of real numbers *}
+subsection \<open>The field of real numbers\<close>
quotient_type real = "nat \<Rightarrow> rat" / partial: realrel
morphisms rep_real Real
@@ -509,7 +509,7 @@
end
-subsection {* Positive reals *}
+subsection \<open>Positive reals\<close>
lift_definition positive :: "real \<Rightarrow> bool"
is "\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
@@ -522,7 +522,7 @@
then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
by fast
obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
- using `0 < r` by (rule obtain_pos_sum)
+ using \<open>0 < r\<close> by (rule obtain_pos_sum)
obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
using vanishesD [OF XY s] ..
have "\<forall>n\<ge>max i j. t < Y n"
@@ -684,7 +684,7 @@
end
-subsection {* Completeness *}
+subsection \<open>Completeness\<close>
lemma not_positive_Real:
assumes X: "cauchy X"
@@ -878,14 +878,14 @@
proof
fix x assume "x \<in> S"
then show "x \<le> Real B"
- using PB [unfolded P_def] `cauchy B`
+ using PB [unfolded P_def] \<open>cauchy B\<close>
by (simp add: le_RealI)
qed
have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
apply clarify
apply (erule contrapos_pp)
apply (simp add: not_le)
- apply (drule less_RealD [OF `cauchy A`], clarify)
+ apply (drule less_RealD [OF \<open>cauchy A\<close>], clarify)
apply (subgoal_tac "\<not> P (A n)")
apply (simp add: P_def not_le, clarify)
apply (erule rev_bexI)
@@ -910,7 +910,7 @@
thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
qed
hence 3: "Real B = Real A"
- by (simp add: eq_Real `cauchy A` `cauchy B` width)
+ by (simp add: eq_Real \<open>cauchy A\<close> \<open>cauchy B\<close> width)
show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
using 1 2 3 by (rule_tac x="Real B" in exI, simp)
qed
@@ -918,7 +918,7 @@
instantiation real :: linear_continuum
begin
-subsection{*Supremum of a set of reals*}
+subsection\<open>Supremum of a set of reals\<close>
definition "Sup X = (LEAST z::real. \<forall>x\<in>X. x \<le> z)"
definition "Inf (X::real set) = - Sup (uminus ` X)"
@@ -954,7 +954,7 @@
end
-subsection {* Hiding implementation details *}
+subsection \<open>Hiding implementation details\<close>
hide_const (open) vanishes cauchy positive Real
@@ -965,10 +965,10 @@
lifting_update real.lifting
lifting_forget real.lifting
-subsection{*More Lemmas*}
+subsection\<open>More Lemmas\<close>
-text {* BH: These lemmas should not be necessary; they should be
-covered by existing simp rules and simplification procedures. *}
+text \<open>BH: These lemmas should not be necessary; they should be
+covered by existing simp rules and simplification procedures.\<close>
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
by simp (* solved by linordered_ring_less_cancel_factor simproc *)
@@ -980,7 +980,7 @@
by simp (* solved by linordered_ring_le_cancel_factor simproc *)
-subsection {* Embedding numbers into the Reals *}
+subsection \<open>Embedding numbers into the Reals\<close>
abbreviation
real_of_nat :: "nat \<Rightarrow> real"
@@ -1167,7 +1167,7 @@
unfolding real_of_int_def by (rule Ints_of_int)
-subsection{*Embedding the Naturals into the Reals*}
+subsection\<open>Embedding the Naturals into the Reals\<close>
lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
by (simp add: real_of_nat_def)
@@ -1306,7 +1306,7 @@
lemma Ints_real_of_nat [simp]: "real (n::nat) \<in> Ints"
unfolding real_of_nat_def by (rule Ints_of_nat)
-subsection {* The Archimedean Property of the Reals *}
+subsection \<open>The Archimedean Property of the Reals\<close>
theorem reals_Archimedean:
assumes x_pos: "0 < x"
@@ -1320,11 +1320,11 @@
lemma reals_Archimedean3:
assumes x_greater_zero: "0 < x"
shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
- unfolding real_of_nat_def using `0 < x`
+ unfolding real_of_nat_def using \<open>0 < x\<close>
by (auto intro: ex_less_of_nat_mult)
-subsection{* Rationals *}
+subsection\<open>Rationals\<close>
lemma Rats_real_nat[simp]: "real(n::nat) \<in> \<rat>"
by (simp add: real_eq_of_nat)
@@ -1338,7 +1338,7 @@
then obtain r where "x = of_rat r" unfolding Rats_def ..
have "of_rat r : ?S"
by (cases r)(auto simp add:of_rat_rat real_eq_of_int)
- thus "x : ?S" using `x = of_rat r` by simp
+ thus "x : ?S" using \<open>x = of_rat r\<close> by simp
qed
next
show "?S \<subseteq> \<rat>"
@@ -1362,7 +1362,7 @@
thus ?thesis by blast
next
assume "~ j>0"
- hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using `j\<noteq>0`
+ hence "real i/real j = real(-i)/real(nat(-j)) \<and> 0<nat(-j)" using \<open>j\<noteq>0\<close>
by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat)
thus ?thesis by blast
qed
@@ -1377,12 +1377,12 @@
obtains m n :: nat
where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd m n = 1"
proof -
- from `x \<in> \<rat>` obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
+ from \<open>x \<in> \<rat>\<close> obtain i::int and n::nat where "n \<noteq> 0" and "x = real i / real n"
by(auto simp add: Rats_eq_int_div_nat)
hence "\<bar>x\<bar> = real(nat(abs i)) / real n" by simp
then obtain m :: nat where x_rat: "\<bar>x\<bar> = real m / real n" by blast
let ?gcd = "gcd m n"
- from `n\<noteq>0` have gcd: "?gcd \<noteq> 0" by simp
+ from \<open>n\<noteq>0\<close> have gcd: "?gcd \<noteq> 0" by simp
let ?k = "m div ?gcd"
let ?l = "n div ?gcd"
let ?gcd' = "gcd ?k ?l"
@@ -1390,7 +1390,7 @@
by (rule dvd_mult_div_cancel)
have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
by (rule dvd_mult_div_cancel)
- from `n \<noteq> 0` and gcd_l
+ from \<open>n \<noteq> 0\<close> and gcd_l
have "?gcd * ?l \<noteq> 0" by simp
then have "?l \<noteq> 0" by (blast dest!: mult_not_zero)
moreover
@@ -1414,17 +1414,17 @@
ultimately show ?thesis ..
qed
-subsection{*Density of the Rational Reals in the Reals*}
+subsection\<open>Density of the Rational Reals in the Reals\<close>
-text{* This density proof is due to Stefan Richter and was ported by TN. The
+text\<open>This density proof is due to Stefan Richter and was ported by TN. The
original source is \emph{Real Analysis} by H.L. Royden.
-It employs the Archimedean property of the reals. *}
+It employs the Archimedean property of the reals.\<close>
lemma Rats_dense_in_real:
fixes x :: real
assumes "x < y" shows "\<exists>r\<in>\<rat>. x < r \<and> r < y"
proof -
- from `x<y` have "0 < y-x" by simp
+ from \<open>x<y\<close> have "0 < y-x" by simp
with reals_Archimedean obtain q::nat
where q: "inverse (real q) < y-x" and "0 < q" by auto
def p \<equiv> "ceiling (y * real q) - 1"
@@ -1432,11 +1432,11 @@
from q have "x < y - inverse (real q)" by simp
also have "y - inverse (real q) \<le> r"
unfolding r_def p_def
- by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`)
+ by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling \<open>0 < q\<close>)
finally have "x < r" .
moreover have "r < y"
unfolding r_def p_def
- by (simp add: divide_less_eq diff_less_eq `0 < q`
+ by (simp add: divide_less_eq diff_less_eq \<open>0 < q\<close>
less_ceiling_iff [symmetric])
moreover from r_def have "r \<in> \<rat>" by simp
ultimately show ?thesis by fast
@@ -1446,18 +1446,18 @@
fixes x y :: real
assumes "x < y"
shows "\<exists>q :: rat. x < of_rat q \<and> of_rat q < y"
-using Rats_dense_in_real [OF `x < y`]
+using Rats_dense_in_real [OF \<open>x < y\<close>]
by (auto elim: Rats_cases)
-subsection{*Numerals and Arithmetic*}
+subsection\<open>Numerals and Arithmetic\<close>
lemma [code_abbrev]:
"real_of_int (numeral k) = numeral k"
"real_of_int (- numeral k) = - numeral k"
by simp_all
-text{*Collapse applications of @{const real} to @{const numeral}*}
+text\<open>Collapse applications of @{const real} to @{const numeral}\<close>
lemma real_numeral [simp]:
"real (numeral v :: int) = numeral v"
"real (- numeral v :: int) = - numeral v"
@@ -1467,7 +1467,7 @@
"real (numeral v :: nat) = numeral v"
by (simp add: real_of_nat_def)
-declaration {*
+declaration \<open>
K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2]
(* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *)
#> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2]
@@ -1482,14 +1482,14 @@
#> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"})
#> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat \<Rightarrow> real"})
#> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \<Rightarrow> real"}))
-*}
+\<close>
-subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
+subsection\<open>Simprules combining x+y and 0: ARE THEY NEEDED?\<close>
lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)"
by arith
-text {* FIXME: redundant with @{text add_eq_0_iff} below *}
+text \<open>FIXME: redundant with @{text add_eq_0_iff} below\<close>
lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)"
by auto
@@ -1505,9 +1505,9 @@
lemma real_0_le_add_iff: "((0::real) \<le> x+y) = (-x \<le> y)"
by auto
-subsection {* Lemmas about powers *}
+subsection \<open>Lemmas about powers\<close>
-text {* FIXME: declare this in Rings.thy or not at all *}
+text \<open>FIXME: declare this in Rings.thy or not at all\<close>
declare abs_mult_self [simp]
(* used by Import/HOL/real.imp *)
@@ -1517,24 +1517,24 @@
lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
by (simp add: of_nat_less_two_power real_of_nat_def)
-text {* TODO: no longer real-specific; rename and move elsewhere *}
+text \<open>TODO: no longer real-specific; rename and move elsewhere\<close>
lemma realpow_Suc_le_self:
fixes r :: "'a::linordered_semidom"
shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
by (insert power_decreasing [of 1 "Suc n" r], simp)
-text {* TODO: no longer real-specific; rename and move elsewhere *}
+text \<open>TODO: no longer real-specific; rename and move elsewhere\<close>
lemma realpow_minus_mult:
fixes x :: "'a::monoid_mult"
shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
by (simp add: power_Suc power_commutes split add: nat_diff_split)
-text {* FIXME: declare this [simp] for all types, or not at all *}
+text \<open>FIXME: declare this [simp] for all types, or not at all\<close>
lemma real_two_squares_add_zero_iff [simp]:
"(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
by (rule sum_squares_eq_zero_iff)
-text {* FIXME: declare this [simp] for all types, or not at all *}
+text \<open>FIXME: declare this [simp] for all types, or not at all\<close>
lemma realpow_two_sum_zero_iff [simp]:
"(x\<^sup>2 + y\<^sup>2 = (0::real)) = (x = 0 & y = 0)"
by (rule sum_power2_eq_zero_iff)
@@ -1606,7 +1606,7 @@
unfolding real_of_int_le_iff[symmetric] by simp
-subsection{*Density of the Reals*}
+subsection\<open>Density of the Reals\<close>
lemma real_lbound_gt_zero:
"[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
@@ -1615,7 +1615,7 @@
done
-text{*Similar results are proved in @{text Fields}*}
+text\<open>Similar results are proved in @{text Fields}\<close>
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
by auto
@@ -1625,7 +1625,7 @@
lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
by simp
-subsection{*Absolute Value Function for the Reals*}
+subsection\<open>Absolute Value Function for the Reals\<close>
lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
by (simp add: abs_if)
@@ -1647,7 +1647,7 @@
by simp
-subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
+subsection\<open>Floor and Ceiling Functions from the Reals to the Integers\<close>
(* FIXME: theorems for negative numerals *)
lemma numeral_less_real_of_int_iff [simp]:
@@ -1776,10 +1776,10 @@
then have "(j - i div b) * b \<le> i mod b" "i mod b < ((j - i div b) + 1) * b"
by (auto simp: field_simps)
then have "(j - i div b) * b < 1 * b" "0 * b < ((j - i div b) + 1) * b"
- using pos_mod_bound[OF `0<b`, of i] pos_mod_sign[OF `0<b`, of i] by linarith+
+ using pos_mod_bound[OF \<open>0<b\<close>, of i] pos_mod_sign[OF \<open>0<b\<close>, of i] by linarith+
then have "j = i div b"
- using `0 < b` unfolding mult_less_cancel_right by auto }
- with `0 < b` show ?thesis
+ using \<open>0 < b\<close> unfolding mult_less_cancel_right by auto }
+ with \<open>0 < b\<close> show ?thesis
by (auto split: floor_split simp: field_simps)
qed auto
@@ -1856,8 +1856,8 @@
"\<lceil>- (numeral a / numeral b :: real)\<rceil> = - (numeral a div numeral b)"
using ceiling_divide_eq_div[of "- numeral a" "numeral b"] by simp
-text{* The following lemmas are remnants of the erstwhile functions natfloor
-and natceiling. *}
+text\<open>The following lemmas are remnants of the erstwhile functions natfloor
+and natceiling.\<close>
lemma nat_floor_neg: "(x::real) <= 0 ==> nat(floor x) = 0"
by linarith
@@ -1886,7 +1886,7 @@
apply simp
done
-subsection {* Exponentiation with floor *}
+subsection \<open>Exponentiation with floor\<close>
lemma floor_power:
assumes "x = real (floor x)"
@@ -1918,9 +1918,9 @@
by (metis ceiling_of_int of_int_numeral of_int_power)
-subsection {* Implementation of rational real numbers *}
+subsection \<open>Implementation of rational real numbers\<close>
-text {* Formal constructor *}
+text \<open>Formal constructor\<close>
definition Ratreal :: "rat \<Rightarrow> real" where
[code_abbrev, simp]: "Ratreal = of_rat"
@@ -1928,7 +1928,7 @@
code_datatype Ratreal
-text {* Numerals *}
+text \<open>Numerals\<close>
lemma [code_abbrev]:
"(of_rat (of_int a) :: real) = of_int a"
@@ -1962,7 +1962,7 @@
by (simp_all add: of_rat_divide of_rat_minus)
-text {* Operations *}
+text \<open>Operations\<close>
lemma zero_real_code [code]:
"0 = Ratreal 0"
@@ -2018,7 +2018,7 @@
by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code)
-text {* Quickcheck *}
+text \<open>Quickcheck\<close>
definition (in term_syntax)
valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
@@ -2071,9 +2071,9 @@
end
-subsection {* Setup for Nitpick *}
+subsection \<open>Setup for Nitpick\<close>
-declaration {*
+declaration \<open>
Nitpick_HOL.register_frac_type @{type_name real}
[(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}),
(@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}),
@@ -2083,7 +2083,7 @@
(@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
(@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}),
(@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
-*}
+\<close>
lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real
ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
@@ -2091,7 +2091,7 @@
zero_real_inst.zero_real
-subsection {* Setup for SMT *}
+subsection \<open>Setup for SMT\<close>
ML_file "Tools/SMT/smt_real.ML"
ML_file "Tools/SMT/z3_real.ML"
--- a/src/HOL/Real_Vector_Spaces.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Johannes Hölzl
*)
-section {* Vector Spaces and Algebras over the Reals *}
+section \<open>Vector Spaces and Algebras over the Reals\<close>
theory Real_Vector_Spaces
imports Real Topological_Spaces
begin
-subsection {* Locale for additive functions *}
+subsection \<open>Locale for additive functions\<close>
locale additive =
fixes f :: "'a::ab_group_add \<Rightarrow> 'b::ab_group_add"
@@ -43,7 +43,7 @@
end
-subsection {* Vector spaces *}
+subsection \<open>Vector spaces\<close>
locale vector_space =
fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b"
@@ -131,7 +131,7 @@
end
-subsection {* Real vector spaces *}
+subsection \<open>Real vector spaces\<close>
class scaleR =
fixes scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*\<^sub>R" 75)
@@ -159,7 +159,7 @@
apply (rule scaleR_one)
done
-text {* Recover original theorem names *}
+text \<open>Recover original theorem names\<close>
lemmas scaleR_left_commute = real_vector.scale_left_commute
lemmas scaleR_zero_left = real_vector.scale_zero_left
@@ -176,7 +176,7 @@
lemmas scaleR_cancel_left = real_vector.scale_cancel_left
lemmas scaleR_cancel_right = real_vector.scale_cancel_right
-text {* Legacy names *}
+text \<open>Legacy names\<close>
lemmas scaleR_left_distrib = scaleR_add_left
lemmas scaleR_right_distrib = scaleR_add_right
@@ -229,8 +229,8 @@
done
-subsection {* Embedding of the Reals into any @{text real_algebra_1}:
-@{term of_real} *}
+subsection \<open>Embedding of the Reals into any @{text real_algebra_1}:
+@{term of_real}\<close>
definition
of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
@@ -303,7 +303,7 @@
by (simp add: of_real_def)
qed
-text{*Collapse nested embeddings*}
+text\<open>Collapse nested embeddings\<close>
lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n"
by (induct n) auto
@@ -322,7 +322,7 @@
lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w"
using of_real_of_int_eq [of "- numeral w"] by simp
-text{*Every real algebra has characteristic zero*}
+text\<open>Every real algebra has characteristic zero\<close>
instance real_algebra_1 < ring_char_0
proof
@@ -333,7 +333,7 @@
instance real_field < field_char_0 ..
-subsection {* The Set of Real Numbers *}
+subsection \<open>The Set of Real Numbers\<close>
definition Reals :: "'a::real_algebra_1 set" where
"Reals = range of_real"
@@ -439,7 +439,7 @@
obtains (of_real) r where "q = of_real r"
unfolding Reals_def
proof -
- from `q \<in> \<real>` have "q \<in> range of_real" unfolding Reals_def .
+ from \<open>q \<in> \<real>\<close> have "q \<in> range of_real" unfolding Reals_def .
then obtain r where "q = of_real r" ..
then show thesis ..
qed
@@ -468,7 +468,7 @@
"q \<in> \<real> \<Longrightarrow> (\<And>r. P (of_real r)) \<Longrightarrow> P q"
by (rule Reals_cases) auto
-subsection {* Ordered real vector spaces *}
+subsection \<open>Ordered real vector spaces\<close>
class ordered_real_vector = real_vector + ordered_ab_group_add +
assumes scaleR_left_mono: "x \<le> y \<Longrightarrow> 0 \<le> a \<Longrightarrow> a *\<^sub>R x \<le> a *\<^sub>R y"
@@ -583,16 +583,16 @@
assume "0 < a"
with lhs have "inverse a *\<^sub>R 0 \<le> inverse a *\<^sub>R (a *\<^sub>R b)"
by (intro scaleR_mono) auto
- hence ?rhs using `0 < a`
+ hence ?rhs using \<open>0 < a\<close>
by simp
} moreover {
assume "0 > a"
with lhs have "- inverse a *\<^sub>R 0 \<le> - inverse a *\<^sub>R (a *\<^sub>R b)"
by (intro scaleR_mono) auto
- hence ?rhs using `0 > a`
+ hence ?rhs using \<open>0 > a\<close>
by simp
- } ultimately show ?rhs using `a \<noteq> 0` by arith
- qed (auto simp: not_le `a \<noteq> 0` intro!: split_scaleR_pos_le)
+ } ultimately show ?rhs using \<open>a \<noteq> 0\<close> by arith
+ qed (auto simp: not_le \<open>a \<noteq> 0\<close> intro!: split_scaleR_pos_le)
qed simp
lemma scaleR_le_0_iff:
@@ -622,7 +622,7 @@
using scaleR_right_mono[of a 1 x] by simp
-subsection {* Real normed vector spaces *}
+subsection \<open>Real normed vector spaces\<close>
class dist =
fixes dist :: "'a \<Rightarrow> 'a \<Rightarrow> real"
@@ -862,7 +862,7 @@
shows "norm (x ^ n) = norm x ^ n"
by (induct n) (simp_all add: norm_mult)
-text{*Despite a superficial resemblance, @{text norm_eq_1} is not relevant.*}
+text\<open>Despite a superficial resemblance, @{text norm_eq_1} is not relevant.\<close>
lemma square_norm_one:
fixes x :: "'a::real_normed_div_algebra"
assumes "x^2 = 1" shows "norm x = 1"
@@ -939,7 +939,7 @@
finally show ?thesis .
qed
-subsection {* Metric spaces *}
+subsection \<open>Metric spaces\<close>
class metric_space = open_dist +
assumes dist_eq_0_iff [simp]: "dist x y = 0 \<longleftrightarrow> x = y"
@@ -1065,7 +1065,7 @@
by blast
qed
-text {* Every normed vector space is a metric space. *}
+text \<open>Every normed vector space is a metric space.\<close>
instance real_normed_vector < metric_space
proof
@@ -1077,7 +1077,7 @@
using norm_triangle_ineq4 [of "x - z" "y - z"] by simp
qed
-subsection {* Class instances for real numbers *}
+subsection \<open>Class instances for real numbers\<close>
instantiation real :: real_normed_field
begin
@@ -1150,24 +1150,24 @@
lemmas closed_real_atLeast = closed_atLeast[where 'a=real]
lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real]
-subsection {* Extra type constraints *}
+subsection \<open>Extra type constraints\<close>
-text {* Only allow @{term "open"} in class @{text topological_space}. *}
+text \<open>Only allow @{term "open"} in class @{text topological_space}.\<close>
-setup {* Sign.add_const_constraint
- (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"}) *}
+setup \<open>Sign.add_const_constraint
+ (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"})\<close>
-text {* Only allow @{term dist} in class @{text metric_space}. *}
+text \<open>Only allow @{term dist} in class @{text metric_space}.\<close>
-setup {* Sign.add_const_constraint
- (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+ (@{const_name dist}, SOME @{typ "'a::metric_space \<Rightarrow> 'a \<Rightarrow> real"})\<close>
-text {* Only allow @{term norm} in class @{text real_normed_vector}. *}
+text \<open>Only allow @{term norm} in class @{text real_normed_vector}.\<close>
-setup {* Sign.add_const_constraint
- (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"}) *}
+setup \<open>Sign.add_const_constraint
+ (@{const_name norm}, SOME @{typ "'a::real_normed_vector \<Rightarrow> real"})\<close>
-subsection {* Sign function *}
+subsection \<open>Sign function\<close>
lemma norm_sgn:
"norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)"
@@ -1225,7 +1225,7 @@
lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b"
by (simp_all add: dist_norm)
-subsection {* Bounded Linear and Bilinear Operators *}
+subsection \<open>Bounded Linear and Bilinear Operators\<close>
locale linear = additive f for f :: "'a::real_vector \<Rightarrow> 'b::real_vector" +
assumes scaleR: "f (scaleR r x) = scaleR r (f x)"
@@ -1492,7 +1492,7 @@
by (clarsimp, rule_tac x="x + of_real (e/2)" in exI, simp)
qed
-subsection {* Filters and Limits on Metric Space *}
+subsection \<open>Filters and Limits on Metric Space\<close>
lemma (in metric_space) nhds_metric: "nhds x = (INF e:{0 <..}. principal {y. dist y x < e})"
unfolding nhds_def
@@ -1552,7 +1552,7 @@
apply (rule_tac c="nat(ceiling (Z + 1))" in eventually_sequentiallyI)
by linarith
-subsubsection {* Limits of Sequences *}
+subsubsection \<open>Limits of Sequences\<close>
lemma lim_sequentially: "X ----> (L::'a::metric_space) \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)"
unfolding tendsto_iff eventually_sequentially ..
@@ -1571,7 +1571,7 @@
by (simp add: lim_sequentially)
-subsubsection {* Limits of Functions *}
+subsubsection \<open>Limits of Functions\<close>
lemma LIM_def: "f -- (a::'a::metric_space) --> (L::'b::metric_space) =
(\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & dist x a < s
@@ -1622,14 +1622,14 @@
shows "(\<lambda>x. g (f x)) -- a --> l"
by (rule metric_LIM_compose2 [OF f g inj])
-subsection {* Complete metric spaces *}
+subsection \<open>Complete metric spaces\<close>
-subsection {* Cauchy sequences *}
+subsection \<open>Cauchy sequences\<close>
definition (in metric_space) Cauchy :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" where
"Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. dist (X m) (X n) < e)"
-subsection {* Cauchy Sequences *}
+subsection \<open>Cauchy Sequences\<close>
lemma metric_CauchyI:
"(\<And>e. 0 < e \<Longrightarrow> \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < e) \<Longrightarrow> Cauchy X"
@@ -1685,7 +1685,7 @@
unfolding convergent_def
by (erule exE, erule LIMSEQ_imp_Cauchy)
-subsubsection {* Cauchy Sequences are Convergent *}
+subsubsection \<open>Cauchy Sequences are Convergent\<close>
class complete_space = metric_space +
assumes Cauchy_convergent: "Cauchy X \<Longrightarrow> convergent X"
@@ -1695,17 +1695,17 @@
shows "Cauchy X = convergent X"
by (fast intro: Cauchy_convergent convergent_Cauchy)
-subsection {* The set of real numbers is a complete metric space *}
+subsection \<open>The set of real numbers is a complete metric space\<close>
-text {*
+text \<open>
Proof that Cauchy sequences converge based on the one from
@{url "http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html"}
-*}
+\<close>
-text {*
+text \<open>
If sequence @{term "X"} is Cauchy, then its limit is the lub of
@{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
-*}
+\<close>
lemma increasing_LIMSEQ:
fixes f :: "nat \<Rightarrow> real"
@@ -1717,9 +1717,9 @@
fix x assume "x < l"
with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x"
by auto
- from en[OF `0 < e`] obtain n where "l - e \<le> f n"
+ from en[OF \<open>0 < e\<close>] obtain n where "l - e \<le> f n"
by (auto simp: field_simps)
- with `e < l - x` `0 < e` have "x < f n" by simp
+ with \<open>e < l - x\<close> \<open>0 < e\<close> have "x < f n" by simp
with incseq_SucI[of f, OF inc] show "eventually (\<lambda>n. x < f n) sequentially"
by (auto simp: eventually_sequentially incseq_def intro: less_le_trans)
qed (insert bdd, auto)
--- a/src/HOL/Record.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Record.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,16 +6,16 @@
Author: Florian Haftmann, TU Muenchen
*)
-section {* Extensible records with structural subtyping *}
+section \<open>Extensible records with structural subtyping\<close>
theory Record
imports Quickcheck_Exhaustive
keywords "record" :: thy_decl
begin
-subsection {* Introduction *}
+subsection \<open>Introduction\<close>
-text {*
+text \<open>
Records are isomorphic to compound tuple types. To implement
efficient records, we make this isomorphism explicit. Consider the
record access/update simplification @{text "alpha (beta_update f
@@ -75,9 +75,9 @@
is valid) left as a rule assumption. All rules are structured to aid
net matching, using either a point-free form or an encapsulating
predicate.
-*}
+\<close>
-subsection {* Operators and lemmas for types isomorphic to tuples *}
+subsection \<open>Operators and lemmas for types isomorphic to tuples\<close>
datatype (dead 'a, dead 'b, dead 'c) tuple_isomorphism =
Tuple_Isomorphism "'a \<Rightarrow> 'b \<times> 'c" "'b \<times> 'c \<Rightarrow> 'a"
@@ -114,7 +114,7 @@
"iso_tuple_cons isom = curry (abst isom)"
-subsection {* Logical infrastructure for records *}
+subsection \<open>Logical infrastructure for records\<close>
definition
iso_tuple_surjective_proof_assist :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
@@ -412,7 +412,7 @@
by (simp add: comp_def)
-subsection {* Concrete record syntax *}
+subsection \<open>Concrete record syntax\<close>
nonterminal
ident and
@@ -452,7 +452,7 @@
"_record_update" :: "'a => field_updates => 'b" ("_/(3\<lparr>_\<rparr>)" [900, 0] 900)
-subsection {* Record package *}
+subsection \<open>Record package\<close>
ML_file "Tools/record.ML"
--- a/src/HOL/Relation.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Relation.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory; Stefan Berghofer, TU Muenchen
*)
-section {* Relations -- as sets of pairs, and binary predicates *}
+section \<open>Relations -- as sets of pairs, and binary predicates\<close>
theory Relation
imports Finite_Set
begin
-text {* A preliminary: classical rules for reasoning on predicates *}
+text \<open>A preliminary: classical rules for reasoning on predicates\<close>
declare predicate1I [Pure.intro!, intro!]
declare predicate1D [Pure.dest, dest]
@@ -51,23 +51,23 @@
declare Sup2_E [elim!]
declare SUP2_E [elim!]
-subsection {* Fundamental *}
+subsection \<open>Fundamental\<close>
-subsubsection {* Relations as sets of pairs *}
+subsubsection \<open>Relations as sets of pairs\<close>
type_synonym 'a rel = "('a * 'a) set"
-lemma subrelI: -- {* Version of @{thm [source] subsetI} for binary relations *}
+lemma subrelI: -- \<open>Version of @{thm [source] subsetI} for binary relations\<close>
"(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
by auto
-lemma lfp_induct2: -- {* Version of @{thm [source] lfp_induct} for binary relations *}
+lemma lfp_induct2: -- \<open>Version of @{thm [source] lfp_induct} for binary relations\<close>
"(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
(\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto
-subsubsection {* Conversions between set and predicate relations *}
+subsubsection \<open>Conversions between set and predicate relations\<close>
lemma pred_equals_eq [pred_set_conv]: "(\<lambda>x. x \<in> R) = (\<lambda>x. x \<in> S) \<longleftrightarrow> R = S"
by (simp add: set_eq_iff fun_eq_iff)
@@ -141,16 +141,16 @@
lemma SUP_Sup_eq2 [pred_set_conv]: "(\<Squnion>i\<in>S. (\<lambda>x y. (x, y) \<in> i)) = (\<lambda>x y. (x, y) \<in> \<Union>S)"
by (simp add: fun_eq_iff)
-subsection {* Properties of relations *}
+subsection \<open>Properties of relations\<close>
-subsubsection {* Reflexivity *}
+subsubsection \<open>Reflexivity\<close>
definition refl_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
where
"refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A \<and> (\<forall>x\<in>A. (x, x) \<in> r)"
abbreviation refl :: "'a rel \<Rightarrow> bool"
-where -- {* reflexivity over a type *}
+where -- \<open>reflexivity over a type\<close>
"refl \<equiv> refl_on UNIV"
definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
@@ -219,7 +219,7 @@
lemma reflp_equality [simp]: "reflp op ="
by(simp add: reflp_def)
-subsubsection {* Irreflexivity *}
+subsubsection \<open>Irreflexivity\<close>
definition irrefl :: "'a rel \<Rightarrow> bool"
where
@@ -246,7 +246,7 @@
by (auto simp add: irrefl_def)
-subsubsection {* Asymmetry *}
+subsubsection \<open>Asymmetry\<close>
inductive asym :: "'a rel \<Rightarrow> bool"
where
@@ -261,7 +261,7 @@
by (auto intro!: asymI asympI elim: asym.cases asymp.cases simp add: irreflp_irrefl_eq)
-subsubsection {* Symmetry *}
+subsubsection \<open>Symmetry\<close>
definition sym :: "'a rel \<Rightarrow> bool"
where
@@ -336,7 +336,7 @@
by (fact sym_UNION [to_pred])
-subsubsection {* Antisymmetry *}
+subsubsection \<open>Antisymmetry\<close>
definition antisym :: "'a rel \<Rightarrow> bool"
where
@@ -362,7 +362,7 @@
lemma antisymP_equality [simp]: "antisymP op ="
by(auto intro: antisymI)
-subsubsection {* Transitivity *}
+subsubsection \<open>Transitivity\<close>
definition trans :: "'a rel \<Rightarrow> bool"
where
@@ -377,7 +377,7 @@
by (simp add: trans_def transp_def)
abbreviation transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where -- {* FIXME drop *}
+where -- \<open>FIXME drop\<close>
"transP r \<equiv> trans {(x, y). r x y}"
lemma transI:
@@ -433,7 +433,7 @@
lemma transp_equality [simp]: "transp op ="
by(auto intro: transpI)
-subsubsection {* Totality *}
+subsubsection \<open>Totality\<close>
definition total_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
where
@@ -445,7 +445,7 @@
by (simp add: total_on_def)
-subsubsection {* Single valued relations *}
+subsubsection \<open>Single valued relations\<close>
definition single_valued :: "('a \<times> 'b) set \<Rightarrow> bool"
where
@@ -470,9 +470,9 @@
by (unfold single_valued_def) blast
-subsection {* Relation operations *}
+subsection \<open>Relation operations\<close>
-subsubsection {* The identity relation *}
+subsubsection \<open>The identity relation\<close>
definition Id :: "'a rel"
where
@@ -491,7 +491,7 @@
by (simp add: refl_on_def)
lemma antisym_Id: "antisym Id"
- -- {* A strange result, since @{text Id} is also symmetric. *}
+ -- \<open>A strange result, since @{text Id} is also symmetric.\<close>
by (simp add: antisym_def)
lemma sym_Id: "sym Id"
@@ -513,7 +513,7 @@
by (simp add: total_on_def)
-subsubsection {* Diagonal: identity over a set *}
+subsubsection \<open>Diagonal: identity over a set\<close>
definition Id_on :: "'a set \<Rightarrow> 'a rel"
where
@@ -530,7 +530,7 @@
lemma Id_onE [elim!]:
"c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
- -- {* The general elimination rule. *}
+ -- \<open>The general elimination rule.\<close>
by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
@@ -559,7 +559,7 @@
by (unfold single_valued_def) blast
-subsubsection {* Composition *}
+subsubsection \<open>Composition\<close>
inductive_set relcomp :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a \<times> 'c) set" (infixr "O" 75)
for r :: "('a \<times> 'b) set" and s :: "('b \<times> 'c) set"
@@ -570,10 +570,10 @@
lemmas relcomppI = relcompp.intros
-text {*
+text \<open>
For historic reasons, the elimination rules are not wholly corresponding.
Feel free to consolidate this.
-*}
+\<close>
inductive_cases relcompEpair: "(a, c) \<in> r O s"
inductive_cases relcomppE [elim!]: "(r OO s) a c"
@@ -677,7 +677,7 @@
by blast
-subsubsection {* Converse *}
+subsubsection \<open>Converse\<close>
inductive_set converse :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'a) set" ("(_^-1)" [1000] 999)
for r :: "('a \<times> 'b) set"
@@ -710,7 +710,7 @@
by (fact converseD [to_pred])
lemma converseE [elim!]:
- -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *}
+ -- \<open>More general than @{text converseD}, as it ``splits'' the member of the relation.\<close>
"yx \<in> r\<inverse> \<Longrightarrow> (\<And>x y. yx = (y, x) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> P) \<Longrightarrow> P"
by (cases yx) (simp, erule converse.cases, iprover)
@@ -826,7 +826,7 @@
by (simp add: set_eq_iff)
-subsubsection {* Domain, range and field *}
+subsubsection \<open>Domain, range and field\<close>
inductive_set Domain :: "('a \<times> 'b) set \<Rightarrow> 'a set"
for r :: "('a \<times> 'b) set"
@@ -980,7 +980,7 @@
by blast
-subsubsection {* Image of a set under a relation *}
+subsubsection \<open>Image of a set under a relation\<close>
definition Image :: "('a \<times> 'b) set \<Rightarrow> 'a set \<Rightarrow> 'b set" (infixr "``" 90)
where
@@ -1003,7 +1003,7 @@
by (unfold Image_def) (iprover elim!: CollectE bexE)
lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
- -- {* This version's more effective when we already have the required @{text a} *}
+ -- \<open>This version's more effective when we already have the required @{text a}\<close>
by blast
lemma Image_empty [simp]: "R``{} = {}"
@@ -1032,7 +1032,7 @@
by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
- -- {* NOT suitable for rewriting *}
+ -- \<open>NOT suitable for rewriting\<close>
by blast
lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
@@ -1047,7 +1047,7 @@
lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
by blast
-text{*Converse inclusion requires some assumptions*}
+text\<open>Converse inclusion requires some assumptions\<close>
lemma Image_INT_eq:
"[|single_valued (r\<inverse>); A\<noteq>{}|] ==> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
apply (rule equalityI)
@@ -1067,7 +1067,7 @@
lemma relcomp_Image: "(X O Y) `` Z = Y `` (X `` Z)"
by auto
-subsubsection {* Inverse image *}
+subsubsection \<open>Inverse image\<close>
definition inv_image :: "'b rel \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a rel"
where
@@ -1099,7 +1099,7 @@
by (simp add: inv_imagep_def)
-subsubsection {* Powerset *}
+subsubsection \<open>Powerset\<close>
definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
where
@@ -1110,7 +1110,7 @@
lemmas Powp_mono [mono] = Pow_mono [to_pred]
-subsubsection {* Expressing relation operations via @{const Finite_Set.fold} *}
+subsubsection \<open>Expressing relation operations via @{const Finite_Set.fold}\<close>
lemma Id_on_fold:
assumes "finite A"
@@ -1152,7 +1152,7 @@
qed
have *: "{x} O S = {(x', z). x' = fst x \<and> (snd x,z) \<in> S}" by (auto simp: relcomp_unfold intro!: exI)
show ?thesis unfolding *
- using `finite S` by (induct S) (auto split: prod.split)
+ using \<open>finite S\<close> by (induct S) (auto split: prod.split)
qed
lemma insert_relcomp_fold:
--- a/src/HOL/Rings.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Rings.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,7 +7,7 @@
Author: Jeremy Avigad
*)
-section {* Rings *}
+section \<open>Rings\<close>
theory Rings
imports Groups
@@ -18,7 +18,7 @@
assumes distrib_left[algebra_simps]: "a * (b + c) = a * b + a * c"
begin
-text{*For the @{text combine_numerals} simproc*}
+text\<open>For the @{text combine_numerals} simproc\<close>
lemma combine_common_factor:
"a * e + (b * e + c) = (a + b) * e + c"
by (simp add: distrib_right ac_simps)
@@ -118,7 +118,7 @@
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
-text {* Abstract divisibility *}
+text \<open>Abstract divisibility\<close>
class dvd = times
begin
@@ -180,8 +180,8 @@
and "c dvd d"
shows "a * c dvd b * d"
proof -
- from `a dvd b` obtain b' where "b = a * b'" ..
- moreover from `c dvd d` obtain d' where "d = c * d'" ..
+ from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
+ moreover from \<open>c dvd d\<close> obtain d' where "d = c * d'" ..
ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps)
then show ?thesis ..
qed
@@ -219,8 +219,8 @@
assumes "a dvd b" and "a dvd c"
shows "a dvd (b + c)"
proof -
- from `a dvd b` obtain b' where "b = a * b'" ..
- moreover from `a dvd c` obtain c' where "c = a * c'" ..
+ from \<open>a dvd b\<close> obtain b' where "b = a * b'" ..
+ moreover from \<open>a dvd c\<close> obtain c' where "c = a * c'" ..
ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left)
then show ?thesis ..
qed
@@ -284,7 +284,7 @@
shows "a dvd b + c \<longleftrightarrow> a dvd c" (is "?P \<longleftrightarrow> ?Q")
proof
assume ?P then obtain d where "b + c = a * d" ..
- moreover from `a dvd b` obtain e where "b = a * e" ..
+ moreover from \<open>a dvd b\<close> obtain e where "b = a * e" ..
ultimately have "a * e + c = a * d" by simp
then have "a * e + c - a * e = a * d - a * e" by simp
then have "c = a * d - a * e" by simp
@@ -306,7 +306,7 @@
subclass semiring_0_cancel ..
-text {* Distribution rules *}
+text \<open>Distribution rules\<close>
lemma minus_mult_left: "- (a * b) = - a * b"
by (rule minus_unique) (simp add: distrib_right [symmetric])
@@ -314,7 +314,7 @@
lemma minus_mult_right: "- (a * b) = a * - b"
by (rule minus_unique) (simp add: distrib_left [symmetric])
-text{*Extract signs from products*}
+text\<open>Extract signs from products\<close>
lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
lemmas mult_minus_right [simp] = minus_mult_right [symmetric]
@@ -543,7 +543,7 @@
end
-text {*
+text \<open>
The theory of partially ordered rings is taken from the books:
\begin{itemize}
\item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
@@ -554,12 +554,12 @@
\item @{url "http://www.mathworld.com"} by Eric Weisstein et. al.
\item \emph{Algebra I} by van der Waerden, Springer.
\end{itemize}
-*}
+\<close>
class divide =
fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "div" 70)
-setup {* Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a \<Rightarrow> 'a \<Rightarrow> 'a"}) *}
+setup \<open>Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a \<Rightarrow> 'a \<Rightarrow> 'a"})\<close>
context semiring
begin
@@ -581,7 +581,7 @@
end
-setup {* Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a::divide \<Rightarrow> 'a \<Rightarrow> 'a"}) *}
+setup \<open>Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a::divide \<Rightarrow> 'a \<Rightarrow> 'a"})\<close>
class semidom_divide = semidom + divide +
assumes nonzero_mult_divide_cancel_right [simp]: "b \<noteq> 0 \<Longrightarrow> (a * b) div b = a"
@@ -761,14 +761,14 @@
proof (rule that)
def b \<equiv> "1 div a"
then show "1 div a = b" by simp
- from b_def `is_unit a` show "is_unit b" by simp
- from `is_unit a` and `is_unit b` show "a \<noteq> 0" and "b \<noteq> 0" by auto
- from b_def `is_unit a` show "a * b = 1" by simp
+ from b_def \<open>is_unit a\<close> show "is_unit b" by simp
+ from \<open>is_unit a\<close> and \<open>is_unit b\<close> show "a \<noteq> 0" and "b \<noteq> 0" by auto
+ from b_def \<open>is_unit a\<close> show "a * b = 1" by simp
then have "1 = a * b" ..
- with b_def `b \<noteq> 0` show "1 div b = a" by simp
- from `is_unit a` have "a dvd c" ..
+ with b_def \<open>b \<noteq> 0\<close> show "1 div b = a" by simp
+ from \<open>is_unit a\<close> have "a dvd c" ..
then obtain d where "c = a * d" ..
- with `a \<noteq> 0` `a * b = 1` show "c div a = c * b"
+ with \<open>a \<noteq> 0\<close> \<open>a * b = 1\<close> show "c div a = c * b"
by (simp add: mult.assoc mult.left_commute [of a])
qed
@@ -802,9 +802,9 @@
assume "a dvd c * b"
with assms have "c * b dvd c * (b * (1 div b))"
by (subst mult_assoc [symmetric]) simp
- also from `is_unit b` have "b * (1 div b) = 1" by (rule is_unitE) simp
+ also from \<open>is_unit b\<close> have "b * (1 div b) = 1" by (rule is_unitE) simp
finally have "c * b dvd c" by simp
- with `a dvd c * b` show "a dvd c" by (rule dvd_trans)
+ with \<open>a dvd c * b\<close> show "a dvd c" by (rule dvd_trans)
next
assume "a dvd c"
then show "a dvd c * b" by simp
@@ -1227,7 +1227,7 @@
lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
using mult_right_mono [of a 0 b] by simp
-text {* Legacy - use @{text mult_nonpos_nonneg} *}
+text \<open>Legacy - use @{text mult_nonpos_nonneg}\<close>
lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
by (drule mult_right_mono [of b 0], auto)
@@ -1303,7 +1303,7 @@
lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
using mult_strict_right_mono [of a 0 b] by simp
-text {* Legacy - use @{text mult_neg_pos} *}
+text \<open>Legacy - use @{text mult_neg_pos}\<close>
lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
by (drule mult_strict_right_mono [of b 0], auto)
@@ -1323,7 +1323,7 @@
apply (auto dest: less_not_sym)
done
-text{*Strict monotonicity in both arguments*}
+text\<open>Strict monotonicity in both arguments\<close>
lemma mult_strict_mono:
assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
shows "a * c < b * d"
@@ -1334,7 +1334,7 @@
apply (erule mult_strict_left_mono, assumption)
done
-text{*This weaker variant has more natural premises*}
+text\<open>This weaker variant has more natural premises\<close>
lemma mult_strict_mono':
assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
shows "a * c < b * d"
@@ -1552,11 +1552,11 @@
apply force
done
-text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
- also with the relations @{text "\<le>"} and equality.*}
+text\<open>Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
+ also with the relations @{text "\<le>"} and equality.\<close>
-text{*These ``disjunction'' versions produce two cases when the comparison is
- an assumption, but effectively four when the comparison is a goal.*}
+text\<open>These ``disjunction'' versions produce two cases when the comparison is
+ an assumption, but effectively four when the comparison is a goal.\<close>
lemma mult_less_cancel_right_disj:
"a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and> b < a"
@@ -1584,8 +1584,8 @@
mult_left_mono_neg)
done
-text{*The ``conjunction of implication'' lemmas produce two cases when the
-comparison is a goal, but give four when the comparison is an assumption.*}
+text\<open>The ``conjunction of implication'' lemmas produce two cases when the
+comparison is a goal, but give four when the comparison is an assumption.\<close>
lemma mult_less_cancel_right:
"a * c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
@@ -1640,7 +1640,7 @@
assumes le_add_diff_inverse2 [simp]: "b \<le> a \<Longrightarrow> a - b + b = a"
begin
-text {* Addition is the inverse of subtraction. *}
+text \<open>Addition is the inverse of subtraction.\<close>
lemma le_add_diff_inverse [simp]: "b \<le> a \<Longrightarrow> b + (a - b) = a"
by (frule le_add_diff_inverse2) (simp add: add.commute)
@@ -1720,7 +1720,7 @@
assumes "x \<noteq> y" obtains "x < y" | "y < x"
using assms by (rule neqE)
-text {* These cancellation simprules also produce two cases when the comparison is a goal. *}
+text \<open>These cancellation simprules also produce two cases when the comparison is a goal.\<close>
lemma mult_le_cancel_right1:
"c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
@@ -1803,9 +1803,9 @@
"\<bar>l\<bar> = \<bar>k\<bar> \<Longrightarrow> l dvd k"
by(subst abs_dvd_iff[symmetric]) simp
-text {* The following lemmas can be proven in more general structures, but
+text \<open>The following lemmas can be proven in more general structures, but
are dangerous as simp rules in absence of @{thm neg_equal_zero},
-@{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. *}
+@{thm neg_less_pos}, @{thm neg_less_eq_nonneg}.\<close>
lemma equation_minus_iff_1 [simp, no_atp]:
"1 = - a \<longleftrightarrow> a = - 1"
@@ -1833,7 +1833,7 @@
end
-text {* Simprules for comparisons where common factors can be cancelled. *}
+text \<open>Simprules for comparisons where common factors can be cancelled.\<close>
lemmas mult_compare_simps =
mult_le_cancel_right mult_le_cancel_left
@@ -1846,7 +1846,7 @@
mult_cancel_right1 mult_cancel_right2
mult_cancel_left1 mult_cancel_left2
-text {* Reasoning about inequalities with division *}
+text \<open>Reasoning about inequalities with division\<close>
context linordered_semidom
begin
@@ -1876,7 +1876,7 @@
end
-text {* Absolute Value *}
+text \<open>Absolute Value\<close>
context linordered_idom
begin
--- a/src/HOL/SAT.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/SAT.thy Sat Jul 18 22:58:50 2015 +0200
@@ -5,7 +5,7 @@
Basic setup for the 'sat' and 'satx' tactics.
*)
-section {* Reconstructing external resolution proofs for propositional logic *}
+section \<open>Reconstructing external resolution proofs for propositional logic\<close>
theory SAT
imports HOL
@@ -15,10 +15,10 @@
ML_file "Tools/sat_solver.ML"
ML_file "Tools/sat.ML"
-method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac) *}
+method_setup sat = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.sat_tac)\<close>
"SAT solver"
-method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o SAT.satx_tac) *}
+method_setup satx = \<open>Scan.succeed (SIMPLE_METHOD' o SAT.satx_tac)\<close>
"SAT solver (with definitional CNF)"
end
--- a/src/HOL/SMT.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/SMT.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,14 +2,14 @@
Author: Sascha Boehme, TU Muenchen
*)
-section {* Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2 *}
+section \<open>Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2\<close>
theory SMT
imports Divides
keywords "smt_status" :: diag
begin
-subsection {* A skolemization tactic and proof method *}
+subsection \<open>A skolemization tactic and proof method\<close>
lemma choices:
"\<And>Q. \<forall>x. \<exists>y ya. Q x y ya \<Longrightarrow> \<exists>f fa. \<forall>x. Q x (f x) (fa x)"
@@ -39,25 +39,25 @@
\<exists>f fa fb fc fd fe ff fg. \<forall>x \<in> S. Q x (f x) (fa x) (fb x) (fc x) (fd x) (fe x) (ff x) (fg x)"
by metis+
-ML {*
+ML \<open>
fun moura_tac ctxt =
Atomize_Elim.atomize_elim_tac ctxt THEN'
SELECT_GOAL (Clasimp.auto_tac (ctxt addSIs @{thms choice choices bchoice bchoices}) THEN
ALLGOALS (Metis_Tactic.metis_tac (take 1 ATP_Proof_Reconstruct.partial_type_encs)
ATP_Proof_Reconstruct.default_metis_lam_trans ctxt [] ORELSE'
blast_tac ctxt))
-*}
+\<close>
-method_setup moura = {*
+method_setup moura = \<open>
Scan.succeed (SIMPLE_METHOD' o moura_tac)
-*} "solve skolemization goals, especially those arising from Z3 proofs"
+\<close> "solve skolemization goals, especially those arising from Z3 proofs"
hide_fact (open) choices bchoices
-subsection {* Triggers for quantifier instantiation *}
+subsection \<open>Triggers for quantifier instantiation\<close>
-text {*
+text \<open>
Some SMT solvers support patterns as a quantifier instantiation
heuristics. Patterns may either be positive terms (tagged by "pat")
triggering quantifier instantiations -- when the solver finds a
@@ -70,7 +70,7 @@
act disjunctively during quantifier instantiation. Each multipattern
should mention at least all quantified variables of the preceding
quantifier block.
-*}
+\<close>
typedecl 'a symb_list
@@ -88,26 +88,26 @@
"trigger _ P = P"
-subsection {* Higher-order encoding *}
+subsection \<open>Higher-order encoding\<close>
-text {*
+text \<open>
Application is made explicit for constants occurring with varying
numbers of arguments. This is achieved by the introduction of the
following constant.
-*}
+\<close>
definition fun_app :: "'a \<Rightarrow> 'a" where "fun_app f = f"
-text {*
+text \<open>
Some solvers support a theory of arrays which can be used to encode
higher-order functions. The following set of lemmas specifies the
properties of such (extensional) arrays.
-*}
+\<close>
lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other fun_upd_upd fun_app_def
-subsection {* Normalization *}
+subsection \<open>Normalization\<close>
lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)"
by simp
@@ -120,12 +120,12 @@
lemmas max_def_raw = max_def[abs_def]
-subsection {* Integer division and modulo for Z3 *}
+subsection \<open>Integer division and modulo for Z3\<close>
-text {*
+text \<open>
The following Z3-inspired definitions are overspecified for the case where @{text "l = 0"}. This
Schönheitsfehler is corrected in the @{text div_as_z3div} and @{text mod_as_z3mod} theorems.
-*}
+\<close>
definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
"z3div k l = (if l \<ge> 0 then k div l else - (k div - l))"
@@ -142,7 +142,7 @@
by (simp add: z3mod_def)
-subsection {* Setup *}
+subsection \<open>Setup\<close>
ML_file "Tools/SMT/smt_util.ML"
ML_file "Tools/SMT/smt_failure.ML"
@@ -171,93 +171,93 @@
ML_file "Tools/SMT/z3_replay.ML"
ML_file "Tools/SMT/smt_systems.ML"
-method_setup smt = {*
+method_setup smt = \<open>
Scan.optional Attrib.thms [] >>
(fn thms => fn ctxt =>
METHOD (fn facts => HEADGOAL (SMT_Solver.smt_tac ctxt (thms @ facts))))
-*} "apply an SMT solver to the current goal"
+\<close> "apply an SMT solver to the current goal"
-subsection {* Configuration *}
+subsection \<open>Configuration\<close>
-text {*
+text \<open>
The current configuration can be printed by the command
@{text smt_status}, which shows the values of most options.
-*}
+\<close>
-subsection {* General configuration options *}
+subsection \<open>General configuration options\<close>
-text {*
+text \<open>
The option @{text smt_solver} can be used to change the target SMT
solver. The possible values can be obtained from the @{text smt_status}
command.
-*}
+\<close>
declare [[smt_solver = z3]]
-text {*
+text \<open>
Since SMT solvers are potentially nonterminating, there is a timeout
(given in seconds) to restrict their runtime.
-*}
+\<close>
declare [[smt_timeout = 20]]
-text {*
+text \<open>
SMT solvers apply randomized heuristics. In case a problem is not
solvable by an SMT solver, changing the following option might help.
-*}
+\<close>
declare [[smt_random_seed = 1]]
-text {*
+text \<open>
In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
solvers are fully trusted without additional checks. The following
option can cause the SMT solver to run in proof-producing mode, giving
a checkable certificate. This is currently only implemented for Z3.
-*}
+\<close>
declare [[smt_oracle = false]]
-text {*
+text \<open>
Each SMT solver provides several commandline options to tweak its
behaviour. They can be passed to the solver by setting the following
options.
-*}
+\<close>
declare [[cvc3_options = ""]]
declare [[cvc4_options = "--full-saturate-quant --inst-when=full-last-call --inst-no-entail --term-db-mode=relevant"]]
declare [[verit_options = ""]]
declare [[z3_options = ""]]
-text {*
+text \<open>
The SMT method provides an inference mechanism to detect simple triggers
in quantified formulas, which might increase the number of problems
solvable by SMT solvers (note: triggers guide quantifier instantiations
in the SMT solver). To turn it on, set the following option.
-*}
+\<close>
declare [[smt_infer_triggers = false]]
-text {*
+text \<open>
Enable the following option to use built-in support for datatypes,
codatatypes, and records in CVC4. Currently, this is implemented only
in oracle mode.
-*}
+\<close>
declare [[cvc4_extensions = false]]
-text {*
+text \<open>
Enable the following option to use built-in support for div/mod, datatypes,
and records in Z3. Currently, this is implemented only in oracle mode.
-*}
+\<close>
declare [[z3_extensions = false]]
-subsection {* Certificates *}
+subsection \<open>Certificates\<close>
-text {*
+text \<open>
By setting the option @{text smt_certificates} to the name of a file,
all following applications of an SMT solver a cached in that file.
Any further application of the same SMT solver (using the very same
@@ -269,11 +269,11 @@
@{text ".certs"} instead of @{text ".thy"}) as the certificates file.
Certificate files should be used at most once in a certain theory context,
to avoid race conditions with other concurrent accesses.
-*}
+\<close>
declare [[smt_certificates = ""]]
-text {*
+text \<open>
The option @{text smt_read_only_certificates} controls whether only
stored certificates are should be used or invocation of an SMT solver
is allowed. When set to @{text true}, no SMT solver will ever be
@@ -281,39 +281,39 @@
cache are used; when set to @{text false} and there is no cached
certificate for some proposition, then the configured SMT solver is
invoked.
-*}
+\<close>
declare [[smt_read_only_certificates = false]]
-subsection {* Tracing *}
+subsection \<open>Tracing\<close>
-text {*
+text \<open>
The SMT method, when applied, traces important information. To
make it entirely silent, set the following option to @{text false}.
-*}
+\<close>
declare [[smt_verbose = true]]
-text {*
+text \<open>
For tracing the generated problem file given to the SMT solver as
well as the returned result of the solver, the option
@{text smt_trace} should be set to @{text true}.
-*}
+\<close>
declare [[smt_trace = false]]
-subsection {* Schematic rules for Z3 proof reconstruction *}
+subsection \<open>Schematic rules for Z3 proof reconstruction\<close>
-text {*
+text \<open>
Several prof rules of Z3 are not very well documented. There are two
lemma groups which can turn failing Z3 proof reconstruction attempts
into succeeding ones: the facts in @{text z3_rule} are tried prior to
any implemented reconstruction procedure for all uncertain Z3 proof
rules; the facts in @{text z3_simp} are only fed to invocations of
the simplifier when reconstructing theory-specific proof steps.
-*}
+\<close>
lemmas [z3_rule] =
refl eq_commute conj_commute disj_commute simp_thms nnf_simps
--- a/src/HOL/Semiring_Normalization.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Semiring_Normalization.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,13 +2,13 @@
Author: Amine Chaieb, TU Muenchen
*)
-section {* Semiring normalization *}
+section \<open>Semiring normalization\<close>
theory Semiring_Normalization
imports Numeral_Simprocs Nat_Transfer
begin
-text {* Prelude *}
+text \<open>Prelude\<close>
class comm_semiring_1_cancel_crossproduct = comm_semiring_1_cancel +
assumes crossproduct_eq: "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z"
@@ -57,15 +57,15 @@
assume "y < z" then have "\<exists>k. z = y + k \<and> k \<noteq> 0" by (intro exI [of _ "z - y"]) auto
then obtain k where "z = y + k" and "k \<noteq> 0" by blast
assume "w * y + x * z = w * z + x * y"
- then have "(w * y + x * y) + x * k = (w * y + x * y) + w * k" by (simp add: `z = y + k` algebra_simps)
+ then have "(w * y + x * y) + x * k = (w * y + x * y) + w * k" by (simp add: \<open>z = y + k\<close> algebra_simps)
then have "x * k = w * k" by simp
- then show "w = x" using `k \<noteq> 0` by simp
+ then show "w = x" using \<open>k \<noteq> 0\<close> by simp
qed
show "w * y + x * z = w * z + x * y \<longleftrightarrow> w = x \<or> y = z"
by (auto simp add: neq_iff dest!: aux)
qed
-text {* Semiring normalization proper *}
+text \<open>Semiring normalization proper\<close>
ML_file "Tools/semiring_normalizer.ML"
--- a/src/HOL/Series.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Series.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,13 +7,13 @@
Additional contributions by Jeremy Avigad
*)
-section {* Infinite Series *}
+section \<open>Infinite Series\<close>
theory Series
imports Limits Inequalities
begin
-subsection {* Definition of infinite summability *}
+subsection \<open>Definition of infinite summability\<close>
definition
sums :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a \<Rightarrow> bool"
@@ -30,7 +30,7 @@
where
"suminf f = (THE s. f sums s)"
-subsection {* Infinite summability on topological monoids *}
+subsection \<open>Infinite summability on topological monoids\<close>
lemma sums_subst[trans]: "f = g \<Longrightarrow> g sums z \<Longrightarrow> f sums z"
by simp
@@ -141,7 +141,7 @@
by (rule sums_zero [THEN sums_unique, symmetric])
-subsection {* Infinite summability on ordered, topological monoids *}
+subsection \<open>Infinite summability on ordered, topological monoids\<close>
lemma sums_le:
fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
@@ -221,7 +221,7 @@
by (auto intro: le_less_trans simp: eventually_sequentially) }
qed
-subsection {* Infinite summability on real normed vector spaces *}
+subsection \<open>Infinite summability on real normed vector spaces\<close>
lemma sums_Suc_iff:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
@@ -303,10 +303,10 @@
fixes r :: real assumes "0 < r" and "summable f"
shows "\<exists>N. \<forall>n\<ge>N. norm (\<Sum>i. f (i + n)) < r"
proof -
- from LIMSEQ_D[OF summable_LIMSEQ[OF `summable f`] `0 < r`]
+ from LIMSEQ_D[OF summable_LIMSEQ[OF \<open>summable f\<close>] \<open>0 < r\<close>]
obtain N :: nat where "\<forall> n \<ge> N. norm (setsum f {..<n} - suminf f) < r" by auto
thus ?thesis
- by (auto simp: norm_minus_commute suminf_minus_initial_segment[OF `summable f`])
+ by (auto simp: norm_minus_commute suminf_minus_initial_segment[OF \<open>summable f\<close>])
qed
lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"
@@ -324,7 +324,7 @@
lemma summable_minus_iff:
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
shows "summable (\<lambda>n. - f n) \<longleftrightarrow> summable f"
- by (auto dest: summable_minus) --{*used two ways, hence must be outside the context above*}
+ by (auto dest: summable_minus) --\<open>used two ways, hence must be outside the context above\<close>
context
@@ -363,7 +363,7 @@
lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right]
lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right]
-subsection {* Infinite summability on real normed algebras *}
+subsection \<open>Infinite summability on real normed algebras\<close>
context
fixes f :: "nat \<Rightarrow> 'a::real_normed_algebra"
@@ -389,7 +389,7 @@
end
-subsection {* Infinite summability on real normed fields *}
+subsection \<open>Infinite summability on real normed fields\<close>
context
fixes c :: "'a::real_normed_field"
@@ -404,7 +404,7 @@
lemma suminf_divide: "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"
by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric])
-text{*Sum of a geometric progression.*}
+text\<open>Sum of a geometric progression.\<close>
lemma geometric_sums: "norm c < 1 \<Longrightarrow> (\<lambda>n. c^n) sums (1 / (1 - c))"
proof -
@@ -439,9 +439,9 @@
by simp
qed
-subsection {* Infinite summability on Banach spaces *}
+subsection \<open>Infinite summability on Banach spaces\<close>
-text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*}
+text\<open>Cauchy-type criterion for convergence of series (c.f. Harrison)\<close>
lemma summable_Cauchy:
fixes f :: "nat \<Rightarrow> 'a::banach"
@@ -465,7 +465,7 @@
fixes f :: "nat \<Rightarrow> 'a::banach"
begin
-text{*Absolute convergence imples normal convergence*}
+text\<open>Absolute convergence imples normal convergence\<close>
lemma summable_norm_cancel: "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"
apply (simp only: summable_Cauchy, safe)
@@ -480,7 +480,7 @@
lemma summable_norm: "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"
by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_setsum)
-text {* Comparison tests *}
+text \<open>Comparison tests\<close>
lemma summable_comparison_test: "\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n \<Longrightarrow> summable g \<Longrightarrow> summable f"
apply (simp add: summable_Cauchy, safe)
@@ -499,7 +499,7 @@
lemma summable_comparison_test': "summable g \<Longrightarrow> (\<And>n. n \<ge> N \<Longrightarrow> norm(f n) \<le> g n) \<Longrightarrow> summable f"
by (rule summable_comparison_test) auto
-subsection {* The Ratio Test*}
+subsection \<open>The Ratio Test\<close>
lemma summable_ratio_test:
assumes "c < 1" "\<And>n. n \<ge> N \<Longrightarrow> norm (f (Suc n)) \<le> c * norm (f n)"
@@ -514,12 +514,12 @@
proof (induct rule: inc_induct)
case (step m)
moreover have "norm (f (Suc m)) / c ^ Suc m * c ^ n \<le> norm (f m) / c ^ m * c ^ n"
- using `0 < c` `c < 1` assms(2)[OF `N \<le> m`] by (simp add: field_simps)
+ using \<open>0 < c\<close> \<open>c < 1\<close> assms(2)[OF \<open>N \<le> m\<close>] by (simp add: field_simps)
ultimately show ?case by simp
- qed (insert `0 < c`, simp)
+ qed (insert \<open>0 < c\<close>, simp)
qed
show "summable (\<lambda>n. norm (f N) / c ^ N * c ^ n)"
- using `0 < c` `c < 1` by (intro summable_mult summable_geometric) simp
+ using \<open>0 < c\<close> \<open>c < 1\<close> by (intro summable_mult summable_geometric) simp
qed
next
assume c: "\<not> 0 < c"
@@ -536,7 +536,7 @@
end
-text{*Relations among convergence and absolute convergence for power series.*}
+text\<open>Relations among convergence and absolute convergence for power series.\<close>
lemma abel_lemma:
fixes a :: "nat \<Rightarrow> 'a::real_normed_vector"
@@ -557,7 +557,7 @@
qed
-text{*Summability of geometric series for real algebras*}
+text\<open>Summability of geometric series for real algebras\<close>
lemma complete_algebra_summable_geometric:
fixes x :: "'a::{real_normed_algebra_1,banach}"
@@ -569,12 +569,12 @@
by (simp add: summable_geometric)
qed
-subsection {* Cauchy Product Formula *}
+subsection \<open>Cauchy Product Formula\<close>
-text {*
+text \<open>
Proof based on Analysis WebNotes: Chapter 07, Class 41
@{url "http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm"}
-*}
+\<close>
lemma Cauchy_product_sums:
fixes a b :: "nat \<Rightarrow> 'a::{real_normed_algebra,banach}"
@@ -655,7 +655,7 @@
using a b
by (rule Cauchy_product_sums [THEN sums_unique])
-subsection {* Series on @{typ real}s *}
+subsection \<open>Series on @{typ real}s\<close>
lemma summable_norm_comparison_test: "\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. norm (f n))"
by (rule summable_comparison_test) auto
@@ -702,7 +702,7 @@
also have "\<dots> \<le> (\<Sum>i<m. f i)"
by (rule setsum_mono3) (auto simp add: pos n[rule_format])
also have "\<dots> \<le> suminf f"
- using `summable f`
+ using \<open>summable f\<close>
by (rule setsum_le_suminf) (simp add: pos)
finally show "(\<Sum>i<n. (f \<circ> g) i) \<le> suminf f" by simp
qed
--- a/src/HOL/Set.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Set.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,12 +1,12 @@
(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *)
-section {* Set theory for higher-order logic *}
+section \<open>Set theory for higher-order logic\<close>
theory Set
imports Lattices
begin
-subsection {* Sets as predicates *}
+subsection \<open>Sets as predicates\<close>
typedecl 'a set
@@ -40,7 +40,7 @@
not_member ("(_/ \<notin> _)" [51, 51] 50)
-text {* Set comprehensions *}
+text \<open>Set comprehensions\<close>
syntax
"_Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})")
@@ -63,12 +63,12 @@
lemma Collect_cong: "(\<And>x. P x = Q x) ==> {x. P x} = {x. Q x}"
by simp
-text {*
+text \<open>
Simproc for pulling @{text "x=t"} in @{text "{x. \<dots> & x=t & \<dots>}"}
to the front (and similarly for @{text "t=x"}):
-*}
-
-simproc_setup defined_Collect ("{x. P x & Q x}") = {*
+\<close>
+
+simproc_setup defined_Collect ("{x. P x & Q x}") = \<open>
fn _ => Quantifier1.rearrange_Collect
(fn ctxt =>
resolve_tac ctxt @{thms Collect_cong} 1 THEN
@@ -76,7 +76,7 @@
ALLGOALS
(EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE},
DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})]))
-*}
+\<close>
lemmas CollectE = CollectD [elim_format]
@@ -92,7 +92,7 @@
"A = B \<longleftrightarrow> (\<forall>x. x \<in> A \<longleftrightarrow> x \<in> B)"
by (auto intro:set_eqI)
-text {* Lifting of predicate class instances *}
+text \<open>Lifting of predicate class instances\<close>
instantiation set :: (type) boolean_algebra
begin
@@ -130,7 +130,7 @@
end
-text {* Set enumerations *}
+text \<open>Set enumerations\<close>
abbreviation empty :: "'a set" ("{}") where
"{} \<equiv> bot"
@@ -145,7 +145,7 @@
"{x}" == "CONST insert x {}"
-subsection {* Subsets and bounded quantifiers *}
+subsection \<open>Subsets and bounded quantifiers\<close>
abbreviation
subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
@@ -256,7 +256,7 @@
"\<exists>A\<subseteq>B. P" => "EX A. A \<subseteq> B & P"
"\<exists>!A\<subseteq>B. P" => "EX! A. A \<subseteq> B & P"
-print_translation {*
+print_translation \<open>
let
val All_binder = Mixfix.binder_name @{const_syntax All};
val Ex_binder = Mixfix.binder_name @{const_syntax Ex};
@@ -287,19 +287,19 @@
in
[tr' All_binder, tr' Ex_binder]
end
-*}
-
-
-text {*
+\<close>
+
+
+text \<open>
\medskip Translate between @{text "{e | x1...xn. P}"} and @{text
"{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is
only translated if @{text "[0..n] subset bvs(e)"}.
-*}
+\<close>
syntax
"_Setcompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})")
-parse_translation {*
+parse_translation \<open>
let
val ex_tr = snd (Syntax_Trans.mk_binder_tr ("EX ", @{const_syntax Ex}));
@@ -314,14 +314,14 @@
in Syntax.const @{const_syntax Collect} $ absdummy dummyT exP end;
in [(@{syntax_const "_Setcompr"}, setcompr_tr)] end;
-*}
-
-print_translation {*
+\<close>
+
+print_translation \<open>
[Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"},
Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"}]
-*} -- {* to avoid eta-contraction of body *}
-
-print_translation {*
+\<close> -- \<open>to avoid eta-contraction of body\<close>
+
+print_translation \<open>
let
val ex_tr' = snd (Syntax_Trans.mk_binder_tr' (@{const_syntax Ex}, "DUMMY"));
@@ -353,21 +353,21 @@
end
end;
in [(@{const_syntax Collect}, setcompr_tr')] end;
-*}
-
-simproc_setup defined_Bex ("EX x:A. P x & Q x") = {*
+\<close>
+
+simproc_setup defined_Bex ("EX x:A. P x & Q x") = \<open>
fn _ => Quantifier1.rearrange_bex
(fn ctxt =>
unfold_tac ctxt @{thms Bex_def} THEN
Quantifier1.prove_one_point_ex_tac ctxt)
-*}
-
-simproc_setup defined_All ("ALL x:A. P x --> Q x") = {*
+\<close>
+
+simproc_setup defined_All ("ALL x:A. P x --> Q x") = \<open>
fn _ => Quantifier1.rearrange_ball
(fn ctxt =>
unfold_tac ctxt @{thms Ball_def} THEN
Quantifier1.prove_one_point_all_tac ctxt)
-*}
+\<close>
lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x"
by (simp add: Ball_def)
@@ -377,16 +377,16 @@
lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x"
by (simp add: Ball_def)
-text {*
+text \<open>
Gives better instantiation for bound:
-*}
-
-setup {*
+\<close>
+
+setup \<open>
map_theory_claset (fn ctxt =>
ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt'))
-*}
-
-ML {*
+\<close>
+
+ML \<open>
structure Simpdata =
struct
@@ -397,22 +397,22 @@
end;
open Simpdata;
-*}
-
-declaration {* fn _ =>
+\<close>
+
+declaration \<open>fn _ =>
Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))
-*}
+\<close>
lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q"
by (unfold Ball_def) blast
lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x"
- -- {* Normally the best argument order: @{prop "P x"} constrains the
- choice of @{prop "x:A"}. *}
+ -- \<open>Normally the best argument order: @{prop "P x"} constrains the
+ choice of @{prop "x:A"}.\<close>
by (unfold Bex_def) blast
lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x"
- -- {* The best argument order when there is only one @{prop "x:A"}. *}
+ -- \<open>The best argument order when there is only one @{prop "x:A"}.\<close>
by (unfold Bex_def) blast
lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x"
@@ -422,11 +422,11 @@
by (unfold Bex_def) blast
lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)"
- -- {* Trival rewrite rule. *}
+ -- \<open>Trival rewrite rule.\<close>
by (simp add: Ball_def)
lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)"
- -- {* Dual form for existentials. *}
+ -- \<open>Dual form for existentials.\<close>
by (simp add: Bex_def)
lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)"
@@ -456,7 +456,7 @@
by blast
-text {* Congruence rules *}
+text \<open>Congruence rules\<close>
lemma ball_cong:
"A = B ==> (!!x. x:B ==> P x = Q x) ==>
@@ -481,34 +481,34 @@
lemma bex1_def: "(\<exists>!x\<in>X. P x) \<longleftrightarrow> (\<exists>x\<in>X. P x) \<and> (\<forall>x\<in>X. \<forall>y\<in>X. P x \<longrightarrow> P y \<longrightarrow> x = y)"
by auto
-subsection {* Basic operations *}
-
-subsubsection {* Subsets *}
+subsection \<open>Basic operations\<close>
+
+subsubsection \<open>Subsets\<close>
lemma subsetI [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B) \<Longrightarrow> A \<subseteq> B"
by (simp add: less_eq_set_def le_fun_def)
-text {*
+text \<open>
\medskip Map the type @{text "'a set => anything"} to just @{typ
'a}; for overloading constants whose first argument has type @{typ
"'a set"}.
-*}
+\<close>
lemma subsetD [elim, intro?]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B"
by (simp add: less_eq_set_def le_fun_def)
- -- {* Rule in Modus Ponens style. *}
+ -- \<open>Rule in Modus Ponens style.\<close>
lemma rev_subsetD [intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
- -- {* The same, with reversed premises for use with @{text erule} --
- cf @{text rev_mp}. *}
+ -- \<open>The same, with reversed premises for use with @{text erule} --
+ cf @{text rev_mp}.\<close>
by (rule subsetD)
-text {*
+text \<open>
\medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
-*}
+\<close>
lemma subsetCE [elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
- -- {* Classical elimination rule. *}
+ -- \<open>Classical elimination rule.\<close>
by (auto simp add: less_eq_set_def le_fun_def)
lemma subset_eq: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
@@ -539,16 +539,16 @@
order_trans_rules set_rev_mp set_mp eq_mem_trans
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
lemma subset_antisym [intro!]: "A \<subseteq> B ==> B \<subseteq> A ==> A = B"
- -- {* Anti-symmetry of the subset relation. *}
+ -- \<open>Anti-symmetry of the subset relation.\<close>
by (iprover intro: set_eqI subsetD)
-text {*
+text \<open>
\medskip Equality rules from ZF set theory -- are they appropriate
here?
-*}
+\<close>
lemma equalityD1: "A = B ==> A \<subseteq> B"
by simp
@@ -556,11 +556,11 @@
lemma equalityD2: "A = B ==> B \<subseteq> A"
by simp
-text {*
+text \<open>
\medskip Be careful when adding this to the claset as @{text
subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{}
\<subseteq> A"} and @{prop "A \<subseteq> {}"} and then back to @{prop "A = {}"}!
-*}
+\<close>
lemma equalityE: "A = B ==> (A \<subseteq> B ==> B \<subseteq> A ==> P) ==> P"
by simp
@@ -576,7 +576,7 @@
by simp
-subsubsection {* The empty set *}
+subsubsection \<open>The empty set\<close>
lemma empty_def:
"{} = {x. False}"
@@ -589,14 +589,14 @@
by simp
lemma empty_subsetI [iff]: "{} \<subseteq> A"
- -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *}
+ -- \<open>One effect is to delete the ASSUMPTION @{prop "{} <= A"}\<close>
by blast
lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
by blast
lemma equals0D: "A = {} ==> a \<notin> A"
- -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *}
+ -- \<open>Use for reasoning about disjointness: @{text "A Int B = {}"}\<close>
by blast
lemma ball_empty [simp]: "Ball {} P = True"
@@ -606,7 +606,7 @@
by (simp add: Bex_def)
-subsubsection {* The universal set -- UNIV *}
+subsubsection \<open>The universal set -- UNIV\<close>
abbreviation UNIV :: "'a set" where
"UNIV \<equiv> top"
@@ -618,7 +618,7 @@
lemma UNIV_I [simp]: "x : UNIV"
by (simp add: UNIV_def)
-declare UNIV_I [intro] -- {* unsafe makes it less likely to cause problems *}
+declare UNIV_I [intro] -- \<open>unsafe makes it less likely to cause problems\<close>
lemma UNIV_witness [intro?]: "EX x. x : UNIV"
by simp
@@ -626,11 +626,11 @@
lemma subset_UNIV: "A \<subseteq> UNIV"
by (fact top_greatest) (* already simp *)
-text {*
+text \<open>
\medskip Eta-contracting these two rules (to remove @{text P})
causes them to be ignored because of their interaction with
congruence rules.
-*}
+\<close>
lemma ball_UNIV [simp]: "Ball UNIV P = All P"
by (simp add: Ball_def)
@@ -647,7 +647,7 @@
lemma empty_not_UNIV[simp]: "{} \<noteq> UNIV"
by blast
-subsubsection {* The Powerset operator -- Pow *}
+subsubsection \<open>The Powerset operator -- Pow\<close>
definition Pow :: "'a set => 'a set set" where
Pow_def: "Pow A = {B. B \<le> A}"
@@ -671,7 +671,7 @@
using Pow_top by blast
-subsubsection {* Set complement *}
+subsubsection \<open>Set complement\<close>
lemma Compl_iff [simp]: "(c \<in> -A) = (c \<notin> A)"
by (simp add: fun_Compl_def uminus_set_def)
@@ -679,10 +679,10 @@
lemma ComplI [intro!]: "(c \<in> A ==> False) ==> c \<in> -A"
by (simp add: fun_Compl_def uminus_set_def) blast
-text {*
+text \<open>
\medskip This form, with negated conclusion, works well with the
Classical prover. Negated assumptions behave like formulae on the
- right side of the notional turnstile ... *}
+ right side of the notional turnstile ...\<close>
lemma ComplD [dest!]: "c : -A ==> c~:A"
by simp
@@ -693,7 +693,7 @@
by blast
-subsubsection {* Binary intersection *}
+subsubsection \<open>Binary intersection\<close>
abbreviation inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
"op Int \<equiv> inf"
@@ -727,7 +727,7 @@
by (fact mono_inf)
-subsubsection {* Binary union *}
+subsubsection \<open>Binary union\<close>
abbreviation union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where
"union \<equiv> sup"
@@ -751,10 +751,10 @@
lemma UnI2 [elim?]: "c:B ==> c : A Un B"
by simp
-text {*
+text \<open>
\medskip Classical introduction rule: no commitment to @{prop A} vs
@{prop B}.
-*}
+\<close>
lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B"
by auto
@@ -769,7 +769,7 @@
by (fact mono_sup)
-subsubsection {* Set difference *}
+subsubsection \<open>Set difference\<close>
lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)"
by (simp add: minus_set_def fun_diff_def)
@@ -792,7 +792,7 @@
by blast
-subsubsection {* Augmenting a set -- @{const insert} *}
+subsubsection \<open>Augmenting a set -- @{const insert}\<close>
lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)"
by (unfold insert_def) blast
@@ -807,7 +807,7 @@
by (unfold insert_def) blast
lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B"
- -- {* Classical introduction rule. *}
+ -- \<open>Classical introduction rule.\<close>
by auto
lemma subset_insert_iff: "(A \<subseteq> insert x B) = (if x:A then A - {x} \<subseteq> B else A \<subseteq> B)"
@@ -833,13 +833,13 @@
assume ?L
show ?R
proof cases
- assume "a=b" with assms `?L` show ?R by (simp add: insert_ident)
+ assume "a=b" with assms \<open>?L\<close> show ?R by (simp add: insert_ident)
next
assume "a\<noteq>b"
let ?C = "A - {b}"
have "A = insert b ?C \<and> b \<notin> ?C \<and> B = insert a ?C \<and> a \<notin> ?C"
- using assms `?L` `a\<noteq>b` by auto
- thus ?R using `a\<noteq>b` by auto
+ using assms \<open>?L\<close> \<open>a\<noteq>b\<close> by auto
+ thus ?R using \<open>a\<noteq>b\<close> by auto
qed
next
assume ?R thus ?L by (auto split: if_splits)
@@ -848,10 +848,10 @@
lemma insert_UNIV: "insert x UNIV = UNIV"
by auto
-subsubsection {* Singletons, using insert *}
+subsubsection \<open>Singletons, using insert\<close>
lemma singletonI [intro!]: "a : {a}"
- -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
+ -- \<open>Redundant? But unlike @{text insertCI}, it proves the subgoal immediately!\<close>
by (rule insertI1)
lemma singletonD [dest!]: "b : {a} ==> b = a"
@@ -897,11 +897,11 @@
by auto
-subsubsection {* Image of a set under a function *}
-
-text {*
+subsubsection \<open>Image of a set under a function\<close>
+
+text \<open>
Frequently @{term b} does not have the syntactic form of @{term "f x"}.
-*}
+\<close>
definition image :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set" (infixr "`" 90)
where
@@ -917,12 +917,12 @@
lemma rev_image_eqI:
"x \<in> A \<Longrightarrow> b = f x \<Longrightarrow> b \<in> f ` A"
- -- {* This version's more effective when we already have the
- required @{term x}. *}
+ -- \<open>This version's more effective when we already have the
+ required @{term x}.\<close>
by (rule image_eqI)
lemma imageE [elim!]:
- assumes "b \<in> (\<lambda>x. f x) ` A" -- {* The eta-expansion gives variable-name preservation. *}
+ assumes "b \<in> (\<lambda>x. f x) ` A" -- \<open>The eta-expansion gives variable-name preservation.\<close>
obtains x where "b = f x" and "x \<in> A"
using assms by (unfold image_def) blast
@@ -940,13 +940,13 @@
lemma image_subsetI:
"(\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` A \<subseteq> B"
- -- {* Replaces the three steps @{text subsetI}, @{text imageE},
- @{text hypsubst}, but breaks too many existing proofs. *}
+ -- \<open>Replaces the three steps @{text subsetI}, @{text imageE},
+ @{text hypsubst}, but breaks too many existing proofs.\<close>
by blast
lemma image_subset_iff:
"f ` A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. f x \<in> B)"
- -- {* This rewrite rule would confuse users if made default. *}
+ -- \<open>This rewrite rule would confuse users if made default.\<close>
by blast
lemma subset_imageE:
@@ -1000,9 +1000,9 @@
lemma image_Collect:
"f ` {x. P x} = {f x | x. P x}"
- -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
+ -- \<open>NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
with its implicit quantifier and conjunction. Also image enjoys better
- equational properties than does the RHS. *}
+ equational properties than does the RHS.\<close>
by blast
lemma if_image_distrib [simp]:
@@ -1036,9 +1036,9 @@
using assms by auto
-text {*
+text \<open>
\medskip Range of a function -- just a translation for image!
-*}
+\<close>
abbreviation range :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b set"
where -- "of function"
@@ -1065,9 +1065,9 @@
by auto
-subsubsection {* Some rules with @{text "if"} *}
-
-text{* Elimination of @{text"{x. \<dots> & x=t & \<dots>}"}. *}
+subsubsection \<open>Some rules with @{text "if"}\<close>
+
+text\<open>Elimination of @{text"{x. \<dots> & x=t & \<dots>}"}.\<close>
lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})"
by auto
@@ -1075,10 +1075,10 @@
lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})"
by auto
-text {*
+text \<open>
Rewrite rules for boolean case-splitting: faster than @{text
"split_if [split]"}.
-*}
+\<close>
lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))"
by (rule split_if)
@@ -1086,10 +1086,10 @@
lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))"
by (rule split_if)
-text {*
+text \<open>
Split ifs on either side of the membership relation. Not for @{text
"[simp]"} -- can cause goals to blow up!
-*}
+\<close>
lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))"
by (rule split_if)
@@ -1109,9 +1109,9 @@
*)
-subsection {* Further operations and lemmas *}
-
-subsubsection {* The ``proper subset'' relation *}
+subsection \<open>Further operations and lemmas\<close>
+
+subsubsection \<open>The ``proper subset'' relation\<close>
lemma psubsetI [intro!]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
by (unfold less_le) blast
@@ -1167,9 +1167,9 @@
using assms by (blast elim: subset_imageE)
-subsubsection {* Derived rules involving subsets. *}
-
-text {* @{text insert}. *}
+subsubsection \<open>Derived rules involving subsets.\<close>
+
+text \<open>@{text insert}.\<close>
lemma subset_insertI: "B \<subseteq> insert a B"
by (rule subsetI) (erule insertI2)
@@ -1181,7 +1181,7 @@
by blast
-text {* \medskip Finite Union -- the least upper bound of two sets. *}
+text \<open>\medskip Finite Union -- the least upper bound of two sets.\<close>
lemma Un_upper1: "A \<subseteq> A \<union> B"
by (fact sup_ge1)
@@ -1193,7 +1193,7 @@
by (fact sup_least)
-text {* \medskip Finite Intersection -- the greatest lower bound of two sets. *}
+text \<open>\medskip Finite Intersection -- the greatest lower bound of two sets.\<close>
lemma Int_lower1: "A \<inter> B \<subseteq> A"
by (fact inf_le1)
@@ -1205,7 +1205,7 @@
by (fact inf_greatest)
-text {* \medskip Set difference. *}
+text \<open>\medskip Set difference.\<close>
lemma Diff_subset: "A - B \<subseteq> A"
by blast
@@ -1214,12 +1214,12 @@
by blast
-subsubsection {* Equalities involving union, intersection, inclusion, etc. *}
-
-text {* @{text "{}"}. *}
+subsubsection \<open>Equalities involving union, intersection, inclusion, etc.\<close>
+
+text \<open>@{text "{}"}.\<close>
lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})"
- -- {* supersedes @{text "Collect_False_empty"} *}
+ -- \<open>supersedes @{text "Collect_False_empty"}\<close>
by auto
lemma subset_empty [simp]: "(A \<subseteq> {}) = (A = {})"
@@ -1250,10 +1250,10 @@
by blast
-text {* \medskip @{text insert}. *}
+text \<open>\medskip @{text insert}.\<close>
lemma insert_is_Un: "insert a A = {a} Un A"
- -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *}
+ -- \<open>NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"}\<close>
by blast
lemma insert_not_empty [simp]: "insert a A \<noteq> {}"
@@ -1263,8 +1263,8 @@
declare empty_not_insert [simp]
lemma insert_absorb: "a \<in> A ==> insert a A = A"
- -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *}
- -- {* with \emph{quadratic} running time *}
+ -- \<open>@{text "[simp]"} causes recursive calls when there are nested inserts\<close>
+ -- \<open>with \emph{quadratic} running time\<close>
by blast
lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A"
@@ -1277,7 +1277,7 @@
by blast
lemma mk_disjoint_insert: "a \<in> A ==> \<exists>B. A = insert a B & a \<notin> B"
- -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *}
+ -- \<open>use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding\<close>
apply (rule_tac x = "A - {a}" in exI, blast)
done
@@ -1298,7 +1298,7 @@
by auto
-text {* \medskip @{text Int} *}
+text \<open>\medskip @{text Int}\<close>
lemma Int_absorb: "A \<inter> A = A"
by (fact inf_idem) (* already simp *)
@@ -1316,7 +1316,7 @@
by (fact inf_assoc)
lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute
- -- {* Intersection is an AC-operator *}
+ -- \<open>Intersection is an AC-operator\<close>
lemma Int_absorb1: "B \<subseteq> A ==> A \<inter> B = B"
by (fact inf_absorb2)
@@ -1358,7 +1358,7 @@
by blast
-text {* \medskip @{text Un}. *}
+text \<open>\medskip @{text Un}.\<close>
lemma Un_absorb: "A \<union> A = A"
by (fact sup_idem) (* already simp *)
@@ -1376,7 +1376,7 @@
by (fact sup_assoc)
lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute
- -- {* Union is an AC-operator *}
+ -- \<open>Union is an AC-operator\<close>
lemma Un_absorb1: "A \<subseteq> B ==> A \<union> B = B"
by (fact sup_absorb2)
@@ -1452,7 +1452,7 @@
by blast
-text {* \medskip Set complement *}
+text \<open>\medskip Set complement\<close>
lemma Compl_disjoint [simp]: "A \<inter> -A = {}"
by (fact inf_compl_bot)
@@ -1479,7 +1479,7 @@
by blast
lemma Un_Int_assoc_eq: "((A \<inter> B) \<union> C = A \<inter> (B \<union> C)) = (C \<subseteq> A)"
- -- {* Halmos, Naive Set Theory, page 16. *}
+ -- \<open>Halmos, Naive Set Theory, page 16.\<close>
by blast
lemma Compl_UNIV_eq: "-UNIV = {}"
@@ -1497,10 +1497,10 @@
lemma Compl_insert: "- insert x A = (-A) - {x}"
by blast
-text {* \medskip Bounded quantifiers.
+text \<open>\medskip Bounded quantifiers.
The following are not added to the default simpset because
- (a) they duplicate the body and (b) there are no similar rules for @{text Int}. *}
+ (a) they duplicate the body and (b) there are no similar rules for @{text Int}.\<close>
lemma ball_Un: "(\<forall>x \<in> A \<union> B. P x) = ((\<forall>x\<in>A. P x) & (\<forall>x\<in>B. P x))"
by blast
@@ -1509,7 +1509,7 @@
by blast
-text {* \medskip Set difference. *}
+text \<open>\medskip Set difference.\<close>
lemma Diff_eq: "A - B = A \<inter> (-B)"
by blast
@@ -1539,11 +1539,11 @@
by blast
lemma Diff_insert: "A - insert a B = A - B - {a}"
- -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
+ -- \<open>NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"}\<close>
by blast
lemma Diff_insert2: "A - insert a B = A - {a} - B"
- -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"} *}
+ -- \<open>NOT SUITABLE FOR REWRITING since @{text "{a} == insert a 0"}\<close>
by blast
lemma insert_Diff_if: "insert x A - B = (if x \<in> B then A - B else insert x (A - B))"
@@ -1601,7 +1601,7 @@
by blast
-text {* \medskip Quantification over type @{typ bool}. *}
+text \<open>\medskip Quantification over type @{typ bool}.\<close>
lemma bool_induct: "P True \<Longrightarrow> P False \<Longrightarrow> P x"
by (cases x) auto
@@ -1618,7 +1618,7 @@
lemma UNIV_bool: "UNIV = {False, True}"
by (auto intro: bool_induct)
-text {* \medskip @{text Pow} *}
+text \<open>\medskip @{text Pow}\<close>
lemma Pow_empty [simp]: "Pow {} = {{}}"
by (auto simp add: Pow_def)
@@ -1642,7 +1642,7 @@
by blast
-text {* \medskip Miscellany. *}
+text \<open>\medskip Miscellany.\<close>
lemma set_eq_subset: "(A = B) = (A \<subseteq> B & B \<subseteq> A)"
by blast
@@ -1684,7 +1684,7 @@
by auto
-subsubsection {* Monotonicity of various operations *}
+subsubsection \<open>Monotonicity of various operations\<close>
lemma image_mono: "A \<subseteq> B ==> f`A \<subseteq> f`B"
by blast
@@ -1707,7 +1707,7 @@
lemma Compl_anti_mono: "A \<subseteq> B ==> -B \<subseteq> -A"
by (fact compl_mono)
-text {* \medskip Monotonicity of implications. *}
+text \<open>\medskip Monotonicity of implications.\<close>
lemma in_mono: "A \<subseteq> B ==> x \<in> A --> x \<in> B"
apply (rule impI)
@@ -1749,7 +1749,7 @@
by iprover
-subsubsection {* Inverse image of a function *}
+subsubsection \<open>Inverse image of a function\<close>
definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where
"f -` B == {x. f x : B}"
@@ -1791,7 +1791,7 @@
by blast
lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)"
- -- {* NOT suitable for rewriting because of the recurrence of @{term "{a}"}. *}
+ -- \<open>NOT suitable for rewriting because of the recurrence of @{term "{a}"}.\<close>
by blast
lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)"
@@ -1801,7 +1801,7 @@
by blast
lemma vimage_mono: "A \<subseteq> B ==> f -` A \<subseteq> f -` B"
- -- {* monotonicity *}
+ -- \<open>monotonicity\<close>
by blast
lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}"
@@ -1832,7 +1832,7 @@
by blast
-subsubsection {* Getting the Contents of a Singleton Set *}
+subsubsection \<open>Getting the Contents of a Singleton Set\<close>
definition the_elem :: "'a set \<Rightarrow> 'a" where
"the_elem X = (THE x. X = {x})"
@@ -1845,20 +1845,20 @@
assumes *: "\<And>y. y \<in> A \<Longrightarrow> f y = f x"
shows "the_elem (f ` A) = f x"
unfolding the_elem_def proof (rule the1_equality)
- from `A \<noteq> {}` obtain y where "y \<in> A" by auto
+ from \<open>A \<noteq> {}\<close> obtain y where "y \<in> A" by auto
with * have "f x = f y" by simp
- with `y \<in> A` have "f x \<in> f ` A" by blast
+ with \<open>y \<in> A\<close> have "f x \<in> f ` A" by blast
with * show "f ` A = {f x}" by auto
then show "\<exists>!x. f ` A = {x}" by auto
qed
-subsubsection {* Least value operator *}
+subsubsection \<open>Least value operator\<close>
lemma Least_mono:
"mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y
==> (LEAST y. y : f ` S) = f (LEAST x. x : S)"
- -- {* Courtesy of Stephan Merz *}
+ -- \<open>Courtesy of Stephan Merz\<close>
apply clarify
apply (erule_tac P = "%x. x : S" in LeastI2_order, fast)
apply (rule LeastI2_order)
@@ -1866,7 +1866,7 @@
done
-subsubsection {* Monad operation *}
+subsubsection \<open>Monad operation\<close>
definition bind :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
"bind A f = {x. \<exists>B \<in> f`A. x \<in> B}"
@@ -1892,7 +1892,7 @@
lemma bind_singleton_conv_image: "Set.bind A (\<lambda>x. {f x}) = f ` A"
by(auto simp add: bind_def)
-subsubsection {* Operations for execution *}
+subsubsection \<open>Operations for execution\<close>
definition is_empty :: "'a set \<Rightarrow> bool" where
[code_abbrev]: "is_empty A \<longleftrightarrow> A = {}"
@@ -1929,13 +1929,13 @@
end
-text {* Misc *}
+text \<open>Misc\<close>
hide_const (open) member not_member
lemmas equalityI = subset_antisym
-ML {*
+ML \<open>
val Ball_def = @{thm Ball_def}
val Bex_def = @{thm Bex_def}
val CollectD = @{thm CollectD}
@@ -1988,7 +1988,7 @@
val vimage_Collect = @{thm vimage_Collect}
val vimage_Int = @{thm vimage_Int}
val vimage_Un = @{thm vimage_Un}
-*}
+\<close>
end
--- a/src/HOL/Set_Interval.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Set_Interval.thy Sat Jul 18 22:58:50 2015 +0200
@@ -11,7 +11,7 @@
Examples: Ico = {_ ..< _} and Ici = {_ ..}
*)
-section {* Set intervals *}
+section \<open>Set intervals\<close>
theory Set_Interval
imports Lattices_Big Nat_Transfer
@@ -55,9 +55,9 @@
end
-text{* A note of warning when using @{term"{..<n}"} on type @{typ
+text\<open>A note of warning when using @{term"{..<n}"} on type @{typ
nat}: it is equivalent to @{term"{0::nat..<n}"} but some lemmas involving
-@{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *}
+@{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well.\<close>
syntax
"_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10)
@@ -84,7 +84,7 @@
"INT i<n. A" == "INT i:{..<n}. A"
-subsection {* Various equivalences *}
+subsection \<open>Various equivalences\<close>
lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i<k)"
by (simp add: lessThan_def)
@@ -128,7 +128,7 @@
lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \<inter> {..< b} = {..< min a b}"
by auto
-subsection {* Logical Equivalences for Set Inclusion and Equality *}
+subsection \<open>Logical Equivalences for Set Inclusion and Equality\<close>
lemma atLeast_subset_iff [iff]:
"(atLeast x \<subseteq> atLeast y) = (y \<le> (x::'a::order))"
@@ -181,7 +181,7 @@
lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \<subseteq> {..< b} \<longleftrightarrow> a < b"
by auto
-subsection {*Two-sided intervals*}
+subsection \<open>Two-sided intervals\<close>
context ord
begin
@@ -202,16 +202,16 @@
"(i : {l..u}) = (l <= i & i <= u)"
by (simp add: atLeastAtMost_def)
-text {* The above four lemmas could be declared as iffs. Unfortunately this
+text \<open>The above four lemmas could be declared as iffs. Unfortunately this
breaks many proofs. Since it only helps blast, it is better to leave them
-alone. *}
+alone.\<close>
lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \<inter> {..< b }"
by auto
end
-subsubsection{* Emptyness, singletons, subset *}
+subsubsection\<open>Emptyness, singletons, subset\<close>
context order
begin
@@ -276,7 +276,7 @@
proof
assume "{a..b} = {c}"
hence *: "\<not> (\<not> a \<le> b)" unfolding atLeastatMost_empty_iff[symmetric] by simp
- with `{a..b} = {c}` have "c \<le> a \<and> b \<le> c" by auto
+ with \<open>{a..b} = {c}\<close> have "c \<le> a \<and> b \<le> c" by auto
with * show "a = b \<and> b = c" by auto
qed simp
@@ -502,7 +502,7 @@
qed
-subsection {* Infinite intervals *}
+subsection \<open>Infinite intervals\<close>
context dense_linorder
begin
@@ -513,16 +513,16 @@
proof
assume fin: "finite {a<..<b}"
moreover have ne: "{a<..<b} \<noteq> {}"
- using `a < b` by auto
+ using \<open>a < b\<close> by auto
ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b"
using Max_in[of "{a <..< b}"] by auto
then obtain x where "Max {a <..< b} < x" "x < b"
using dense[of "Max {a<..<b}" b] by auto
then have "x \<in> {a <..< b}"
- using `a < Max {a <..< b}` by auto
+ using \<open>a < Max {a <..< b}\<close> by auto
then have "x \<le> Max {a <..< b}"
using fin by auto
- with `Max {a <..< b} < x` show False by auto
+ with \<open>Max {a <..< b} < x\<close> show False by auto
qed
lemma infinite_Icc: "a < b \<Longrightarrow> \<not> finite {a .. b}"
@@ -550,11 +550,11 @@
obtain y where "y < Min {..< a}"
using lt_ex by auto
also have "Min {..< a} \<le> x"
- using `x < a` by fact
- also note `x < a`
+ using \<open>x < a\<close> by fact
+ also note \<open>x < a\<close>
finally have "Min {..< a} \<le> y"
by fact
- with `y < Min {..< a}` show False by auto
+ with \<open>y < Min {..< a}\<close> show False by auto
qed
lemma infinite_Iic: "\<not> finite {.. a :: 'a :: {no_bot, linorder}}"
@@ -574,17 +574,17 @@
using gt_ex by auto
also then have "x \<le> Max {a <..}"
by fact
- also note `Max {a <..} < y`
+ also note \<open>Max {a <..} < y\<close>
finally have "y \<le> Max { a <..}"
by fact
- with `Max {a <..} < y` show False by auto
+ with \<open>Max {a <..} < y\<close> show False by auto
qed
lemma infinite_Ici: "\<not> finite {a :: 'a :: {no_top, linorder} ..}"
using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"]
by (auto simp: subset_eq less_imp_le)
-subsubsection {* Intersection *}
+subsubsection \<open>Intersection\<close>
context linorder
begin
@@ -652,9 +652,9 @@
and Inf_greaterThanLessThan[simp]: "x < y \<Longrightarrow> Inf { x <..< y} = x"
by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded)
-subsection {* Intervals of natural numbers *}
+subsection \<open>Intervals of natural numbers\<close>
-subsubsection {* The Constant @{term lessThan} *}
+subsubsection \<open>The Constant @{term lessThan}\<close>
lemma lessThan_0 [simp]: "lessThan (0::nat) = {}"
by (simp add: lessThan_def)
@@ -662,9 +662,9 @@
lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)"
by (simp add: lessThan_def less_Suc_eq, blast)
-text {* The following proof is convenient in induction proofs where
+text \<open>The following proof is convenient in induction proofs where
new elements get indices at the beginning. So it is used to transform
-@{term "{..<Suc n}"} to @{term "0::nat"} and @{term "{..< n}"}. *}
+@{term "{..<Suc n}"} to @{term "0::nat"} and @{term "{..< n}"}.\<close>
lemma zero_notin_Suc_image: "0 \<notin> Suc ` A"
by auto
@@ -681,7 +681,7 @@
lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV"
by blast
-subsubsection {* The Constant @{term greaterThan} *}
+subsubsection \<open>The Constant @{term greaterThan}\<close>
lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc"
apply (simp add: greaterThan_def)
@@ -696,7 +696,7 @@
lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}"
by blast
-subsubsection {* The Constant @{term atLeast} *}
+subsubsection \<open>The Constant @{term atLeast}\<close>
lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV"
by (unfold atLeast_def UNIV_def, simp)
@@ -713,7 +713,7 @@
lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV"
by blast
-subsubsection {* The Constant @{term atMost} *}
+subsubsection \<open>The Constant @{term atMost}\<close>
lemma atMost_0 [simp]: "atMost (0::nat) = {0}"
by (simp add: atMost_def)
@@ -726,14 +726,14 @@
lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV"
by blast
-subsubsection {* The Constant @{term atLeastLessThan} *}
+subsubsection \<open>The Constant @{term atLeastLessThan}\<close>
-text{*The orientation of the following 2 rules is tricky. The lhs is
+text\<open>The orientation of the following 2 rules is tricky. The lhs is
defined in terms of the rhs. Hence the chosen orientation makes sense
in this theory --- the reverse orientation complicates proofs (eg
nontermination). But outside, when the definition of the lhs is rarely
used, the opposite orientation seems preferable because it reduces a
-specific concept to a more general one. *}
+specific concept to a more general one.\<close>
lemma atLeast0LessThan: "{0::nat..<n} = {..<n}"
by(simp add:lessThan_def atLeastLessThan_def)
@@ -747,9 +747,9 @@
lemma atLeastLessThan0: "{m..<0::nat} = {}"
by (simp add: atLeastLessThan_def)
-subsubsection {* Intervals of nats with @{term Suc} *}
+subsubsection \<open>Intervals of nats with @{term Suc}\<close>
-text{*Not a simprule because the RHS is too messy.*}
+text\<open>Not a simprule because the RHS is too messy.\<close>
lemma atLeastLessThanSuc:
"{m..<Suc n} = (if m \<le> n then insert n {m..<n} else {})"
by (auto simp add: atLeastLessThan_def)
@@ -780,7 +780,7 @@
lemma atLeastAtMost_insertL: "m \<le> n \<Longrightarrow> insert m {Suc m..n} = {m ..n}"
by auto
-text {* The analogous result is useful on @{typ int}: *}
+text \<open>The analogous result is useful on @{typ int}:\<close>
(* here, because we don't have an own int section *)
lemma atLeastAtMostPlus1_int_conv:
"m <= 1+n \<Longrightarrow> {m..1+n} = insert (1+n) {m..n::int}"
@@ -791,23 +791,23 @@
apply (simp_all add: atLeastLessThanSuc)
done
-subsubsection {* Intervals and numerals *}
+subsubsection \<open>Intervals and numerals\<close>
-lemma lessThan_nat_numeral: --{*Evaluation for specific numerals*}
+lemma lessThan_nat_numeral: --\<open>Evaluation for specific numerals\<close>
"lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))"
by (simp add: numeral_eq_Suc lessThan_Suc)
-lemma atMost_nat_numeral: --{*Evaluation for specific numerals*}
+lemma atMost_nat_numeral: --\<open>Evaluation for specific numerals\<close>
"atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))"
by (simp add: numeral_eq_Suc atMost_Suc)
-lemma atLeastLessThan_nat_numeral: --{*Evaluation for specific numerals*}
+lemma atLeastLessThan_nat_numeral: --\<open>Evaluation for specific numerals\<close>
"atLeastLessThan m (numeral k :: nat) =
(if m \<le> (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k))
else {})"
by (simp add: numeral_eq_Suc atLeastLessThanSuc)
-subsubsection {* Image *}
+subsubsection \<open>Image\<close>
lemma image_add_atLeastAtMost:
fixes k ::"'a::linordered_semidom"
@@ -898,7 +898,7 @@
next
fix y assume "y \<le> -x"
have "- (-y) \<in> uminus ` {x..}"
- by (rule imageI) (insert `y \<le> -x`[THEN le_imp_neg_le], simp)
+ by (rule imageI) (insert \<open>y \<le> -x\<close>[THEN le_imp_neg_le], simp)
thus "y \<in> uminus ` {x..}" by simp
qed simp_all
@@ -924,7 +924,7 @@
greaterThanLessThan_def image_Int[OF inj_uminus] Int_commute)
end
-subsubsection {* Finiteness *}
+subsubsection \<open>Finiteness\<close>
lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..<k}"
by (induct k) (simp_all add: lessThan_Suc)
@@ -948,21 +948,21 @@
fixes l :: nat shows "finite {l..u}"
by (simp add: atLeastAtMost_def)
-text {* A bounded set of natural numbers is finite. *}
+text \<open>A bounded set of natural numbers is finite.\<close>
lemma bounded_nat_set_is_finite:
"(ALL i:N. i < (n::nat)) ==> finite N"
apply (rule finite_subset)
apply (rule_tac [2] finite_lessThan, auto)
done
-text {* A set of natural numbers is finite iff it is bounded. *}
+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")
proof
assume f:?F show ?B
- using Max_ge[OF `?F`, simplified less_Suc_eq_le[symmetric]] by blast
+ using Max_ge[OF \<open>?F\<close>, simplified less_Suc_eq_le[symmetric]] by blast
next
- assume ?B show ?F using `?B` by(blast intro:bounded_nat_set_is_finite)
+ assume ?B show ?F using \<open>?B\<close> by(blast intro:bounded_nat_set_is_finite)
qed
lemma finite_nat_set_iff_bounded_le:
@@ -976,8 +976,8 @@
by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans)
-text{* Any subset of an interval of natural numbers the size of the
-subset is exactly that interval. *}
+text\<open>Any subset of an interval of natural numbers the size of the
+subset is exactly that interval.\<close>
lemma subset_card_intvl_is_intvl:
assumes "A \<subseteq> {k..<k + card A}"
@@ -1000,7 +1000,7 @@
qed
-subsubsection {* Proving Inclusions and Equalities between Unions *}
+subsubsection \<open>Proving Inclusions and Equalities between Unions\<close>
lemma UN_le_eq_Un0:
"(\<Union>i\<le>n::nat. M i) = (\<Union>i\<in>{1..n}. M i) \<union> M 0" (is "?A = ?B")
@@ -1056,7 +1056,7 @@
done
-subsubsection {* Cardinality *}
+subsubsection \<open>Cardinality\<close>
lemma card_lessThan [simp]: "card {..<u} = u"
by (induct u, simp_all add: lessThan_Suc)
@@ -1141,7 +1141,7 @@
ultimately show "(\<exists>f. inj_on f A \<and> f ` A \<le> B)" by blast
qed (insert assms, auto)
-subsection {* Intervals of integers *}
+subsection \<open>Intervals of integers\<close>
lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..<u+1} = {l..(u::int)}"
by (auto simp add: atLeastAtMost_def atLeastLessThan_def)
@@ -1153,7 +1153,7 @@
"{l+1..<u} = {l<..<u::int}"
by (auto simp add: atLeastLessThan_def greaterThanLessThan_def)
-subsubsection {* Finiteness *}
+subsubsection \<open>Finiteness\<close>
lemma image_atLeastZeroLessThan_int: "0 \<le> u ==>
{(0::int)..<u} = int ` {..<nat u}"
@@ -1188,7 +1188,7 @@
by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
-subsubsection {* Cardinality *}
+subsubsection \<open>Cardinality\<close>
lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u"
apply (cases "0 \<le> u")
@@ -1265,13 +1265,13 @@
qed
-subsection {*Lemmas useful with the summation operator setsum*}
+subsection \<open>Lemmas useful with the summation operator setsum\<close>
-text {* For examples, see Algebra/poly/UnivPoly2.thy *}
+text \<open>For examples, see Algebra/poly/UnivPoly2.thy\<close>
-subsubsection {* Disjoint Unions *}
+subsubsection \<open>Disjoint Unions\<close>
-text {* Singletons and open intervals *}
+text \<open>Singletons and open intervals\<close>
lemma ivl_disj_un_singleton:
"{l::'a::linorder} Un {l<..} = {l..}"
@@ -1282,7 +1282,7 @@
"(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}"
by auto
-text {* One- and two-sided intervals *}
+text \<open>One- and two-sided intervals\<close>
lemma ivl_disj_un_one:
"(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}"
@@ -1295,7 +1295,7 @@
"(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}"
by auto
-text {* Two- and two-sided intervals *}
+text \<open>Two- and two-sided intervals\<close>
lemma ivl_disj_un_two:
"[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
@@ -1317,9 +1317,9 @@
lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch
-subsubsection {* Disjoint Intersections *}
+subsubsection \<open>Disjoint Intersections\<close>
-text {* One- and two-sided intervals *}
+text \<open>One- and two-sided intervals\<close>
lemma ivl_disj_int_one:
"{..l::'a::order} Int {l<..<u} = {}"
@@ -1332,7 +1332,7 @@
"{l..<u} Int {u..} = {}"
by auto
-text {* Two- and two-sided intervals *}
+text \<open>Two- and two-sided intervals\<close>
lemma ivl_disj_int_two:
"{l::'a::order<..<m} Int {m..<u} = {}"
@@ -1347,7 +1347,7 @@
lemmas ivl_disj_int = ivl_disj_int_one ivl_disj_int_two
-subsubsection {* Some Differences *}
+subsubsection \<open>Some Differences\<close>
lemma ivl_diff[simp]:
"i \<le> n \<Longrightarrow> {i..<m} - {i..<n} = {n..<(m::'a::linorder)}"
@@ -1358,7 +1358,7 @@
by auto
-subsubsection {* Some Subset Conditions *}
+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))"
@@ -1370,7 +1370,7 @@
done
-subsection {* Summation indexed over intervals *}
+subsection \<open>Summation indexed over intervals\<close>
syntax
"_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
@@ -1403,7 +1403,7 @@
"\<Sum>i\<le>n. t" == "CONST setsum (\<lambda>i. t) {..n}"
"\<Sum>i<n. t" == "CONST setsum (\<lambda>i. t) {..<n}"
-text{* The above introduces some pretty alternative syntaxes for
+text\<open>The above introduces some pretty alternative syntaxes for
summation over intervals:
\begin{center}
\begin{tabular}{lll}
@@ -1425,12 +1425,12 @@
Note that for uniformity on @{typ nat} it is better to use
@{term"\<Sum>x::nat=0..<n. e"} rather than @{text"\<Sum>x<n. e"}: @{text setsum} may
not provide all lemmas available for @{term"{m..<n}"} also in the
-special form for @{term"{..<n}"}. *}
+special form for @{term"{..<n}"}.\<close>
-text{* This congruence rule should be used for sums over intervals as
+text\<open>This congruence rule should be used for sums over intervals as
the standard theorem @{text[source]setsum.cong} does not work well
with the simplifier who adds the unsimplified premise @{term"x:B"} to
-the context. *}
+the context.\<close>
lemma setsum_ivl_cong:
"\<lbrakk>a = c; b = d; !!x. \<lbrakk> c \<le> x; x < d \<rbrakk> \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow>
@@ -1486,7 +1486,7 @@
lemma setsum_ub_add_nat: assumes "(m::nat) \<le> n + 1"
shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}"
proof-
- have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using `m \<le> n+1` by auto
+ have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using \<open>m \<le> n+1\<close> by auto
thus ?thesis by (auto simp: ivl_disj_int setsum.union_disjoint
atLeastSucAtMost_greaterThanAtMost)
qed
@@ -1532,7 +1532,7 @@
lemma nat_diff_setsum_reindex: "(\<Sum>i<n. f (n - Suc i)) = (\<Sum>i<n. f i)"
by (rule setsum.reindex_bij_witness[where i="\<lambda>i. n - Suc i" and j="\<lambda>i. n - Suc i"]) auto
-subsection{* Shifting bounds *}
+subsection\<open>Shifting bounds\<close>
lemma setsum_shift_bounds_nat_ivl:
"setsum f {m+k..<n+k} = setsum (%i. f(i + k)){m..<n::nat}"
@@ -1603,7 +1603,7 @@
by auto
-subsection {* The formula for geometric sums *}
+subsection \<open>The formula for geometric sums\<close>
lemma geometric_sum:
assumes "x \<noteq> 1"
@@ -1611,7 +1611,7 @@
proof -
from assms obtain y where "y = x - 1" and "y \<noteq> 0" by simp_all
moreover have "(\<Sum>i<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
- by (induct n) (simp_all add: power_Suc field_simps `y \<noteq> 0`)
+ by (induct n) (simp_all add: power_Suc field_simps \<open>y \<noteq> 0\<close>)
ultimately show ?thesis by simp
qed
@@ -1635,7 +1635,7 @@
finally show ?case .
qed simp
-corollary power_diff_sumr2: --{* @{text COMPLEX_POLYFUN} in HOL Light *}
+corollary power_diff_sumr2: --\<open>@{text COMPLEX_POLYFUN} in HOL Light\<close>
fixes x :: "'a::{comm_ring,monoid_mult}"
shows "x^n - y^n = (x - y) * (\<Sum>i<n. y^(n - Suc i) * x^i)"
using diff_power_eq_setsum[of x "n - 1" y]
@@ -1659,7 +1659,7 @@
by (metis one_diff_power_eq' [of n x] nat_diff_setsum_reindex)
-subsection {* The formula for arithmetic sums *}
+subsection \<open>The formula for arithmetic sums\<close>
lemma gauss_sum:
"(2::'a::comm_semiring_1)*(\<Sum>i\<in>{1..n}. of_nat i) = of_nat n*((of_nat n)+1)"
@@ -1721,7 +1721,7 @@
lemma sum_diff_distrib: "\<forall>x. Q x \<le> P x \<Longrightarrow> (\<Sum>x<n. P x) - (\<Sum>x<n. Q x) = (\<Sum>x<n. P x - Q x :: nat)"
by (subst setsum_subtractf_nat) auto
-subsection {* Products indexed over intervals *}
+subsection \<open>Products indexed over intervals\<close>
syntax
"_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10)
@@ -1754,7 +1754,7 @@
"\<Prod>i\<le>n. t" == "CONST setprod (\<lambda>i. t) {..n}"
"\<Prod>i<n. t" == "CONST setprod (\<lambda>i. t) {..<n}"
-subsection {* Transfer setup *}
+subsection \<open>Transfer setup\<close>
lemma transfer_nat_int_set_functions:
"{..n} = nat ` {0..int n}"
--- a/src/HOL/Sledgehammer.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Sledgehammer.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* Sledgehammer: Isabelle--ATP Linkup *}
+section \<open>Sledgehammer: Isabelle--ATP Linkup\<close>
theory Sledgehammer
imports Presburger SMT
--- a/src/HOL/String.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/String.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,12 +1,12 @@
(* Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *)
-section {* Character and string types *}
+section \<open>Character and string types\<close>
theory String
imports Enum
begin
-subsection {* Characters and strings *}
+subsection \<open>Characters and strings\<close>
datatype nibble =
Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7
@@ -246,7 +246,7 @@
"nat_of_char (Char x y) = nat_of_nibble x * 16 + nat_of_nibble y"
by (simp add: nat_of_char_def)
-setup {*
+setup \<open>
let
val nibbles = map_range (Thm.cterm_of @{context} o HOLogic.mk_nibble) 16;
val simpset =
@@ -261,7 +261,7 @@
[((@{binding nat_of_char_simps}, []), [(code_eqns, [])])]
#> snd
end
-*}
+\<close>
declare nat_of_char_simps [code]
@@ -281,7 +281,7 @@
with Char show ?thesis by (simp add: nat_of_char_def add.commute)
qed
-code_datatype Char -- {* drop case certificate for char *}
+code_datatype Char -- \<open>drop case certificate for char\<close>
lemma case_char_code [code]:
"case_char f c = (let n = nat_of_char c in f (nibble_of_nat (n div 16)) (nibble_of_nat n))"
@@ -354,7 +354,7 @@
qed
-subsection {* Strings as dedicated type *}
+subsection \<open>Strings as dedicated type\<close>
typedef literal = "UNIV :: string set"
morphisms explode STR ..
@@ -399,7 +399,7 @@
lifting_update literal.lifting
lifting_forget literal.lifting
-subsection {* Code generator *}
+subsection \<open>Code generator\<close>
ML_file "Tools/string_code.ML"
@@ -414,9 +414,9 @@
and (Haskell) "String"
and (Scala) "String"
-setup {*
+setup \<open>
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"]
-*}
+\<close>
code_printing
class_instance literal :: equal \<rightharpoonup>
@@ -427,7 +427,7 @@
and (Haskell) infix 4 "=="
and (Scala) infixl 5 "=="
-setup {* Sign.map_naming (Name_Space.mandatory_path "Code") *}
+setup \<open>Sign.map_naming (Name_Space.mandatory_path "Code")\<close>
definition abort :: "literal \<Rightarrow> (unit \<Rightarrow> 'a) \<Rightarrow> 'a"
where [simp, code del]: "abort _ f = f ()"
@@ -435,9 +435,9 @@
lemma abort_cong: "msg = msg' ==> Code.abort msg f = Code.abort msg' f"
by simp
-setup {* Sign.map_naming Name_Space.parent_path *}
+setup \<open>Sign.map_naming Name_Space.parent_path\<close>
-setup {* Code_Simp.map_ss (Simplifier.add_cong @{thm Code.abort_cong}) *}
+setup \<open>Code_Simp.map_ss (Simplifier.add_cong @{thm Code.abort_cong})\<close>
code_printing constant Code.abort \<rightharpoonup>
(SML) "!(raise/ Fail/ _)"
--- a/src/HOL/Sum_Type.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Sum_Type.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1992 University of Cambridge
*)
-section{*The Disjoint Sum of Two Types*}
+section\<open>The Disjoint Sum of Two Types\<close>
theory Sum_Type
imports Typedef Inductive Fun
begin
-subsection {* Construction of the sum type and its basic abstract operations *}
+subsection \<open>Construction of the sum type and its basic abstract operations\<close>
definition Inl_Rep :: "'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool \<Rightarrow> bool" where
"Inl_Rep a x y p \<longleftrightarrow> x = a \<and> p"
@@ -90,9 +90,9 @@
| Inr projr
by (erule sumE, assumption) (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)
-text {* Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}. *}
+text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
-setup {* Sign.mandatory_path "old" *}
+setup \<open>Sign.mandatory_path "old"\<close>
old_rep_datatype Inl Inr
proof -
@@ -102,11 +102,11 @@
then show "P s" by (auto intro: sumE [of s])
qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
-text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
+text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>
-setup {* Sign.mandatory_path "sum" *}
+setup \<open>Sign.mandatory_path "sum"\<close>
declare
old.sum.inject[iff del]
@@ -117,7 +117,7 @@
lemmas rec = old.sum.rec
lemmas simps = sum.inject sum.distinct sum.case sum.rec
-setup {* Sign.parent_path *}
+setup \<open>Sign.parent_path\<close>
primrec map_sum :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd" where
"map_sum f1 f2 (Inl a) = Inl (f1 a)"
@@ -147,7 +147,7 @@
lemma split_sum_ex: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P (Inl x)) \<or> (\<exists>x. P (Inr x))"
using split_sum_all[of "\<lambda>x. \<not>P x"] by blast
-subsection {* Projections *}
+subsection \<open>Projections\<close>
lemma case_sum_KK [simp]: "case_sum (\<lambda>x. a) (\<lambda>x. a) = (\<lambda>x. a)"
by (rule ext) (simp split: sum.split)
@@ -201,7 +201,7 @@
qed
-subsection {* The Disjoint Sum of Sets *}
+subsection \<open>The Disjoint Sum of Sets\<close>
definition Plus :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a + 'b) set" (infixr "<+>" 65) where
"A <+> B = Inl ` A \<union> Inr ` B"
@@ -214,7 +214,7 @@
lemma InrI [intro!]: "b \<in> B \<Longrightarrow> Inr b \<in> A <+> B"
by (simp add: Plus_def)
-text {* Exhaustion rule for sums, a degenerate form of induction *}
+text \<open>Exhaustion rule for sums, a degenerate form of induction\<close>
lemma PlusE [elim!]:
"u \<in> A <+> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> u = Inl x \<Longrightarrow> P) \<Longrightarrow> (\<And>y. y \<in> B \<Longrightarrow> u = Inr y \<Longrightarrow> P) \<Longrightarrow> P"
--- a/src/HOL/Taylor.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Taylor.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,16 +2,16 @@
Author: Lukas Bulwahn, Bernhard Haeupler, Technische Universitaet Muenchen
*)
-section {* Taylor series *}
+section \<open>Taylor series\<close>
theory Taylor
imports MacLaurin
begin
-text {*
+text \<open>
We use MacLaurin and the translation of the expansion point @{text c} to @{text 0}
to prove Taylor's theorem.
-*}
+\<close>
lemma taylor_up:
assumes INIT: "n>0" "diff 0 = f"
--- a/src/HOL/Topological_Spaces.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Topological_Spaces.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,7 +3,7 @@
Author: Johannes Hölzl
*)
-section {* Topological Spaces *}
+section \<open>Topological Spaces\<close>
theory Topological_Spaces
imports Main Conditionally_Complete_Lattices
@@ -12,7 +12,7 @@
named_theorems continuous_intros "structural introduction rules for continuity"
-subsection {* Topological space *}
+subsection \<open>Topological space\<close>
class "open" =
fixes "open" :: "'a set \<Rightarrow> bool"
@@ -131,7 +131,7 @@
end
-subsection{* Hausdorff and other separation properties *}
+subsection\<open>Hausdorff and other separation properties\<close>
class t0_space = topological_space +
assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
@@ -172,7 +172,7 @@
shows "finite S \<Longrightarrow> closed S"
by (induct set: finite, simp_all)
-text {* T2 spaces are also known as Hausdorff spaces. *}
+text \<open>T2 spaces are also known as Hausdorff spaces.\<close>
class t2_space = topological_space +
assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
@@ -190,13 +190,13 @@
shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
using t0_space[of x y] by blast
-text {* A perfect space is a topological space with no isolated points. *}
+text \<open>A perfect space is a topological space with no isolated points.\<close>
class perfect_space = topological_space +
assumes not_open_singleton: "\<not> open {x}"
-subsection {* Generators for toplogies *}
+subsection \<open>Generators for toplogies\<close>
inductive generate_topology for S where
UNIV: "generate_topology S UNIV"
@@ -214,7 +214,7 @@
"class.topological_space (generate_topology S)"
by default (auto intro: generate_topology.intros)
-subsection {* Order topologies *}
+subsection \<open>Order topologies\<close>
class order_topology = order + "open" +
assumes open_generated_order: "open = generate_topology (range (\<lambda>a. {..< a}) \<union> range (\<lambda>a. {a <..}))"
@@ -262,7 +262,7 @@
then show ?thesis by blast
next
case False
- with `x < y` have "x \<in> {..< y} \<and> y \<in> {x <..} \<and> {x <..} \<inter> {..< y} = {}"
+ with \<open>x < y\<close> have "x \<in> {..< y} \<and> y \<in> {x <..} \<and> {x <..} \<inter> {..< y} = {}"
by auto
then show ?thesis by blast
qed
@@ -297,9 +297,9 @@
case (Basis S) then show ?case by (fastforce intro: exI[of _ y] lt_ex)
qed blast+
-subsubsection {* Boolean is an order topology *}
+subsubsection \<open>Boolean is an order topology\<close>
-text {* It also is a discrete topology, but don't have a type class for it (yet). *}
+text \<open>It also is a discrete topology, but don't have a type class for it (yet).\<close>
instantiation bool :: order_topology
begin
@@ -322,7 +322,7 @@
by auto
qed
-subsubsection {* Topological filters *}
+subsubsection \<open>Topological filters\<close>
definition (in topological_space) nhds :: "'a \<Rightarrow> 'a filter"
where "nhds a = (INF S:{S. open S \<and> a \<in> S}. principal S)"
@@ -458,7 +458,7 @@
"eventually P (at (x::'a::linorder_topology)) \<longleftrightarrow> eventually P (at_left x) \<and> eventually P (at_right x)"
by (subst at_eq_sup_left_right) (simp add: eventually_sup)
-subsubsection {* Tendsto *}
+subsubsection \<open>Tendsto\<close>
abbreviation (in topological_space)
tendsto :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" (infixr "--->" 55) where
@@ -471,12 +471,12 @@
by simp
named_theorems tendsto_intros "introduction rules for tendsto"
-setup {*
+setup \<open>
Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros},
fn context =>
Named_Theorems.get (Context.proof_of context) @{named_theorems tendsto_intros}
|> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm])))
-*}
+\<close>
lemma (in topological_space) tendsto_def:
"(f ---> l) F \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) F)"
@@ -562,20 +562,20 @@
proof (rule ccontr)
assume "a \<noteq> b"
obtain U V where "open U" "open V" "a \<in> U" "b \<in> V" "U \<inter> V = {}"
- using hausdorff [OF `a \<noteq> b`] by fast
+ using hausdorff [OF \<open>a \<noteq> b\<close>] by fast
have "eventually (\<lambda>x. f x \<in> U) F"
- using `(f ---> a) F` `open U` `a \<in> U` by (rule topological_tendstoD)
+ using \<open>(f ---> a) F\<close> \<open>open U\<close> \<open>a \<in> U\<close> by (rule topological_tendstoD)
moreover
have "eventually (\<lambda>x. f x \<in> V) F"
- using `(f ---> b) F` `open V` `b \<in> V` by (rule topological_tendstoD)
+ using \<open>(f ---> b) F\<close> \<open>open V\<close> \<open>b \<in> V\<close> by (rule topological_tendstoD)
ultimately
have "eventually (\<lambda>x. False) F"
proof eventually_elim
case (elim x)
hence "f x \<in> U \<inter> V" by simp
- with `U \<inter> V = {}` show ?case by simp
+ with \<open>U \<inter> V = {}\<close> show ?case by simp
qed
- with `\<not> trivial_limit F` show "False"
+ with \<open>\<not> trivial_limit F\<close> show "False"
by (simp add: trivial_limit_def)
qed
@@ -642,7 +642,7 @@
shows "a \<ge> x"
by (rule tendsto_le [OF F tendsto_const x a])
-subsubsection {* Rules about @{const Lim} *}
+subsubsection \<open>Rules about @{const Lim}\<close>
lemma tendsto_Lim:
"\<not>(trivial_limit net) \<Longrightarrow> (f ---> l) net \<Longrightarrow> Lim net f = l"
@@ -666,10 +666,10 @@
fix z assume "z \<le> x"
with x have "P z" by auto
have "eventually (\<lambda>x. x \<le> g z) (at_right a)"
- using bound[OF bij(2)[OF `P z`]]
- unfolding eventually_at_right[OF bound[OF bij(2)[OF `P z`]]] by (auto intro!: exI[of _ "g z"])
+ using bound[OF bij(2)[OF \<open>P z\<close>]]
+ unfolding eventually_at_right[OF bound[OF bij(2)[OF \<open>P z\<close>]]] by (auto intro!: exI[of _ "g z"])
with Q show "eventually (\<lambda>x. f x \<le> z) (at_right a)"
- by eventually_elim (metis bij `P z` mono)
+ by eventually_elim (metis bij \<open>P z\<close> mono)
qed
qed
@@ -688,10 +688,10 @@
fix z assume "x \<le> z"
with x have "P z" by auto
have "eventually (\<lambda>x. g z \<le> x) (at_left a)"
- using bound[OF bij(2)[OF `P z`]]
- unfolding eventually_at_left[OF bound[OF bij(2)[OF `P z`]]] by (auto intro!: exI[of _ "g z"])
+ using bound[OF bij(2)[OF \<open>P z\<close>]]
+ unfolding eventually_at_left[OF bound[OF bij(2)[OF \<open>P z\<close>]]] by (auto intro!: exI[of _ "g z"])
with Q show "eventually (\<lambda>x. z \<le> f x) (at_left a)"
- by eventually_elim (metis bij `P z` mono)
+ by eventually_elim (metis bij \<open>P z\<close> mono)
qed
qed
@@ -710,7 +710,7 @@
unfolding eventually_nhds
proof safe
fix S :: "'a set" assume "open S" "top \<in> S"
- note open_left[OF this `b < top`]
+ note open_left[OF this \<open>b < top\<close>]
moreover assume "\<forall>s\<in>S. P s"
ultimately show "\<exists>b<top. \<forall>z>b. P z"
by (auto simp: subset_eq Ball_def)
@@ -725,7 +725,7 @@
unfolding tendsto_def eventually_at_filter eventually_inf_principal
by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1)
-subsection {* Limits on sequences *}
+subsection \<open>Limits on sequences\<close>
abbreviation (in topological_space)
LIMSEQ :: "[nat \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
@@ -741,14 +741,14 @@
lemma lim_def: "lim X = (THE L. X ----> L)"
unfolding Lim_def ..
-subsubsection {* Monotone sequences and subsequences *}
+subsubsection \<open>Monotone sequences and subsequences\<close>
definition
monoseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
- --{*Definition of monotonicity.
+ --\<open>Definition of monotonicity.
The use of disjunction here complicates proofs considerably.
One alternative is to add a Boolean argument to indicate the direction.
- Another is to develop the notions of increasing and decreasing first.*}
+ Another is to develop the notions of increasing and decreasing first.\<close>
"monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) \<or> (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
abbreviation incseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
@@ -765,7 +765,7 @@
definition
subseq :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where
- --{*Definition of subsequence*}
+ --\<open>Definition of subsequence\<close>
"subseq f \<longleftrightarrow> (\<forall>m. \<forall>n>m. f m < f n)"
lemma incseq_SucI:
@@ -831,11 +831,11 @@
thus ?thesis by (rule monoI2)
next
case False
- hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
+ hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using \<open>monoseq a\<close>[unfolded monoseq_def] by auto
thus ?thesis by (rule monoI1)
qed
-text{*Subsequence (alternative definition, (e.g. Hoskins)*}
+text\<open>Subsequence (alternative definition, (e.g. Hoskins)\<close>
lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
apply (simp add: subseq_def)
@@ -844,7 +844,7 @@
apply (auto intro: less_trans)
done
-text{* for any sequence, there is a monotonic subsequence *}
+text\<open>for any sequence, there is a monotonic subsequence\<close>
lemma seq_monosub:
fixes s :: "nat => 'a::linorder"
shows "\<exists>f. subseq f \<and> monoseq (\<lambda>n. (s (f n)))"
@@ -927,7 +927,7 @@
show "x \<in> F i"
proof cases
from x have "x \<in> F n" by auto
- also assume "i \<le> n" with `decseq F` have "F n \<subseteq> F i"
+ also assume "i \<le> n" with \<open>decseq F\<close> have "F n \<subseteq> F i"
unfolding decseq_def by simp
finally show ?thesis .
qed (insert x, simp)
@@ -1015,7 +1015,7 @@
lemma lim_le: "convergent f \<Longrightarrow> (\<And>n. f n \<le> (x::'a::linorder_topology)) \<Longrightarrow> lim f \<le> x"
using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff)
-subsubsection{*Increasing and Decreasing Series*}
+subsubsection\<open>Increasing and Decreasing Series\<close>
lemma incseq_le: "incseq X \<Longrightarrow> X ----> L \<Longrightarrow> X n \<le> (L::'a::linorder_topology)"
by (metis incseq_def LIMSEQ_le_const)
@@ -1023,7 +1023,7 @@
lemma decseq_le: "decseq X \<Longrightarrow> X ----> L \<Longrightarrow> (L::'a::linorder_topology) \<le> X n"
by (metis decseq_def LIMSEQ_le_const2)
-subsection {* First countable topologies *}
+subsection \<open>First countable topologies\<close>
class first_countable_topology = topological_space +
assumes first_countable_basis:
@@ -1139,7 +1139,7 @@
unfolding filterlim_def[of _ "nhds a"] le_filter_def eventually_filtermap at_within_def eventually_nhds_within_iff_sequentially comp_def
by metis
-subsection {* Function limit at a point *}
+subsection \<open>Function limit at a point\<close>
abbreviation
LIM :: "('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
@@ -1168,7 +1168,7 @@
shows "f -- a --> L \<Longrightarrow> f -- a --> M \<Longrightarrow> L = M"
using at_neq_bot by (rule tendsto_unique)
-text {* Limits are equal for functions equal except at limit point *}
+text \<open>Limits are equal for functions equal except at limit point\<close>
lemma LIM_equal: "\<forall>x. x \<noteq> a --> (f x = g x) \<Longrightarrow> (f -- a --> l) \<longleftrightarrow> (g -- a --> l)"
unfolding tendsto_def eventually_at_topological by simp
@@ -1205,7 +1205,7 @@
lemma tendsto_compose_filtermap: "((g \<circ> f) ---> T) F \<longleftrightarrow> (g ---> T) (filtermap f F)"
by (simp add: filterlim_def filtermap_filtermap comp_def)
-subsubsection {* Relation of LIM and LIMSEQ *}
+subsubsection \<open>Relation of LIM and LIMSEQ\<close>
lemma (in first_countable_topology) sequentially_imp_eventually_within:
"(\<forall>f. (\<forall>n. f n \<in> s \<and> f n \<noteq> a) \<and> f ----> a \<longrightarrow> eventually (\<lambda>n. P (f n)) sequentially) \<Longrightarrow>
@@ -1259,7 +1259,7 @@
qed
then guess s ..
then have "\<And>n. b < X (s n)" "\<And>n. X (s n) < a" "incseq (\<lambda>n. X (s n))" "(\<lambda>n. X (s n)) ----> a" "\<And>n. \<not> P (X (s n))"
- using X by (auto simp: subseq_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF `X ----> a`, unfolded comp_def])
+ using X by (auto simp: subseq_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \<open>X ----> a\<close>, unfolded comp_def])
from *[OF this(1,2,3,4)] this(5) show False by auto
qed
qed
@@ -1297,7 +1297,7 @@
qed
then guess s ..
then have "\<And>n. a < X (s n)" "\<And>n. X (s n) < b" "decseq (\<lambda>n. X (s n))" "(\<lambda>n. X (s n)) ----> a" "\<And>n. \<not> P (X (s n))"
- using X by (auto simp: subseq_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF `X ----> a`, unfolded comp_def])
+ using X by (auto simp: subseq_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \<open>X ----> a\<close>, unfolded comp_def])
from *[OF this(1,2,3,4)] this(5) show False by auto
qed
qed
@@ -1310,9 +1310,9 @@
using assms unfolding tendsto_def [where l=L]
by (simp add: sequentially_imp_eventually_at_right)
-subsection {* Continuity *}
+subsection \<open>Continuity\<close>
-subsubsection {* Continuity on a set *}
+subsubsection \<open>Continuity on a set\<close>
definition continuous_on :: "'a set \<Rightarrow> ('a :: topological_space \<Rightarrow> 'b :: topological_space) \<Rightarrow> bool" where
"continuous_on s f \<longleftrightarrow> (\<forall>x\<in>s. (f ---> f x) (at x within s))"
@@ -1343,7 +1343,7 @@
proof safe
fix x B assume "x \<in> s" "open B" "f x \<in> B"
with B obtain A where A: "open A" "A \<inter> s = f -` B \<inter> s" by auto
- with `x \<in> s` `f x \<in> B` show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)"
+ with \<open>x \<in> s\<close> \<open>f x \<in> B\<close> show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)"
by (intro exI[of _ A]) auto
qed
qed
@@ -1478,7 +1478,7 @@
(auto intro: less_le_trans[OF _ mono] less_imp_le)
qed
-subsubsection {* Continuity at a point *}
+subsubsection \<open>Continuity at a point\<close>
definition continuous :: "'a::t2_space filter \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool" where
"continuous F f \<longleftrightarrow> (f ---> f (Lim F (\<lambda>x. x))) F"
@@ -1566,7 +1566,7 @@
"continuous (at (x::'a::linorder_topology)) f = (continuous (at_left x) f \<and> continuous (at_right x) f)"
by (simp add: continuous_within filterlim_at_split)
-subsubsection{* Open-cover compactness *}
+subsubsection\<open>Open-cover compactness\<close>
context topological_space
begin
@@ -1599,10 +1599,10 @@
shows "compact (s \<inter> t)"
proof (rule compactI)
fix C assume C: "\<forall>c\<in>C. open c" and cover: "s \<inter> t \<subseteq> \<Union>C"
- from C `closed t` have "\<forall>c\<in>C \<union> {-t}. open c" by auto
+ from C \<open>closed t\<close> have "\<forall>c\<in>C \<union> {-t}. open c" by auto
moreover from cover have "s \<subseteq> \<Union>(C \<union> {-t})" by auto
ultimately have "\<exists>D\<subseteq>C \<union> {-t}. finite D \<and> s \<subseteq> \<Union>D"
- using `compact s` unfolding compact_eq_heine_borel by auto
+ using \<open>compact s\<close> unfolding compact_eq_heine_borel by auto
then obtain D where "D \<subseteq> C \<union> {- t} \<and> finite D \<and> s \<subseteq> \<Union>D" ..
then show "\<exists>D\<subseteq>C. finite D \<and> s \<inter> t \<subseteq> \<Union>D"
by (intro exI[of _ "D - {-t}"]) auto
@@ -1622,7 +1622,7 @@
and fi: "\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}"
from A have "(\<forall>a\<in>uminus`A. open a) \<and> U \<subseteq> \<Union>(uminus`A)"
by auto
- with `compact U` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<subseteq> \<Union>(uminus`B)"
+ with \<open>compact U\<close> obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<subseteq> \<Union>(uminus`B)"
unfolding compact_eq_heine_borel by (metis subset_image_iff)
with fi[THEN spec, of B] show False
by (auto dest: finite_imageD intro: inj_setminus)
@@ -1632,7 +1632,7 @@
assume "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A"
then have "U \<inter> \<Inter>(uminus`A) = {}" "\<forall>a\<in>uminus`A. closed a"
by auto
- with `?R` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>(uminus`B) = {}"
+ with \<open>?R\<close> obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>(uminus`B) = {}"
by (metis subset_image_iff)
then show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD)
@@ -1649,7 +1649,7 @@
and Q: "\<And>I'. finite I' \<Longrightarrow> I' \<subseteq> I \<Longrightarrow> (s \<inter> (\<Inter>i\<in>I'. f i) \<noteq> {})"
shows "s \<inter> (\<Inter>i\<in>I. f i) \<noteq> {}"
proof -
- note `compact s`
+ note \<open>compact s\<close>
moreover from P have "\<forall>i \<in> f ` I. closed i" by blast
moreover have "\<forall>A. finite A \<and> A \<subseteq> f ` I \<longrightarrow> (s \<inter> (\<Inter>A) \<noteq> {})"
proof (rule, rule, erule conjE)
@@ -1672,23 +1672,23 @@
proof (rule openI)
fix y assume "y \<in> - s"
let ?C = "\<Union>x\<in>s. {u. open u \<and> x \<in> u \<and> eventually (\<lambda>y. y \<notin> u) (nhds y)}"
- note `compact s`
+ note \<open>compact s\<close>
moreover have "\<forall>u\<in>?C. open u" by simp
moreover have "s \<subseteq> \<Union>?C"
proof
fix x assume "x \<in> s"
- with `y \<in> - s` have "x \<noteq> y" by clarsimp
+ with \<open>y \<in> - s\<close> have "x \<noteq> y" by clarsimp
hence "\<exists>u v. open u \<and> open v \<and> x \<in> u \<and> y \<in> v \<and> u \<inter> v = {}"
by (rule hausdorff)
- with `x \<in> s` show "x \<in> \<Union>?C"
+ with \<open>x \<in> s\<close> show "x \<in> \<Union>?C"
unfolding eventually_nhds by auto
qed
ultimately obtain D where "D \<subseteq> ?C" and "finite D" and "s \<subseteq> \<Union>D"
by (rule compactE)
- from `D \<subseteq> ?C` have "\<forall>x\<in>D. eventually (\<lambda>y. y \<notin> x) (nhds y)" by auto
- with `finite D` have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
+ from \<open>D \<subseteq> ?C\<close> have "\<forall>x\<in>D. eventually (\<lambda>y. y \<notin> x) (nhds y)" by auto
+ with \<open>finite D\<close> have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
by (simp add: eventually_ball_finite)
- with `s \<subseteq> \<Union>D` have "eventually (\<lambda>y. y \<notin> s) (nhds y)"
+ with \<open>s \<subseteq> \<Union>D\<close> have "eventually (\<lambda>y. y \<notin> s) (nhds y)"
by (auto elim!: eventually_mono [rotated])
thus "\<exists>t. open t \<and> y \<in> t \<and> t \<subseteq> - s"
by (simp add: eventually_nhds subset_eq)
@@ -1721,10 +1721,10 @@
have 1: "\<forall>x\<in>s. f x \<in> f ` (s - B) \<longleftrightarrow> x \<in> s - B"
using assms(3) by (auto, metis)
have "continuous_on (s - B) f"
- using `continuous_on s f` Diff_subset
+ using \<open>continuous_on s f\<close> Diff_subset
by (rule continuous_on_subset)
moreover have "compact (s - B)"
- using `open B` and `compact s`
+ using \<open>open B\<close> and \<open>compact s\<close>
unfolding Diff_eq by (intro compact_inter_closed closed_Compl)
ultimately have "compact (f ` (s - B))"
by (rule compact_continuous_image)
@@ -1733,7 +1733,7 @@
hence "open (- f ` (s - B))"
by (rule open_Compl)
moreover have "f x \<in> - f ` (s - B)"
- using `x \<in> s` and `x \<in> B` by (simp add: 1)
+ using \<open>x \<in> s\<close> and \<open>x \<in> B\<close> by (simp add: 1)
moreover have "\<forall>y\<in>s. f y \<in> - f ` (s - B) \<longrightarrow> y \<in> B"
by (simp add: 1)
ultimately show "\<exists>A. open A \<and> f x \<in> A \<and> (\<forall>y\<in>s. f y \<in> A \<longrightarrow> y \<in> B)"
@@ -1755,13 +1755,13 @@
by (metis not_le)
then have "\<forall>s\<in>S. open {..< t s}" "S \<subseteq> (\<Union>s\<in>S. {..< t s})"
by auto
- with `compact S` obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {..< t s})"
+ with \<open>compact S\<close> obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {..< t s})"
by (erule compactE_image)
- with `S \<noteq> {}` have Max: "Max (t`C) \<in> t`C" and "\<forall>s\<in>t`C. s \<le> Max (t`C)"
+ with \<open>S \<noteq> {}\<close> have Max: "Max (t`C) \<in> t`C" and "\<forall>s\<in>t`C. s \<le> Max (t`C)"
by (auto intro!: Max_in)
with C have "S \<subseteq> {..< Max (t`C)}"
by (auto intro: less_le_trans simp: subset_eq)
- with t Max `C \<subseteq> S` show ?thesis
+ with t Max \<open>C \<subseteq> S\<close> show ?thesis
by fastforce
qed
@@ -1774,13 +1774,13 @@
by (metis not_le)
then have "\<forall>s\<in>S. open {t s <..}" "S \<subseteq> (\<Union>s\<in>S. {t s <..})"
by auto
- with `compact S` obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {t s <..})"
+ with \<open>compact S\<close> obtain C where "C \<subseteq> S" "finite C" and C: "S \<subseteq> (\<Union>s\<in>C. {t s <..})"
by (erule compactE_image)
- with `S \<noteq> {}` have Min: "Min (t`C) \<in> t`C" and "\<forall>s\<in>t`C. Min (t`C) \<le> s"
+ with \<open>S \<noteq> {}\<close> have Min: "Min (t`C) \<in> t`C" and "\<forall>s\<in>t`C. Min (t`C) \<le> s"
by (auto intro!: Min_in)
with C have "S \<subseteq> {Min (t`C) <..}"
by (auto intro: le_less_trans simp: subset_eq)
- with t Min `C \<subseteq> S` show ?thesis
+ with t Min \<open>C \<subseteq> S\<close> show ?thesis
by fastforce
qed
@@ -1794,7 +1794,7 @@
shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f \<Longrightarrow> (\<exists>x\<in>s. \<forall>y\<in>s. f x \<le> f y)"
using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto
-subsection {* Connectedness *}
+subsection \<open>Connectedness\<close>
context topological_space
begin
@@ -1827,7 +1827,7 @@
obtain t f where "open t" "open f" and *: "f \<inter> S = P -` {False} \<inter> S" "t \<inter> S = P -` {True} \<inter> S"
by auto
then have "t \<inter> S = {} \<or> f \<inter> S = {}"
- by (intro connectedD[OF `connected S`]) auto
+ by (intro connectedD[OF \<open>connected S\<close>]) auto
then show "\<exists>c. \<forall>s\<in>S. P s = c"
proof (rule disjE)
assume "t \<inter> S = {}" then show ?thesis
@@ -1876,17 +1876,17 @@
let ?P = "\<Union>b\<in>{b\<in>A. f a = f b}. S b" and ?N = "\<Union>b\<in>{b\<in>A. f a \<noteq> f b}. S b"
have "?P \<inter> A = {} \<or> ?N \<inter> A = {}"
- using `connected A` S `a\<in>A`
+ using \<open>connected A\<close> S \<open>a\<in>A\<close>
by (intro connectedD) (auto, metis)
then show "f a = f b"
proof
assume "?N \<inter> A = {}"
then have "\<forall>x\<in>A. f a = f x"
using S(1) by auto
- with `b\<in>A` show ?thesis by auto
+ with \<open>b\<in>A\<close> show ?thesis by auto
next
assume "?P \<inter> A = {}" then show ?thesis
- using `a \<in> A` S(1)[of a] by auto
+ using \<open>a \<in> A\<close> S(1)[of a] by auto
qed
qed
@@ -1915,11 +1915,11 @@
fix P :: "'b \<Rightarrow> bool" assume "continuous_on (f ` s) P"
then have "continuous_on s (P \<circ> f)"
by (rule continuous_on_compose[OF *])
- from connectedD_const[OF `connected s` this] show "\<exists>c. \<forall>s\<in>f ` s. P s = c"
+ from connectedD_const[OF \<open>connected s\<close> this] show "\<exists>c. \<forall>s\<in>f ` s. P s = c"
by auto
qed
-section {* Connectedness *}
+section \<open>Connectedness\<close>
class linear_continuum_topology = linorder_topology + linear_continuum
begin
@@ -1934,7 +1934,7 @@
with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A"
by (auto simp: subset_eq)
then show False
- using cInf_lower[OF `c \<in> A`] bnd by (metis not_le less_imp_le bdd_belowI)
+ using cInf_lower[OF \<open>c \<in> A\<close>] bnd by (metis not_le less_imp_le bdd_belowI)
qed
lemma Sup_notin_open:
@@ -1947,7 +1947,7 @@
with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A"
by (auto simp: subset_eq)
then show False
- using cSup_upper[OF `c \<in> A`] bnd by (metis less_imp_le not_le bdd_aboveI)
+ using cSup_upper[OF \<open>c \<in> A\<close>] bnd by (metis less_imp_le not_le bdd_aboveI)
qed
end
@@ -1973,30 +1973,30 @@
let ?z = "Inf (B \<inter> {x <..})"
have "x \<le> ?z" "?z \<le> y"
- using `y \<in> B` `x < y` by (auto intro: cInf_lower cInf_greatest)
- with `x \<in> U` `y \<in> U` have "?z \<in> U"
+ using \<open>y \<in> B\<close> \<open>x < y\<close> by (auto intro: cInf_lower cInf_greatest)
+ with \<open>x \<in> U\<close> \<open>y \<in> U\<close> have "?z \<in> U"
by (rule *)
moreover have "?z \<notin> B \<inter> {x <..}"
- using `open B` by (intro Inf_notin_open) auto
+ using \<open>open B\<close> by (intro Inf_notin_open) auto
ultimately have "?z \<in> A"
- using `x \<le> ?z` `A \<inter> B \<inter> U = {}` `x \<in> A` `U \<subseteq> A \<union> B` by auto
+ using \<open>x \<le> ?z\<close> \<open>A \<inter> B \<inter> U = {}\<close> \<open>x \<in> A\<close> \<open>U \<subseteq> A \<union> B\<close> by auto
{ assume "?z < y"
obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
- using open_right[OF `open A` `?z \<in> A` `?z < y`] by auto
+ using open_right[OF \<open>open A\<close> \<open>?z \<in> A\<close> \<open>?z < y\<close>] by auto
moreover obtain b where "b \<in> B" "x < b" "b < min a y"
- using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
+ using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] \<open>?z < a\<close> \<open>?z < y\<close> \<open>x < y\<close> \<open>y \<in> B\<close>
by (auto intro: less_imp_le)
moreover have "?z \<le> b"
- using `b \<in> B` `x < b`
+ using \<open>b \<in> B\<close> \<open>x < b\<close>
by (intro cInf_lower) auto
moreover have "b \<in> U"
- using `x \<le> ?z` `?z \<le> b` `b < min a y`
- by (intro *[OF `x \<in> U` `y \<in> U`]) (auto simp: less_imp_le)
+ using \<open>x \<le> ?z\<close> \<open>?z \<le> b\<close> \<open>b < min a y\<close>
+ by (intro *[OF \<open>x \<in> U\<close> \<open>y \<in> U\<close>]) (auto simp: less_imp_le)
ultimately have "\<exists>b\<in>B. b \<in> A \<and> b \<in> U"
by (intro bexI[of _ b]) auto }
then have False
- using `?z \<le> y` `?z \<in> A` `y \<in> B` `y \<in> U` `A \<inter> B \<inter> U = {}` unfolding le_less by blast }
+ using \<open>?z \<le> y\<close> \<open>?z \<in> A\<close> \<open>y \<in> B\<close> \<open>y \<in> U\<close> \<open>A \<inter> B \<inter> U = {}\<close> unfolding le_less by blast }
note not_disjoint = this
fix A B assume AB: "open A" "open B" "U \<subseteq> A \<union> B" "A \<inter> B \<inter> U = {}"
@@ -2043,7 +2043,7 @@
assumes A: "connected A" "a \<in> A" "b \<in> A" shows "{a <..< b} \<subseteq> A"
using connectedD_interval[OF A] by (simp add: subset_eq Ball_def less_imp_le)
-subsection {* Intermediate Value Theorem *}
+subsection \<open>Intermediate Value Theorem\<close>
lemma IVT':
fixes f :: "'a :: linear_continuum_topology \<Rightarrow> 'b :: linorder_topology"
@@ -2117,29 +2117,29 @@
show "f (Sup S) \<le> (SUP s:S. f s)"
proof cases
assume "Sup S \<in> S" then show ?thesis
- by (rule cSUP_upper) (auto intro: bdd_above_image_mono S `mono f`)
+ by (rule cSUP_upper) (auto intro: bdd_above_image_mono S \<open>mono f\<close>)
next
assume "Sup S \<notin> S"
- from `S \<noteq> {}` obtain s where "s \<in> S"
+ from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
by auto
- with `Sup S \<notin> S` S have "s < Sup S"
+ with \<open>Sup S \<notin> S\<close> S have "s < Sup S"
unfolding less_le by (blast intro: cSup_upper)
show ?thesis
proof (rule ccontr)
assume "\<not> ?thesis"
with order_tendstoD(1)[OF f, of "SUP s:S. f s"] obtain b where "b < Sup S"
and *: "\<And>y. b < y \<Longrightarrow> y < Sup S \<Longrightarrow> (SUP s:S. f s) < f y"
- by (auto simp: not_le eventually_at_left[OF `s < Sup S`])
- with `S \<noteq> {}` obtain c where "c \<in> S" "b < c"
+ by (auto simp: not_le eventually_at_left[OF \<open>s < Sup S\<close>])
+ with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "b < c"
using less_cSupD[of S b] by auto
- with `Sup S \<notin> S` S have "c < Sup S"
+ with \<open>Sup S \<notin> S\<close> S have "c < Sup S"
unfolding less_le by (blast intro: cSup_upper)
- from *[OF `b < c` `c < Sup S`] cSUP_upper[OF `c \<in> S` bdd_above_image_mono[of f]]
+ from *[OF \<open>b < c\<close> \<open>c < Sup S\<close>] cSUP_upper[OF \<open>c \<in> S\<close> bdd_above_image_mono[of f]]
show False
by (auto simp: assms)
qed
qed
-qed (intro cSUP_least `mono f`[THEN monoD] cSup_upper S)
+qed (intro cSUP_least \<open>mono f\<close>[THEN monoD] cSup_upper S)
lemma continuous_at_Sup_antimono:
fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
@@ -2154,29 +2154,29 @@
show "(INF s:S. f s) \<le> f (Sup S)"
proof cases
assume "Sup S \<in> S" then show ?thesis
- by (intro cINF_lower) (auto intro: bdd_below_image_antimono S `antimono f`)
+ by (intro cINF_lower) (auto intro: bdd_below_image_antimono S \<open>antimono f\<close>)
next
assume "Sup S \<notin> S"
- from `S \<noteq> {}` obtain s where "s \<in> S"
+ from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
by auto
- with `Sup S \<notin> S` S have "s < Sup S"
+ with \<open>Sup S \<notin> S\<close> S have "s < Sup S"
unfolding less_le by (blast intro: cSup_upper)
show ?thesis
proof (rule ccontr)
assume "\<not> ?thesis"
with order_tendstoD(2)[OF f, of "INF s:S. f s"] obtain b where "b < Sup S"
and *: "\<And>y. b < y \<Longrightarrow> y < Sup S \<Longrightarrow> f y < (INF s:S. f s)"
- by (auto simp: not_le eventually_at_left[OF `s < Sup S`])
- with `S \<noteq> {}` obtain c where "c \<in> S" "b < c"
+ by (auto simp: not_le eventually_at_left[OF \<open>s < Sup S\<close>])
+ with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "b < c"
using less_cSupD[of S b] by auto
- with `Sup S \<notin> S` S have "c < Sup S"
+ with \<open>Sup S \<notin> S\<close> S have "c < Sup S"
unfolding less_le by (blast intro: cSup_upper)
- from *[OF `b < c` `c < Sup S`] cINF_lower[OF bdd_below_image_antimono, of f S c] `c \<in> S`
+ from *[OF \<open>b < c\<close> \<open>c < Sup S\<close>] cINF_lower[OF bdd_below_image_antimono, of f S c] \<open>c \<in> S\<close>
show False
by (auto simp: assms)
qed
qed
-qed (intro cINF_greatest `antimono f`[THEN antimonoD] cSup_upper S)
+qed (intro cINF_greatest \<open>antimono f\<close>[THEN antimonoD] cSup_upper S)
lemma continuous_at_Inf_mono:
fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
@@ -2191,29 +2191,29 @@
show "(INF s:S. f s) \<le> f (Inf S)"
proof cases
assume "Inf S \<in> S" then show ?thesis
- by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S `mono f`)
+ by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S \<open>mono f\<close>)
next
assume "Inf S \<notin> S"
- from `S \<noteq> {}` obtain s where "s \<in> S"
+ from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
by auto
- with `Inf S \<notin> S` S have "Inf S < s"
+ with \<open>Inf S \<notin> S\<close> S have "Inf S < s"
unfolding less_le by (blast intro: cInf_lower)
show ?thesis
proof (rule ccontr)
assume "\<not> ?thesis"
with order_tendstoD(2)[OF f, of "INF s:S. f s"] obtain b where "Inf S < b"
and *: "\<And>y. Inf S < y \<Longrightarrow> y < b \<Longrightarrow> f y < (INF s:S. f s)"
- by (auto simp: not_le eventually_at_right[OF `Inf S < s`])
- with `S \<noteq> {}` obtain c where "c \<in> S" "c < b"
+ by (auto simp: not_le eventually_at_right[OF \<open>Inf S < s\<close>])
+ with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "c < b"
using cInf_lessD[of S b] by auto
- with `Inf S \<notin> S` S have "Inf S < c"
+ with \<open>Inf S \<notin> S\<close> S have "Inf S < c"
unfolding less_le by (blast intro: cInf_lower)
- from *[OF `Inf S < c` `c < b`] cINF_lower[OF bdd_below_image_mono[of f] `c \<in> S`]
+ from *[OF \<open>Inf S < c\<close> \<open>c < b\<close>] cINF_lower[OF bdd_below_image_mono[of f] \<open>c \<in> S\<close>]
show False
by (auto simp: assms)
qed
qed
-qed (intro cINF_greatest `mono f`[THEN monoD] cInf_lower `bdd_below S` `S \<noteq> {}`)
+qed (intro cINF_greatest \<open>mono f\<close>[THEN monoD] cInf_lower \<open>bdd_below S\<close> \<open>S \<noteq> {}\<close>)
lemma continuous_at_Inf_antimono:
fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder} \<Rightarrow> 'b :: {linorder_topology, conditionally_complete_linorder}"
@@ -2228,28 +2228,28 @@
show "f (Inf S) \<le> (SUP s:S. f s)"
proof cases
assume "Inf S \<in> S" then show ?thesis
- by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S `antimono f`)
+ by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S \<open>antimono f\<close>)
next
assume "Inf S \<notin> S"
- from `S \<noteq> {}` obtain s where "s \<in> S"
+ from \<open>S \<noteq> {}\<close> obtain s where "s \<in> S"
by auto
- with `Inf S \<notin> S` S have "Inf S < s"
+ with \<open>Inf S \<notin> S\<close> S have "Inf S < s"
unfolding less_le by (blast intro: cInf_lower)
show ?thesis
proof (rule ccontr)
assume "\<not> ?thesis"
with order_tendstoD(1)[OF f, of "SUP s:S. f s"] obtain b where "Inf S < b"
and *: "\<And>y. Inf S < y \<Longrightarrow> y < b \<Longrightarrow> (SUP s:S. f s) < f y"
- by (auto simp: not_le eventually_at_right[OF `Inf S < s`])
- with `S \<noteq> {}` obtain c where "c \<in> S" "c < b"
+ by (auto simp: not_le eventually_at_right[OF \<open>Inf S < s\<close>])
+ with \<open>S \<noteq> {}\<close> obtain c where "c \<in> S" "c < b"
using cInf_lessD[of S b] by auto
- with `Inf S \<notin> S` S have "Inf S < c"
+ with \<open>Inf S \<notin> S\<close> S have "Inf S < c"
unfolding less_le by (blast intro: cInf_lower)
- from *[OF `Inf S < c` `c < b`] cSUP_upper[OF `c \<in> S` bdd_above_image_antimono[of f]]
+ from *[OF \<open>Inf S < c\<close> \<open>c < b\<close>] cSUP_upper[OF \<open>c \<in> S\<close> bdd_above_image_antimono[of f]]
show False
by (auto simp: assms)
qed
qed
-qed (intro cSUP_least `antimono f`[THEN antimonoD] cInf_lower S)
+qed (intro cSUP_least \<open>antimono f\<close>[THEN antimonoD] cInf_lower S)
end
--- a/src/HOL/Transcendental.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Transcendental.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Author: Jeremy Avigad
*)
-section{*Power Series, Transcendental Functions etc.*}
+section\<open>Power Series, Transcendental Functions etc.\<close>
theory Transcendental
imports Binomial Series Deriv NthRoot
@@ -33,9 +33,9 @@
proof -
have "0 \<le> x"
by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1])
- from `x < 1` obtain z where z: "x < z" "z < 1"
+ from \<open>x < 1\<close> obtain z where z: "x < z" "z < 1"
by (metis dense)
- from f `x < z`
+ from f \<open>x < z\<close>
have "eventually (\<lambda>n. root n (norm (f n)) < z) sequentially"
by (rule order_tendstoD)
then have "eventually (\<lambda>n. norm (f n) \<le> z^n) sequentially"
@@ -48,10 +48,10 @@
qed
then show "summable f"
unfolding eventually_sequentially
- using z `0 \<le> x` by (auto intro!: summable_comparison_test[OF _ summable_geometric])
+ using z \<open>0 \<le> x\<close> by (auto intro!: summable_comparison_test[OF _ summable_geometric])
qed
-subsection {* Properties of Power Series *}
+subsection \<open>Properties of Power Series\<close>
lemma powser_zero:
fixes f :: "nat \<Rightarrow> 'a::real_normed_algebra_1"
@@ -68,8 +68,8 @@
using sums_finite [of "{0}" "\<lambda>n. a n * 0 ^ n"]
by simp
-text{*Power series has a circle or radius of convergence: if it sums for @{term
- x}, then it sums absolutely for @{term z} with @{term "\<bar>z\<bar> < \<bar>x\<bar>"}.*}
+text\<open>Power series has a circle or radius of convergence: if it sums for @{term
+ x}, then it sums absolutely for @{term z} with @{term "\<bar>z\<bar> < \<bar>x\<bar>"}.\<close>
lemma powser_insidea:
fixes x z :: "'a::real_normed_div_algebra"
@@ -207,7 +207,7 @@
proof (rule LIMSEQ_I)
fix r :: real
assume "0 < r"
- from `g sums x`[unfolded sums_def, THEN LIMSEQ_D, OF this]
+ from \<open>g sums x\<close>[unfolded sums_def, THEN LIMSEQ_D, OF this]
obtain no where no_eq: "\<And> n. n \<ge> no \<Longrightarrow> (norm (setsum g {..<n} - x) < r)" by blast
let ?SUM = "\<lambda> m. \<Sum>i<m. if even i then 0 else g ((i - 1) div 2)"
@@ -218,7 +218,7 @@
have sum_eq: "?SUM (2 * (m div 2)) = setsum g {..< m div 2}"
using sum_split_even_odd by auto
hence "(norm (?SUM (2 * (m div 2)) - x) < r)"
- using no_eq unfolding sum_eq using `m div 2 \<ge> no` by auto
+ using no_eq unfolding sum_eq using \<open>m div 2 \<ge> no\<close> by auto
moreover
have "?SUM (2 * (m div 2)) = ?SUM m"
proof (cases "even m")
@@ -227,9 +227,9 @@
next
case False
then have eq: "Suc (2 * (m div 2)) = m" by simp
- hence "even (2 * (m div 2))" using `odd m` by auto
+ hence "even (2 * (m div 2))" using \<open>odd m\<close> by auto
have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq ..
- also have "\<dots> = ?SUM (2 * (m div 2))" using `even (2 * (m div 2))` by auto
+ also have "\<dots> = ?SUM (2 * (m div 2))" using \<open>even (2 * (m div 2))\<close> by auto
finally show ?thesis by auto
qed
ultimately have "(norm (?SUM m - x) < r)" by auto
@@ -249,11 +249,11 @@
by (cases B) auto
} note if_sum = this
have g_sums: "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x"
- using sums_if'[OF `g sums x`] .
+ using sums_if'[OF \<open>g sums x\<close>] .
{
have if_eq: "\<And>B T E. (if \<not> B then T else E) = (if B then E else T)" by auto
- have "?s sums y" using sums_if'[OF `f sums y`] .
+ have "?s sums y" using sums_if'[OF \<open>f sums y\<close>] .
from this[unfolded sums_def, THEN LIMSEQ_Suc]
have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
by (simp add: lessThan_Suc_eq_insert_0 image_iff setsum.reindex if_eq sums_def cong del: if_cong)
@@ -261,7 +261,7 @@
from sums_add[OF g_sums this] show ?thesis unfolding if_sum .
qed
-subsection {* Alternating series test / Leibniz formula *}
+subsection \<open>Alternating series test / Leibniz formula\<close>
lemma sums_alternating_upper_lower:
fixes a :: "nat \<Rightarrow> real"
@@ -293,7 +293,7 @@
proof (rule LIMSEQ_I)
fix r :: real
assume "0 < r"
- with `a ----> 0`[THEN LIMSEQ_D] obtain N where "\<And> n. n \<ge> N \<Longrightarrow> norm (a n - 0) < r"
+ with \<open>a ----> 0\<close>[THEN LIMSEQ_D] obtain N where "\<And> n. n \<ge> N \<Longrightarrow> norm (a n - 0) < r"
by auto
hence "\<forall>n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
thus "\<exists>N. \<forall>n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
@@ -327,10 +327,10 @@
proof (rule LIMSEQ_I)
fix r :: real
assume "0 < r"
- with `?f ----> l`[THEN LIMSEQ_D]
+ with \<open>?f ----> l\<close>[THEN LIMSEQ_D]
obtain f_no where f: "\<And> n. n \<ge> f_no \<Longrightarrow> norm (?f n - l) < r" by auto
- from `0 < r` `?g ----> l`[THEN LIMSEQ_D]
+ from \<open>0 < r\<close> \<open>?g ----> l\<close>[THEN LIMSEQ_D]
obtain g_no where g: "\<And> n. n \<ge> g_no \<Longrightarrow> norm (?g n - l) < r" by auto
{
@@ -341,7 +341,7 @@
proof (cases "even n")
case True
then have n_eq: "2 * (n div 2) = n" by (simp add: even_two_times_div_two)
- with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no"
+ with \<open>n \<ge> 2 * f_no\<close> have "n div 2 \<ge> f_no"
by auto
from f[OF this] show ?thesis
unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
@@ -353,7 +353,7 @@
hence range_eq: "n - 1 + 1 = n"
using odd_pos[OF False] by auto
- from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no"
+ from n_eq \<open>n \<ge> 2 * g_no\<close> have "(n - 1) div 2 \<ge> g_no"
by auto
from g[OF this] show ?thesis
unfolding n_eq range_eq .
@@ -373,9 +373,9 @@
show "?f n \<le> suminf ?S"
unfolding sums_unique[OF sums_l, symmetric] using below_l by auto
show "?g ----> suminf ?S"
- using `?g ----> l` `l = suminf ?S` by auto
+ using \<open>?g ----> l\<close> \<open>l = suminf ?S\<close> by auto
show "?f ----> suminf ?S"
- using `?f ----> l` `l = suminf ?S` by auto
+ using \<open>?f ----> l\<close> \<open>l = suminf ?S\<close> by auto
qed
theorem summable_Leibniz:
@@ -399,13 +399,13 @@
have "a (Suc n) \<le> a n"
using ord[where n="Suc n" and m=n] by auto
} note mono = this
- note leibniz = summable_Leibniz'[OF `a ----> 0` ge0]
+ note leibniz = summable_Leibniz'[OF \<open>a ----> 0\<close> ge0]
from leibniz[OF mono]
- show ?thesis using `0 \<le> a 0` by auto
+ show ?thesis using \<open>0 \<le> a 0\<close> by auto
next
let ?a = "\<lambda> n. - a n"
case False
- with monoseq_le[OF `monoseq a` `a ----> 0`]
+ with monoseq_le[OF \<open>monoseq a\<close> \<open>a ----> 0\<close>]
have "(\<forall> n. a n \<le> 0) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)" by auto
hence ord: "\<And>n m. m \<le> n \<Longrightarrow> ?a n \<le> ?a m" and ge0: "\<And> n. 0 \<le> ?a n"
by auto
@@ -416,7 +416,7 @@
} note monotone = this
note leibniz =
summable_Leibniz'[OF _ ge0, of "\<lambda>x. x",
- OF tendsto_minus[OF `a ----> 0`, unfolded minus_zero] monotone]
+ OF tendsto_minus[OF \<open>a ----> 0\<close>, unfolded minus_zero] monotone]
have "summable (\<lambda> n. (-1)^n * ?a n)"
using leibniz(1) by auto
then obtain l where "(\<lambda> n. (-1)^n * ?a n) sums l"
@@ -432,7 +432,7 @@
have move_minus: "(\<Sum>n. - ((- 1) ^ n * a n)) = - (\<Sum>n. (- 1) ^ n * a n)"
by auto
- have ?pos using `0 \<le> ?a 0` by auto
+ have ?pos using \<open>0 \<le> ?a 0\<close> by auto
moreover have ?neg
using leibniz(2,4)
unfolding mult_minus_right setsum_negf move_minus neg_le_iff_le
@@ -446,12 +446,12 @@
by safe
qed
-subsection {* Term-by-Term Differentiability of Power Series *}
+subsection \<open>Term-by-Term Differentiability of Power Series\<close>
definition diffs :: "(nat \<Rightarrow> 'a::ring_1) \<Rightarrow> nat \<Rightarrow> 'a"
where "diffs c = (\<lambda>n. of_nat (Suc n) * c (Suc n))"
-text{*Lemma about distributing negation over it*}
+text\<open>Lemma about distributing negation over it\<close>
lemma diffs_minus: "diffs (\<lambda>n. - c n) = (\<lambda>n. - diffs c n)"
by (simp add: diffs_def)
@@ -603,7 +603,7 @@
qed
-text{* FIXME: Long proofs*}
+text\<open>FIXME: Long proofs\<close>
lemma termdiffs_aux:
fixes x :: "'a::{real_normed_field,banach}"
@@ -702,7 +702,7 @@
qed
qed
-subsection {* The Derivative of a Power Series Has the Same Radius of Convergence *}
+subsection \<open>The Derivative of a Power Series Has the Same Radius of Convergence\<close>
lemma termdiff_converges:
fixes x :: "'a::{real_normed_field,banach}"
@@ -829,7 +829,7 @@
qed
-subsection {* Derivability of power series *}
+subsection \<open>Derivability of power series\<close>
lemma DERIV_series':
fixes f :: "real \<Rightarrow> nat \<Rightarrow> real"
@@ -845,10 +845,10 @@
assume "0 < r" hence "0 < r/3" by auto
obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3"
- using suminf_exist_split[OF `0 < r/3` `summable L`] by auto
+ using suminf_exist_split[OF \<open>0 < r/3\<close> \<open>summable L\<close>] by auto
obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3"
- using suminf_exist_split[OF `0 < r/3` `summable (f' x0)`] by auto
+ using suminf_exist_split[OF \<open>0 < r/3\<close> \<open>summable (f' x0)\<close>] by auto
let ?N = "Suc (max N_L N_f')"
have "\<bar> \<Sum> i. f' x0 (i + ?N) \<bar> < r/3" (is "?f'_part < r/3") and
@@ -857,7 +857,7 @@
let ?diff = "\<lambda>i x. (f (x0 + x) i - f x0 i) / x"
let ?r = "r / (3 * real ?N)"
- from `0 < r` have "0 < ?r" by simp
+ from \<open>0 < r\<close> have "0 < ?r" by simp
let ?s = "\<lambda>n. SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
def S' \<equiv> "Min (?s ` {..< ?N })"
@@ -870,17 +870,17 @@
assume "x \<in> ?s ` {..<?N}"
then obtain n where "x = ?s n" and "n \<in> {..<?N}"
using image_iff[THEN iffD1] by blast
- from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def]
+ from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \<open>0 < ?r\<close>, unfolded real_norm_def]
obtain s where s_bound: "0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < s \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r)"
by auto
have "0 < ?s n" by (rule someI2[where a=s]) (auto simp add: s_bound)
- thus "0 < x" unfolding `x = ?s n` .
+ thus "0 < x" unfolding \<open>x = ?s n\<close> .
qed
qed auto
def S \<equiv> "min (min (x0 - a) (b - x0)) S'"
hence "0 < S" and S_a: "S \<le> x0 - a" and S_b: "S \<le> b - x0"
- and "S \<le> S'" using x0_in_I and `0 < S'`
+ and "S \<le> S'" using x0_in_I and \<open>0 < S'\<close>
by auto
{
@@ -891,21 +891,21 @@
note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
note div_smbl = summable_divide[OF diff_smbl]
- note all_smbl = summable_diff[OF div_smbl `summable (f' x0)`]
+ note all_smbl = summable_diff[OF div_smbl \<open>summable (f' x0)\<close>]
note ign = summable_ignore_initial_segment[where k="?N"]
note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]]
note div_shft_smbl = summable_divide[OF diff_shft_smbl]
- note all_shft_smbl = summable_diff[OF div_smbl ign[OF `summable (f' x0)`]]
+ note all_shft_smbl = summable_diff[OF div_smbl ign[OF \<open>summable (f' x0)\<close>]]
{ fix n
have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>"
using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero]
unfolding abs_divide .
hence "\<bar> (\<bar>?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)"
- using `x \<noteq> 0` by auto }
- note 1 = this and 2 = summable_rabs_comparison_test[OF _ ign[OF `summable L`]]
+ using \<open>x \<noteq> 0\<close> by auto }
+ note 1 = this and 2 = summable_rabs_comparison_test[OF _ ign[OF \<open>summable L\<close>]]
then have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> (\<Sum> i. L (i + ?N))"
- by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF `summable L`]]])
+ by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF \<open>summable L\<close>]]])
then have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> r / 3" (is "?L_part \<le> r/3")
using L_estimate by auto
@@ -914,19 +914,19 @@
proof (rule setsum_strict_mono)
fix n
assume "n \<in> {..< ?N}"
- have "\<bar>x\<bar> < S" using `\<bar>x\<bar> < S` .
- also have "S \<le> S'" using `S \<le> S'` .
+ have "\<bar>x\<bar> < S" using \<open>\<bar>x\<bar> < S\<close> .
+ also have "S \<le> S'" using \<open>S \<le> S'\<close> .
also have "S' \<le> ?s n" unfolding S'_def
proof (rule Min_le_iff[THEN iffD2])
have "?s n \<in> (?s ` {..<?N}) \<and> ?s n \<le> ?s n"
- using `n \<in> {..< ?N}` by auto
+ using \<open>n \<in> {..< ?N}\<close> by auto
thus "\<exists> a \<in> (?s ` {..<?N}). a \<le> ?s n" by blast
qed auto
finally have "\<bar>x\<bar> < ?s n" .
- from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2]
+ from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \<open>0 < ?r\<close>, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2]
have "\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < ?s n \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r" .
- with `x \<noteq> 0` and `\<bar>x\<bar> < ?s n` show "\<bar>?diff n x - f' x0 n\<bar> < ?r"
+ with \<open>x \<noteq> 0\<close> and \<open>\<bar>x\<bar> < ?s n\<close> show "\<bar>?diff n x - f' x0 n\<bar> < ?r"
by blast
qed auto
also have "\<dots> = of_nat (card {..<?N}) * ?r"
@@ -939,24 +939,24 @@
from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
have "\<bar>(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\<bar> =
\<bar>\<Sum>n. ?diff n x - f' x0 n\<bar>"
- unfolding suminf_diff[OF div_smbl `summable (f' x0)`, symmetric]
+ unfolding suminf_diff[OF div_smbl \<open>summable (f' x0)\<close>, symmetric]
using suminf_divide[OF diff_smbl, symmetric] by auto
also have "\<dots> \<le> ?diff_part + \<bar> (\<Sum>n. ?diff (n + ?N) x) - (\<Sum> n. f' x0 (n + ?N)) \<bar>"
unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"]
- unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]]
+ unfolding suminf_diff[OF div_shft_smbl ign[OF \<open>summable (f' x0)\<close>]]
apply (subst (5) add.commute)
by (rule abs_triangle_ineq)
also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part"
using abs_triangle_ineq4 by auto
also have "\<dots> < r /3 + r/3 + r/3"
- using `?diff_part < r/3` `?L_part \<le> r/3` and `?f'_part < r/3`
+ using \<open>?diff_part < r/3\<close> \<open>?L_part \<le> r/3\<close> and \<open>?f'_part < r/3\<close>
by (rule add_strict_mono [OF add_less_le_mono])
finally have "\<bar>(suminf (f (x0 + x)) - suminf (f x0)) / x - suminf (f' x0)\<bar> < r"
by auto
}
thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow>
norm (((\<Sum>n. f (x0 + x) n) - (\<Sum>n. f x0 n)) / x - (\<Sum>n. f' x0 n)) < r"
- using `0 < S` unfolding real_norm_def diff_0_right by blast
+ using \<open>0 < S\<close> unfolding real_norm_def diff_0_right by blast
qed
lemma DERIV_power_series':
@@ -976,11 +976,11 @@
show "summable (\<lambda> n. \<bar>f n * real (Suc n) * R'^n\<bar>)"
proof -
have "(R' + R) / 2 < R" and "0 < (R' + R) / 2"
- using `0 < R'` `0 < R` `R' < R` by auto
+ using \<open>0 < R'\<close> \<open>0 < R\<close> \<open>R' < R\<close> by auto
hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}"
- using `R' < R` by auto
+ using \<open>R' < R\<close> by auto
have "norm R' < norm ((R' + R) / 2)"
- using `0 < R'` `0 < R` `R' < R` by auto
+ using \<open>0 < R'\<close> \<open>0 < R\<close> \<open>R' < R\<close> by auto
from powser_insidea[OF converges[OF in_Rball] this] show ?thesis
by auto
qed
@@ -1010,16 +1010,16 @@
hence "\<bar>x^n\<bar> \<le> R'^n"
unfolding power_abs by (rule power_mono, auto)
}
- from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
+ from mult_mono[OF this[OF \<open>x \<in> {-R'<..<R'}\<close>, of p] this[OF \<open>y \<in> {-R'<..<R'}\<close>, of "n-p"]] \<open>0 < R'\<close>
have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)"
unfolding abs_mult by auto
thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n"
- unfolding power_add[symmetric] using `p \<le> n` by auto
+ unfolding power_add[symmetric] using \<open>p \<le> n\<close> by auto
qed
also have "\<dots> = real (Suc n) * R' ^ n"
unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
finally show "\<bar>\<Sum>p<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>"
- unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
+ unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \<open>0 < R'\<close>]]] .
show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>"
unfolding abs_mult[symmetric] by auto
qed
@@ -1037,7 +1037,7 @@
fix x
assume "x \<in> {-R' <..< R'}"
hence "R' \<in> {-R <..< R}" and "norm x < norm R'"
- using assms `R' < R` by auto
+ using assms \<open>R' < R\<close> by auto
have "summable (\<lambda> n. f n * x^n)"
proof (rule summable_comparison_test, intro exI allI impI)
fix n
@@ -1046,14 +1046,14 @@
show "norm (f n * x^n) \<le> norm (f n * real (Suc n) * x^n)"
unfolding real_norm_def abs_mult
by (rule mult_right_mono) (auto simp add: le[unfolded mult_1_right])
- qed (rule powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`])
+ qed (rule powser_insidea[OF converges[OF \<open>R' \<in> {-R <..< R}\<close>] \<open>norm x < norm R'\<close>])
from this[THEN summable_mult2[where c=x], unfolded mult.assoc, unfolded mult.commute]
show "summable (?f x)" by auto
}
show "summable (?f' x0)"
- using converges[OF `x0 \<in> {-R <..< R}`] .
+ using converges[OF \<open>x0 \<in> {-R <..< R}\<close>] .
show "x0 \<in> {-R' <..< R'}"
- using `x0 \<in> {-R' <..< R'}` .
+ using \<open>x0 \<in> {-R' <..< R'}\<close> .
qed
} note for_subinterval = this
let ?R = "(R + \<bar>x0\<bar>) / 2"
@@ -1061,7 +1061,7 @@
hence "- ?R < x0"
proof (cases "x0 < 0")
case True
- hence "- x0 < ?R" using `\<bar>x0\<bar> < ?R` by auto
+ hence "- x0 < ?R" using \<open>\<bar>x0\<bar> < ?R\<close> by auto
thus ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto
next
case False
@@ -1076,7 +1076,7 @@
qed
-subsection {* Exponential Function *}
+subsection \<open>Exponential Function\<close>
definition exp :: "'a \<Rightarrow> 'a::{real_normed_algebra_1,banach}"
where "exp = (\<lambda>x. \<Sum>n. x^n /\<^sub>R fact n)"
@@ -1187,7 +1187,7 @@
unfolding continuous_on_def by (auto intro: tendsto_exp)
-subsubsection {* Properties of the Exponential Function *}
+subsubsection \<open>Properties of the Exponential Function\<close>
lemma exp_zero [simp]: "exp 0 = 1"
unfolding exp_def by (simp add: scaleR_conv_of_real powser_zero)
@@ -1301,11 +1301,11 @@
by (induction I rule: finite_induct) (auto simp: exp_add_commuting mult.commute)
-subsubsection {* Properties of the Exponential Function on Reals *}
-
-text {* Comparisons of @{term "exp x"} with zero. *}
-
-text{*Proof: because every exponential can be seen as a square.*}
+subsubsection \<open>Properties of the Exponential Function on Reals\<close>
+
+text \<open>Comparisons of @{term "exp x"} with zero.\<close>
+
+text\<open>Proof: because every exponential can be seen as a square.\<close>
lemma exp_ge_zero [simp]: "0 \<le> exp (x::real)"
proof -
have "0 \<le> exp (x/2) * exp (x/2)" by simp
@@ -1324,7 +1324,7 @@
lemma abs_exp_cancel [simp]: "\<bar>exp x::real\<bar> = exp x"
by simp
-text {* Strict monotonicity of exponential. *}
+text \<open>Strict monotonicity of exponential.\<close>
lemma exp_ge_add_one_self_aux:
assumes "0 \<le> (x::real)" shows "1+x \<le> exp(x)"
@@ -1335,7 +1335,7 @@
by (auto simp add: numeral_2_eq_2)
also have "... \<le> (\<Sum>n. inverse (fact n) * x^n)"
apply (rule setsum_le_suminf [OF summable_exp])
- using `0 < x`
+ using \<open>0 < x\<close>
apply (auto simp add: zero_le_mult_iff)
done
finally show "1+x \<le> exp x"
@@ -1360,7 +1360,7 @@
assumes "x < y"
shows "exp x < exp y"
proof -
- from `x < y` have "0 < y - x" by simp
+ from \<open>x < y\<close> have "0 < y - x" by simp
hence "1 < exp (y - x)" by (rule exp_gt_one)
hence "1 < exp y / exp x" by (simp only: exp_diff)
thus "exp x < exp y" by simp
@@ -1379,7 +1379,7 @@
lemma exp_inj_iff [iff]: "exp (x::real) = exp y \<longleftrightarrow> x = y"
by (simp add: order_eq_iff)
-text {* Comparisons of @{term "exp x"} with one. *}
+text \<open>Comparisons of @{term "exp x"} with one.\<close>
lemma one_less_exp_iff [simp]: "1 < exp (x::real) \<longleftrightarrow> 0 < x"
using exp_less_cancel_iff [where x=0 and y=x] by simp
@@ -1417,14 +1417,14 @@
qed
-subsection {* Natural Logarithm *}
+subsection \<open>Natural Logarithm\<close>
class ln = real_normed_algebra_1 + banach +
fixes ln :: "'a \<Rightarrow> 'a"
assumes ln_one [simp]: "ln 1 = 0"
definition powr :: "['a,'a] => 'a::ln" (infixr "powr" 80)
- -- {*exponentation via ln and exp*}
+ -- \<open>exponentation via ln and exp\<close>
where [code del]: "x powr a \<equiv> if x = 0 then 0 else exp(a * ln x)"
lemma powr_0 [simp]: "0 powr z = 0"
@@ -1553,7 +1553,7 @@
ultimately show ?thesis
by simp
next
- assume "\<not> 0 < x" with `x \<noteq> 0` show "isCont ln x"
+ assume "\<not> 0 < x" with \<open>x \<noteq> 0\<close> show "isCont ln x"
unfolding isCont_def
by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\<lambda>_. ln 0"])
(auto simp: ln_neg_is_const not_less eventually_at dist_real_def
@@ -1606,28 +1606,28 @@
assume "x \<in> {0 <..< 2}"
hence "0 < x" and "x < 2" by auto
have "norm (1 - x) < 1"
- using `0 < x` and `x < 2` by auto
+ using \<open>0 < x\<close> and \<open>x < 2\<close> by auto
have "1 / x = 1 / (1 - (1 - x))" by auto
also have "\<dots> = (\<Sum> n. (1 - x)^n)"
- using geometric_sums[OF `norm (1 - x) < 1`] by (rule sums_unique)
+ using geometric_sums[OF \<open>norm (1 - x) < 1\<close>] by (rule sums_unique)
also have "\<dots> = suminf (?f' x)"
unfolding power_mult_distrib[symmetric]
by (rule arg_cong[where f=suminf], rule arg_cong[where f="op ^"], auto)
finally have "DERIV ln x :> suminf (?f' x)"
- using DERIV_ln[OF `0 < x`] unfolding divide_inverse by auto
+ using DERIV_ln[OF \<open>0 < x\<close>] unfolding divide_inverse by auto
moreover
have repos: "\<And> h x :: real. h - 1 + x = h + x - 1" by auto
have "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :>
(\<Sum>n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)"
proof (rule DERIV_power_series')
show "x - 1 \<in> {- 1<..<1}" and "(0 :: real) < 1"
- using `0 < x` `x < 2` by auto
+ using \<open>0 < x\<close> \<open>x < 2\<close> by auto
fix x :: real
assume "x \<in> {- 1<..<1}"
hence "norm (-x) < 1" by auto
show "summable (\<lambda>n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)"
unfolding One_nat_def
- by (auto simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
+ by (auto simp add: power_mult_distrib[symmetric] summable_geometric[OF \<open>norm (-x) < 1\<close>])
qed
hence "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> suminf (?f' x)"
unfolding One_nat_def by auto
@@ -1945,33 +1945,33 @@
show ?thesis
proof (cases rule: linorder_cases)
assume "x < 1"
- from dense[OF `x < 1`] obtain a where "x < a" "a < 1" by blast
- from `x < a` have "?l x < ?l a"
+ from dense[OF \<open>x < 1\<close>] obtain a where "x < a" "a < 1" by blast
+ from \<open>x < a\<close> have "?l x < ?l a"
proof (rule DERIV_pos_imp_increasing, safe)
fix y
assume "x \<le> y" "y \<le> a"
- with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y"
+ with \<open>0 < x\<close> \<open>a < 1\<close> have "0 < 1 / y - 1" "0 < y"
by (auto simp: field_simps)
with D show "\<exists>z. DERIV ?l y :> z \<and> 0 < z"
by auto
qed
also have "\<dots> \<le> 0"
- using ln_le_minus_one `0 < x` `x < a` by (auto simp: field_simps)
+ using ln_le_minus_one \<open>0 < x\<close> \<open>x < a\<close> by (auto simp: field_simps)
finally show "x = 1" using assms by auto
next
assume "1 < x"
from dense[OF this] obtain a where "1 < a" "a < x" by blast
- from `a < x` have "?l x < ?l a"
+ from \<open>a < x\<close> have "?l x < ?l a"
proof (rule DERIV_neg_imp_decreasing, safe)
fix y
assume "a \<le> y" "y \<le> x"
- with `1 < a` have "1 / y - 1 < 0" "0 < y"
+ with \<open>1 < a\<close> have "1 / y - 1 < 0" "0 < y"
by (auto simp: field_simps)
with D show "\<exists>z. DERIV ?l y :> z \<and> z < 0"
by blast
qed
also have "\<dots> \<le> 0"
- using ln_le_minus_one `1 < a` by (auto simp: field_simps)
+ using ln_le_minus_one \<open>1 < a\<close> by (auto simp: field_simps)
finally show "x = 1" using assms by auto
next
assume "x = 1"
@@ -1988,7 +1988,7 @@
assume "x < ln r"
then have "exp x < exp (ln r)"
by simp
- with `0 < r` have "exp x < r"
+ with \<open>0 < r\<close> have "exp x < r"
by simp
}
then show "\<exists>k. \<forall>n<k. exp n < r" by auto
@@ -2048,7 +2048,7 @@
definition log :: "[real,real] => real"
- -- {*logarithm of @{term x} to base @{term a}*}
+ -- \<open>logarithm of @{term x} to base @{term a}\<close>
where "log a x = ln x / ln a"
@@ -2170,7 +2170,7 @@
proof -
def lb \<equiv> "1 / ln b"
moreover have "DERIV (\<lambda>y. lb * ln y) x :> lb / x"
- using `x > 0` by (auto intro!: derivative_eq_intros)
+ using \<open>x > 0\<close> by (auto intro!: derivative_eq_intros)
ultimately show ?thesis
by (simp add: log_def)
qed
@@ -2193,7 +2193,7 @@
log a x = (ln b/ln a) * log b x"
by (simp add: log_def divide_inverse)
-text{*Base 10 logarithms*}
+text\<open>Base 10 logarithms\<close>
lemma log_base_10_eq1: "0 < x \<Longrightarrow> log 10 x = (ln (exp 1) / ln 10) * ln x"
by (simp add: log_def)
@@ -2242,11 +2242,11 @@
then show ?thesis by simp
next
assume "x < y" hence "log b x < log b y"
- using log_less_cancel_iff[OF `1 < b`] pos by simp
+ using log_less_cancel_iff[OF \<open>1 < b\<close>] pos by simp
then show ?thesis using * by simp
next
assume "y < x" hence "log b y < log b x"
- using log_less_cancel_iff[OF `1 < b`] pos by simp
+ using log_less_cancel_iff[OF \<open>1 < b\<close>] pos by simp
then show ?thesis using * by simp
qed
qed
@@ -2333,7 +2333,7 @@
proof (cases "i < 0")
case True
have r: "x powr i = 1 / x powr (-i)" by (simp add: powr_minus field_simps)
- show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powr_realpow[symmetric])
+ show ?thesis using \<open>i < 0\<close> \<open>x > 0\<close> by (simp add: r field_simps powr_realpow[symmetric])
next
case False
then show ?thesis by (simp add: assms powr_realpow[symmetric])
@@ -2513,7 +2513,7 @@
elim: eventually_elim1 intro: tendsto_mono inf_le1)
then have "((\<lambda>x. exp (g x * ln (f x))) ---> 0) (inf F (principal {x. f x \<noteq> 0}))"
by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_0]
- filterlim_tendsto_pos_mult_at_bot[OF _ `0 < b`]
+ filterlim_tendsto_pos_mult_at_bot[OF _ \<open>0 < b\<close>]
intro: tendsto_mono inf_le1 g) }
then show "((\<lambda>x. exp (g x * ln (f x))) ---> (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x \<noteq> 0}))"
using f g by (auto intro!: tendsto_intros intro: tendsto_mono inf_le1)
@@ -2577,7 +2577,7 @@
proof (rule filterlim_mono_eventually)
show "eventually (\<lambda>xa. exp (ln (1 + x * xa) / xa) = (1 + x * xa) powr (1 / xa)) (at_right 0)"
unfolding eventually_at_right[OF zero_less_one]
- using `x \<noteq> 0`
+ using \<open>x \<noteq> 0\<close>
apply (intro exI[of _ "1 / \<bar>x\<bar>"])
apply (auto simp: field_simps powr_def abs_if)
by (metis add_less_same_cancel1 mult_less_0_iff not_less_iff_gr_or_eq zero_less_one)
@@ -2610,7 +2610,7 @@
by (rule filterlim_compose [OF tendsto_exp_limit_at_top filterlim_real_sequentially])
qed auto
-subsection {* Sine and Cosine *}
+subsection \<open>Sine and Cosine\<close>
definition sin_coeff :: "nat \<Rightarrow> real" where
"sin_coeff = (\<lambda>n. if even n then 0 else (- 1) ^ ((n - Suc 0) div 2) / (fact n))"
@@ -2710,7 +2710,7 @@
lemma diffs_cos_coeff: "diffs cos_coeff = (\<lambda>n. - sin_coeff n)"
by (simp add: diffs_def cos_coeff_Suc real_of_nat_def del: of_nat_Suc)
-text{*Now at last we can get the derivatives of exp, sin and cos*}
+text\<open>Now at last we can get the derivatives of exp, sin and cos\<close>
lemma DERIV_sin [simp]:
fixes x :: "'a::{real_normed_field,banach}"
@@ -2803,7 +2803,7 @@
shows "continuous (at z within s) cos"
by (simp add: continuous_within tendsto_cos)
-subsection {* Properties of Sine and Cosine *}
+subsection \<open>Properties of Sine and Cosine\<close>
lemma sin_zero [simp]: "sin 0 = 0"
unfolding sin_def sin_coeff_def by (simp add: scaleR_conv_of_real powser_zero)
@@ -2819,9 +2819,9 @@
"DERIV g x :> m \<Longrightarrow> DERIV (\<lambda>x. cos(g x)) x :> -sin(g x) * m"
by (auto intro!: derivative_eq_intros simp: real_of_nat_def)
-subsection {*Deriving the Addition Formulas*}
-
-text{*The the product of two cosine series*}
+subsection \<open>Deriving the Addition Formulas\<close>
+
+text\<open>The the product of two cosine series\<close>
lemma cos_x_cos_y:
fixes x :: "'a::{real_normed_field,banach}"
shows "(\<lambda>p. \<Sum>n\<le>p.
@@ -2835,7 +2835,7 @@
by (metis div_add power_add le_add_diff_inverse odd_add)
have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) =
(if even p \<and> even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)"
- using `n\<le>p`
+ using \<open>n\<le>p\<close>
by (auto simp: * algebra_simps cos_coeff_def binomial_fact real_of_nat_def)
}
then have "(\<lambda>p. \<Sum>n\<le>p. if even p \<and> even n
@@ -2850,7 +2850,7 @@
finally show ?thesis .
qed
-text{*The product of two sine series*}
+text\<open>The product of two sine series\<close>
lemma sin_x_sin_y:
fixes x :: "'a::{real_normed_field,banach}"
shows "(\<lambda>p. \<Sum>n\<le>p.
@@ -2861,12 +2861,12 @@
{ fix n p::nat
assume "n\<le>p"
{ assume np: "odd n" "even p"
- with `n\<le>p` have "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \<le> p"
+ with \<open>n\<le>p\<close> have "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \<le> p"
by arith+
moreover have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0"
by simp
ultimately have *: "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))"
- using np `n\<le>p`
+ using np \<open>n\<le>p\<close>
apply (simp add: power_add [symmetric] div_add [symmetric] del: div_add)
apply (metis (no_types) One_nat_def Suc_1 le_div_geq minus_minus mult.left_neutral mult_minus_left power.simps(2) zero_less_Suc)
done
@@ -2874,7 +2874,7 @@
have "(sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) =
(if even p \<and> odd n
then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)"
- using `n\<le>p`
+ using \<open>n\<le>p\<close>
by (auto simp: algebra_simps sin_coeff_def binomial_fact real_of_nat_def)
}
then have "(\<lambda>p. \<Sum>n\<le>p. if even p \<and> odd n
@@ -3053,13 +3053,13 @@
"DERIV g x :> m ==> DERIV (\<lambda>x. exp(g x)) x :> exp(g x) * m"
by (auto intro!: derivative_intros)
-subsection {* The Constant Pi *}
+subsection \<open>The Constant Pi\<close>
definition pi :: real
where "pi = 2 * (THE x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)"
-text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"};
- hence define pi.*}
+text\<open>Show that there's a least positive @{term x} with @{term "cos(x) = 0"};
+ hence define pi.\<close>
lemma sin_paired:
fixes x :: real
@@ -3086,7 +3086,7 @@
have "x * x < ?k2 * ?k3"
using assms by (intro mult_strict_mono', simp_all)
hence "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)"
- by (intro mult_strict_right_mono zero_less_power `0 < x`)
+ by (intro mult_strict_right_mono zero_less_power \<open>0 < x\<close>)
thus "0 < ?f n"
by (simp add: real_of_nat_def divide_simps mult_ac del: mult_Suc)
qed
@@ -3439,7 +3439,7 @@
assume "\<not> 2 \<le> pi" hence "pi < 2" by auto
have "\<exists>y > pi. y < 2 \<and> y < 2*pi"
proof (cases "2 < 2*pi")
- case True with dense[OF `pi < 2`] show ?thesis by auto
+ case True with dense[OF \<open>pi < 2\<close>] show ?thesis by auto
next
case False have "pi < 2*pi" by auto
from dense[OF this] and False show ?thesis by auto
@@ -3447,7 +3447,7 @@
then obtain y where "pi < y" and "y < 2" and "y < 2*pi" by blast
hence "0 < sin y" using sin_gt_zero_02 by auto
moreover
- have "sin y < 0" using sin_gt_zero[of "y - pi"] `pi < y` and `y < 2*pi` sin_periodic_pi[of "y - pi"] by auto
+ have "sin y < 0" using sin_gt_zero[of "y - pi"] \<open>pi < y\<close> and \<open>y < 2*pi\<close> sin_periodic_pi[of "y - pi"] by auto
ultimately show False by auto
qed
@@ -3458,8 +3458,8 @@
using sin_ge_zero [of "x-pi"]
by (simp add: sin_diff)
-text {* FIXME: This proof is almost identical to lemma @{text cos_is_zero}.
- It should be possible to factor out some of the common parts. *}
+text \<open>FIXME: This proof is almost identical to lemma @{text cos_is_zero}.
+ It should be possible to factor out some of the common parts.\<close>
lemma cos_total: "\<lbrakk>-1 \<le> y; y \<le> 1\<rbrakk> \<Longrightarrow> EX! x. 0 \<le> x & x \<le> pi & (cos x = y)"
proof (rule ex_ex1I)
@@ -3625,14 +3625,14 @@
proof -
have "- (x - y) < 0" using assms by auto
- from MVT2[OF `y < x` DERIV_cos[THEN impI, THEN allI]]
+ from MVT2[OF \<open>y < x\<close> DERIV_cos[THEN impI, THEN allI]]
obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z"
by auto
hence "0 < z" and "z < pi" using assms by auto
hence "0 < sin z" using sin_gt_zero by auto
hence "cos x - cos y < 0"
unfolding cos_diff minus_mult_commute[symmetric]
- using `- (x - y) < 0` by (rule mult_pos_neg2)
+ using \<open>- (x - y) < 0\<close> by (rule mult_pos_neg2)
thus ?thesis by auto
qed
@@ -3642,10 +3642,10 @@
proof (cases "y < x")
case True
show ?thesis
- using cos_monotone_0_pi[OF `0 \<le> y` True `x \<le> pi`] by auto
+ using cos_monotone_0_pi[OF \<open>0 \<le> y\<close> True \<open>x \<le> pi\<close>] by auto
next
case False
- hence "y = x" using `y \<le> x` by auto
+ hence "y = x" using \<open>y \<le> x\<close> by auto
thus ?thesis by auto
qed
@@ -3664,11 +3664,11 @@
shows "cos y \<le> cos x"
proof (cases "y < x")
case True
- show ?thesis using cos_monotone_minus_pi_0[OF `-pi \<le> y` True `x \<le> 0`]
+ show ?thesis using cos_monotone_minus_pi_0[OF \<open>-pi \<le> y\<close> True \<open>x \<le> 0\<close>]
by auto
next
case False
- hence "y = x" using `y \<le> x` by auto
+ hence "y = x" using \<open>y \<le> x\<close> by auto
thus ?thesis by auto
qed
@@ -3716,7 +3716,7 @@
by (auto simp: abs_real_def)
-subsection {* More Corollaries about Sine and Cosine *}
+subsection \<open>More Corollaries about Sine and Cosine\<close>
lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n"
proof -
@@ -3780,7 +3780,7 @@
assume n: "even n" "x = real n * (pi/2)"
then obtain m where m: "n = 2 * m"
using dvdE by blast
- then have me: "even m" using `?lhs` n
+ then have me: "even m" using \<open>?lhs\<close> n
by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one)
show ?rhs
using m me n
@@ -3790,7 +3790,7 @@
assume n: "even n" "x = - (real n * (pi/2))"
then obtain m where m: "n = 2 * m"
using dvdE by blast
- then have me: "even m" using `?lhs` n
+ then have me: "even m" using \<open>?lhs\<close> n
by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one)
show ?rhs
using m me n
@@ -3803,7 +3803,7 @@
qed
lemma cos_one_2pi_int: "cos(x) = 1 \<longleftrightarrow> (\<exists>n::int. x = n * 2*pi)"
- apply auto --{*FIXME simproc bug*}
+ apply auto --\<open>FIXME simproc bug\<close>
apply (auto simp: cos_one_2pi)
apply (metis real_of_int_of_nat_eq)
apply (metis mult_minus_right real_of_int_minus real_of_int_of_nat_eq)
@@ -3901,7 +3901,7 @@
done
-subsection {* Tangent *}
+subsection \<open>Tangent\<close>
definition tan :: "'a \<Rightarrow> 'a::{real_normed_field,banach}"
where "tan = (\<lambda>x. sin x / cos x)"
@@ -4063,7 +4063,7 @@
apply (rule_tac [2] Rolle)
apply (auto del: exI intro!: DERIV_tan DERIV_isCont exI
simp add: real_differentiable_def)
- txt{*Now, simulate TRYALL*}
+ txt\<open>Now, simulate TRYALL\<close>
apply (rule_tac [!] DERIV_tan asm_rl)
apply (auto dest!: DERIV_unique [OF _ DERIV_tan]
simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym])
@@ -4082,13 +4082,13 @@
have "cos x' \<noteq> 0" by auto
thus "DERIV tan x' :> inverse ((cos x')\<^sup>2)" by (rule DERIV_tan)
qed
- from MVT2[OF `y < x` this]
+ from MVT2[OF \<open>y < x\<close> this]
obtain z where "y < z" and "z < x"
and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto
hence "- (pi / 2) < z" and "z < pi / 2" using assms by auto
hence "0 < cos z" using cos_gt_zero_pi by auto
hence inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto
- have "0 < x - y" using `y < x` by auto
+ have "0 < x - y" using \<open>y < x\<close> by auto
with inv_pos have "0 < tan x - tan y" unfolding tan_diff by auto
thus ?thesis by auto
qed
@@ -4102,7 +4102,7 @@
proof
assume "y < x"
thus "tan y < tan x"
- using tan_monotone and `- (pi / 2) < y` and `x < pi / 2` by auto
+ using tan_monotone and \<open>- (pi / 2) < y\<close> and \<open>x < pi / 2\<close> by auto
next
assume "tan y < tan x"
show "y < x"
@@ -4112,10 +4112,10 @@
proof (cases "x = y")
case True thus ?thesis by auto
next
- case False hence "x < y" using `x \<le> y` by auto
- from tan_monotone[OF `- (pi/2) < x` this `y < pi / 2`] show ?thesis by auto
+ case False hence "x < y" using \<open>x \<le> y\<close> by auto
+ from tan_monotone[OF \<open>- (pi/2) < x\<close> this \<open>y < pi / 2\<close>] show ?thesis by auto
qed
- thus False using `tan y < tan x` by auto
+ thus False using \<open>tan y < tan x\<close> by auto
qed
qed
@@ -4197,7 +4197,7 @@
lemma tan_cot: "tan(pi/2 - x) = inverse(tan x)"
by (simp add: tan_def sin_diff cos_diff)
-subsection {* Inverse Trigonometric Functions *}
+subsection \<open>Inverse Trigonometric Functions\<close>
definition arcsin :: "real => real"
where "arcsin y = (THE x. -(pi/2) \<le> x & x \<le> pi/2 & sin x = y)"
@@ -4393,7 +4393,7 @@
have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1"
unfolding tan_def by (simp add: distrib_left power_divide)
thus "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2"
- using `0 < 1 + x\<^sup>2` by (simp add: arctan power_divide eq_divide_eq)
+ using \<open>0 < 1 + x\<^sup>2\<close> by (simp add: arctan power_divide eq_divide_eq)
qed
lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)"
@@ -4553,7 +4553,7 @@
assume "0 < e"
def y \<equiv> "pi/2 - min (pi/2) e"
then have y: "0 \<le> y" "y < pi/2" "pi/2 \<le> e + y"
- using `0 < e` by auto
+ using \<open>0 < e\<close> by auto
show "eventually (\<lambda>x. dist (arctan x) (pi / 2) < e) at_top"
proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI)
@@ -4563,7 +4563,7 @@
by (simp add: arctan_less_iff)
with y have "y < arctan x"
by (subst (asm) arctan_tan) simp_all
- with arctan_ubound[of x, arith] y `0 < e`
+ with arctan_ubound[of x, arith] y \<open>0 < e\<close>
show "dist (arctan x) (pi / 2) < e"
by (simp add: dist_real_def)
qed
@@ -4574,7 +4574,7 @@
by (intro tendsto_minus tendsto_arctan_at_top)
-subsection{* Prove Totality of the Trigonometric Functions *}
+subsection\<open>Prove Totality of the Trigonometric Functions\<close>
lemma cos_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> cos (arccos y) = y"
by (simp add: abs_le_iff)
@@ -4639,7 +4639,7 @@
then obtain t where "t\<ge>0" "t \<le> pi/2" "-x = cos t" "y = sin t"
using sincos_total_pi_half assms
apply auto
- by (metis `0 \<le> - x` power2_minus)
+ by (metis \<open>0 \<le> - x\<close> power2_minus)
then show ?thesis
by (rule_tac x="pi-t" in exI, auto)
qed
@@ -4658,7 +4658,7 @@
then obtain t where "t\<ge>0" "t \<le> pi" "x = cos t" "-y = sin t"
using sincos_total_pi assms
apply auto
- by (metis `0 \<le> - y` power2_minus)
+ by (metis \<open>0 \<le> - y\<close> power2_minus)
then show ?thesis
by (rule_tac x="2*pi-t" in exI, auto)
qed
@@ -4711,7 +4711,7 @@
lemma arccos_eq_iff: "abs x \<le> 1 & abs y \<le> 1 \<Longrightarrow> (arccos x = arccos y \<longleftrightarrow> x = y)"
using cos_arccos_abs by fastforce
-subsection {* Machins formula *}
+subsection \<open>Machins formula\<close>
lemma arctan_one: "arctan 1 = pi / 4"
by (rule arctan_unique, simp_all add: tan_45 m2pi_less_pi)
@@ -4796,7 +4796,7 @@
12 * arctan(1/18) + 8 * arctan(1/57) - 5 * arctan(1/239) = pi/4*)
-subsection {* Introducing the inverse tangent power series *}
+subsection \<open>Introducing the inverse tangent power series\<close>
lemma monoseq_arctan_series:
fixes x :: real
@@ -4822,23 +4822,23 @@
show "0 \<le> 1 / real (Suc (n * 2))"
by auto
show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)"
- by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
+ by (rule power_decreasing) (simp_all add: \<open>0 \<le> x\<close> \<open>x \<le> 1\<close>)
show "0 \<le> x ^ Suc (Suc n * 2)"
- by (rule zero_le_power) (simp add: `0 \<le> x`)
+ by (rule zero_le_power) (simp add: \<open>0 \<le> x\<close>)
qed
} note mono = this
show ?thesis
proof (cases "0 \<le> x")
- case True from mono[OF this `x \<le> 1`, THEN allI]
+ case True from mono[OF this \<open>x \<le> 1\<close>, THEN allI]
show ?thesis unfolding Suc_eq_plus1[symmetric]
by (rule mono_SucI2)
next
case False
- hence "0 \<le> -x" and "-x \<le> 1" using `-1 \<le> x` by auto
+ hence "0 \<le> -x" and "-x \<le> 1" using \<open>-1 \<le> x\<close> by auto
from mono[OF this]
have "\<And>n. 1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<ge>
- 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" using `0 \<le> -x` by auto
+ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" using \<open>0 \<le> -x\<close> by auto
thus ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI])
qed
qed
@@ -4859,13 +4859,13 @@
proof (cases "\<bar>x\<bar> < 1")
case True
hence "norm x < 1" by auto
- from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF `norm x < 1`, THEN LIMSEQ_Suc]]
+ from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF \<open>norm x < 1\<close>, THEN LIMSEQ_Suc]]
have "(\<lambda>n. 1 / real (n + 1) * x ^ (n + 1)) ----> 0"
unfolding inverse_eq_divide Suc_eq_plus1 by simp
then show ?thesis using pos2 by (rule LIMSEQ_linear)
next
case False
- hence "x = -1 \<or> x = 1" using `\<bar>x\<bar> \<le> 1` by auto
+ hence "x = -1 \<or> x = 1" using \<open>\<bar>x\<bar> \<le> 1\<close> by auto
hence n_eq: "\<And> n. x ^ (n * 2 + 1) = x"
unfolding One_nat_def by auto
from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]]
@@ -4873,7 +4873,7 @@
qed
qed
-text{*FIXME: generalise from the reals via type classes?*}
+text\<open>FIXME: generalise from the reals via type classes?\<close>
lemma summable_arctan_series:
fixes x :: real and n :: nat
assumes "\<bar>x\<bar> \<le> 1"
@@ -4899,7 +4899,7 @@
assume "\<bar>x\<bar> < 1"
hence "x\<^sup>2 < 1" by (simp add: abs_square_less_1)
have "summable (\<lambda> n. (- 1) ^ n * (x\<^sup>2) ^n)"
- by (rule summable_Leibniz(1), auto intro!: LIMSEQ_realpow_zero monoseq_realpow `x\<^sup>2 < 1` order_less_imp_le[OF `x\<^sup>2 < 1`])
+ by (rule summable_Leibniz(1), auto intro!: LIMSEQ_realpow_zero monoseq_realpow \<open>x\<^sup>2 < 1\<close> order_less_imp_le[OF \<open>x\<^sup>2 < 1\<close>])
hence "summable (\<lambda> n. (- 1) ^ n * x^(2*n))" unfolding power_mult .
} note summable_Integral = this
@@ -4938,7 +4938,7 @@
have "DERIV (\<lambda> x. \<Sum> n. ?f n * x^(Suc n)) x :> (\<Sum> n. ?f n * real (Suc n) * x^n)"
proof (rule DERIV_power_series')
- show "x \<in> {- 1 <..< 1}" using `\<bar> x \<bar> < 1` by auto
+ show "x \<in> {- 1 <..< 1}" using \<open>\<bar> x \<bar> < 1\<close> by auto
{
fix x' :: real
assume x'_bounds: "x' \<in> {- 1 <..< 1}"
@@ -4969,7 +4969,7 @@
{
fix r x :: real
assume "0 < r" and "r < 1" and "\<bar> x \<bar> < r"
- have "\<bar>x\<bar> < 1" using `r < 1` and `\<bar>x\<bar> < r` by auto
+ have "\<bar>x\<bar> < 1" using \<open>r < 1\<close> and \<open>\<bar>x\<bar> < r\<close> by auto
from DERIV_arctan_series[OF this] have "DERIV (\<lambda> x. suminf (?c x)) x :> (suminf (?c' x))" .
} note DERIV_arctan_suminf = this
@@ -4985,7 +4985,7 @@
have "arctan x = (\<Sum>k. ?c x k)"
proof -
obtain r where "\<bar>x\<bar> < r" and "r < 1"
- using dense[OF `\<bar>x\<bar> < 1`] by blast
+ using dense[OF \<open>\<bar>x\<bar> < 1\<close>] by blast
hence "0 < r" and "-r < x" and "x < r" by auto
have suminf_eq_arctan_bounded: "\<And>x a b. \<lbrakk> -r < a ; b < r ; a < b ; a \<le> x ; x \<le> b \<rbrakk> \<Longrightarrow>
@@ -4997,15 +4997,15 @@
show "suminf (?c x) - arctan x = suminf (?c a) - arctan a"
proof (rule DERIV_isconst2[of "a" "b"])
show "a < b" and "a \<le> x" and "x \<le> b"
- using `a < b` `a \<le> x` `x \<le> b` by auto
+ using \<open>a < b\<close> \<open>a \<le> x\<close> \<open>x \<le> b\<close> by auto
have "\<forall>x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
proof (rule allI, rule impI)
fix x
assume "-r < x \<and> x < r"
hence "\<bar>x\<bar> < r" by auto
- hence "\<bar>x\<bar> < 1" using `r < 1` by auto
+ hence "\<bar>x\<bar> < 1" using \<open>r < 1\<close> by auto
have "\<bar> - (x\<^sup>2) \<bar> < 1"
- using abs_square_less_1 `\<bar>x\<bar> < 1` by auto
+ using abs_square_less_1 \<open>\<bar>x\<bar> < 1\<close> by auto
hence "(\<lambda> n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))"
unfolding real_norm_def[symmetric] by (rule geometric_sums)
hence "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))"
@@ -5014,15 +5014,15 @@
using sums_unique unfolding inverse_eq_divide by auto
have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))"
unfolding suminf_c'_eq_geom
- by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
+ by (rule DERIV_arctan_suminf[OF \<open>0 < r\<close> \<open>r < 1\<close> \<open>\<bar>x\<bar> < r\<close>])
from DERIV_diff [OF this DERIV_arctan]
show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
by auto
qed
hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0"
- using `-r < a` `b < r` by auto
+ using \<open>-r < a\<close> \<open>b < r\<close> by auto
thus "\<forall> y. a < y \<and> y < b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0"
- using `\<bar>x\<bar> < r` by auto
+ using \<open>\<bar>x\<bar> < r\<close> by auto
show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y"
using DERIV_in_rball DERIV_isCont by auto
qed
@@ -5041,11 +5041,11 @@
hence "0 < \<bar>x\<bar>" and "- \<bar>x\<bar> < \<bar>x\<bar>" by auto
have "suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>) = suminf (?c 0) - arctan 0"
by (rule suminf_eq_arctan_bounded[where x1="0" and a1="-\<bar>x\<bar>" and b1="\<bar>x\<bar>", symmetric])
- (simp_all only: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>` neg_less_iff_less)
+ (simp_all only: \<open>\<bar>x\<bar> < r\<close> \<open>-\<bar>x\<bar> < \<bar>x\<bar>\<close> neg_less_iff_less)
moreover
have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
by (rule suminf_eq_arctan_bounded[where x1="x" and a1="-\<bar>x\<bar>" and b1="\<bar>x\<bar>"])
- (simp_all only: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>` neg_less_iff_less)
+ (simp_all only: \<open>\<bar>x\<bar> < r\<close> \<open>-\<bar>x\<bar> < \<bar>x\<bar>\<close> neg_less_iff_less)
ultimately
show ?thesis using suminf_arctan_zero by auto
qed
@@ -5059,7 +5059,7 @@
thus ?thesis by (rule when_less_one)
next
case False
- hence "\<bar>x\<bar> = 1" using `\<bar>x\<bar> \<le> 1` by auto
+ hence "\<bar>x\<bar> = 1" using \<open>\<bar>x\<bar> \<le> 1\<close> by auto
let ?a = "\<lambda>x n. \<bar>1 / real (n*2+1) * x^(n*2+1)\<bar>"
let ?diff = "\<lambda> x n. \<bar> arctan x - (\<Sum> i<n. ?c x i)\<bar>"
{
@@ -5070,18 +5070,18 @@
fix x :: real
assume "0 < x" and "x < 1"
hence "\<bar>x\<bar> \<le> 1" and "\<bar>x\<bar> < 1" by auto
- from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)"
+ from \<open>0 < x\<close> have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)"
by auto
- note bounds = mp[OF arctan_series_borders(2)[OF `\<bar>x\<bar> \<le> 1`] this, unfolded when_less_one[OF `\<bar>x\<bar> < 1`, symmetric], THEN spec]
+ note bounds = mp[OF arctan_series_borders(2)[OF \<open>\<bar>x\<bar> \<le> 1\<close>] this, unfolded when_less_one[OF \<open>\<bar>x\<bar> < 1\<close>, symmetric], THEN spec]
have "0 < 1 / real (n*2+1) * x^(n*2+1)"
- by (rule mult_pos_pos, auto simp only: zero_less_power[OF `0 < x`], auto)
+ by (rule mult_pos_pos, auto simp only: zero_less_power[OF \<open>0 < x\<close>], auto)
hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)"
by (rule abs_of_pos)
have "?diff x n \<le> ?a x n"
proof (cases "even n")
case True
hence sgn_pos: "(-1)^n = (1::real)" by auto
- from `even n` obtain m where "n = 2 * m" ..
+ from \<open>even n\<close> obtain m where "n = 2 * m" ..
then have "2 * m = n" ..
from bounds[of m, unfolded this atLeastAtMost_iff]
have "\<bar>arctan x - (\<Sum>i<n. (?c x i))\<bar> \<le> (\<Sum>i<n + 1. (?c x i)) - (\<Sum>i<n. (?c x i))"
@@ -5092,7 +5092,7 @@
next
case False
hence sgn_neg: "(-1)^n = (-1::real)" by auto
- from `odd n` obtain m where "n = 2 * m + 1" ..
+ from \<open>odd n\<close> obtain m where "n = 2 * m + 1" ..
then have m_def: "2 * m + 1 = n" ..
hence m_plus: "2 * (m + 1) = n + 1" by auto
from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
@@ -5122,10 +5122,10 @@
fix r :: real
assume "0 < r"
obtain N :: nat where N_I: "\<And>n. N \<le> n \<Longrightarrow> ?a 1 n < r"
- using LIMSEQ_D[OF `?a 1 ----> 0` `0 < r`] by auto
+ using LIMSEQ_D[OF \<open>?a 1 ----> 0\<close> \<open>0 < r\<close>] by auto
{
fix n
- assume "N \<le> n" from `?diff 1 n \<le> ?a 1 n` N_I[OF this]
+ assume "N \<le> n" from \<open>?diff 1 n \<le> ?a 1 n\<close> N_I[OF this]
have "norm (?diff 1 n - 0) < r" by auto
}
thus "\<exists> N. \<forall> n \<ge> N. norm (?diff 1 n - 0) < r" by blast
@@ -5137,10 +5137,10 @@
show ?thesis
proof (cases "x = 1")
case True
- then show ?thesis by (simp add: `arctan 1 = (\<Sum> i. ?c 1 i)`)
+ then show ?thesis by (simp add: \<open>arctan 1 = (\<Sum> i. ?c 1 i)\<close>)
next
case False
- hence "x = -1" using `\<bar>x\<bar> = 1` by auto
+ hence "x = -1" using \<open>\<bar>x\<bar> = 1\<close> by auto
have "- (pi / 2) < 0" using pi_gt_zero by auto
have "- (2 * pi) < 0" using pi_gt_zero by auto
@@ -5151,17 +5151,17 @@
have "arctan (- 1) = arctan (tan (-(pi / 4)))"
unfolding tan_45 tan_minus ..
also have "\<dots> = - (pi / 4)"
- by (rule arctan_tan, auto simp add: order_less_trans[OF `- (pi / 2) < 0` pi_gt_zero])
+ by (rule arctan_tan, auto simp add: order_less_trans[OF \<open>- (pi / 2) < 0\<close> pi_gt_zero])
also have "\<dots> = - (arctan (tan (pi / 4)))"
- unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric], auto simp add: order_less_trans[OF `- (2 * pi) < 0` pi_gt_zero])
+ unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric], auto simp add: order_less_trans[OF \<open>- (2 * pi) < 0\<close> pi_gt_zero])
also have "\<dots> = - (arctan 1)"
unfolding tan_45 ..
also have "\<dots> = - (\<Sum> i. ?c 1 i)"
- using `arctan 1 = (\<Sum> i. ?c 1 i)` by auto
+ using \<open>arctan 1 = (\<Sum> i. ?c 1 i)\<close> by auto
also have "\<dots> = (\<Sum> i. ?c (- 1) i)"
- using suminf_minus[OF sums_summable[OF `(?c 1) sums (arctan 1)`]]
+ using suminf_minus[OF sums_summable[OF \<open>(?c 1) sums (arctan 1)\<close>]]
unfolding c_minus_minus by auto
- finally show ?thesis using `x = -1` by auto
+ finally show ?thesis using \<open>x = -1\<close> by auto
qed
qed
qed
@@ -5182,21 +5182,21 @@
have "1 + (tan y)\<^sup>2 = 1 + (sin y)\<^sup>2 / (cos y)\<^sup>2"
unfolding tan_def power_divide ..
also have "\<dots> = (cos y)\<^sup>2 / (cos y)\<^sup>2 + (sin y)\<^sup>2 / (cos y)\<^sup>2"
- using `cos y \<noteq> 0` by auto
+ using \<open>cos y \<noteq> 0\<close> by auto
also have "\<dots> = 1 / (cos y)\<^sup>2"
unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 ..
finally have "1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2" .
have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)"
- unfolding tan_def using `cos y \<noteq> 0` by (simp add: field_simps)
+ unfolding tan_def using \<open>cos y \<noteq> 0\<close> by (simp add: field_simps)
also have "\<dots> = tan y / (1 + 1 / cos y)"
- using `cos y \<noteq> 0` unfolding add_divide_distrib by auto
+ using \<open>cos y \<noteq> 0\<close> unfolding add_divide_distrib by auto
also have "\<dots> = tan y / (1 + 1 / sqrt ((cos y)\<^sup>2))"
unfolding cos_sqrt ..
also have "\<dots> = tan y / (1 + sqrt (1 / (cos y)\<^sup>2))"
unfolding real_sqrt_divide by auto
finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)\<^sup>2))"
- unfolding `1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2` .
+ unfolding \<open>1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2\<close> .
have "arctan x = y"
using arctan_tan low high y_eq by auto
@@ -5205,7 +5205,7 @@
also have "\<dots> = 2 * (arctan (sin y / (cos y + 1)))"
unfolding tan_half by auto
finally show ?thesis
- unfolding eq `tan y = x` .
+ unfolding eq \<open>tan y = x\<close> .
qed
lemma arctan_monotone: "x < y \<Longrightarrow> arctan x < arctan y"
@@ -5243,7 +5243,7 @@
qed
-subsection {* Existence of Polar Coordinates *}
+subsection \<open>Existence of Polar Coordinates\<close>
lemma cos_x_y_le_one: "\<bar>x / sqrt (x\<^sup>2 + y\<^sup>2)\<bar> \<le> 1"
apply (rule power2_le_imp_le [OF _ zero_le_one])
@@ -5279,7 +5279,7 @@
qed
-subsection{*Basics about polynomial functions: products, extremal behaviour and root counts*}
+subsection\<open>Basics about polynomial functions: products, extremal behaviour and root counts\<close>
lemma pairs_le_eq_Sigma:
fixes m::nat
@@ -5435,7 +5435,7 @@
then have wnz: "\<And>w. w \<noteq> 0 \<Longrightarrow> (\<Sum>i\<le>n. c (Suc i) * w^i) = 0"
using Suc by auto
then have "(\<lambda>h. \<Sum>i\<le>n. c (Suc i) * h^i) -- 0 --> 0"
- by (simp cong: LIM_cong) --{*the case @{term"w=0"} by continuity*}
+ by (simp cong: LIM_cong) --\<open>the case @{term"w=0"} by continuity\<close>
then have "(\<Sum>i\<le>n. c (Suc i) * 0^i) = 0"
using isCont_polynom [of 0 "\<lambda>i. c (Suc i)" n] LIM_unique
by (force simp add: Limits.isCont_iff)
@@ -5444,7 +5444,7 @@
then have "\<And>i. i\<le>n \<Longrightarrow> c (Suc i) = 0"
using Suc.IH [of "\<lambda>i. c (Suc i)"]
by blast
- then show ?case using `k \<le> Suc n`
+ then show ?case using \<open>k \<le> Suc n\<close>
by (cases k) auto
qed
--- a/src/HOL/Transfer.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Transfer.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,13 +3,13 @@
Author: Ondrej Kuncar, TU Muenchen
*)
-section {* Generic theorem transfer using relations *}
+section \<open>Generic theorem transfer using relations\<close>
theory Transfer
imports Basic_BNF_LFPs Hilbert_Choice Metis
begin
-subsection {* Relator for function space *}
+subsection \<open>Relator for function space\<close>
locale lifting_syntax
begin
@@ -38,15 +38,15 @@
by (simp add: rel_fun_def)
-subsection {* Transfer method *}
+subsection \<open>Transfer method\<close>
-text {* Explicit tag for relation membership allows for
- backward proof methods. *}
+text \<open>Explicit tag for relation membership allows for
+ backward proof methods.\<close>
definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
where "Rel r \<equiv> r"
-text {* Handling of equality relations *}
+text \<open>Handling of equality relations\<close>
definition is_equality :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
where "is_equality R \<longleftrightarrow> R = (op =)"
@@ -54,12 +54,12 @@
lemma is_equality_eq: "is_equality (op =)"
unfolding is_equality_def by simp
-text {* Reverse implication for monotonicity rules *}
+text \<open>Reverse implication for monotonicity rules\<close>
definition rev_implies where
"rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)"
-text {* Handling of meta-logic connectives *}
+text \<open>Handling of meta-logic connectives\<close>
definition transfer_forall where
"transfer_forall \<equiv> All"
@@ -105,7 +105,7 @@
shows "Rel (A ===> B) (\<lambda>x. f x) (\<lambda>y. g y)"
using assms unfolding Rel_def rel_fun_def by fast
-subsection {* Predicates on relations, i.e. ``class constraints'' *}
+subsection \<open>Predicates on relations, i.e. ``class constraints''\<close>
definition left_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
where "left_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y)"
@@ -229,7 +229,7 @@
end
-subsection {* Equality restricted by a predicate *}
+subsection \<open>Equality restricted by a predicate\<close>
definition eq_onp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "eq_onp R = (\<lambda>x y. R x \<and> x = y)"
@@ -261,7 +261,7 @@
begin
interpretation lifting_syntax .
-text {* Handling of domains *}
+text \<open>Handling of domains\<close>
lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)"
by auto
@@ -275,7 +275,7 @@
"Domainp (op= ===> T) = (\<lambda>f. \<forall>x. (Domainp T) (f x))"
by (auto intro: choice simp: Domainp_iff rel_fun_def fun_eq_iff)
-text {* Properties are preserved by relation composition. *}
+text \<open>Properties are preserved by relation composition.\<close>
lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)"
by auto
@@ -301,7 +301,7 @@
unfolding left_unique_def OO_def by blast
-subsection {* Properties of relators *}
+subsection \<open>Properties of relators\<close>
lemma left_total_eq[transfer_rule]: "left_total op="
unfolding left_total_def by blast
@@ -384,7 +384,7 @@
declare pred_fun_def [simp]
declare rel_fun_eq [relator_eq]
-subsection {* Transfer rules *}
+subsection \<open>Transfer rules\<close>
context
begin
@@ -398,7 +398,7 @@
unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff
by fast
-text {* Transfer rules using implication instead of equality on booleans. *}
+text \<open>Transfer rules using implication instead of equality on booleans.\<close>
lemma transfer_forall_transfer [transfer_rule]:
"bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall"
@@ -425,7 +425,7 @@
"right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)"
unfolding right_unique_alt_def2 .
-text {* Transfer rules using equality. *}
+text \<open>Transfer rules using equality.\<close>
lemma left_unique_transfer [transfer_rule]:
assumes "right_total A"
@@ -587,13 +587,13 @@
thus "R'\<^sup>*\<^sup>* x' y'"
proof(induction arbitrary: y')
case base
- with `bi_unique A` `A x x'` have "x' = y'" by(rule bi_uniqueDr)
+ with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x' = y'" by(rule bi_uniqueDr)
thus ?case by simp
next
case (step y z z')
- from `bi_total A` obtain y' where "A y y'" unfolding bi_total_def by blast
+ from \<open>bi_total A\<close> obtain y' where "A y y'" unfolding bi_total_def by blast
hence "R'\<^sup>*\<^sup>* x' y'" by(rule step.IH)
- moreover from R `A y y'` `A z z'` `R y z`
+ moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R y z\<close>
have "R' y' z'" by(auto dest: rel_funD)
ultimately show ?case ..
qed
@@ -602,13 +602,13 @@
thus "R\<^sup>*\<^sup>* x y"
proof(induction arbitrary: y)
case base
- with `bi_unique A` `A x x'` have "x = y" by(rule bi_uniqueDl)
+ with \<open>bi_unique A\<close> \<open>A x x'\<close> have "x = y" by(rule bi_uniqueDl)
thus ?case by simp
next
case (step y' z' z)
- from `bi_total A` obtain y where "A y y'" unfolding bi_total_def by blast
+ from \<open>bi_total A\<close> obtain y where "A y y'" unfolding bi_total_def by blast
hence "R\<^sup>*\<^sup>* x y" by(rule step.IH)
- moreover from R `A y y'` `A z z'` `R' y' z'`
+ moreover from R \<open>A y y'\<close> \<open>A z z'\<close> \<open>R' y' z'\<close>
have "R y z" by(auto dest: rel_funD)
ultimately show ?case ..
qed
--- a/src/HOL/Transitive_Closure.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Transitive_Closure.thy Sat Jul 18 22:58:50 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Reflexive and Transitive closure of a relation *}
+section \<open>Reflexive and Transitive closure of a relation\<close>
theory Transitive_Closure
imports Relation
@@ -11,14 +11,14 @@
ML_file "~~/src/Provers/trancl.ML"
-text {*
+text \<open>
@{text rtrancl} is reflexive/transitive closure,
@{text trancl} is transitive closure,
@{text reflcl} is reflexive closure.
These postfix operators have \emph{maximum priority}, forcing their
operands to be atomic.
-*}
+\<close>
inductive_set
rtrancl :: "('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" ("(_^*)" [1000] 999)
@@ -68,7 +68,7 @@
reflcl ("(_\<^sup>=)" [1000] 999)
-subsection {* Reflexive closure *}
+subsection \<open>Reflexive closure\<close>
lemma refl_reflcl[simp]: "refl(r^=)"
by(simp add:refl_on_def)
@@ -82,23 +82,23 @@
lemma reflclp_idemp [simp]: "(P^==)^== = P^=="
by blast
-subsection {* Reflexive-transitive closure *}
+subsection \<open>Reflexive-transitive closure\<close>
lemma reflcl_set_eq [pred_set_conv]: "(sup (\<lambda>x y. (x, y) \<in> r) op =) = (\<lambda>x y. (x, y) \<in> r \<union> Id)"
by (auto simp add: fun_eq_iff)
lemma r_into_rtrancl [intro]: "!!p. p \<in> r ==> p \<in> r^*"
- -- {* @{text rtrancl} of @{text r} contains @{text r} *}
+ -- \<open>@{text rtrancl} of @{text r} contains @{text r}\<close>
apply (simp only: split_tupled_all)
apply (erule rtrancl_refl [THEN rtrancl_into_rtrancl])
done
lemma r_into_rtranclp [intro]: "r x y ==> r^** x y"
- -- {* @{text rtrancl} of @{text r} contains @{text r} *}
+ -- \<open>@{text rtrancl} of @{text r} contains @{text r}\<close>
by (erule rtranclp.rtrancl_refl [THEN rtranclp.rtrancl_into_rtrancl])
lemma rtranclp_mono: "r \<le> s ==> r^** \<le> s^**"
- -- {* monotonicity of @{text rtrancl} *}
+ -- \<open>monotonicity of @{text rtrancl}\<close>
apply (rule predicate2I)
apply (erule rtranclp.induct)
apply (rule_tac [2] rtranclp.rtrancl_into_rtrancl, blast+)
@@ -129,7 +129,7 @@
lemma refl_rtrancl: "refl (r^*)"
by (unfold refl_on_def) fast
-text {* Transitivity of transitive closure. *}
+text \<open>Transitivity of transitive closure.\<close>
lemma trans_rtrancl: "trans (r^*)"
proof (rule transI)
fix x y z
@@ -141,7 +141,7 @@
show "(x, y) \<in> r\<^sup>*" by fact
next
case (step u v)
- from `(x, u) \<in> r\<^sup>*` and `(u, v) \<in> r`
+ from \<open>(x, u) \<in> r\<^sup>*\<close> and \<open>(u, v) \<in> r\<close>
show "(x, v) \<in> r\<^sup>*" ..
qed
qed
@@ -159,7 +159,7 @@
obtains
(base) "a = b"
| (step) y where "(a, y) : r^*" and "(y, b) : r"
- -- {* elimination of @{text rtrancl} -- by induction on a special formula *}
+ -- \<open>elimination of @{text rtrancl} -- by induction on a special formula\<close>
apply (subgoal_tac "(a::'a) = b | (EX y. (a,y) : r^* & (y,b) : r)")
apply (rule_tac [2] major [THEN rtrancl_induct])
prefer 2 apply blast
@@ -179,9 +179,9 @@
lemmas converse_rtrancl_into_rtrancl = converse_rtranclp_into_rtranclp [to_set]
-text {*
+text \<open>
\medskip More @{term "r^*"} equations and inclusions.
-*}
+\<close>
lemma rtranclp_idemp [simp]: "(r^**)^** = r^**"
apply (auto intro!: order_antisym)
@@ -336,7 +336,7 @@
qed
-subsection {* Transitive closure *}
+subsection \<open>Transitive closure\<close>
lemma trancl_mono: "!!p. p \<in> r^+ ==> r \<subseteq> s ==> p \<in> s^+"
apply (simp add: split_tupled_all)
@@ -347,9 +347,9 @@
lemma r_into_trancl': "!!p. p : r ==> p : r^+"
by (simp only: split_tupled_all) (erule r_into_trancl)
-text {*
+text \<open>
\medskip Conversions between @{text trancl} and @{text rtrancl}.
-*}
+\<close>
lemma tranclp_into_rtranclp: "r^++ a b ==> r^** a b"
by (erule tranclp.induct) iprover+
@@ -363,7 +363,7 @@
lemmas rtrancl_into_trancl1 = rtranclp_into_tranclp1 [to_set]
lemma rtranclp_into_tranclp2: "[| r a b; r^** b c |] ==> r^++ a c"
- -- {* intro rule from @{text r} and @{text rtrancl} *}
+ -- \<open>intro rule from @{text r} and @{text rtrancl}\<close>
apply (erule rtranclp.cases)
apply iprover
apply (rule rtranclp_trans [THEN rtranclp_into_tranclp1])
@@ -372,7 +372,7 @@
lemmas rtrancl_into_trancl2 = rtranclp_into_tranclp2 [to_set]
-text {* Nice induction rule for @{text trancl} *}
+text \<open>Nice induction rule for @{text trancl}\<close>
lemma tranclp_induct [consumes 1, case_names base step, induct pred: tranclp]:
assumes a: "r^++ a b"
and cases: "!!y. r a y ==> P y"
@@ -395,7 +395,7 @@
and cases: "!!x y. r x y ==> P x y"
"!!x y z. [| r^++ x y; P x y; r^++ y z; P y z |] ==> P x z"
shows "P x y"
- -- {* Another induction rule for trancl, incorporating transitivity *}
+ -- \<open>Another induction rule for trancl, incorporating transitivity\<close>
by (iprover intro: major [THEN tranclp_induct] cases)
lemmas trancl_trans_induct = tranclp_trans_induct [to_set]
@@ -418,7 +418,7 @@
lemma trancl_unfold: "r^+ = r Un r^+ O r"
by (auto intro: trancl_into_trancl elim: tranclE)
-text {* Transitivity of @{term "r^+"} *}
+text \<open>Transitivity of @{term "r^+"}\<close>
lemma trans_trancl [simp]: "trans (r^+)"
proof (rule transI)
fix x y z
@@ -427,11 +427,11 @@
then show "(x, z) \<in> r^+"
proof induct
case (base u)
- from `(x, y) \<in> r^+` and `(y, u) \<in> r`
+ from \<open>(x, y) \<in> r^+\<close> and \<open>(y, u) \<in> r\<close>
show "(x, u) \<in> r^+" ..
next
case (step u v)
- from `(x, u) \<in> r^+` and `(u, v) \<in> r`
+ from \<open>(x, u) \<in> r^+\<close> and \<open>(u, v) \<in> r\<close>
show "(x, v) \<in> r^+" ..
qed
qed
@@ -521,12 +521,12 @@
from this(2) show P
proof (cases rule: rtranclp.cases)
case rtrancl_refl
- with `r x y` base show P by iprover
+ with \<open>r x y\<close> base show P by iprover
next
case rtrancl_into_rtrancl
from this have "tranclp r y z"
by (iprover intro: rtranclp_into_tranclp1)
- with `r x y` step show P by iprover
+ with \<open>r x y\<close> step show P by iprover
qed
qed
@@ -596,7 +596,7 @@
lemma trancl_insert:
"(insert (y, x) r)^+ = r^+ \<union> {(a, b). (a, y) \<in> r^* \<and> (x, b) \<in> r^*}"
- -- {* primitive recursion for @{text trancl} over finite relations *}
+ -- \<open>primitive recursion for @{text trancl} over finite relations\<close>
apply (rule equalityI)
apply (rule subsetI)
apply (simp only: split_tupled_all)
@@ -617,7 +617,7 @@
by(simp add: rtrancl_trancl_reflcl del: reflcl_trancl) blast
-text {* Simplifying nested closures *}
+text \<open>Simplifying nested closures\<close>
lemma rtrancl_trancl_absorb[simp]: "(R^*)^+ = R^*"
by (simp add: trans_rtrancl)
@@ -629,7 +629,7 @@
by auto
-text {* @{text Domain} and @{text Range} *}
+text \<open>@{text Domain} and @{text Range}\<close>
lemma Domain_rtrancl [simp]: "Domain (R^*) = UNIV"
by blast
@@ -673,8 +673,8 @@
apply (auto simp add: finite_Field)
done
-text {* More about converse @{text rtrancl} and @{text trancl}, should
- be merged with main body. *}
+text \<open>More about converse @{text rtrancl} and @{text trancl}, should
+ be merged with main body.\<close>
lemma single_valued_confluent:
"\<lbrakk> single_valued r; (x,y) \<in> r^*; (x,z) \<in> r^* \<rbrakk>
@@ -720,9 +720,9 @@
declare trancl_into_rtrancl [elim]
-subsection {* The power operation on relations *}
+subsection \<open>The power operation on relations\<close>
-text {* @{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R} *}
+text \<open>@{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R}\<close>
overloading
relpow == "compow :: nat \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set"
@@ -744,7 +744,7 @@
shows "(\<lambda>x y. (x, y) \<in> R) ^^ n = (\<lambda>x y. (x, y) \<in> R ^^ n)"
by (induct n) (simp_all add: relcompp_relcomp_eq)
-text {* for code generation *}
+text \<open>for code generation\<close>
definition relpow :: "nat \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" where
relpow_code_def [code_abbrev]: "relpow = compow"
@@ -973,7 +973,7 @@
"(P^**) x y \<Longrightarrow> \<exists>n. (P ^^ n) x y"
by (auto dest: rtranclp_imp_Sup_relpowp)
-text{* By Sternagel/Thiemann: *}
+text\<open>By Sternagel/Thiemann:\<close>
lemma relpow_fun_conv:
"((a,b) \<in> R ^^ n) = (\<exists>f. f 0 = a \<and> f n = b \<and> (\<forall>i<n. (f i, f(Suc i)) \<in> R))"
proof (induct n arbitrary: b)
@@ -1011,19 +1011,19 @@
proof(induct k arbitrary: b)
case 0
hence "R \<noteq> {}" by auto
- with card_0_eq[OF `finite R`] have "card R >= Suc 0" by auto
+ with card_0_eq[OF \<open>finite R\<close>] have "card R >= Suc 0" by auto
thus ?case using 0 by force
next
case (Suc k)
then obtain a' where "(a,a') : R^^(Suc k)" and "(a',b) : R" by auto
- from Suc(1)[OF `(a,a') : R^^(Suc k)`]
+ from Suc(1)[OF \<open>(a,a') : R^^(Suc k)\<close>]
obtain n where "n \<le> card R" and "(a,a') \<in> R ^^ n" by auto
- have "(a,b) : R^^(Suc n)" using `(a,a') \<in> R^^n` and `(a',b)\<in> R` by auto
+ have "(a,b) : R^^(Suc n)" using \<open>(a,a') \<in> R^^n\<close> and \<open>(a',b)\<in> R\<close> by auto
{ assume "n < card R"
- hence ?case using `(a,b): R^^(Suc n)` Suc_leI[OF `n < card R`] by blast
+ hence ?case using \<open>(a,b): R^^(Suc n)\<close> Suc_leI[OF \<open>n < card R\<close>] by blast
} moreover
{ assume "n = card R"
- from `(a,b) \<in> R ^^ (Suc n)`[unfolded relpow_fun_conv]
+ from \<open>(a,b) \<in> R ^^ (Suc n)\<close>[unfolded relpow_fun_conv]
obtain f where "f 0 = a" and "f(Suc n) = b"
and steps: "\<And>i. i <= n \<Longrightarrow> (f i, f (Suc i)) \<in> R" by auto
let ?p = "%i. (f i, f(Suc i))"
@@ -1031,7 +1031,7 @@
have "?p ` ?N <= R" using steps by auto
from card_mono[OF assms(1) this]
have "card(?p ` ?N) <= card R" .
- also have "\<dots> < card ?N" using `n = card R` by simp
+ also have "\<dots> < card ?N" using \<open>n = card R\<close> by simp
finally have "~ inj_on ?p ?N" by(rule pigeonhole)
then obtain i j where i: "i <= n" and j: "j <= n" and ij: "i \<noteq> j" and
pij: "?p i = ?p j" by(auto simp: inj_on_def)
@@ -1045,7 +1045,7 @@
let ?n = "Suc(n - (j - i))"
have abl: "(a,b) \<in> R ^^ ?n" unfolding relpow_fun_conv
proof (rule exI[of _ ?g], intro conjI impI allI)
- show "?g ?n = b" using `f(Suc n) = b` j ij by auto
+ show "?g ?n = b" using \<open>f(Suc n) = b\<close> j ij by auto
next
fix k assume "k < ?n"
show "(?g k, ?g (Suc k)) \<in> R"
@@ -1062,19 +1062,19 @@
thus ?thesis using ij pij steps[OF i] by simp
next
case False
- with `i \<le> k` have "i < k" by auto
- hence small: "k + (j - i) <= n" using `k<?n` by arith
- show ?thesis using steps[OF small] `i<k` by auto
+ with \<open>i \<le> k\<close> have "i < k" by auto
+ hence small: "k + (j - i) <= n" using \<open>k<?n\<close> by arith
+ show ?thesis using steps[OF small] \<open>i<k\<close> by auto
qed
qed
- qed (simp add: `f 0 = a`)
+ qed (simp add: \<open>f 0 = a\<close>)
moreover have "?n <= n" using i j ij by arith
- ultimately have ?case using `n = card R` by blast
+ ultimately have ?case using \<open>n = card R\<close> by blast
}
- ultimately show ?case using `n \<le> card R` by force
+ ultimately show ?case using \<open>n \<le> card R\<close> by force
qed
}
- thus ?thesis using gr0_implies_Suc[OF `k>0`] by auto
+ thus ?thesis using gr0_implies_Suc[OF \<open>k>0\<close>] by auto
qed
lemma relpow_finite_bounded:
@@ -1121,7 +1121,7 @@
done
-subsection {* Bounded transitive closure *}
+subsection \<open>Bounded transitive closure\<close>
definition ntrancl :: "nat \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set"
where
@@ -1151,15 +1151,15 @@
have "(a, b) \<in> ntrancl n R O (Id \<union> R)"
proof (cases "i = 1")
case True
- from this `(a, b) \<in> R ^^ i` show ?thesis
+ from this \<open>(a, b) \<in> R ^^ i\<close> show ?thesis
unfolding ntrancl_def by auto
next
case False
- from this `0 < i` obtain j where j: "i = Suc j" "0 < j"
+ from this \<open>0 < i\<close> obtain j where j: "i = Suc j" "0 < j"
by (cases i) auto
- from this `(a, b) \<in> R ^^ i` obtain c where c1: "(a, c) \<in> R ^^ j" and c2:"(c, b) \<in> R"
+ from this \<open>(a, b) \<in> R ^^ i\<close> obtain c where c1: "(a, c) \<in> R ^^ j" and c2:"(c, b) \<in> R"
by auto
- from c1 j `i \<le> Suc (Suc n)` have "(a, c) \<in> ntrancl n R"
+ from c1 j \<open>i \<le> Suc (Suc n)\<close> have "(a, c) \<in> ntrancl n R"
unfolding ntrancl_def by fastforce
from this c2 show ?thesis by fastforce
qed
@@ -1180,7 +1180,7 @@
by (cases "card R") (auto simp add: trancl_finite_eq_relpow relpow_empty ntrancl_def)
-subsection {* Acyclic relations *}
+subsection \<open>Acyclic relations\<close>
definition acyclic :: "('a * 'a) set => bool" where
"acyclic r \<longleftrightarrow> (!x. (x,x) ~: r^+)"
@@ -1235,9 +1235,9 @@
done
-subsection {* Setup of transitivity reasoner *}
+subsection \<open>Setup of transitivity reasoner\<close>
-ML {*
+ML \<open>
structure Trancl_Tac = Trancl_Tac
(
@@ -1284,30 +1284,30 @@
in dec t end
| decomp _ = NONE;
);
-*}
+\<close>
-setup {*
+setup \<open>
map_theory_simpset (fn ctxt => ctxt
addSolver (mk_solver "Trancl" Trancl_Tac.trancl_tac)
addSolver (mk_solver "Rtrancl" Trancl_Tac.rtrancl_tac)
addSolver (mk_solver "Tranclp" Tranclp_Tac.trancl_tac)
addSolver (mk_solver "Rtranclp" Tranclp_Tac.rtrancl_tac))
-*}
+\<close>
-text {* Optional methods. *}
+text \<open>Optional methods.\<close>
method_setup trancl =
- {* Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.trancl_tac) *}
- {* simple transitivity reasoner *}
+ \<open>Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.trancl_tac)\<close>
+ \<open>simple transitivity reasoner\<close>
method_setup rtrancl =
- {* Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.rtrancl_tac) *}
- {* simple transitivity reasoner *}
+ \<open>Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.rtrancl_tac)\<close>
+ \<open>simple transitivity reasoner\<close>
method_setup tranclp =
- {* Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.trancl_tac) *}
- {* simple transitivity reasoner (predicate version) *}
+ \<open>Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.trancl_tac)\<close>
+ \<open>simple transitivity reasoner (predicate version)\<close>
method_setup rtranclp =
- {* Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.rtrancl_tac) *}
- {* simple transitivity reasoner (predicate version) *}
+ \<open>Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.rtrancl_tac)\<close>
+ \<open>simple transitivity reasoner (predicate version)\<close>
end
--- a/src/HOL/Typedef.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Typedef.thy Sat Jul 18 22:58:50 2015 +0200
@@ -2,7 +2,7 @@
Author: Markus Wenzel, TU Munich
*)
-section {* HOL type definitions *}
+section \<open>HOL type definitions\<close>
theory Typedef
imports Set
@@ -14,7 +14,7 @@
assumes Rep: "Rep x \<in> A"
and Rep_inverse: "Abs (Rep x) = x"
and Abs_inverse: "y \<in> A ==> Rep (Abs y) = y"
- -- {* This will be axiomatized for each typedef! *}
+ -- \<open>This will be axiomatized for each typedef!\<close>
begin
lemma Rep_inject:
--- a/src/HOL/Typerep.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Typerep.thy Sat Jul 18 22:58:50 2015 +0200
@@ -1,6 +1,6 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Reflecting Pure types into HOL *}
+section \<open>Reflecting Pure types into HOL\<close>
theory Typerep
imports String
@@ -20,7 +20,7 @@
syntax
"_TYPEREP" :: "type => logic" ("(1TYPEREP/(1'(_')))")
-parse_translation {*
+parse_translation \<open>
let
fun typerep_tr (*"_TYPEREP"*) [ty] =
Syntax.const @{const_syntax typerep} $
@@ -28,9 +28,9 @@
(Syntax.const @{type_syntax itself} $ ty))
| typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts);
in [(@{syntax_const "_TYPEREP"}, K typerep_tr)] end
-*}
+\<close>
-typed_print_translation {*
+typed_print_translation \<open>
let
fun typerep_tr' ctxt (*"typerep"*)
(Type (@{type_name fun}, [Type (@{type_name itself}, [T]), _]))
@@ -39,9 +39,9 @@
(Syntax.const @{syntax_const "_TYPEREP"} $ Syntax_Phases.term_of_typ ctxt T, ts)
| typerep_tr' _ T ts = raise Match;
in [(@{const_syntax typerep}, typerep_tr')] end
-*}
+\<close>
-setup {*
+setup \<open>
let
fun add_typerep tyco thy =
@@ -76,7 +76,7 @@
#> Code.abstype_interpretation (ensure_typerep o fst)
end
-*}
+\<close>
lemma [code]:
"HOL.equal (Typerep tyco1 tys1) (Typerep tyco2 tys2) \<longleftrightarrow> HOL.equal tyco1 tyco2
--- a/src/HOL/Wellfounded.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Wellfounded.thy Sat Jul 18 22:58:50 2015 +0200
@@ -6,13 +6,13 @@
Author: Andrei Popescu, TU Muenchen
*)
-section {*Well-founded Recursion*}
+section \<open>Well-founded Recursion\<close>
theory Wellfounded
imports Transitive_Closure
begin
-subsection {* Basic Definitions *}
+subsection \<open>Basic Definitions\<close>
definition wf :: "('a * 'a) set => bool" where
"wf r \<longleftrightarrow> (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))"
@@ -29,8 +29,8 @@
lemmas wfPUNIVI = wfUNIVI [to_pred]
-text{*Restriction to domain @{term A} and range @{term B}. If @{term r} is
- well-founded over their intersection, then @{term "wf r"}*}
+text\<open>Restriction to domain @{term A} and range @{term B}. If @{term r} is
+ well-founded over their intersection, then @{term "wf r"}\<close>
lemma wfI:
"[| r \<subseteq> A <*> B;
!!x P. [|\<forall>x. (\<forall>y. (y,x) : r --> P y) --> P x; x : A; x : B |] ==> P x |]
@@ -75,9 +75,9 @@
unfolding wf_def by (blast intro: less_induct)
-subsection {* Basic Results *}
+subsection \<open>Basic Results\<close>
-text {* Point-free characterization of well-foundedness *}
+text \<open>Point-free characterization of well-foundedness\<close>
lemma wfE_pf:
assumes wf: "wf R"
@@ -105,7 +105,7 @@
with a show "P x" by blast
qed
-text{*Minimal-element characterization of well-foundedness*}
+text\<open>Minimal-element characterization of well-foundedness\<close>
lemma wfE_min:
assumes wf: "wf R" and Q: "x \<in> Q"
@@ -131,7 +131,7 @@
lemmas wfP_eq_minimal = wf_eq_minimal [to_pred]
-text{* Well-foundedness of transitive closure *}
+text\<open>Well-foundedness of transitive closure\<close>
lemma wf_trancl:
assumes "wf r"
@@ -143,17 +143,17 @@
have "P x"
proof (rule induct_step)
fix y assume "(y, x) : r^+"
- with `wf r` show "P y"
+ with \<open>wf r\<close> show "P y"
proof (induct x arbitrary: y)
case (less x)
- note hyp = `\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'`
- from `(y, x) : r^+` show "P y"
+ note hyp = \<open>\<And>x' y'. (x', x) : r ==> (y', x') : r^+ ==> P y'\<close>
+ from \<open>(y, x) : r^+\<close> show "P y"
proof cases
case base
show "P y"
proof (rule induct_step)
fix y' assume "(y', y) : r^+"
- with `(y, x) : r` show "P y'" by (rule hyp [of y y'])
+ with \<open>(y, x) : r\<close> show "P y'" by (rule hyp [of y y'])
qed
next
case step
@@ -172,7 +172,7 @@
apply (erule wf_trancl)
done
-text {* Well-foundedness of subsets *}
+text \<open>Well-foundedness of subsets\<close>
lemma wf_subset: "[| wf(r); p<=r |] ==> wf(p)"
apply (simp (no_asm_use) add: wf_eq_minimal)
@@ -181,7 +181,7 @@
lemmas wfP_subset = wf_subset [to_pred]
-text {* Well-foundedness of the empty relation *}
+text \<open>Well-foundedness of the empty relation\<close>
lemma wf_empty [iff]: "wf {}"
by (simp add: wf_def)
@@ -203,7 +203,7 @@
apply (rule Int_lower2)
done
-text {* Exponentiation *}
+text \<open>Exponentiation\<close>
lemma wf_exp:
assumes "wf (R ^^ n)"
@@ -211,11 +211,11 @@
proof (rule wfI_pf)
fix A assume "A \<subseteq> R `` A"
then have "A \<subseteq> (R ^^ n) `` A" by (induct n) force+
- with `wf (R ^^ n)`
+ with \<open>wf (R ^^ n)\<close>
show "A = {}" by (rule wfE_pf)
qed
-text {* Well-foundedness of insert *}
+text \<open>Well-foundedness of insert\<close>
lemma wf_insert [iff]: "wf(insert (y,x) r) = (wf(r) & (x,y) ~: r^*)"
apply (rule iffI)
@@ -233,12 +233,12 @@
apply (rule_tac x = "{z. z:Q & (z,y) : r^*}" in allE)
apply assumption
apply (erule_tac V = "ALL Q. (EX x. x : Q) --> P Q" for P in thin_rl)
- --{*essential for speed*}
-txt{*Blast with new substOccur fails*}
+ --\<open>essential for speed\<close>
+txt\<open>Blast with new substOccur fails\<close>
apply (fast intro: converse_rtrancl_into_rtrancl)
done
-text{*Well-foundedness of image*}
+text\<open>Well-foundedness of image\<close>
lemma wf_map_prod_image: "[| wf r; inj f |] ==> wf (map_prod f f ` r)"
apply (simp only: wf_eq_minimal, clarify)
@@ -248,7 +248,7 @@
done
-subsection {* Well-Foundedness Results for Unions *}
+subsection \<open>Well-Foundedness Results for Unions\<close>
lemma wf_union_compatible:
assumes "wf R" "wf S"
@@ -259,8 +259,8 @@
let ?Q' = "{x \<in> Q. \<forall>y. (y, x) \<in> R \<longrightarrow> y \<notin> Q}"
assume "x \<in> Q"
obtain a where "a \<in> ?Q'"
- by (rule wfE_min [OF `wf R` `x \<in> Q`]) blast
- with `wf S`
+ by (rule wfE_min [OF \<open>wf R\<close> \<open>x \<in> Q\<close>]) blast
+ with \<open>wf S\<close>
obtain z where "z \<in> ?Q'" and zmin: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> ?Q'" by (erule wfE_min)
{
fix y assume "(y, z) \<in> S"
@@ -269,19 +269,19 @@
have "y \<notin> Q"
proof
assume "y \<in> Q"
- with `y \<notin> ?Q'`
+ with \<open>y \<notin> ?Q'\<close>
obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto
- from `(w, y) \<in> R` `(y, z) \<in> S` have "(w, z) \<in> R O S" by (rule relcompI)
- with `R O S \<subseteq> R` have "(w, z) \<in> R" ..
- with `z \<in> ?Q'` have "w \<notin> Q" by blast
- with `w \<in> Q` show False by contradiction
+ from \<open>(w, y) \<in> R\<close> \<open>(y, z) \<in> S\<close> have "(w, z) \<in> R O S" by (rule relcompI)
+ with \<open>R O S \<subseteq> R\<close> have "(w, z) \<in> R" ..
+ with \<open>z \<in> ?Q'\<close> have "w \<notin> Q" by blast
+ with \<open>w \<in> Q\<close> show False by contradiction
qed
}
- with `z \<in> ?Q'` show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<union> S \<longrightarrow> y \<notin> Q" by blast
+ with \<open>z \<in> ?Q'\<close> show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> R \<union> S \<longrightarrow> y \<notin> Q" by blast
qed
-text {* Well-foundedness of indexed union with disjoint domains and ranges *}
+text \<open>Well-foundedness of indexed union with disjoint domains and ranges\<close>
lemma wf_UN: "[| ALL i:I. wf(r i);
ALL i:I. ALL j:I. r i ~= r j --> Domain(r i) Int Range(r j) = {}
@@ -340,7 +340,7 @@
fix Q :: "'a set" and x
assume "x \<in> Q"
- with `wf ?B`
+ with \<open>wf ?B\<close>
obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q"
by (erule wfE_min)
then have A1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q"
@@ -351,7 +351,7 @@
show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> ?A \<longrightarrow> y \<notin> Q"
proof (cases "\<forall>y. (y, z) \<in> R \<longrightarrow> y \<notin> Q")
case True
- with `z \<in> Q` A3 show ?thesis by blast
+ with \<open>z \<in> Q\<close> A3 show ?thesis by blast
next
case False
then obtain z' where "z'\<in>Q" "(z', z) \<in> R" by blast
@@ -362,24 +362,24 @@
then show "y \<notin> Q"
proof
assume "(y, z') \<in> R"
- then have "(y, z) \<in> R O R" using `(z', z) \<in> R` ..
+ then have "(y, z) \<in> R O R" using \<open>(z', z) \<in> R\<close> ..
with A1 show "y \<notin> Q" .
next
assume "(y, z') \<in> S"
- then have "(y, z) \<in> S O R" using `(z', z) \<in> R` ..
+ then have "(y, z) \<in> S O R" using \<open>(z', z) \<in> R\<close> ..
with A2 show "y \<notin> Q" .
qed
qed
- with `z' \<in> Q` show ?thesis ..
+ with \<open>z' \<in> Q\<close> show ?thesis ..
qed
qed
qed
-lemma wf_comp_self: "wf R = wf (R O R)" -- {* special case *}
+lemma wf_comp_self: "wf R = wf (R O R)" -- \<open>special case\<close>
by (rule wf_union_merge [where S = "{}", simplified])
-subsection {* Well-Foundedness of Composition *}
+subsection \<open>Well-Foundedness of Composition\<close>
text \<open>Bachmair and Dershowitz 1986, Lemma 2. [Provided by Tjark Weber]\<close>
@@ -437,7 +437,7 @@
qed
-subsection {* Acyclic relations *}
+subsection \<open>Acyclic relations\<close>
lemma wf_acyclic: "wf r ==> acyclic r"
apply (simp add: acyclic_def)
@@ -446,7 +446,7 @@
lemmas wfP_acyclicP = wf_acyclic [to_pred]
-text{* Wellfoundedness of finite acyclic relations*}
+text\<open>Wellfoundedness of finite acyclic relations\<close>
lemma finite_acyclic_wf [rule_format]: "finite r ==> acyclic r --> wf r"
apply (erule finite_induct, blast)
@@ -463,7 +463,7 @@
by (blast intro: finite_acyclic_wf wf_acyclic)
-subsection {* @{typ nat} is well-founded *}
+subsection \<open>@{typ nat} is well-founded\<close>
lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)^++"
proof (rule ext, rule ext, rule iffI)
@@ -517,12 +517,12 @@
by (rule Wellfounded.wellorder_class.wf)
-subsection {* Accessible Part *}
+subsection \<open>Accessible Part\<close>
-text {*
+text \<open>
Inductive definition of the accessible part @{term "acc r"} of a
relation; see also @{cite "paulin-tlca"}.
-*}
+\<close>
inductive_set
acc :: "('a * 'a) set => 'a set"
@@ -545,7 +545,7 @@
by (simp add: acc_def)
-text {* Induction rules *}
+text \<open>Induction rules\<close>
theorem accp_induct:
assumes major: "accp r a"
@@ -613,7 +613,7 @@
done
-text {* Smaller relations have bigger accessible parts: *}
+text \<open>Smaller relations have bigger accessible parts:\<close>
lemma accp_subset:
assumes sub: "R1 \<le> R2"
@@ -630,8 +630,8 @@
qed
-text {* This is a generalized induction theorem that works on
- subsets of the accessible part. *}
+text \<open>This is a generalized induction theorem that works on
+ subsets of the accessible part.\<close>
lemma accp_subset_induct:
assumes subset: "D \<le> accp R"
@@ -640,9 +640,9 @@
and istep: "\<And>x. \<lbrakk>D x; (\<And>z. R z x \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x"
shows "P x"
proof -
- from subset and `D x`
+ from subset and \<open>D x\<close>
have "accp R x" ..
- then show "P x" using `D x`
+ then show "P x" using \<open>D x\<close>
proof (induct x)
fix x
assume "D x"
@@ -652,7 +652,7 @@
qed
-text {* Set versions of the above theorems *}
+text \<open>Set versions of the above theorems\<close>
lemmas acc_induct = accp_induct [to_set]
@@ -677,9 +677,9 @@
lemmas acc_subset_induct = accp_subset_induct [to_set]
-subsection {* Tools for building wellfounded relations *}
+subsection \<open>Tools for building wellfounded relations\<close>
-text {* Inverse Image *}
+text \<open>Inverse Image\<close>
lemma wf_inv_image [simp,intro!]: "wf(r) ==> wf(inv_image r (f::'a=>'b))"
apply (simp (no_asm_use) add: inv_image_def wf_eq_minimal)
@@ -691,7 +691,7 @@
apply blast
done
-text {* Measure functions into @{typ nat} *}
+text \<open>Measure functions into @{typ nat}\<close>
definition measure :: "('a => nat) => ('a * 'a)set"
where "measure = inv_image less_than"
@@ -713,7 +713,7 @@
done
-text{* Lexicographic combinations *}
+text\<open>Lexicographic combinations\<close>
definition lex_prod :: "('a \<times>'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set" (infixr "<*lex*>" 80) where
"ra <*lex*> rb = {((a, b), (a', b')). (a, a') \<in> ra \<or> a = a' \<and> (b, b') \<in> rb}"
@@ -731,13 +731,13 @@
"(((a,b),(a',b')): r <*lex*> s) = ((a,a'): r \<or> (a = a' \<and> (b, b') : s))"
by (auto simp:lex_prod_def)
-text{* @{term "op <*lex*>"} preserves transitivity *}
+text\<open>@{term "op <*lex*>"} preserves transitivity\<close>
lemma trans_lex_prod [intro!]:
"[| trans R1; trans R2 |] ==> trans (R1 <*lex*> R2)"
by (unfold trans_def lex_prod_def, blast)
-text {* lexicographic combinations with measure functions *}
+text \<open>lexicographic combinations with measure functions\<close>
definition
mlex_prod :: "('a \<Rightarrow> nat) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" (infixr "<*mlex*>" 80)
@@ -754,7 +754,7 @@
lemma mlex_leq: "f x \<le> f y \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> (x, y) \<in> f <*mlex*> R"
unfolding mlex_prod_def by auto
-text {* proper subset relation on finite sets *}
+text \<open>proper subset relation on finite sets\<close>
definition finite_psubset :: "('a set * 'a set) set"
where "finite_psubset = {(A,B). A < B & finite B}"
@@ -772,7 +772,7 @@
lemma in_finite_psubset[simp]: "(A, B) \<in> finite_psubset = (A < B & finite B)"
unfolding finite_psubset_def by auto
-text {* max- and min-extension of order to finite sets *}
+text \<open>max- and min-extension of order to finite sets\<close>
inductive_set max_ext :: "('a \<times> 'a) set \<Rightarrow> ('a set \<times> 'a set) set"
for R :: "('a \<times> 'a) set"
@@ -828,9 +828,9 @@
next
assume "M \<noteq> {}"
from asm1 finites have N2: "(?N2, M) \<in> max_ext r"
- by (rule_tac max_extI[OF _ _ `M \<noteq> {}`]) auto
+ by (rule_tac max_extI[OF _ _ \<open>M \<noteq> {}\<close>]) auto
- with `M \<in> ?W` show "?N2 \<in> ?W" by (rule acc_downward)
+ with \<open>M \<in> ?W\<close> show "?N2 \<in> ?W" by (rule acc_downward)
qed
with finites have "?N1 \<union> ?N2 \<in> ?W"
by (rule add_less) simp
@@ -869,22 +869,22 @@
with nonempty
obtain e x where "x \<in> Q" "e \<in> x" by force
then have eU: "e \<in> \<Union>Q" by auto
- with `wf r`
+ with \<open>wf r\<close>
obtain z where z: "z \<in> \<Union>Q" "\<And>y. (y, z) \<in> r \<Longrightarrow> y \<notin> \<Union>Q"
by (erule wfE_min)
from z obtain m where "m \<in> Q" "z \<in> m" by auto
- from `m \<in> Q`
+ from \<open>m \<in> Q\<close>
show ?thesis
proof (rule, intro bexI allI impI)
fix n
assume smaller: "(n, m) \<in> min_ext r"
- with `z \<in> m` obtain y where y: "y \<in> n" "(y, z) \<in> r" by (auto simp: min_ext_def)
+ with \<open>z \<in> m\<close> obtain y where y: "y \<in> n" "(y, z) \<in> r" by (auto simp: min_ext_def)
then show "n \<notin> Q" using z(2) by auto
qed
qed
qed
-text{* Bounded increase must terminate: *}
+text\<open>Bounded increase must terminate:\<close>
lemma wf_bounded_measure:
fixes ub :: "'a \<Rightarrow> nat" and f :: "'a \<Rightarrow> nat"
--- a/src/HOL/Wfrec.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Wfrec.thy Sat Jul 18 22:58:50 2015 +0200
@@ -4,7 +4,7 @@
Author: Konrad Slind
*)
-section {* Well-Founded Recursion Combinator *}
+section \<open>Well-Founded Recursion Combinator\<close>
theory Wfrec
imports Wellfounded
@@ -28,20 +28,20 @@
lemma cut_apply: "(x, a) \<in> R \<Longrightarrow> cut f R a x = f x"
by (simp add: cut_def)
-text{*Inductive characterization of wfrec combinator; for details see:
-John Harrison, "Inductive definitions: automation and application"*}
+text\<open>Inductive characterization of wfrec combinator; for details see:
+John Harrison, "Inductive definitions: automation and application"\<close>
lemma theI_unique: "\<exists>!x. P x \<Longrightarrow> P x \<longleftrightarrow> x = The P"
by (auto intro: the_equality[symmetric] theI)
lemma wfrec_unique: assumes "adm_wf R F" "wf R" shows "\<exists>!y. wfrec_rel R F x y"
- using `wf R`
+ using \<open>wf R\<close>
proof induct
def f \<equiv> "\<lambda>y. THE z. wfrec_rel R F y z"
case (less x)
then have "\<And>y z. (y, x) \<in> R \<Longrightarrow> wfrec_rel R F y z \<longleftrightarrow> z = f y"
unfolding f_def by (rule theI_unique)
- with `adm_wf R F` show ?case
+ with \<open>adm_wf R F\<close> show ?case
by (subst wfrec_rel.simps) (auto simp: adm_wf_def)
qed
@@ -57,12 +57,12 @@
done
-text{** This form avoids giant explosions in proofs. NOTE USE OF ==*}
+text\<open>* This form avoids giant explosions in proofs. NOTE USE OF ==\<close>
lemma def_wfrec: "f \<equiv> wfrec R F \<Longrightarrow> wf R \<Longrightarrow> f a = F (cut f R a) a"
by (auto intro: wfrec)
-subsubsection {* Well-founded recursion via genuine fixpoints *}
+subsubsection \<open>Well-founded recursion via genuine fixpoints\<close>
lemma wfrec_fixpoint:
assumes WF: "wf R" and ADM: "adm_wf R F"
@@ -79,12 +79,12 @@
finally show "wfrec R F x = F (wfrec R F) x" .
qed
-subsection {* Wellfoundedness of @{text same_fst} *}
+subsection \<open>Wellfoundedness of @{text same_fst}\<close>
definition same_fst :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> ('b \<times> 'b) set) \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set" where
"same_fst P R = {((x', y'), (x, y)) . x' = x \<and> P x \<and> (y',y) \<in> R x}"
- --{*For @{const wfrec} declarations where the first n parameters
- stay unchanged in the recursive call. *}
+ --\<open>For @{const wfrec} declarations where the first n parameters
+ stay unchanged in the recursive call.\<close>
lemma same_fstI [intro!]: "P x \<Longrightarrow> (y', y) \<in> R x \<Longrightarrow> ((x, y'), (x, y)) \<in> same_fst P R"
by (simp add: same_fst_def)
--- a/src/HOL/Zorn.thy Sat Jul 18 21:44:18 2015 +0200
+++ b/src/HOL/Zorn.thy Sat Jul 18 22:58:50 2015 +0200
@@ -7,17 +7,17 @@
The well-ordering theorem.
*)
-section {* Zorn's Lemma *}
+section \<open>Zorn's Lemma\<close>
theory Zorn
imports Order_Relation Hilbert_Choice
begin
-subsection {* Zorn's Lemma for the Subset Relation *}
+subsection \<open>Zorn's Lemma for the Subset Relation\<close>
-subsubsection {* Results that do not require an order *}
+subsubsection \<open>Results that do not require an order\<close>
-text {*Let @{text P} be a binary predicate on the set @{text A}.*}
+text \<open>Let @{text P} be a binary predicate on the set @{text A}.\<close>
locale pred_on =
fixes A :: "'a set"
and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubset>" 50)
@@ -26,21 +26,21 @@
abbreviation Peq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) where
"x \<sqsubseteq> y \<equiv> P\<^sup>=\<^sup>= x y"
-text {*A chain is a totally ordered subset of @{term A}.*}
+text \<open>A chain is a totally ordered subset of @{term A}.\<close>
definition chain :: "'a set \<Rightarrow> bool" where
"chain C \<longleftrightarrow> C \<subseteq> A \<and> (\<forall>x\<in>C. \<forall>y\<in>C. x \<sqsubseteq> y \<or> y \<sqsubseteq> x)"
-text {*We call a chain that is a proper superset of some set @{term X},
-but not necessarily a chain itself, a superchain of @{term X}.*}
+text \<open>We call a chain that is a proper superset of some set @{term X},
+but not necessarily a chain itself, a superchain of @{term X}.\<close>
abbreviation superchain :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "<c" 50) where
"X <c C \<equiv> chain C \<and> X \<subset> C"
-text {*A maximal chain is a chain that does not have a superchain.*}
+text \<open>A maximal chain is a chain that does not have a superchain.\<close>
definition maxchain :: "'a set \<Rightarrow> bool" where
"maxchain C \<longleftrightarrow> chain C \<and> \<not> (\<exists>S. C <c S)"
-text {*We define the successor of a set to be an arbitrary
-superchain, if such exists, or the set itself, otherwise.*}
+text \<open>We define the successor of a set to be an arbitrary
+superchain, if such exists, or the set itself, otherwise.\<close>
definition suc :: "'a set \<Rightarrow> 'a set" where
"suc C = (if \<not> chain C \<or> maxchain C then C else (SOME D. C <c D))"
@@ -76,21 +76,21 @@
assumes "X \<subseteq> Y" shows "X \<subseteq> suc Y"
using assms by (rule subset_trans) (rule suc_subset)
-text {*We build a set @{term \<C>} that is closed under applications
-of @{term suc} and contains the union of all its subsets.*}
+text \<open>We build a set @{term \<C>} that is closed under applications
+of @{term suc} and contains the union of all its subsets.\<close>
inductive_set suc_Union_closed ("\<C>") where
suc: "X \<in> \<C> \<Longrightarrow> suc X \<in> \<C>" |
Union [unfolded Pow_iff]: "X \<in> Pow \<C> \<Longrightarrow> \<Union>X \<in> \<C>"
-text {*Since the empty set as well as the set itself is a subset of
+text \<open>Since the empty set as well as the set itself is a subset of
every set, @{term \<C>} contains at least @{term "{} \<in> \<C>"} and
-@{term "\<Union>\<C> \<in> \<C>"}.*}
+@{term "\<Union>\<C> \<in> \<C>"}.\<close>
lemma
suc_Union_closed_empty: "{} \<in> \<C>" and
suc_Union_closed_Union: "\<Union>\<C> \<in> \<C>"
using Union [of "{}"] and Union [of "\<C>"] by simp+
-text {*Thus closure under @{term suc} will hit a maximal chain
-eventually, as is shown below.*}
+text \<open>Thus closure under @{term suc} will hit a maximal chain
+eventually, as is shown below.\<close>
lemma suc_Union_closed_induct [consumes 1, case_names suc Union,
induct pred: suc_Union_closed]:
@@ -108,7 +108,7 @@
shows "Q"
using assms by (cases) simp+
-text {*On chains, @{term suc} yields a chain.*}
+text \<open>On chains, @{term suc} yields a chain.\<close>
lemma chain_suc:
assumes "chain X" shows "chain (suc X)"
using assms
@@ -118,7 +118,7 @@
lemma chain_sucD:
assumes "chain X" shows "suc X \<subseteq> A \<and> chain (suc X)"
proof -
- from `chain X` have *: "chain (suc X)" by (rule chain_suc)
+ from \<open>chain X\<close> have *: "chain (suc X)" by (rule chain_suc)
then have "suc X \<subseteq> A" unfolding chain_def by blast
with * show ?thesis by blast
qed
@@ -127,7 +127,7 @@
assumes "X \<in> \<C>" and "Y \<in> \<C>"
and *: "\<And>Z. Z \<in> \<C> \<Longrightarrow> Z \<subseteq> Y \<Longrightarrow> Z = Y \<or> suc Z \<subseteq> Y"
shows "X \<subseteq> Y \<or> suc Y \<subseteq> X"
- using `X \<in> \<C>`
+ using \<open>X \<in> \<C>\<close>
proof (induct)
case (suc X)
with * show ?case by (blast del: subsetI intro: subset_suc)
@@ -139,13 +139,13 @@
using assms(2-, 1)
proof (induct arbitrary: Y)
case (suc X)
- note * = `\<And>Y. \<lbrakk>Y \<in> \<C>; Y \<subseteq> X\<rbrakk> \<Longrightarrow> X = Y \<or> suc Y \<subseteq> X`
- with suc_Union_closed_total' [OF `Y \<in> \<C>` `X \<in> \<C>`]
+ note * = \<open>\<And>Y. \<lbrakk>Y \<in> \<C>; Y \<subseteq> X\<rbrakk> \<Longrightarrow> X = Y \<or> suc Y \<subseteq> X\<close>
+ with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>X \<in> \<C>\<close>]
have "Y \<subseteq> X \<or> suc X \<subseteq> Y" by blast
then show ?case
proof
assume "Y \<subseteq> X"
- with * and `Y \<in> \<C>` have "X = Y \<or> suc Y \<subseteq> X" by blast
+ with * and \<open>Y \<in> \<C>\<close> have "X = Y \<or> suc Y \<subseteq> X" by blast
then show ?thesis
proof
assume "X = Y" then show ?thesis by simp
@@ -156,43 +156,43 @@
qed
next
assume "suc X \<subseteq> Y"
- with `Y \<subseteq> suc X` show ?thesis by blast
+ with \<open>Y \<subseteq> suc X\<close> show ?thesis by blast
qed
next
case (Union X)
show ?case
proof (rule ccontr)
assume "\<not> ?thesis"
- with `Y \<subseteq> \<Union>X` obtain x y z
+ with \<open>Y \<subseteq> \<Union>X\<close> obtain x y z
where "\<not> suc Y \<subseteq> \<Union>X"
and "x \<in> X" and "y \<in> x" and "y \<notin> Y"
and "z \<in> suc Y" and "\<forall>x\<in>X. z \<notin> x" by blast
- with `X \<subseteq> \<C>` have "x \<in> \<C>" by blast
- from Union and `x \<in> X`
+ with \<open>X \<subseteq> \<C>\<close> have "x \<in> \<C>" by blast
+ from Union and \<open>x \<in> X\<close>
have *: "\<And>y. \<lbrakk>y \<in> \<C>; y \<subseteq> x\<rbrakk> \<Longrightarrow> x = y \<or> suc y \<subseteq> x" by blast
- with suc_Union_closed_total' [OF `Y \<in> \<C>` `x \<in> \<C>`]
+ with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>x \<in> \<C>\<close>]
have "Y \<subseteq> x \<or> suc x \<subseteq> Y" by blast
then show False
proof
assume "Y \<subseteq> x"
- with * [OF `Y \<in> \<C>`] have "x = Y \<or> suc Y \<subseteq> x" by blast
+ with * [OF \<open>Y \<in> \<C>\<close>] have "x = Y \<or> suc Y \<subseteq> x" by blast
then show False
proof
- assume "x = Y" with `y \<in> x` and `y \<notin> Y` show False by blast
+ assume "x = Y" with \<open>y \<in> x\<close> and \<open>y \<notin> Y\<close> show False by blast
next
assume "suc Y \<subseteq> x"
- with `x \<in> X` have "suc Y \<subseteq> \<Union>X" by blast
- with `\<not> suc Y \<subseteq> \<Union>X` show False by contradiction
+ with \<open>x \<in> X\<close> have "suc Y \<subseteq> \<Union>X" by blast
+ with \<open>\<not> suc Y \<subseteq> \<Union>X\<close> show False by contradiction
qed
next
assume "suc x \<subseteq> Y"
- moreover from suc_subset and `y \<in> x` have "y \<in> suc x" by blast
- ultimately show False using `y \<notin> Y` by blast
+ moreover from suc_subset and \<open>y \<in> x\<close> have "y \<in> suc x" by blast
+ ultimately show False using \<open>y \<notin> Y\<close> by blast
qed
qed
qed
-text {*The elements of @{term \<C>} are totally ordered by the subset relation.*}
+text \<open>The elements of @{term \<C>} are totally ordered by the subset relation.\<close>
lemma suc_Union_closed_total:
assumes "X \<in> \<C>" and "Y \<in> \<C>"
shows "X \<subseteq> Y \<or> Y \<subseteq> X"
@@ -205,20 +205,20 @@
case False
then obtain Z
where "Z \<in> \<C>" and "Z \<subseteq> Y" and "Z \<noteq> Y" and "\<not> suc Z \<subseteq> Y" by blast
- with suc_Union_closed_subsetD and `Y \<in> \<C>` show ?thesis by blast
+ with suc_Union_closed_subsetD and \<open>Y \<in> \<C>\<close> show ?thesis by blast
qed
-text {*Once we hit a fixed point w.r.t. @{term suc}, all other elements
-of @{term \<C>} are subsets of this fixed point.*}
+text \<open>Once we hit a fixed point w.r.t. @{term suc}, all other elements
+of @{term \<C>} are subsets of this fixed point.\<close>
lemma suc_Union_closed_suc:
assumes "X \<in> \<C>" and "Y \<in> \<C>" and "suc Y = Y"
shows "X \<subseteq> Y"
-using `X \<in> \<C>`
+using \<open>X \<in> \<C>\<close>
proof (induct)
case (suc X)
- with `Y \<in> \<C>` and suc_Union_closed_subsetD
+ with \<open>Y \<in> \<C>\<close> and suc_Union_closed_subsetD
have "X = Y \<or> suc X \<subseteq> Y" by blast
- then show ?case by (auto simp: `suc Y = Y`)
+ then show ?case by (auto simp: \<open>suc Y = Y\<close>)
qed blast
lemma eq_suc_Union:
@@ -226,11 +226,11 @@
shows "suc X = X \<longleftrightarrow> X = \<Union>\<C>"
proof
assume "suc X = X"
- with suc_Union_closed_suc [OF suc_Union_closed_Union `X \<in> \<C>`]
+ with suc_Union_closed_suc [OF suc_Union_closed_Union \<open>X \<in> \<C>\<close>]
have "\<Union>\<C> \<subseteq> X" .
- with `X \<in> \<C>` show "X = \<Union>\<C>" by blast
+ with \<open>X \<in> \<C>\<close> show "X = \<Union>\<C>" by blast
next
- from `X \<in> \<C>` have "suc X \<in> \<C>" by (rule suc)
+ from \<open>X \<in> \<C>\<close> have "suc X \<in> \<C>" by (rule suc)
then have "suc X \<subseteq> \<Union>\<C>" by blast
moreover assume "X = \<Union>\<C>"
ultimately have "suc X \<subseteq> X" by simp
@@ -251,7 +251,7 @@
using assms
by (induct) (auto dest: suc_in_carrier)
-text {*All elements of @{term \<C>} are chains.*}
+text \<open>All elements of @{term \<C>} are chains.\<close>
lemma suc_Union_closed_chain:
assumes "X \<in> \<C>"
shows "chain X"
@@ -271,27 +271,27 @@
then show "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
proof
assume "u \<subseteq> v"
- from `chain v` show ?thesis
+ from \<open>chain v\<close> show ?thesis
proof (rule chain_total)
show "y \<in> v" by fact
- show "x \<in> v" using `u \<subseteq> v` and `x \<in> u` by blast
+ show "x \<in> v" using \<open>u \<subseteq> v\<close> and \<open>x \<in> u\<close> by blast
qed
next
assume "v \<subseteq> u"
- from `chain u` show ?thesis
+ from \<open>chain u\<close> show ?thesis
proof (rule chain_total)
show "x \<in> u" by fact
- show "y \<in> u" using `v \<subseteq> u` and `y \<in> v` by blast
+ show "y \<in> u" using \<open>v \<subseteq> u\<close> and \<open>y \<in> v\<close> by blast
qed
qed
qed
ultimately show ?case unfolding chain_def ..
qed
-subsubsection {* Hausdorff's Maximum Principle *}
+subsubsection \<open>Hausdorff's Maximum Principle\<close>
-text {*There exists a maximal totally ordered subset of @{term A}. (Note that we do not
-require @{term A} to be partially ordered.)*}
+text \<open>There exists a maximal totally ordered subset of @{term A}. (Note that we do not
+require @{term A} to be partially ordered.)\<close>
theorem Hausdorff: "\<exists>C. maxchain C"
proof -
@@ -309,7 +309,7 @@
then show ?thesis by blast
qed
-text {*Make notation @{term \<C>} available again.*}
+text \<open>Make notation @{term \<C>} available again.\<close>
no_notation suc_Union_closed ("\<C>")
lemma chain_extend:
@@ -322,8 +322,8 @@
end
-text {*Hide constant @{const pred_on.suc_Union_closed}, which was just needed
-for the proof of Hausforff's maximum principle.*}
+text \<open>Hide constant @{const pred_on.suc_Union_closed}, which was just needed
+for the proof of Hausforff's maximum principle.\<close>
hide_const pred_on.suc_Union_closed
lemma chain_mono:
@@ -332,7 +332,7 @@
shows "pred_on.chain A Q C"
using assms unfolding pred_on.chain_def by blast
-subsubsection {* Results for the proper subset relation *}
+subsubsection \<open>Results for the proper subset relation\<close>
interpretation subset: pred_on "A" "op \<subset>" for A .
@@ -341,20 +341,20 @@
shows "\<Union>C = X"
proof (rule ccontr)
let ?C = "{X} \<union> C"
- from `subset.maxchain A C` have "subset.chain A C"
+ from \<open>subset.maxchain A C\<close> have "subset.chain A C"
and *: "\<And>S. subset.chain A S \<Longrightarrow> \<not> C \<subset> S"
by (auto simp: subset.maxchain_def)
- moreover have "\<forall>x\<in>C. x \<subseteq> X" using `\<Union>C \<subseteq> X` by auto
+ moreover have "\<forall>x\<in>C. x \<subseteq> X" using \<open>\<Union>C \<subseteq> X\<close> by auto
ultimately have "subset.chain A ?C"
- using subset.chain_extend [of A C X] and `X \<in> A` by auto
+ using subset.chain_extend [of A C X] and \<open>X \<in> A\<close> by auto
moreover assume **: "\<Union>C \<noteq> X"
- moreover from ** have "C \<subset> ?C" using `\<Union>C \<subseteq> X` by auto
+ moreover from ** have "C \<subset> ?C" using \<open>\<Union>C \<subseteq> X\<close> by auto
ultimately show False using * by blast
qed
-subsubsection {* Zorn's lemma *}
+subsubsection \<open>Zorn's lemma\<close>
-text {*If every chain has an upper bound, then there is a maximal set.*}
+text \<open>If every chain has an upper bound, then there is a maximal set.\<close>
lemma subset_Zorn:
assumes "\<And>C. subset.chain A C \<Longrightarrow> \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U"
shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
@@ -369,18 +369,18 @@
show "Y = X"
proof (rule ccontr)
assume "Y \<noteq> X"
- with `Y \<subseteq> X` have "\<not> X \<subseteq> Y" by blast
- from subset.chain_extend [OF `subset.chain A M` `X \<in> A`] and `\<forall>X\<in>M. X \<subseteq> Y`
- have "subset.chain A ({X} \<union> M)" using `Y \<subseteq> X` by auto
- moreover have "M \<subset> {X} \<union> M" using `\<forall>X\<in>M. X \<subseteq> Y` and `\<not> X \<subseteq> Y` by auto
+ with \<open>Y \<subseteq> X\<close> have "\<not> X \<subseteq> Y" by blast
+ from subset.chain_extend [OF \<open>subset.chain A M\<close> \<open>X \<in> A\<close>] and \<open>\<forall>X\<in>M. X \<subseteq> Y\<close>
+ have "subset.chain A ({X} \<union> M)" using \<open>Y \<subseteq> X\<close> by auto
+ moreover have "M \<subset> {X} \<union> M" using \<open>\<forall>X\<in>M. X \<subseteq> Y\<close> and \<open>\<not> X \<subseteq> Y\<close> by auto
ultimately show False
- using `subset.maxchain A M` by (auto simp: subset.maxchain_def)
+ using \<open>subset.maxchain A M\<close> by (auto simp: subset.maxchain_def)
qed
qed
ultimately show ?thesis by blast
qed
-text{*Alternative version of Zorn's lemma for the subset relation.*}
+text\<open>Alternative version of Zorn's lemma for the subset relation.\<close>
lemma subset_Zorn':
assumes "\<And>C. subset.chain A C \<Longrightarrow> \<Union>C \<in> A"
shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
@@ -392,16 +392,16 @@
proof (intro ballI impI)
fix Z
assume "Z \<in> A" and "\<Union>M \<subseteq> Z"
- with subset_maxchain_max [OF `subset.maxchain A M`]
+ with subset_maxchain_max [OF \<open>subset.maxchain A M\<close>]
show "\<Union>M = Z" .
qed
ultimately show ?thesis by blast
qed
-subsection {* Zorn's Lemma for Partial Orders *}
+subsection \<open>Zorn's Lemma for Partial Orders\<close>
-text {*Relate old to new definitions.*}
+text \<open>Relate old to new definitions.\<close>
(* Define globally? In Set.thy? *)
definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^sub>\<subseteq>") where
@@ -450,7 +450,7 @@
"\<forall>C\<in>chains A. \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
using subset_Zorn [of A] by (auto simp: chains_alt_def)
-text{*Various other lemmas*}
+text\<open>Various other lemmas\<close>
lemma chainsD: "[| c \<in> chains S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
by (unfold chains_def chain_subset_def) blast
@@ -464,7 +464,7 @@
shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
proof -
have "Preorder r" using po by (simp add: partial_order_on_def)
---{* Mirror r in the set of subsets below (wrt r) elements of A*}
+--\<open>Mirror r in the set of subsets below (wrt r) elements of A\<close>
let ?B = "%x. r\<inverse> `` {x}" let ?S = "?B ` Field r"
{
fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
@@ -476,7 +476,7 @@
assume "a \<in> Field r" and "?B a \<in> C" and "b \<in> Field r" and "?B b \<in> C"
hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
thus "(a, b) \<in> r \<or> (b, a) \<in> r"
- using `Preorder r` and `a \<in> Field r` and `b \<in> Field r`
+ using \<open>Preorder r\<close> and \<open>a \<in> Field r\<close> and \<open>b \<in> Field r\<close>
by (simp add:subset_Image1_Image1_iff)
qed
then obtain u where uA: "u \<in> Field r" "\<forall>a\<in>?A. (a, u) \<in> r" using u by auto
@@ -484,10 +484,10 @@
proof auto
fix a B assume aB: "B \<in> C" "a \<in> B"
with 1 obtain x where "x \<in> Field r" and "B = r\<inverse> `` {x}" by auto
- thus "(a, u) \<in> r" using uA and aB and `Preorder r`
+ thus "(a, u) \<in> r" using uA and aB and \<open>Preorder r\<close>
unfolding preorder_on_def refl_on_def by simp (fast dest: transD)
qed
- then have "\<exists>u\<in>Field r. ?P u" using `u \<in> Field r` by blast
+ then have "\<exists>u\<in>Field r. ?P u" using \<open>u \<in> Field r\<close> by blast
}
then have "\<forall>C\<in>chains ?S. \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U"
by (auto simp: chains_def chain_subset_def)
@@ -496,13 +496,13 @@
and "\<forall>x\<in>Field r. B \<subseteq> r\<inverse> `` {x} \<longrightarrow> r\<inverse> `` {x} = B"
by auto
hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
- using po and `Preorder r` and `m \<in> Field r`
+ using po and \<open>Preorder r\<close> and \<open>m \<in> Field r\<close>
by (auto simp: subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
- thus ?thesis using `m \<in> Field r` by blast
+ thus ?thesis using \<open>m \<in> Field r\<close> by blast
qed
-subsection {* The Well Ordering Theorem *}
+subsection \<open>The Well Ordering Theorem\<close>
(* The initial segment of a relation appears generally useful.
Move to Relation.thy?
@@ -541,7 +541,7 @@
with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1" unfolding chain_subset_def by blast
moreover assume "(x, y) \<in> S1" "(y, z) \<in> S2"
ultimately have "((x, y) \<in> S1 \<and> (y, z) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, z) \<in> S2)" by blast
- with `S1 \<in> R` `S2 \<in> R` assms(2) show "(x, z) \<in> \<Union>R" by (auto elim: transE)
+ with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "(x, z) \<in> \<Union>R" by (auto elim: transE)
qed
lemma chain_subset_antisym_Union:
@@ -553,7 +553,7 @@
with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1" unfolding chain_subset_def by blast
moreover assume "(x, y) \<in> S1" "(y, x) \<in> S2"
ultimately have "((x, y) \<in> S1 \<and> (y, x) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, x) \<in> S2)" by blast
- with `S1 \<in> R` `S2 \<in> R` assms(2) show "x = y" unfolding antisym_def by auto
+ with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "x = y" unfolding antisym_def by auto
qed
lemma chain_subset_Total_Union:
@@ -561,17 +561,17 @@
shows "Total (\<Union>R)"
proof (simp add: total_on_def Ball_def, auto del: disjCI)
fix r s a b assume A: "r \<in> R" "s \<in> R" "a \<in> Field r" "b \<in> Field s" "a \<noteq> b"
- from `chain\<^sub>\<subseteq> R` and `r \<in> R` and `s \<in> R` have "r \<subseteq> s \<or> s \<subseteq> r"
+ from \<open>chain\<^sub>\<subseteq> R\<close> and \<open>r \<in> R\<close> and \<open>s \<in> R\<close> have "r \<subseteq> s \<or> s \<subseteq> r"
by (auto simp add: chain_subset_def)
thus "(\<exists>r\<in>R. (a, b) \<in> r) \<or> (\<exists>r\<in>R. (b, a) \<in> r)"
proof
assume "r \<subseteq> s" hence "(a, b) \<in> s \<or> (b, a) \<in> s" using assms(2) A mono_Field[of r s]
by (auto simp add: total_on_def)
- thus ?thesis using `s \<in> R` by blast
+ thus ?thesis using \<open>s \<in> R\<close> by blast
next
assume "s \<subseteq> r" hence "(a, b) \<in> r \<or> (b, a) \<in> r" using assms(2) A mono_Field[of s r]
by (fastforce simp add: total_on_def)
- thus ?thesis using `r \<in> R` by blast
+ thus ?thesis using \<open>r \<in> R\<close> by blast
qed
qed
@@ -589,11 +589,11 @@
then obtain s where s: "s \<in> R" "(f (Suc (Suc i)), f(Suc i)) \<in> s"
using 1 by auto
then have "s initial_segment_of r \<or> r initial_segment_of s"
- using assms(1) `r \<in> R` by (simp add: Chains_def)
+ using assms(1) \<open>r \<in> R\<close> by (simp add: Chains_def)
with Suc s show ?case by (simp add: init_seg_of_def) blast
qed
}
- thus False using assms(2) and `r \<in> R`
+ thus False using assms(2) and \<open>r \<in> R\<close>
by (simp add: wf_iff_no_infinite_down_chain) blast
qed
@@ -607,7 +607,7 @@
theorem well_ordering: "\<exists>r::'a rel. Well_order r \<and> Field r = UNIV"
proof -
--- {*The initial segment relation on well-orders: *}
+-- \<open>The initial segment relation on well-orders:\<close>
let ?WO = "{r::'a rel. Well_order r}"
def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
have I_init: "I \<subseteq> init_seg_of" by (auto simp: I_def)
@@ -619,42 +619,42 @@
hence 0: "Partial_order I"
by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
trans_def I_def elim!: trans_init_seg_of)
--- {*I-chains have upper bounds in ?WO wrt I: their Union*}
+-- \<open>I-chains have upper bounds in ?WO wrt I: their Union\<close>
{ fix R assume "R \<in> Chains I"
hence Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
- have subch: "chain\<^sub>\<subseteq> R" using `R : Chains I` I_init
+ have subch: "chain\<^sub>\<subseteq> R" using \<open>R : Chains I\<close> I_init
by (auto simp: init_seg_of_def chain_subset_def Chains_def)
have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r"
and "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)"
- using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
- have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
+ using Chains_wo [OF \<open>R \<in> Chains I\<close>] by (simp_all add: order_on_defs)
+ have "Refl (\<Union>R)" using \<open>\<forall>r\<in>R. Refl r\<close> unfolding refl_on_def by fastforce
moreover have "trans (\<Union>R)"
- by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
+ by (rule chain_subset_trans_Union [OF subch \<open>\<forall>r\<in>R. trans r\<close>])
moreover have "antisym (\<Union>R)"
- by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
+ by (rule chain_subset_antisym_Union [OF subch \<open>\<forall>r\<in>R. antisym r\<close>])
moreover have "Total (\<Union>R)"
- by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
+ by (rule chain_subset_Total_Union [OF subch \<open>\<forall>r\<in>R. Total r\<close>])
moreover have "wf ((\<Union>R) - Id)"
proof -
have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
- with `\<forall>r\<in>R. wf (r - Id)` and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
+ with \<open>\<forall>r\<in>R. wf (r - Id)\<close> and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
show ?thesis by fastforce
qed
ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs)
moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
by(simp add: Chains_init_seg_of_Union)
ultimately have "\<Union>R \<in> ?WO \<and> (\<forall>r\<in>R. (r, \<Union>R) \<in> I)"
- using mono_Chains [OF I_init] Chains_wo[of R] and `R \<in> Chains I`
+ using mono_Chains [OF I_init] Chains_wo[of R] and \<open>R \<in> Chains I\<close>
unfolding I_def by blast
}
hence 1: "\<forall>R \<in> Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
---{*Zorn's Lemma yields a maximal well-order m:*}
+--\<open>Zorn's Lemma yields a maximal well-order m:\<close>
then obtain m::"'a rel" where "Well_order m" and
max: "\<forall>r. Well_order r \<and> (m, r) \<in> I \<longrightarrow> r = m"
using Zorns_po_lemma[OF 0 1] unfolding FI by fastforce
---{*Now show by contradiction that m covers the whole type:*}
+--\<open>Now show by contradiction that m covers the whole type:\<close>
{ fix x::'a assume "x \<notin> Field m"
---{*We assume that x is not covered and extend m at the top with x*}
+--\<open>We assume that x is not covered and extend m at the top with x\<close>
have "m \<noteq> {}"
proof
assume "m = {}"
@@ -664,40 +664,40 @@
by (auto simp: I_def init_seg_of_def simp del: Field_insert)
qed
hence "Field m \<noteq> {}" by(auto simp:Field_def)
- moreover have "wf (m - Id)" using `Well_order m`
+ moreover have "wf (m - Id)" using \<open>Well_order m\<close>
by (simp add: well_order_on_def)
---{*The extension of m by x:*}
+--\<open>The extension of m by x:\<close>
let ?s = "{(a, x) | a. a \<in> Field m}"
let ?m = "insert (x, x) m \<union> ?s"
have Fm: "Field ?m = insert x (Field m)"
by (auto simp: Field_def)
have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
- using `Well_order m` by (simp_all add: order_on_defs)
---{*We show that the extension is a well-order*}
- have "Refl ?m" using `Refl m` Fm unfolding refl_on_def by blast
- moreover have "trans ?m" using `trans m` and `x \<notin> Field m`
+ using \<open>Well_order m\<close> by (simp_all add: order_on_defs)
+--\<open>We show that the extension is a well-order\<close>
+ have "Refl ?m" using \<open>Refl m\<close> Fm unfolding refl_on_def by blast
+ moreover have "trans ?m" using \<open>trans m\<close> and \<open>x \<notin> Field m\<close>
unfolding trans_def Field_def by blast
- moreover have "antisym ?m" using `antisym m` and `x \<notin> Field m`
+ moreover have "antisym ?m" using \<open>antisym m\<close> and \<open>x \<notin> Field m\<close>
unfolding antisym_def Field_def by blast
- moreover have "Total ?m" using `Total m` and Fm by (auto simp: total_on_def)
+ moreover have "Total ?m" using \<open>Total m\<close> and Fm by (auto simp: total_on_def)
moreover have "wf (?m - Id)"
proof -
- have "wf ?s" using `x \<notin> Field m` unfolding wf_eq_minimal Field_def
+ have "wf ?s" using \<open>x \<notin> Field m\<close> unfolding wf_eq_minimal Field_def
by (auto simp: Bex_def)
- thus ?thesis using `wf (m - Id)` and `x \<notin> Field m`
- wf_subset [OF `wf ?s` Diff_subset]
+ thus ?thesis using \<open>wf (m - Id)\<close> and \<open>x \<notin> Field m\<close>
+ wf_subset [OF \<open>wf ?s\<close> Diff_subset]
unfolding Un_Diff Field_def by (auto intro: wf_Un)
qed
ultimately have "Well_order ?m" by (simp add: order_on_defs)
---{*We show that the extension is above m*}
- moreover have "(m, ?m) \<in> I" using `Well_order ?m` and `Well_order m` and `x \<notin> Field m`
+--\<open>We show that the extension is above m\<close>
+ moreover have "(m, ?m) \<in> I" using \<open>Well_order ?m\<close> and \<open>Well_order m\<close> and \<open>x \<notin> Field m\<close>
by (fastforce simp: I_def init_seg_of_def Field_def)
ultimately
---{*This contradicts maximality of m:*}
- have False using max and `x \<notin> Field m` unfolding Field_def by blast
+--\<open>This contradicts maximality of m:\<close>
+ have False using max and \<open>x \<notin> Field m\<close> unfolding Field_def by blast
}
hence "Field m = UNIV" by auto
- with `Well_order m` show ?thesis by blast
+ with \<open>Well_order m\<close> show ?thesis by blast
qed
corollary well_order_on: "\<exists>r::'a rel. well_order_on A r"
@@ -708,14 +708,14 @@
have 1: "Field ?r = A" using wo univ
by (fastforce simp: Field_def order_on_defs refl_on_def)
have "Refl r" and "trans r" and "antisym r" and "Total r" and "wf (r - Id)"
- using `Well_order r` by (simp_all add: order_on_defs)
- have "Refl ?r" using `Refl r` by (auto simp: refl_on_def 1 univ)
- moreover have "trans ?r" using `trans r`
+ using \<open>Well_order r\<close> by (simp_all add: order_on_defs)
+ have "Refl ?r" using \<open>Refl r\<close> by (auto simp: refl_on_def 1 univ)
+ moreover have "trans ?r" using \<open>trans r\<close>
unfolding trans_def by blast
- moreover have "antisym ?r" using `antisym r`
+ moreover have "antisym ?r" using \<open>antisym r\<close>
unfolding antisym_def by blast
- moreover have "Total ?r" using `Total r` by (simp add:total_on_def 1 univ)
- moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf (r - Id)`]) blast
+ moreover have "Total ?r" using \<open>Total r\<close> by (simp add:total_on_def 1 univ)
+ moreover have "wf (?r - Id)" by (rule wf_subset [OF \<open>wf (r - Id)\<close>]) blast
ultimately have "Well_order ?r" by (simp add: order_on_defs)
with 1 show ?thesis by auto
qed
@@ -733,11 +733,11 @@
proof (intro exI allI)
fix x
def f \<equiv> "wfrec R (\<lambda>f x. SOME r. P f x r)"
- from `wf R` show "P f x (f x)"
+ from \<open>wf R\<close> show "P f x (f x)"
proof (induct x)
fix x assume "\<And>y. (y, x) \<in> R \<Longrightarrow> P f y (f y)"
show "P f x (f x)"
- proof (subst (2) wfrec_def_adm[OF f_def `wf R`])
+ proof (subst (2) wfrec_def_adm[OF f_def \<open>wf R\<close>])
show "adm_wf R (\<lambda>f x. SOME r. P f x r)"
by (auto simp add: adm_wf_def intro!: arg_cong[where f=Eps] ext adm)
show "P f x (Eps (P f x))"