--- 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)"