--- a/src/Doc/Classes/Classes.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Classes/Classes.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,7 +10,7 @@
of overloading\footnote{throughout this tutorial, we are referring
to classical Haskell 1.0 type classes, not considering later
additions in expressiveness}. As a canonical example, a polymorphic
- equality function @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on
+ equality function @{text "eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on
different types for @{text "\<alpha>"}, which is achieved by splitting
introduction of the @{text eq} function from its overloaded
definitions by means of @{text class} and @{text instance}
@@ -20,20 +20,20 @@
\begin{quote}
\noindent@{text "class eq where"} \\
- \hspace*{2ex}@{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
+ \hspace*{2ex}@{text "eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
- \medskip\noindent@{text "instance nat \<Colon> eq where"} \\
+ \medskip\noindent@{text "instance nat :: eq where"} \\
\hspace*{2ex}@{text "eq 0 0 = True"} \\
\hspace*{2ex}@{text "eq 0 _ = False"} \\
\hspace*{2ex}@{text "eq _ 0 = False"} \\
\hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"}
- \medskip\noindent@{text "instance (\<alpha>\<Colon>eq, \<beta>\<Colon>eq) pair \<Colon> eq where"} \\
+ \medskip\noindent@{text "instance (\<alpha>::eq, \<beta>::eq) pair :: eq where"} \\
\hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \<and> eq y1 y2"}
\medskip\noindent@{text "class ord extends eq where"} \\
- \hspace*{2ex}@{text "less_eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
- \hspace*{2ex}@{text "less \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
+ \hspace*{2ex}@{text "less_eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
+ \hspace*{2ex}@{text "less :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"}
\end{quote}
@@ -57,7 +57,7 @@
\begin{quote}
\noindent@{text "class eq where"} \\
- \hspace*{2ex}@{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
+ \hspace*{2ex}@{text "eq :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} \\
@{text "satisfying"} \\
\hspace*{2ex}@{text "refl: eq x x"} \\
\hspace*{2ex}@{text "sym: eq x y \<longleftrightarrow> eq x y"} \\
@@ -111,9 +111,9 @@
"fixes"}), the \qn{logical} part specifies properties on them
(@{element "assumes"}). The local @{element "fixes"} and @{element
"assumes"} are lifted to the theory toplevel, yielding the global
- parameter @{term [source] "mult \<Colon> \<alpha>\<Colon>semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the
- global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y z \<Colon>
- \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
+ parameter @{term [source] "mult :: \<alpha>::semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the
+ global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y z ::
+ \<alpha>::semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
*}
@@ -130,7 +130,7 @@
begin
definition %quote
- mult_int_def: "i \<otimes> j = i + (j\<Colon>int)"
+ mult_int_def: "i \<otimes> j = i + (j::int)"
instance %quote proof
fix i j k :: int have "(i + j) + k = i + (j + k)" by simp
@@ -163,7 +163,7 @@
begin
primrec %quote mult_nat where
- "(0\<Colon>nat) \<otimes> n = n"
+ "(0::nat) \<otimes> n = n"
| "Suc m \<otimes> n = Suc (m \<otimes> n)"
instance %quote proof
@@ -197,7 +197,7 @@
mult_prod_def: "p\<^sub>1 \<otimes> p\<^sub>2 = (fst p\<^sub>1 \<otimes> fst p\<^sub>2, snd p\<^sub>1 \<otimes> snd p\<^sub>2)"
instance %quote proof
- fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\<alpha>\<Colon>semigroup \<times> \<beta>\<Colon>semigroup"
+ fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\<alpha>::semigroup \<times> \<beta>::semigroup"
show "p\<^sub>1 \<otimes> p\<^sub>2 \<otimes> p\<^sub>3 = p\<^sub>1 \<otimes> (p\<^sub>2 \<otimes> p\<^sub>3)"
unfolding mult_prod_def by (simp add: assoc)
qed
@@ -237,10 +237,10 @@
begin
definition %quote
- neutral_nat_def: "\<one> = (0\<Colon>nat)"
+ neutral_nat_def: "\<one> = (0::nat)"
definition %quote
- neutral_int_def: "\<one> = (0\<Colon>int)"
+ neutral_int_def: "\<one> = (0::int)"
instance %quote proof
fix n :: nat
@@ -261,7 +261,7 @@
neutral_prod_def: "\<one> = (\<one>, \<one>)"
instance %quote proof
- fix p :: "\<alpha>\<Colon>monoidl \<times> \<beta>\<Colon>monoidl"
+ fix p :: "\<alpha>::monoidl \<times> \<beta>::monoidl"
show "\<one> \<otimes> p = p"
unfolding neutral_prod_def mult_prod_def by (simp add: neutl)
qed
@@ -295,7 +295,7 @@
begin
instance %quote proof
- fix p :: "\<alpha>\<Colon>monoid \<times> \<beta>\<Colon>monoid"
+ fix p :: "\<alpha>::monoid \<times> \<beta>::monoid"
show "p \<otimes> \<one> = p"
unfolding neutral_prod_def mult_prod_def by (simp add: neutr)
qed
@@ -315,7 +315,7 @@
begin
definition %quote
- inverse_int_def: "i\<div> = - (i\<Colon>int)"
+ inverse_int_def: "i\<div> = - (i::int)"
instance %quote proof
fix i :: int
@@ -361,7 +361,7 @@
*}
interpretation %quote idem_class:
- idem "f \<Colon> (\<alpha>\<Colon>idem) \<Rightarrow> \<alpha>"
+ idem "f :: (\<alpha>::idem) \<Rightarrow> \<alpha>"
(*<*)sorry(*>*)
text {*
@@ -394,10 +394,10 @@
\noindent Here the \qt{@{keyword "in"} @{class group}} target
specification indicates that the result is recorded within that
context for later use. This local theorem is also lifted to the
- global one @{fact "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon>
- \<alpha>\<Colon>group. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}. Since type @{text "int"} has been
+ global one @{fact "group.left_cancel:"} @{prop [source] "\<And>x y z ::
+ \<alpha>::group. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}. Since type @{text "int"} has been
made an instance of @{text "group"} before, we may refer to that
- fact as well: @{prop [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y =
+ fact as well: @{prop [source] "\<And>x y z :: int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y =
z"}.
*}
@@ -415,7 +415,7 @@
text {*
\noindent If the locale @{text group} is also a class, this local
definition is propagated onto a global definition of @{term [source]
- "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"} with corresponding theorems
+ "pow_nat :: nat \<Rightarrow> \<alpha>::monoid \<Rightarrow> \<alpha>::monoid"} with corresponding theorems
@{thm pow_nat.simps [no_vars]}.
@@ -542,8 +542,8 @@
else (pow_nat (nat (- k)) x)\<div>)"
text {*
- \noindent yields the global definition of @{term [source] "pow_int \<Colon>
- int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"} with the corresponding theorem @{thm
+ \noindent yields the global definition of @{term [source] "pow_int ::
+ int \<Rightarrow> \<alpha>::group \<Rightarrow> \<alpha>::group"} with the corresponding theorem @{thm
pow_int_def [no_vars]}.
*}
@@ -559,7 +559,7 @@
begin
term %quote "x \<otimes> y" -- {* example 1 *}
-term %quote "(x\<Colon>nat) \<otimes> y" -- {* example 2 *}
+term %quote "(x::nat) \<otimes> y" -- {* example 2 *}
end %quote
@@ -570,7 +570,7 @@
operation @{text "mult [\<alpha>]"}, whereas in example 2 the type
constraint enforces the global class operation @{text "mult [nat]"}.
In the global context in example 3, the reference is to the
- polymorphic global class operation @{text "mult [?\<alpha> \<Colon> semigroup]"}.
+ polymorphic global class operation @{text "mult [?\<alpha> :: semigroup]"}.
*}
section {* Further issues *}
--- a/src/Doc/Classes/Setup.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Classes/Setup.thy Tue Sep 01 22:32:58 2015 +0200
@@ -9,9 +9,9 @@
syntax
"_alpha" :: "type" ("\<alpha>")
- "_alpha_ofsort" :: "sort \<Rightarrow> type" ("\<alpha>()\<Colon>_" [0] 1000)
+ "_alpha_ofsort" :: "sort \<Rightarrow> type" ("\<alpha>()::_" [0] 1000)
"_beta" :: "type" ("\<beta>")
- "_beta_ofsort" :: "sort \<Rightarrow> type" ("\<beta>()\<Colon>_" [0] 1000)
+ "_beta_ofsort" :: "sort \<Rightarrow> type" ("\<beta>()::_" [0] 1000)
parse_ast_translation {*
let
--- a/src/Doc/Codegen/Adaptation.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Codegen/Adaptation.thy Tue Sep 01 22:32:58 2015 +0200
@@ -174,12 +174,12 @@
\item[@{text "Code_Binary_Nat"}] implements type
@{typ nat} using a binary rather than a linear representation,
which yields a considerable speedup for computations.
- Pattern matching with @{term "0\<Colon>nat"} / @{const "Suc"} is eliminated
+ Pattern matching with @{term "0::nat"} / @{const "Suc"} is eliminated
by a preprocessor.\label{abstract_nat}
\item[@{text "Code_Target_Nat"}] implements type @{typ nat}
by @{typ integer} and thus by target-language built-in integers.
- Pattern matching with @{term "0\<Colon>nat"} / @{const "Suc"} is eliminated
+ Pattern matching with @{term "0::nat"} / @{const "Suc"} is eliminated
by a preprocessor.
\item[@{text "Code_Target_Numeral"}] is a convenience theory
@@ -344,7 +344,7 @@
instantiation %quote bar :: equal
begin
-definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y"
+definition %quote "HOL.equal (x::bar) y \<longleftrightarrow> x = y"
instance %quote by default (simp add: equal_bar_def)
--- a/src/Doc/Codegen/Foundations.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Codegen/Foundations.thy Tue Sep 01 22:32:58 2015 +0200
@@ -144,7 +144,7 @@
\emph{Function transformers} provide a very general
interface, transforming a list of function theorems to another list
of function theorems, provided that neither the heading constant nor
- its type change. The @{term "0\<Colon>nat"} / @{const Suc} pattern
+ its type change. The @{term "0::nat"} / @{const Suc} pattern
used in theory @{text Code_Abstract_Nat} (see \secref{abstract_nat})
uses this interface.
--- a/src/Doc/Codegen/Introduction.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Codegen/Introduction.thy Tue Sep 01 22:32:58 2015 +0200
@@ -127,7 +127,7 @@
begin
primrec %quote mult_nat where
- "0 \<otimes> n = (0\<Colon>nat)"
+ "0 \<otimes> n = (0::nat)"
| "Suc m \<otimes> n = n + m \<otimes> n"
definition %quote neutral_nat where
--- a/src/Doc/Datatypes/Datatypes.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Datatypes/Datatypes.thy Tue Sep 01 22:32:58 2015 +0200
@@ -192,8 +192,8 @@
text {*
\noindent
-The constructors are @{text "None \<Colon> 'a option"} and
-@{text "Some \<Colon> 'a \<Rightarrow> 'a option"}.
+The constructors are @{text "None :: 'a option"} and
+@{text "Some :: 'a \<Rightarrow> 'a option"}.
The next example has three type parameters:
*}
@@ -203,7 +203,7 @@
text {*
\noindent
The constructor is
-@{text "Triple \<Colon> 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a, 'b, 'c) triple"}.
+@{text "Triple :: 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a, 'b, 'c) triple"}.
Unlike in Standard ML, curried constructors are supported. The uncurried variant
is also possible:
*}
@@ -381,21 +381,21 @@
\begin{tabular}{@ {}ll@ {}}
Constructors: &
- @{text "Nil \<Colon> 'a list"} \\
+ @{text "Nil :: 'a list"} \\
&
- @{text "Cons \<Colon> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
+ @{text "Cons :: 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
Discriminator: &
- @{text "null \<Colon> 'a list \<Rightarrow> bool"} \\
+ @{text "null :: 'a list \<Rightarrow> bool"} \\
Selectors: &
- @{text "hd \<Colon> 'a list \<Rightarrow> 'a"} \\
+ @{text "hd :: 'a list \<Rightarrow> 'a"} \\
&
- @{text "tl \<Colon> 'a list \<Rightarrow> 'a list"} \\
+ @{text "tl :: 'a list \<Rightarrow> 'a list"} \\
Set function: &
- @{text "set \<Colon> 'a list \<Rightarrow> 'a set"} \\
+ @{text "set :: 'a list \<Rightarrow> 'a set"} \\
Map function: &
- @{text "map \<Colon> ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
+ @{text "map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
Relator: &
- @{text "list_all2 \<Colon> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
+ @{text "list_all2 :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
\end{tabular}
\medskip
@@ -1327,7 +1327,7 @@
map function @{const map_option}:
*}
- primrec (*<*)(in early) (*>*)sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" where
+ primrec (*<*)(in early) (*>*)sum_btree :: "('a::{zero,plus}) btree \<Rightarrow> 'a" where
"sum_btree (BNode a lt rt) =
a + the_default 0 (map_option sum_btree lt) +
the_default 0 (map_option sum_btree rt)"
@@ -1427,7 +1427,7 @@
*}
primrec
- sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" and
+ sum_btree :: "('a::{zero,plus}) btree \<Rightarrow> 'a" and
sum_btree_option :: "'a btree option \<Rightarrow> 'a"
where
"sum_btree (BNode a lt rt) =
@@ -1590,7 +1590,7 @@
\setlength{\itemsep}{0pt}
\item
-Introduce a fully unspecified constant @{text "un_D\<^sub>0 \<Colon> 'a"} using
+Introduce a fully unspecified constant @{text "un_D\<^sub>0 :: 'a"} using
@{command consts}.
\item
@@ -1627,7 +1627,7 @@
text {* \blankline *}
overloading
- termi\<^sub>0 \<equiv> "termi\<^sub>0 \<Colon> ('a, 'b) tlist \<Rightarrow> 'b"
+ termi\<^sub>0 \<equiv> "termi\<^sub>0 :: ('a, 'b) tlist \<Rightarrow> 'b"
begin
primrec termi\<^sub>0 :: "('a, 'b) tlist \<Rightarrow> 'b" where
"termi\<^sub>0 (TNil y) = y" |
@@ -2322,12 +2322,12 @@
text {*
\noindent
Since there is no sequentiality, we can apply the equation for @{const Choice}
-without having first to discharge @{term "n mod (4\<Colon>int) \<noteq> 0"},
-@{term "n mod (4\<Colon>int) \<noteq> 1"}, and
-@{term "n mod (4\<Colon>int) \<noteq> 2"}.
+without having first to discharge @{term "n mod (4::int) \<noteq> 0"},
+@{term "n mod (4::int) \<noteq> 1"}, and
+@{term "n mod (4::int) \<noteq> 2"}.
The price to pay for this elegance is that we must discharge exclusiveness proof
obligations, one for each pair of conditions
-@{term "(n mod (4\<Colon>int) = i, n mod (4\<Colon>int) = j)"}
+@{term "(n mod (4::int) = i, n mod (4::int) = j)"}
with @{term "i < j"}. If we prefer not to discharge any obligations, we can
enable the @{text "sequential"} option. This pushes the problem to the users of
the generated properties.
@@ -2649,7 +2649,7 @@
(functorial action), $n$ set functions (natural transformations),
and an infinite cardinal bound that satisfy certain properties.
For example, @{typ "'a llist"} is a unary BNF.
-Its relator @{text "llist_all2 \<Colon>
+Its relator @{text "llist_all2 ::
('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow>
'a llist \<Rightarrow> 'b llist \<Rightarrow> bool"}
extends binary predicates over elements to binary predicates over parallel
@@ -2680,7 +2680,7 @@
set function, and relator.
*}
- typedef ('d, 'a) fn = "UNIV \<Colon> ('d \<Rightarrow> 'a) set"
+ typedef ('d, 'a) fn = "UNIV :: ('d \<Rightarrow> 'a) set"
by simp
text {* \blankline *}
@@ -2723,22 +2723,22 @@
show "set_fn \<circ> map_fn f = op ` f \<circ> set_fn"
by transfer (auto simp add: comp_def)
next
- show "card_order (natLeq +c |UNIV \<Colon> 'd set| )"
+ show "card_order (natLeq +c |UNIV :: 'd set| )"
apply (rule card_order_csum)
apply (rule natLeq_card_order)
by (rule card_of_card_order_on)
next
- show "cinfinite (natLeq +c |UNIV \<Colon> 'd set| )"
+ show "cinfinite (natLeq +c |UNIV :: 'd set| )"
apply (rule cinfinite_csum)
apply (rule disjI1)
by (rule natLeq_cinfinite)
next
fix F :: "('d, 'a) fn"
- have "|set_fn F| \<le>o |UNIV \<Colon> 'd set|" (is "_ \<le>o ?U")
+ have "|set_fn F| \<le>o |UNIV :: 'd set|" (is "_ \<le>o ?U")
by transfer (rule card_of_image)
also have "?U \<le>o natLeq +c ?U"
by (rule ordLeq_csum2) (rule card_of_Card_order)
- finally show "|set_fn F| \<le>o natLeq +c |UNIV \<Colon> 'd set|" .
+ finally show "|set_fn F| \<le>o natLeq +c |UNIV :: 'd set|" .
next
fix R :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and S :: "'b \<Rightarrow> 'c \<Rightarrow> bool"
show "rel_fn R OO rel_fn S \<le> rel_fn (R OO S)"
@@ -3122,7 +3122,7 @@
text {*
For each datatype, the \hthm{size} plugin generates a generic size
function @{text "t.size_t"} as well as a specific instance
-@{text "size \<Colon> t \<Rightarrow> nat"} belonging to the @{text size} type
+@{text "size :: t \<Rightarrow> nat"} belonging to the @{text size} type
class. The \keyw{fun} command relies on @{const size} to prove termination of
recursive functions on datatypes.
--- a/src/Doc/Tutorial/Types/Axioms.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Doc/Tutorial/Types/Axioms.thy Tue Sep 01 22:32:58 2015 +0200
@@ -20,13 +20,13 @@
text {* \noindent This @{command class} specification requires that
all instances of @{class semigroup} obey @{fact "assoc:"}~@{prop
-[source] "\<And>x y z \<Colon> 'a\<Colon>semigroup. (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"}.
+[source] "\<And>x y z :: 'a::semigroup. (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"}.
We can use this class axiom to derive further abstract theorems
relative to class @{class semigroup}: *}
lemma assoc_left:
- fixes x y z :: "'a\<Colon>semigroup"
+ fixes x y z :: "'a::semigroup"
shows "x \<oplus> (y \<oplus> z) = (x \<oplus> y) \<oplus> z"
using assoc by (rule sym)
@@ -63,7 +63,7 @@
begin
instance proof
- fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a\<Colon>semigroup \<times> 'b\<Colon>semigroup"
+ fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a::semigroup \<times> 'b::semigroup"
show "p\<^sub>1 \<oplus> p\<^sub>2 \<oplus> p\<^sub>3 = p\<^sub>1 \<oplus> (p\<^sub>2 \<oplus> p\<^sub>3)"
by (cases p\<^sub>1, cases p\<^sub>2, cases p\<^sub>3) (simp add: assoc)
@@ -96,7 +96,7 @@
begin
definition
- neutral_nat_def: "\<zero> = (0\<Colon>nat)"
+ neutral_nat_def: "\<zero> = (0::nat)"
instance proof
fix n :: nat
@@ -119,7 +119,7 @@
neutral_prod_def: "\<zero> = (\<zero>, \<zero>)"
instance proof
- fix p :: "'a\<Colon>monoidl \<times> 'b\<Colon>monoidl"
+ fix p :: "'a::monoidl \<times> 'b::monoidl"
show "\<zero> \<oplus> p = p"
by (cases p) (simp add: neutral_prod_def neutl)
qed
@@ -149,7 +149,7 @@
proofs relative to type classes: *}
lemma left_cancel:
- fixes x y z :: "'a\<Colon>group"
+ fixes x y z :: "'a::group"
shows "x \<oplus> y = x \<oplus> z \<longleftrightarrow> y = z"
proof
assume "x \<oplus> y = x \<oplus> z"
@@ -255,7 +255,7 @@
Further note that classes may contain axioms but \emph{no} operations.
An example is class @{class finite} from theory @{theory Finite_Set}
-which specifies a type to be finite: @{lemma [source] "finite (UNIV \<Colon> 'a\<Colon>finite
+which specifies a type to be finite: @{lemma [source] "finite (UNIV :: 'a::finite
set)" by (fact finite_UNIV)}. *}
(*<*)end(*>*)
--- a/src/HOL/Bali/AxCompl.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Bali/AxCompl.thy Tue Sep 01 22:32:58 2015 +0200
@@ -859,7 +859,7 @@
assumes wf: "wf_prog G"
and mgf_c1: "G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
and mgf_c2: "G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- shows "G,(A\<Colon>state triple set)\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
+ shows "G,(A::state triple set)\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp)
fix T L accC C
assume wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>In1r (c1 Finally c2)\<Colon>T"
@@ -956,7 +956,7 @@
assumes wf: "wf_prog G"
and mgf_init: "G,A\<turnstile>{=:n} \<langle>Init D\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
and mgf_c: "G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- shows "G,(A\<Colon>state triple set)\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
+ shows "G,(A::state triple set)\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp)
fix T L accC E
assume wt: "\<lparr>prg=G, cls=accC,lcl=L\<rparr>\<turnstile>\<langle>Body D c\<rangle>\<^sub>e\<Colon>T"
--- a/src/HOL/Basic_BNF_LFPs.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Basic_BNF_LFPs.thy Tue Sep 01 22:32:58 2015 +0200
@@ -94,7 +94,7 @@
ML_file "~~/src/HOL/Tools/Old_Datatype/old_size.ML"
-lemma size_bool[code]: "size (b\<Colon>bool) = 0"
+lemma size_bool[code]: "size (b::bool) = 0"
by (cases b) auto
declare prod.size[no_atp]
--- a/src/HOL/Binomial.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Binomial.thy Tue Sep 01 22:32:58 2015 +0200
@@ -448,7 +448,7 @@
case False
have *: "finite {1 .. n}" "0 \<notin> {1 .. n}" by auto
have eq: "insert 0 {1 .. n} = {0..n}" by auto
- have **: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) = (\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
+ have **: "(\<Prod>n\<in>{1::nat..n}. a + of_nat n) = (\<Prod>n\<in>{0::nat..n - 1}. a + 1 + of_nat n)"
apply (rule setprod.reindex_cong [where l = Suc])
using False
apply (auto simp add: fun_eq_iff field_simps)
@@ -564,7 +564,7 @@
next
case False
from this setprod_constant[of "{0 .. n - 1}" "- (1:: 'a)"]
- have eq: "(- (1\<Colon>'a)) ^ n = setprod (\<lambda>i. - 1) {0 .. n - 1}"
+ have eq: "(- (1::'a)) ^ n = setprod (\<lambda>i. - 1) {0 .. n - 1}"
by auto
from False show ?thesis
by (simp add: pochhammer_def gbinomial_def field_simps
@@ -687,7 +687,7 @@
(of_nat h * (\<Prod>i = 0..h. a - of_nat i) + 2 * (\<Prod>i = 0..h. a - of_nat i))"
by (simp add: field_simps)
also have "... =
- ((a gchoose Suc h) * (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0\<Colon>nat..Suc h}. a - of_nat i)"
+ ((a gchoose Suc h) * (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0::nat..Suc h}. a - of_nat i)"
unfolding gbinomial_mult_fact'
by (simp add: comm_semiring_class.distrib field_simps Suc)
also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
--- a/src/HOL/Code_Evaluation.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Code_Evaluation.thy Tue Sep 01 22:32:58 2015 +0200
@@ -93,7 +93,7 @@
begin
definition
- "term_of (f \<Colon> 'a \<Rightarrow> 'b) =
+ "term_of (f :: 'a \<Rightarrow> 'b) =
Const (STR ''Pure.dummy_pattern'')
(Typerep.Typerep (STR ''fun'') [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
@@ -119,8 +119,8 @@
by (subst term_of_anything) rule
code_printing
- constant "term_of \<Colon> integer \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT"
-| constant "term_of \<Colon> String.literal \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_literal"
+ constant "term_of :: integer \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT"
+| constant "term_of :: String.literal \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_literal"
declare [[code drop: "term_of :: integer \<Rightarrow> _"]]
--- a/src/HOL/Code_Numeral.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Code_Numeral.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,7 +10,7 @@
subsection \<open>Type of target language integers\<close>
-typedef integer = "UNIV \<Colon> int set"
+typedef integer = "UNIV :: int set"
morphisms int_of_integer integer_of_int ..
setup_lifting type_definition_integer
@@ -615,7 +615,7 @@
subsection \<open>Type of target language naturals\<close>
-typedef natural = "UNIV \<Colon> nat set"
+typedef natural = "UNIV :: nat set"
morphisms nat_of_natural natural_of_nat ..
setup_lifting type_definition_natural
--- a/src/HOL/Complex.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Complex.thy Tue Sep 01 22:32:58 2015 +0200
@@ -70,7 +70,7 @@
"Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)"
| "Im (inverse x) = - Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)"
-definition "x div (y\<Colon>complex) = x * inverse y"
+definition "x div (y::complex) = x * inverse y"
instance
by intro_classes
--- a/src/HOL/Datatype_Examples/Misc_Primcorec.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Primcorec.thy Tue Sep 01 22:32:58 2015 +0200
@@ -30,7 +30,7 @@
else if ys = MyNil then xs
else MyCons (myhd xs) (myapp (mytl xs) ys))"
-primcorec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
+primcorec shuffle_sp :: "('a :: ord, 'b :: ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
"shuffle_sp sp =
(case sp of
SP1 sp' \<Rightarrow> SP1 (shuffle_sp sp')
@@ -48,7 +48,7 @@
| Let SL l \<Rightarrow> Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l))"
primcorec
- j1_sum :: "('a\<Colon>{zero,one,plus}) \<Rightarrow> 'a J1" and
+ j1_sum :: "('a::{zero,one,plus}) \<Rightarrow> 'a J1" and
j2_sum :: "'a \<Rightarrow> 'a J2"
where
"n = 0 \<Longrightarrow> is_J11 (j1_sum n)" |
--- a/src/HOL/Datatype_Examples/Misc_Primrec.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Primrec.thy Tue Sep 01 22:32:58 2015 +0200
@@ -35,7 +35,7 @@
"myrev MyNil = MyNil" |
"myrev (MyCons x xs) = myapp (myrev xs) (MyCons x MyNil)"
-primrec shuffle_sp :: "('a \<Colon> ord, 'b \<Colon> ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
+primrec shuffle_sp :: "('a :: ord, 'b :: ord, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
"shuffle_sp (SP1 sp) = SP1 (shuffle_sp sp)" |
"shuffle_sp (SP2 a) = SP3 a" |
"shuffle_sp (SP3 b) = SP4 b" |
@@ -54,7 +54,7 @@
"rename_lam f (Let SL l) = Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l)"
primrec
- sum_i1 :: "('a\<Colon>{zero,plus}) I1 \<Rightarrow> 'a" and
+ sum_i1 :: "('a::{zero,plus}) I1 \<Rightarrow> 'a" and
sum_i2 :: "'a I2 \<Rightarrow> 'a"
where
"sum_i1 (I11 n i) = n + sum_i1 i" |
--- a/src/HOL/Decision_Procs/MIR.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Tue Sep 01 22:32:58 2015 +0200
@@ -48,7 +48,7 @@
(* The Divisibility relation between reals *)
definition rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
- where "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)"
+ where "x rdvd y \<longleftrightarrow> (\<exists>k::int. y = x * real k)"
lemma int_rdvd_real:
"real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r")
@@ -60,7 +60,7 @@
hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject)
thus ?r using th' by (simp add: dvd_def)
next
- assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
+ assume "?r" hence "(i::int) dvd \<lfloor>x::real\<rfloor>" ..
hence "\<exists> k. real (floor x) = real (i*k)"
by (simp only: real_of_int_inject) (simp add: dvd_def)
thus ?l using \<open>?r\<close> by (simp add: rdvd_def)
@@ -2438,7 +2438,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -2456,7 +2456,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<le> (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -2474,7 +2474,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -2492,7 +2492,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<ge> (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -2510,7 +2510,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -2528,7 +2528,7 @@
have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
by simp
- hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
+ hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0::real)) =
(real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
by simp
also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<noteq> (real (l div c)) * 0)" by (simp add: algebra_simps)
@@ -3431,7 +3431,7 @@
and n0: "n = 0"
and s_def: "s = (Add (Floor s') (C j))"
and jp: "0 \<le> j" and jn: "j \<le> n'"
- from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
+ from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \<longrightarrow>
Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
numbound0 s' \<and> isrlfm p'" by blast
hence nb: "numbound0 s'" by simp
@@ -3457,7 +3457,7 @@
and n0: "n = 0"
and s_def: "s = (Add (Floor s') (C j))"
and jp: "n' \<le> j" and jn: "j \<le> 0"
- from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
+ from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \<longrightarrow>
Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
numbound0 s' \<and> isrlfm p'" by blast
hence nb: "numbound0 s'" by simp
--- a/src/HOL/Divides.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Divides.thy Tue Sep 01 22:32:58 2015 +0200
@@ -814,8 +814,8 @@
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).
+ @{term "m::nat"}, @{term "n::nat"} and two output arguments
+ @{term "q::nat"}(uotient) and @{term "r::nat"}(emainder).
\<close>
definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where
@@ -835,7 +835,7 @@
have "\<exists>q r. m = q * n + r \<and> r < n"
proof (induct m)
case 0 with \<open>n \<noteq> 0\<close>
- have "(0\<Colon>nat) = 0 * n + 0 \<and> 0 < n" by simp
+ have "(0::nat) = 0 * n + 0 \<and> 0 < n" by simp
then show ?case by blast
next
case (Suc m) then obtain q' r'
@@ -868,7 +868,7 @@
(simp add: divmod_nat_rel_def)
next
case False
- have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q\<Colon>nat)"
+ have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q::nat)"
apply (rule leI)
apply (subst less_iff_Suc_add)
apply (auto simp add: add_mult_distrib)
@@ -1115,10 +1115,10 @@
then show "m div n * n + m mod n \<le> m" by auto
qed
-lemma mod_geq: "\<not> m < (n\<Colon>nat) \<Longrightarrow> m mod n = (m - n) mod n"
+lemma mod_geq: "\<not> m < (n::nat) \<Longrightarrow> m mod n = (m - n) mod n"
by (simp add: le_mod_geq linorder_not_less)
-lemma mod_if: "m mod (n\<Colon>nat) = (if m < n then m else (m - n) mod n)"
+lemma mod_if: "m mod (n::nat) = (if m < n then m else (m - n) mod n)"
by (simp add: le_mod_geq)
lemma mod_1 [simp]: "m mod Suc 0 = 0"
@@ -1313,7 +1313,7 @@
lemma split_div_lemma:
assumes "0 < n"
- shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m\<Colon>nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs")
+ shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m::nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?rhs
with mult_div_cancel have nq: "n * q = m - (m mod n)" by simp
@@ -1486,7 +1486,7 @@
lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
by (simp add: mult_2 [symmetric])
-lemma mod2_gr_0 [simp]: "0 < (m\<Colon>nat) mod 2 \<longleftrightarrow> m mod 2 = 1"
+lemma mod2_gr_0 [simp]: "0 < (m::nat) mod 2 \<longleftrightarrow> m mod 2 = 1"
proof -
{ fix n :: nat have "(n::nat) < 2 \<Longrightarrow> n = 0 \<or> n = 1" by (cases n) simp_all }
moreover have "m mod 2 < 2" by simp
@@ -1999,7 +1999,7 @@
lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1]
lemma zmod_zdiv_equality' [nitpick_unfold]:
- "(m\<Colon>int) mod n = m - (m div n) * n"
+ "(m::int) mod n = m - (m div n) * n"
using mod_div_equality [of m n] by arith
--- a/src/HOL/Enum.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Enum.thy Tue Sep 01 22:32:58 2015 +0200
@@ -144,7 +144,7 @@
by (fact equal_refl)
lemma order_fun [code]:
- fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order"
+ fixes f g :: "'a::enum \<Rightarrow> 'b::order"
shows "f \<le> g \<longleftrightarrow> enum_all (\<lambda>x. f x \<le> g x)"
and "f < g \<longleftrightarrow> f \<le> g \<and> enum_ex (\<lambda>x. f x \<noteq> g x)"
by (simp_all add: fun_eq_iff le_fun_def order_less_le)
@@ -244,32 +244,32 @@
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)"
- shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"
+ assumes "length ys = length (enum :: 'a::enum list)"
+ shows "\<exists>y. map_of (zip (enum :: 'a::enum list) ys) x = Some y"
proof -
- from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow>
- (\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)"
+ from assms have "x \<in> set (enum :: 'a::enum list) \<longleftrightarrow>
+ (\<exists>y. map_of (zip (enum :: 'a::enum list) ys) x = Some y)"
by (auto intro!: map_of_zip_is_Some)
then show ?thesis using enum_UNIV by auto
qed
lemma map_of_zip_enum_inject:
- fixes xs ys :: "'b\<Colon>enum list"
- assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)"
- "length ys = length (enum \<Colon> 'a\<Colon>enum list)"
- and map_of: "the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys)"
+ fixes xs ys :: "'b::enum list"
+ assumes length: "length xs = length (enum :: 'a::enum list)"
+ "length ys = length (enum :: 'a::enum list)"
+ and map_of: "the \<circ> map_of (zip (enum :: 'a::enum list) xs) = the \<circ> map_of (zip (enum :: 'a::enum list) ys)"
shows "xs = ys"
proof -
- have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)"
+ have "map_of (zip (enum :: 'a list) xs) = map_of (zip (enum :: 'a list) ys)"
proof
fix x :: 'a
from length map_of_zip_enum_is_Some obtain y1 y2
- where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1"
- and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast
+ where "map_of (zip (enum :: 'a list) xs) x = Some y1"
+ and "map_of (zip (enum :: 'a list) ys) x = Some y2" by blast
moreover from map_of
- have "the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x) = the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x)"
+ have "the (map_of (zip (enum :: 'a::enum list) xs) x) = the (map_of (zip (enum :: 'a::enum list) ys) x)"
by (auto dest: fun_cong)
- ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x"
+ ultimately show "map_of (zip (enum :: 'a::enum list) xs) x = map_of (zip (enum :: 'a::enum list) ys) x"
by simp
qed
with length enum_distinct show "xs = ys" by (rule map_of_zip_inject)
@@ -297,7 +297,7 @@
begin
definition
- "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (List.n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)"
+ "enum = map (\<lambda>ys. the o map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)"
definition
"enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
@@ -306,17 +306,17 @@
"enum_ex P = ex_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))"
instance proof
- show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
+ show "UNIV = set (enum :: ('a \<Rightarrow> 'b) list)"
proof (rule UNIV_eq_I)
fix f :: "'a \<Rightarrow> 'b"
- have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
+ have "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
then show "f \<in> set enum"
by (auto simp add: enum_fun_def set_n_lists intro: in_enum)
qed
next
from map_of_zip_enum_inject
- show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
+ show "distinct (enum :: ('a \<Rightarrow> 'b) list)"
by (auto intro!: inj_onI simp add: enum_fun_def
distinct_map distinct_n_lists enum_distinct set_n_lists)
next
@@ -327,7 +327,7 @@
show "Ball UNIV P"
proof
fix f :: "'a \<Rightarrow> 'b"
- have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
+ have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_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
@@ -352,9 +352,9 @@
next
assume "Bex UNIV P"
from this obtain f where "P f" ..
- have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"
+ have f: "f = the \<circ> map_of (zip (enum :: 'a::enum list) (map f enum))"
by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum)
- from \<open>P f\<close> 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 :: 'a::enum list) (map f enum)))"
by auto
from this show "enum_ex P"
unfolding enum_ex_fun_def ex_n_lists_def
@@ -367,7 +367,7 @@
end
-lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
+lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list)
in map (\<lambda>ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))"
by (simp add: enum_fun_def Let_def)
--- a/src/HOL/Fields.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Fields.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1161,7 +1161,7 @@
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
+ obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1::'a"] 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 .
--- a/src/HOL/Finite_Set.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Finite_Set.thy Tue Sep 01 22:32:58 2015 +0200
@@ -561,13 +561,13 @@
subsection \<open>Class @{text finite}\<close>
class finite =
- assumes finite_UNIV: "finite (UNIV \<Colon> 'a set)"
+ assumes finite_UNIV: "finite (UNIV :: 'a set)"
begin
-lemma finite [simp]: "finite (A \<Colon> 'a set)"
+lemma finite [simp]: "finite (A :: 'a set)"
by (rule subset_UNIV finite_UNIV finite_subset)+
-lemma finite_code [code]: "finite (A \<Colon> 'a set) \<longleftrightarrow> True"
+lemma finite_code [code]: "finite (A :: 'a set) \<longleftrightarrow> True"
by simp
end
--- a/src/HOL/Groups.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Groups.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1392,7 +1392,7 @@
subsection \<open>Tools setup\<close>
lemma add_mono_thms_linordered_semiring:
- fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
+ fixes i j k :: "'a::ordered_ab_semigroup_add"
shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
and "i \<le> j \<and> k = l \<Longrightarrow> i + k \<le> j + l"
@@ -1400,7 +1400,7 @@
by (rule add_mono, clarify+)+
lemma add_mono_thms_linordered_field:
- fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"
+ fixes i j k :: "'a::ordered_cancel_ab_semigroup_add"
shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"
and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"
and "i < j \<and> k \<le> l \<Longrightarrow> i + k < j + l"
--- a/src/HOL/HOL.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/HOL.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1792,7 +1792,7 @@
"equal TYPE('a) TYPE('a) \<longleftrightarrow> True"
by (simp add: equal)
-setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})\<close>
+setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a::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
@@ -1806,7 +1806,7 @@
qed (simp add: \<open>PROP ?equal\<close>)
qed
-setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>equal \<Rightarrow> 'a \<Rightarrow> bool"})\<close>
+setup \<open>Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a::equal \<Rightarrow> 'a \<Rightarrow> bool"})\<close>
setup \<open>Nbe.add_const_alias @{thm equal_alias_cert}\<close>
--- a/src/HOL/HOLCF/Lift.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/HOLCF/Lift.thy Tue Sep 01 22:32:58 2015 +0200
@@ -29,7 +29,7 @@
apply (simp add: Def_def)
done
-old_rep_datatype "\<bottom>\<Colon>'a lift" Def
+old_rep_datatype "\<bottom>::'a lift" Def
by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject inst_lift_pcpo)
text {* @{term bottom} and @{term Def} *}
--- a/src/HOL/Hahn_Banach/Linearform.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Hahn_Banach/Linearform.thy Tue Sep 01 22:32:58 2015 +0200
@@ -14,7 +14,7 @@
\<close>
locale linearform =
- fixes V :: "'a\<Colon>{minus, plus, zero, uminus} set" and f
+ fixes V :: "'a::{minus, plus, zero, uminus} set" and f
assumes add [iff]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> f (x + y) = f x + f y"
and mult [iff]: "x \<in> V \<Longrightarrow> f (a \<cdot> x) = a * f x"
--- a/src/HOL/Hahn_Banach/Normed_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Hahn_Banach/Normed_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -17,7 +17,7 @@
\<close>
locale seminorm =
- fixes V :: "'a\<Colon>{minus, plus, zero, uminus} set"
+ fixes V :: "'a::{minus, plus, zero, uminus} set"
fixes norm :: "'a \<Rightarrow> real" ("\<parallel>_\<parallel>")
assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
--- a/src/HOL/Hahn_Banach/Subspace.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Hahn_Banach/Subspace.thy Tue Sep 01 22:32:58 2015 +0200
@@ -17,7 +17,7 @@
\<close>
locale subspace =
- fixes U :: "'a\<Colon>{minus, plus, zero, uminus} set" and V
+ fixes U :: "'a::{minus, plus, zero, uminus} set" and V
assumes non_empty [iff, intro]: "U \<noteq> {}"
and subset [iff]: "U \<subseteq> V"
and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
--- a/src/HOL/Hilbert_Choice.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Hilbert_Choice.thy Tue Sep 01 22:32:58 2015 +0200
@@ -287,7 +287,7 @@
proof (rule UNIV_eq_I)
fix x :: 'a
from b1b2 have "x = inv (\<lambda>y. if y = x then b1 else b2) b1" by (simp add: inv_into_def)
- thus "x \<in> range (\<lambda>f\<Colon>'a \<Rightarrow> 'b. inv f b1)" by blast
+ thus "x \<in> range (\<lambda>f::'a \<Rightarrow> 'b. inv f b1)" by blast
qed
ultimately show "finite (UNIV :: 'a set)" by simp
qed
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Tue Sep 01 22:32:58 2015 +0200
@@ -494,7 +494,7 @@
definition "m_c m c = (let as = annos c in \<Sum>i=0..<size as. m(as!i))"
-lemma measure_m_c: "finite X \<Longrightarrow> {(c, c \<nabla>\<^sub>c c') |c c'\<Colon>ivl st option acom.
+lemma measure_m_c: "finite X \<Longrightarrow> {(c, c \<nabla>\<^sub>c c') |c c'::ivl st option acom.
strip c' = strip c \<and> c : Com X \<and> c' : Com X \<and> \<not> c' \<sqsubseteq> c}\<inverse>
\<subseteq> measure(m_c(m_o (m_st m_ivl) (2*card(X))))"
apply(auto simp: m_c_def Let_def Com_def)
--- a/src/HOL/Imperative_HOL/Array.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,63 +10,63 @@
subsection {* Primitives *}
-definition present :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> bool" where
+definition present :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> bool" where
"present h a \<longleftrightarrow> addr_of_array a < lim h"
-definition get :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a list" where
+definition get :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> 'a list" where
"get h a = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))"
-definition set :: "'a\<Colon>heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
+definition set :: "'a::heap array \<Rightarrow> 'a list \<Rightarrow> heap \<Rightarrow> heap" where
"set a x = arrays_update (\<lambda>h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))"
-definition alloc :: "'a list \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap array \<times> heap" where
+definition alloc :: "'a list \<Rightarrow> heap \<Rightarrow> 'a::heap array \<times> heap" where
"alloc xs h = (let
l = lim h;
r = Array l;
h'' = set r xs (h\<lparr>lim := l + 1\<rparr>)
in (r, h''))"
-definition length :: "heap \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> nat" where
+definition length :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> nat" where
"length h a = List.length (get h a)"
-definition update :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+definition update :: "'a::heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
"update a i x h = set a ((get h a)[i:=x]) h"
-definition noteq :: "'a\<Colon>heap array \<Rightarrow> 'b\<Colon>heap array \<Rightarrow> bool" (infix "=!!=" 70) where
+definition noteq :: "'a::heap array \<Rightarrow> 'b::heap array \<Rightarrow> bool" (infix "=!!=" 70) where
"r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
subsection {* Monad operations *}
-definition new :: "nat \<Rightarrow> 'a\<Colon>heap \<Rightarrow> 'a array Heap" where
+definition new :: "nat \<Rightarrow> 'a::heap \<Rightarrow> 'a array Heap" where
[code del]: "new n x = Heap_Monad.heap (alloc (replicate n x))"
-definition of_list :: "'a\<Colon>heap list \<Rightarrow> 'a array Heap" where
+definition of_list :: "'a::heap list \<Rightarrow> 'a array Heap" where
[code del]: "of_list xs = Heap_Monad.heap (alloc xs)"
-definition make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a\<Colon>heap) \<Rightarrow> 'a array Heap" where
+definition make :: "nat \<Rightarrow> (nat \<Rightarrow> 'a::heap) \<Rightarrow> 'a array Heap" where
[code del]: "make n f = Heap_Monad.heap (alloc (map f [0 ..< n]))"
-definition len :: "'a\<Colon>heap array \<Rightarrow> nat Heap" where
+definition len :: "'a::heap array \<Rightarrow> nat Heap" where
[code del]: "len a = Heap_Monad.tap (\<lambda>h. length h a)"
-definition nth :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" where
+definition nth :: "'a::heap array \<Rightarrow> nat \<Rightarrow> 'a Heap" where
[code del]: "nth a i = Heap_Monad.guard (\<lambda>h. i < length h a)
(\<lambda>h. (get h a ! i, h))"
-definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap" where
+definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a::heap array \<Rightarrow> 'a::heap array Heap" where
[code del]: "upd i x a = Heap_Monad.guard (\<lambda>h. i < length h a)
(\<lambda>h. (a, update a i x h))"
-definition map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where
+definition map_entry :: "nat \<Rightarrow> ('a::heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where
[code del]: "map_entry i f a = Heap_Monad.guard (\<lambda>h. i < length h a)
(\<lambda>h. (a, update a i (f (get h a ! i)) h))"
-definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap" where
+definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a::heap array \<Rightarrow> 'a Heap" where
[code del]: "swap i x a = Heap_Monad.guard (\<lambda>h. i < length h a)
(\<lambda>h. (get h a ! i, update a i x h))"
-definition freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap" where
+definition freeze :: "'a::heap array \<Rightarrow> 'a list Heap" where
[code del]: "freeze a = Heap_Monad.tap (\<lambda>h. get h a)"
--- a/src/HOL/Imperative_HOL/Heap.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/Heap.thy Tue Sep 01 22:32:58 2015 +0200
@@ -79,10 +79,10 @@
text {* Syntactic convenience *}
setup {*
- Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap array"})
- #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>heap ref"})
- #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\<Colon>heap array \<Rightarrow> nat"})
- #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\<Colon>heap ref \<Rightarrow> nat"})
+ Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a::heap array"})
+ #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a::heap ref"})
+ #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a::heap array \<Rightarrow> nat"})
+ #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a::heap ref \<Rightarrow> nat"})
*}
hide_const (open) empty
--- a/src/HOL/Imperative_HOL/Ref.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/Ref.thy Tue Sep 01 22:32:58 2015 +0200
@@ -16,38 +16,38 @@
subsection {* Primitives *}
-definition present :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> bool" where
+definition present :: "heap \<Rightarrow> 'a::heap ref \<Rightarrow> bool" where
"present h r \<longleftrightarrow> addr_of_ref r < lim h"
-definition get :: "heap \<Rightarrow> 'a\<Colon>heap ref \<Rightarrow> 'a" where
+definition get :: "heap \<Rightarrow> 'a::heap ref \<Rightarrow> 'a" where
"get h = from_nat \<circ> refs h TYPEREP('a) \<circ> addr_of_ref"
-definition set :: "'a\<Colon>heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+definition set :: "'a::heap ref \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
"set r x = refs_update
(\<lambda>h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))"
-definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a\<Colon>heap ref \<times> heap" where
+definition alloc :: "'a \<Rightarrow> heap \<Rightarrow> 'a::heap ref \<times> heap" where
"alloc x h = (let
l = lim h;
r = Ref l
in (r, set r x (h\<lparr>lim := l + 1\<rparr>)))"
-definition noteq :: "'a\<Colon>heap ref \<Rightarrow> 'b\<Colon>heap ref \<Rightarrow> bool" (infix "=!=" 70) where
+definition noteq :: "'a::heap ref \<Rightarrow> 'b::heap ref \<Rightarrow> bool" (infix "=!=" 70) where
"r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
subsection {* Monad operations *}
-definition ref :: "'a\<Colon>heap \<Rightarrow> 'a ref Heap" where
+definition ref :: "'a::heap \<Rightarrow> 'a ref Heap" where
[code del]: "ref v = Heap_Monad.heap (alloc v)"
-definition lookup :: "'a\<Colon>heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
+definition lookup :: "'a::heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
[code del]: "lookup r = Heap_Monad.tap (\<lambda>h. get h r)"
-definition update :: "'a ref \<Rightarrow> 'a\<Colon>heap \<Rightarrow> unit Heap" ("_ := _" 62) where
+definition update :: "'a ref \<Rightarrow> 'a::heap \<Rightarrow> unit Heap" ("_ := _" 62) where
[code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))"
-definition change :: "('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
+definition change :: "('a::heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
"change f r = do {
x \<leftarrow> ! r;
let y = f x;
--- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Tue Sep 01 22:32:58 2015 +0200
@@ -8,7 +8,7 @@
imports Subarray "~~/src/HOL/Imperative_HOL/Imperative_HOL"
begin
-fun swap :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" where
+fun swap :: "'a::heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" where
"swap a i j = do {
x \<leftarrow> Array.nth a i;
y \<leftarrow> Array.nth a j;
@@ -17,7 +17,7 @@
return ()
}"
-fun rev :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" where
+fun rev :: "'a::heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" where
"rev a i j = (if (i < j) then do {
swap a i j;
rev a (i + 1) (j - 1)
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,25 +10,25 @@
section {* Definition of Linked Lists *}
-setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>type ref"}) *}
+setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a::type ref"}) *}
datatype 'a node = Empty | Node 'a "'a node ref"
primrec
- node_encode :: "'a\<Colon>countable node \<Rightarrow> nat"
+ node_encode :: "'a::countable node \<Rightarrow> nat"
where
"node_encode Empty = 0"
| "node_encode (Node x r) = Suc (to_nat (x, r))"
instance node :: (countable) countable
proof (rule countable_classI [of "node_encode"])
- fix x y :: "'a\<Colon>countable node"
+ fix x y :: "'a::countable node"
show "node_encode x = node_encode y \<Longrightarrow> x = y"
by (induct x, auto, induct y, auto, induct y, auto)
qed
instance node :: (heap) heap ..
-primrec make_llist :: "'a\<Colon>heap list \<Rightarrow> 'a node Heap"
+primrec make_llist :: "'a::heap list \<Rightarrow> 'a node Heap"
where
[simp del]: "make_llist [] = return Empty"
| "make_llist (x#xs) = do { tl \<leftarrow> make_llist xs;
@@ -37,7 +37,7 @@
}"
-partial_function (heap) traverse :: "'a\<Colon>heap node \<Rightarrow> 'a list Heap"
+partial_function (heap) traverse :: "'a::heap node \<Rightarrow> 'a list Heap"
where
[code del]: "traverse l =
(case l of Empty \<Rightarrow> return []
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Tue Sep 01 22:32:58 2015 +0200
@@ -119,7 +119,7 @@
text {* Specific definition for derived clauses in the Array *}
definition
- array_ran :: "('a\<Colon>heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
+ array_ran :: "('a::heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
"array_ran a h = {e. Some e \<in> set (Array.get h a)}"
-- {*FIXME*}
--- a/src/HOL/Import/HOL_Light_Maps.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Import/HOL_Light_Maps.thy Tue Sep 01 22:32:58 2015 +0200
@@ -16,19 +16,19 @@
by simp
lemma [import_const "/\\"]:
- "(op \<and>) = (\<lambda>p q. (\<lambda>f. f p q \<Colon> bool) = (\<lambda>f. f True True))"
+ "(op \<and>) = (\<lambda>p q. (\<lambda>f. f p q :: bool) = (\<lambda>f. f True True))"
by metis
lemma [import_const "==>"]:
- "(op \<longrightarrow>) = (\<lambda>(p\<Colon>bool) q\<Colon>bool. (p \<and> q) = p)"
+ "(op \<longrightarrow>) = (\<lambda>(p::bool) q::bool. (p \<and> q) = p)"
by auto
lemma [import_const "!"]:
- "All = (\<lambda>P\<Colon>'A \<Rightarrow> bool. P = (\<lambda>x\<Colon>'A. True))"
+ "All = (\<lambda>P::'A \<Rightarrow> bool. P = (\<lambda>x::'A. True))"
by auto
lemma [import_const "?"]:
- "Ex = (\<lambda>P\<Colon>'A \<Rightarrow> bool. All (\<lambda>q\<Colon>bool. (All (\<lambda>x\<Colon>'A\<Colon>type. (P x) \<longrightarrow> q)) \<longrightarrow> q))"
+ "Ex = (\<lambda>P::'A \<Rightarrow> bool. All (\<lambda>q::bool. (All (\<lambda>x::'A::type. (P x) \<longrightarrow> q)) \<longrightarrow> q))"
by auto
lemma [import_const "\\/"]:
@@ -44,16 +44,16 @@
by simp
lemma [import_const "?!"]:
- "Ex1 = (\<lambda>P\<Colon>'A \<Rightarrow> bool. Ex P \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y))"
+ "Ex1 = (\<lambda>P::'A \<Rightarrow> bool. Ex P \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y))"
by auto
lemma [import_const "_FALSITY_"]: "False = False"
by simp
-lemma hl_ax1: "\<forall>t\<Colon>'A \<Rightarrow> 'B. (\<lambda>x. t x) = t"
+lemma hl_ax1: "\<forall>t::'A \<Rightarrow> 'B. (\<lambda>x. t x) = t"
by metis
-lemma hl_ax2: "\<forall>P x\<Colon>'A. P x \<longrightarrow> P (Eps P)"
+lemma hl_ax2: "\<forall>P x::'A. P x \<longrightarrow> P (Eps P)"
by (auto intro: someI)
lemma [import_const COND]:
@@ -61,51 +61,51 @@
unfolding fun_eq_iff by auto
lemma [import_const o]:
- "(op \<circ>) = (\<lambda>(f\<Colon>'B \<Rightarrow> 'C) g x\<Colon>'A. f (g x))"
+ "(op \<circ>) = (\<lambda>(f::'B \<Rightarrow> 'C) g x::'A. f (g x))"
unfolding fun_eq_iff by simp
-lemma [import_const I]: "id = (\<lambda>x\<Colon>'A. x)"
+lemma [import_const I]: "id = (\<lambda>x::'A. x)"
by auto
lemma [import_type 1 one_ABS one_REP]:
"type_definition Rep_unit Abs_unit (Collect (\<lambda>b. b))"
by (metis (full_types) Collect_cong singleton_conv2 type_definition_unit)
-lemma [import_const one]: "() = (SOME x\<Colon>unit. True)"
+lemma [import_const one]: "() = (SOME x::unit. True)"
by auto
lemma [import_const mk_pair]:
- "Pair_Rep = (\<lambda>(x\<Colon>'A) (y\<Colon>'B) (a\<Colon>'A) b\<Colon>'B. a = x \<and> b = y)"
+ "Pair_Rep = (\<lambda>(x::'A) (y::'B) (a::'A) b::'B. a = x \<and> b = y)"
by (simp add: Pair_Rep_def fun_eq_iff)
lemma [import_type prod ABS_prod REP_prod]:
- "type_definition Rep_prod Abs_prod (Collect (\<lambda>x\<Colon>'A \<Rightarrow> 'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b))"
+ "type_definition Rep_prod Abs_prod (Collect (\<lambda>x::'A \<Rightarrow> 'B \<Rightarrow> bool. \<exists>a b. x = Pair_Rep a b))"
using type_definition_prod[unfolded Product_Type.prod_def] by simp
-lemma [import_const ","]: "Pair = (\<lambda>(x\<Colon>'A) y\<Colon>'B. Abs_prod (Pair_Rep x y))"
+lemma [import_const ","]: "Pair = (\<lambda>(x::'A) y::'B. Abs_prod (Pair_Rep x y))"
by (metis Pair_def)
lemma [import_const FST]:
- "fst = (\<lambda>p\<Colon>'A \<times> 'B. SOME x\<Colon>'A. \<exists>y\<Colon>'B. p = (x, y))"
+ "fst = (\<lambda>p::'A \<times> 'B. SOME x::'A. \<exists>y::'B. p = (x, y))"
by auto
lemma [import_const SND]:
- "snd = (\<lambda>p\<Colon>'A \<times> 'B. SOME y\<Colon>'B. \<exists>x\<Colon>'A. p = (x, y))"
+ "snd = (\<lambda>p::'A \<times> 'B. SOME y::'B. \<exists>x::'A. p = (x, y))"
by auto
lemma CURRY_DEF[import_const CURRY]:
- "curry = (\<lambda>(f\<Colon>'A \<times> 'B \<Rightarrow> 'C) x y. f (x, y))"
+ "curry = (\<lambda>(f::'A \<times> 'B \<Rightarrow> 'C) x y. f (x, y))"
using curry_def .
lemma [import_const ONE_ONE : inj]:
- "inj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>x1 x2. f x1 = f x2 \<longrightarrow> x1 = x2)"
+ "inj = (\<lambda>(f::'A \<Rightarrow> 'B). \<forall>x1 x2. f x1 = f x2 \<longrightarrow> x1 = x2)"
by (auto simp add: fun_eq_iff inj_on_def)
lemma [import_const ONTO : surj]:
- "surj = (\<lambda>(f\<Colon>'A \<Rightarrow> 'B). \<forall>y. \<exists>x. y = f x)"
+ "surj = (\<lambda>(f::'A \<Rightarrow> 'B). \<forall>y. \<exists>x. y = f x)"
by (auto simp add: fun_eq_iff)
-lemma hl_ax3: "\<exists>f\<Colon>ind \<Rightarrow> ind. inj f \<and> \<not> surj f"
+lemma hl_ax3: "\<exists>f::ind \<Rightarrow> ind. inj f \<and> \<not> surj f"
by (rule_tac x="Suc_Rep" in exI)
(metis Suc_Rep_inject' injI Suc_Rep_not_Zero_Rep surjD)
@@ -142,7 +142,7 @@
definition [simp]: "pred n = n - 1"
lemma PRE[import_const PRE : pred]:
- "pred (id (0\<Colon>nat)) = id (0\<Colon>nat) \<and> (\<forall>n\<Colon>nat. pred (Suc n) = n)"
+ "pred (id (0::nat)) = id (0::nat) \<and> (\<forall>n::nat. pred (Suc n) = n)"
by simp
lemma ADD[import_const "+" : plus]:
@@ -186,11 +186,11 @@
"even = Parity.even"
lemma EVEN[import_const "EVEN" : even]:
- "even (id 0\<Colon>nat) = True \<and> (\<forall>n. even (Suc n) = (\<not> even n))"
+ "even (id 0::nat) = True \<and> (\<forall>n. even (Suc n) = (\<not> even n))"
by (simp add: even_def)
lemma SUB[import_const "-" : minus]:
- "(\<forall>m\<Colon>nat. m - (id 0) = m) \<and> (\<forall>m n. m - (Suc n) = pred (m - n))"
+ "(\<forall>m::nat. m - (id 0) = m) \<and> (\<forall>m n. m - (Suc n) = pred (m - n))"
by simp
lemma FACT[import_const "FACT" : fact]:
@@ -201,7 +201,7 @@
import_const_map DIV : divide
lemma DIVISION_0:
- "\<forall>m n\<Colon>nat. if n = id 0 then m div n = id 0 \<and> m mod n = m else m = m div n * n + m mod n \<and> m mod n < n"
+ "\<forall>m n::nat. if n = id 0 then m div n = id 0 \<and> m mod n = m else m = m div n * n + m mod n \<and> m mod n < n"
by simp
lemmas [import_type sum "_dest_sum" "_mk_sum"] = type_definition_sum[where 'a="'A" and 'b="'B"]
@@ -229,40 +229,40 @@
import_const_map CONS : Cons
lemma list_INDUCT:
- "\<forall>P\<Colon>'A list \<Rightarrow> bool. P [] \<and> (\<forall>a0 a1. P a1 \<longrightarrow> P (a0 # a1)) \<longrightarrow> (\<forall>x. P x)"
+ "\<forall>P::'A list \<Rightarrow> bool. P [] \<and> (\<forall>a0 a1. P a1 \<longrightarrow> P (a0 # a1)) \<longrightarrow> (\<forall>x. P x)"
using list.induct by auto
lemma list_RECURSION:
- "\<forall>nil' cons'. \<exists>fn\<Colon>'A list \<Rightarrow> 'Z. fn [] = nil' \<and> (\<forall>(a0\<Colon>'A) a1\<Colon>'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))"
+ "\<forall>nil' cons'. \<exists>fn::'A list \<Rightarrow> 'Z. fn [] = nil' \<and> (\<forall>(a0::'A) a1::'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))"
by (intro allI, rule_tac x="rec_list nil' cons'" in exI) auto
lemma HD[import_const HD : hd]:
- "hd ((h\<Colon>'A) # t) = h"
+ "hd ((h::'A) # t) = h"
by simp
lemma TL[import_const TL : tl]:
- "tl ((h\<Colon>'A) # t) = t"
+ "tl ((h::'A) # t) = t"
by simp
lemma APPEND[import_const APPEND : append]:
- "(\<forall>l\<Colon>'A list. [] @ l = l) \<and> (\<forall>(h\<Colon>'A) t l. (h # t) @ l = h # t @ l)"
+ "(\<forall>l::'A list. [] @ l = l) \<and> (\<forall>(h::'A) t l. (h # t) @ l = h # t @ l)"
by simp
lemma REVERSE[import_const REVERSE : rev]:
- "rev [] = ([] :: 'A list) \<and> rev ((x\<Colon>'A) # l) = rev l @ [x]"
+ "rev [] = ([] :: 'A list) \<and> rev ((x::'A) # l) = rev l @ [x]"
by simp
lemma LENGTH[import_const LENGTH : length]:
- "length ([] :: 'A list) = id 0 \<and> (\<forall>(h\<Colon>'A) t. length (h # t) = Suc (length t))"
+ "length ([] :: 'A list) = id 0 \<and> (\<forall>(h::'A) t. length (h # t) = Suc (length t))"
by simp
lemma MAP[import_const MAP : map]:
- "(\<forall>f\<Colon>'A \<Rightarrow> 'B. map f [] = []) \<and>
- (\<forall>(f\<Colon>'A \<Rightarrow> 'B) h t. map f (h # t) = f h # map f t)"
+ "(\<forall>f::'A \<Rightarrow> 'B. map f [] = []) \<and>
+ (\<forall>(f::'A \<Rightarrow> 'B) h t. map f (h # t) = f h # map f t)"
by simp
lemma LAST[import_const LAST : last]:
- "last ((h\<Colon>'A) # t) = (if t = [] then h else last t)"
+ "last ((h::'A) # t) = (if t = [] then h else last t)"
by simp
lemma BUTLAST[import_const BUTLAST : butlast]:
@@ -271,43 +271,43 @@
by simp
lemma REPLICATE[import_const REPLICATE : replicate]:
- "replicate (id (0\<Colon>nat)) (x\<Colon>'t18358) = [] \<and>
+ "replicate (id (0::nat)) (x::'t18358) = [] \<and>
replicate (Suc n) x = x # replicate n x"
by simp
lemma NULL[import_const NULL : List.null]:
- "List.null ([]\<Colon>'t18373 list) = True \<and> List.null ((h\<Colon>'t18373) # t) = False"
+ "List.null ([]::'t18373 list) = True \<and> List.null ((h::'t18373) # t) = False"
unfolding null_def by simp
lemma ALL[import_const ALL : list_all]:
- "list_all (P\<Colon>'t18393 \<Rightarrow> bool) [] = True \<and>
+ "list_all (P::'t18393 \<Rightarrow> bool) [] = True \<and>
list_all P (h # t) = (P h \<and> list_all P t)"
by simp
lemma EX[import_const EX : list_ex]:
- "list_ex (P\<Colon>'t18414 \<Rightarrow> bool) [] = False \<and>
+ "list_ex (P::'t18414 \<Rightarrow> bool) [] = False \<and>
list_ex P (h # t) = (P h \<or> list_ex P t)"
by simp
lemma ITLIST[import_const ITLIST : foldr]:
- "foldr (f\<Colon>'t18437 \<Rightarrow> 't18436 \<Rightarrow> 't18436) [] b = b \<and>
+ "foldr (f::'t18437 \<Rightarrow> 't18436 \<Rightarrow> 't18436) [] b = b \<and>
foldr f (h # t) b = f h (foldr f t b)"
by simp
lemma ALL2_DEF[import_const ALL2 : list_all2]:
- "list_all2 (P\<Colon>'t18495 \<Rightarrow> 't18502 \<Rightarrow> bool) [] (l2\<Colon>'t18502 list) = (l2 = []) \<and>
- list_all2 P ((h1\<Colon>'t18495) # (t1\<Colon>'t18495 list)) l2 =
+ "list_all2 (P::'t18495 \<Rightarrow> 't18502 \<Rightarrow> bool) [] (l2::'t18502 list) = (l2 = []) \<and>
+ list_all2 P ((h1::'t18495) # (t1::'t18495 list)) l2 =
(if l2 = [] then False else P h1 (hd l2) \<and> list_all2 P t1 (tl l2))"
by simp (induct_tac l2, simp_all)
lemma FILTER[import_const FILTER : filter]:
- "filter (P\<Colon>'t18680 \<Rightarrow> bool) [] = [] \<and>
- filter P ((h\<Colon>'t18680) # t) = (if P h then h # filter P t else filter P t)"
+ "filter (P::'t18680 \<Rightarrow> bool) [] = [] \<and>
+ filter P ((h::'t18680) # t) = (if P h then h # filter P t else filter P t)"
by simp
lemma ZIP[import_const ZIP : zip]:
"zip [] [] = ([] :: ('t18824 \<times> 't18825) list) \<and>
- zip ((h1\<Colon>'t18849) # t1) ((h2\<Colon>'t18850) # t2) = (h1, h2) # zip t1 t2"
+ zip ((h1::'t18849) # t1) ((h2::'t18850) # t2) = (h1, h2) # zip t1 t2"
by simp
lemma WF[import_const WF : wfP]:
--- a/src/HOL/Inductive.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Inductive.thy Tue Sep 01 22:32:58 2015 +0200
@@ -57,7 +57,7 @@
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"
+ fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
assumes mono: "mono f"
and P_f: "\<And>S. P S \<Longrightarrow> S \<le> lfp f \<Longrightarrow> P (f S)"
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Sup M)"
@@ -177,7 +177,7 @@
by (blast dest: gfp_lemma2 mono_Un)
lemma gfp_ordinal_induct[case_names mono step union]:
- fixes f :: "'a\<Colon>complete_lattice \<Rightarrow> 'a"
+ fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
assumes mono: "mono f"
and P_f: "\<And>S. P S \<Longrightarrow> gfp f \<le> S \<Longrightarrow> P (f S)"
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Inf M)"
--- a/src/HOL/Int.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Int.thy Tue Sep 01 22:32:58 2015 +0200
@@ -107,10 +107,10 @@
begin
definition
- "(inf \<Colon> int \<Rightarrow> int \<Rightarrow> int) = min"
+ "(inf :: int \<Rightarrow> int \<Rightarrow> int) = min"
definition
- "(sup \<Colon> int \<Rightarrow> int \<Rightarrow> int) = max"
+ "(sup :: int \<Rightarrow> int \<Rightarrow> int) = max"
instance
by intro_classes
@@ -161,10 +161,10 @@
begin
definition
- zabs_def: "\<bar>i\<Colon>int\<bar> = (if i < 0 then - i else i)"
+ zabs_def: "\<bar>i::int\<bar> = (if i < 0 then - i else i)"
definition
- zsgn_def: "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
+ zsgn_def: "sgn (i::int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
instance proof
fix i j k :: int
@@ -172,17 +172,17 @@
by (rule zmult_zless_mono2)
show "\<bar>i\<bar> = (if i < 0 then -i else i)"
by (simp only: zabs_def)
- show "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
+ show "sgn (i::int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
by (simp only: zsgn_def)
qed
end
-lemma zless_imp_add1_zle: "w < z \<Longrightarrow> w + (1\<Colon>int) \<le> z"
+lemma zless_imp_add1_zle: "w < z \<Longrightarrow> w + (1::int) \<le> z"
by transfer clarsimp
lemma zless_iff_Suc_zadd:
- "(w \<Colon> int) < z \<longleftrightarrow> (\<exists>n. z = w + int (Suc n))"
+ "(w :: int) < z \<longleftrightarrow> (\<exists>n. z = w + int (Suc n))"
apply transfer
apply auto
apply (rename_tac a b c d)
@@ -438,7 +438,7 @@
subsection\<open>Lemmas about the Function @{term of_nat} and Orderings\<close>
-lemma negative_zless_0: "- (int (Suc n)) < (0 \<Colon> int)"
+lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)"
by (simp add: order_less_le del: of_nat_Suc)
lemma negative_zless [iff]: "- (int (Suc n)) < int m"
@@ -1117,9 +1117,9 @@
lemma infinite_UNIV_int: "\<not> finite (UNIV::int set)"
proof
assume "finite (UNIV::int set)"
- moreover have "inj (\<lambda>i\<Colon>int. 2 * i)"
+ moreover have "inj (\<lambda>i::int. 2 * i)"
by (rule injI) simp
- ultimately have "surj (\<lambda>i\<Colon>int. 2 * i)"
+ ultimately have "surj (\<lambda>i::int. 2 * i)"
by (rule finite_UNIV_inj_surj)
then obtain i :: int where "1 = 2 * i" by (rule surjE)
then show False by (simp add: pos_zmult_eq_1_iff)
--- a/src/HOL/Lattice/Orders.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Lattice/Orders.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,7 +10,7 @@
text {*
We define several classes of ordered structures over some type @{typ
- 'a} with relation @{text "\<sqsubseteq> \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> bool"}. For a
+ 'a} with relation @{text "\<sqsubseteq> :: 'a \<Rightarrow> 'a \<Rightarrow> bool"}. For a
\emph{quasi-order} that relation is required to be reflexive and
transitive, for a \emph{partial order} it also has to be
anti-symmetric, while for a \emph{linear order} all elements are
--- a/src/HOL/Lattices.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Lattices.thy Tue Sep 01 22:32:58 2015 +0200
@@ -222,7 +222,7 @@
by (fast intro: inf_greatest le_infI1 le_infI2)
lemma mono_inf:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>semilattice_inf"
+ fixes f :: "'a \<Rightarrow> 'b::semilattice_inf"
shows "mono f \<Longrightarrow> f (A \<sqinter> B) \<sqsubseteq> f A \<sqinter> f B"
by (auto simp add: mono_def intro: Lattices.inf_greatest)
@@ -259,7 +259,7 @@
by (fast intro: sup_least le_supI1 le_supI2)
lemma mono_sup:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>semilattice_sup"
+ fixes f :: "'a \<Rightarrow> 'b::semilattice_sup"
shows "mono f \<Longrightarrow> f A \<squnion> f B \<sqsubseteq> f (A \<squnion> B)"
by (auto simp add: mono_def intro: Lattices.sup_least)
@@ -770,21 +770,21 @@
by (simp add: max_def)
lemma min_of_mono:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>linorder"
+ fixes f :: "'a \<Rightarrow> 'b::linorder"
shows "mono f \<Longrightarrow> min (f m) (f n) = f (min m n)"
by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym)
lemma max_of_mono:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>linorder"
+ fixes f :: "'a \<Rightarrow> 'b::linorder"
shows "mono f \<Longrightarrow> max (f m) (f n) = f (max m n)"
by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym)
end
-lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{semilattice_inf, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
+lemma inf_min: "inf = (min :: 'a::{semilattice_inf,linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
by (auto intro: antisym simp add: min_def fun_eq_iff)
-lemma sup_max: "sup = (max \<Colon> 'a\<Colon>{semilattice_sup, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
+lemma sup_max: "sup = (max :: 'a::{semilattice_sup,linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
by (auto intro: antisym simp add: max_def fun_eq_iff)
--- a/src/HOL/Library/Cardinality.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Cardinality.thy Tue Sep 01 22:32:58 2015 +0200
@@ -41,7 +41,7 @@
syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
-translations "CARD('t)" => "CONST card (CONST UNIV \<Colon> 't set)"
+translations "CARD('t)" => "CONST card (CONST UNIV :: 't set)"
print_translation \<open>
let
--- a/src/HOL/Library/Char_ord.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Char_ord.thy Tue Sep 01 22:32:58 2015 +0200
@@ -22,8 +22,8 @@
instantiation nibble :: distrib_lattice
begin
-definition "(inf \<Colon> nibble \<Rightarrow> _) = min"
-definition "(sup \<Colon> nibble \<Rightarrow> _) = max"
+definition "(inf :: nibble \<Rightarrow> _) = min"
+definition "(sup :: nibble \<Rightarrow> _) = max"
instance
by standard (auto simp add: inf_nibble_def sup_nibble_def max_min_distrib2)
@@ -74,8 +74,8 @@
instantiation char :: distrib_lattice
begin
-definition "(inf \<Colon> char \<Rightarrow> _) = min"
-definition "(sup \<Colon> char \<Rightarrow> _) = max"
+definition "(inf :: char \<Rightarrow> _) = min"
+definition "(sup :: char \<Rightarrow> _) = max"
instance
by standard (auto simp add: inf_char_def sup_char_def max_min_distrib2)
--- a/src/HOL/Library/Code_Binary_Nat.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Code_Binary_Nat.thy Tue Sep 01 22:32:58 2015 +0200
@@ -27,10 +27,10 @@
by (simp_all add: nat_of_num_inverse)
lemma [code]:
- "(1\<Colon>nat) = Numeral1"
+ "(1::nat) = Numeral1"
by simp
-lemma [code_abbrev]: "Numeral1 = (1\<Colon>nat)"
+lemma [code_abbrev]: "Numeral1 = (1::nat)"
by simp
lemma [code]:
--- a/src/HOL/Library/Code_Char.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Code_Char.thy Tue Sep 01 22:32:58 2015 +0200
@@ -28,7 +28,7 @@
and (OCaml) "!((_ : char) = _)"
and (Haskell) infix 4 "=="
and (Scala) infixl 5 "=="
-| constant "Code_Evaluation.term_of \<Colon> char \<Rightarrow> term" \<rightharpoonup>
+| constant "Code_Evaluation.term_of :: char \<Rightarrow> term" \<rightharpoonup>
(Eval) "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))"
code_reserved SML
--- a/src/HOL/Library/Countable.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Countable.thy Tue Sep 01 22:32:58 2015 +0200
@@ -13,7 +13,7 @@
subsection \<open>The class of countable types\<close>
class countable =
- assumes ex_inj: "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat"
+ assumes ex_inj: "\<exists>to_nat :: 'a \<Rightarrow> nat. inj to_nat"
lemma countable_classI:
fixes f :: "'a \<Rightarrow> nat"
@@ -27,11 +27,11 @@
subsection \<open>Conversion functions\<close>
-definition to_nat :: "'a\<Colon>countable \<Rightarrow> nat" where
+definition to_nat :: "'a::countable \<Rightarrow> nat" where
"to_nat = (SOME f. inj f)"
-definition from_nat :: "nat \<Rightarrow> 'a\<Colon>countable" where
- "from_nat = inv (to_nat \<Colon> 'a \<Rightarrow> nat)"
+definition from_nat :: "nat \<Rightarrow> 'a::countable" where
+ "from_nat = inv (to_nat :: 'a \<Rightarrow> nat)"
lemma inj_to_nat [simp]: "inj to_nat"
by (rule exE_some [OF ex_inj]) (simp add: to_nat_def)
@@ -54,13 +54,13 @@
subclass (in finite) countable
proof
- have "finite (UNIV\<Colon>'a set)" by (rule finite_UNIV)
+ have "finite (UNIV::'a set)" by (rule finite_UNIV)
with finite_conv_nat_seg_image [of "UNIV::'a set"]
obtain n and f :: "nat \<Rightarrow> 'a"
where "UNIV = f ` {i. i < n}" by auto
then have "surj f" unfolding surj_def by auto
then have "inj (inv f)" by (rule surj_imp_inj_inv)
- then show "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat" by (rule exI[of inj])
+ then show "\<exists>to_nat :: 'a \<Rightarrow> nat. inj to_nat" by (rule exI[of inj])
qed
--- a/src/HOL/Library/Extended_Nat.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Extended_Nat.thy Tue Sep 01 22:32:58 2015 +0200
@@ -112,8 +112,8 @@
by (simp add: zero_enat_def)
lemma zero_one_enat_neq [simp]:
- "\<not> 0 = (1\<Colon>enat)"
- "\<not> 1 = (0\<Colon>enat)"
+ "\<not> 0 = (1::enat)"
+ "\<not> 1 = (0::enat)"
unfolding zero_enat_def one_enat_def by simp_all
lemma infinity_ne_i1 [simp]: "(\<infinity>::enat) \<noteq> 1"
@@ -380,14 +380,14 @@
a generalize linordered_semidom to a new class that includes enat? *)
lemma enat_ord_number [simp]:
- "(numeral m \<Colon> enat) \<le> numeral n \<longleftrightarrow> (numeral m \<Colon> nat) \<le> numeral n"
- "(numeral m \<Colon> enat) < numeral n \<longleftrightarrow> (numeral m \<Colon> nat) < numeral n"
+ "(numeral m :: enat) \<le> numeral n \<longleftrightarrow> (numeral m :: nat) \<le> numeral n"
+ "(numeral m :: enat) < numeral n \<longleftrightarrow> (numeral m :: nat) < numeral n"
by (simp_all add: numeral_eq_enat)
-lemma i0_lb [simp]: "(0\<Colon>enat) \<le> n"
+lemma i0_lb [simp]: "(0::enat) \<le> n"
by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
-lemma ile0_eq [simp]: "n \<le> (0\<Colon>enat) \<longleftrightarrow> n = 0"
+lemma ile0_eq [simp]: "n \<le> (0::enat) \<longleftrightarrow> n = 0"
by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
lemma infinity_ileE [elim!]: "\<infinity> \<le> enat m \<Longrightarrow> R"
@@ -396,10 +396,10 @@
lemma infinity_ilessE [elim!]: "\<infinity> < enat m \<Longrightarrow> R"
by simp
-lemma not_iless0 [simp]: "\<not> n < (0\<Colon>enat)"
+lemma not_iless0 [simp]: "\<not> n < (0::enat)"
by (simp add: zero_enat_def less_enat_def split: enat.splits)
-lemma i0_less [simp]: "(0\<Colon>enat) < n \<longleftrightarrow> n \<noteq> 0"
+lemma i0_less [simp]: "(0::enat) < n \<longleftrightarrow> n \<noteq> 0"
by (simp add: zero_enat_def less_enat_def split: enat.splits)
lemma eSuc_ile_mono [simp]: "eSuc n \<le> eSuc m \<longleftrightarrow> n \<le> m"
@@ -623,7 +623,7 @@
instance enat :: wellorder
proof
fix P and n
- assume hyp: "(\<And>n\<Colon>enat. (\<And>m\<Colon>enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
+ assume hyp: "(\<And>n::enat. (\<And>m::enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
show "P n" by (blast intro: enat_less_induct hyp)
qed
--- a/src/HOL/Library/Fraction_Field.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Fraction_Field.thy Tue Sep 01 22:32:58 2015 +0200
@@ -416,10 +416,10 @@
eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split)
definition inf_fract_def:
- "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
+ "(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
definition sup_fract_def:
- "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
+ "(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
instance
by intro_classes
--- a/src/HOL/Library/List_lexord.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/List_lexord.thy Tue Sep 01 22:32:58 2015 +0200
@@ -69,9 +69,9 @@
instantiation list :: (linorder) distrib_lattice
begin
-definition "(inf \<Colon> 'a list \<Rightarrow> _) = min"
+definition "(inf :: 'a list \<Rightarrow> _) = min"
-definition "(sup \<Colon> 'a list \<Rightarrow> _) = max"
+definition "(sup :: 'a list \<Rightarrow> _) = max"
instance
by standard (auto simp add: inf_list_def sup_list_def max_min_distrib2)
@@ -107,15 +107,15 @@
end
lemma less_list_code [code]:
- "xs < ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
- "[] < (x\<Colon>'a\<Colon>{equal, order}) # xs \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{equal, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
+ "xs < ([]::'a::{equal, order} list) \<longleftrightarrow> False"
+ "[] < (x::'a::{equal, order}) # xs \<longleftrightarrow> True"
+ "(x::'a::{equal, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
by simp_all
lemma less_eq_list_code [code]:
- "x # xs \<le> ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
- "[] \<le> (xs\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{equal, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
+ "x # xs \<le> ([]::'a::{equal, order} list) \<longleftrightarrow> False"
+ "[] \<le> (xs::'a::{equal, order} list) \<longleftrightarrow> True"
+ "(x::'a::{equal, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
by simp_all
end
--- a/src/HOL/Library/Mapping.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Mapping.thy Tue Sep 01 22:32:58 2015 +0200
@@ -138,7 +138,7 @@
subsection \<open>Derived operations\<close>
-definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list"
+definition ordered_keys :: "('a::linorder, 'b) mapping \<Rightarrow> 'a list"
where
"ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])"
--- a/src/HOL/Library/Multiset.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Multiset.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1205,7 +1205,7 @@
end
-lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + \<Colon> 'a multiset \<Rightarrow> _ \<Rightarrow> _)"
+lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + :: 'a multiset \<Rightarrow> _ \<Rightarrow> _)"
by standard (simp add: add_ac comp_def)
declare comp_fun_commute.fold_mset_insert[OF comp_fun_commute_plus_mset, simp]
@@ -1240,7 +1240,7 @@
end
lemma msetsum_diff:
- fixes M N :: "('a \<Colon> ordered_cancel_comm_monoid_diff) multiset"
+ fixes M N :: "('a :: ordered_cancel_comm_monoid_diff) multiset"
shows "N \<le># M \<Longrightarrow> msetsum (M - N) = msetsum M - msetsum N"
by (metis add_diff_cancel_right' msetsum.union subset_mset.diff_add)
@@ -1756,10 +1756,10 @@
subsubsection \<open>Partial-order properties\<close>
-definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#<#" 50) where
+definition less_multiset :: "'a::order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#<#" 50) where
"M' #<# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
-definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#<=#" 50) where
+definition le_multiset :: "'a::order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#<=#" 50) where
"M' #<=# M \<longleftrightarrow> M' #<# M \<or> M' = M"
notation (xsymbols) less_multiset (infix "#\<subset>#" 50)
@@ -2210,7 +2210,7 @@
text \<open>Quickcheck generators\<close>
definition (in term_syntax)
- msetify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
+ msetify :: "'a::typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
\<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
[code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\<cdot>} xs"
--- a/src/HOL/Library/Multiset_Order.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Multiset_Order.thy Tue Sep 01 22:32:58 2015 +0200
@@ -178,10 +178,10 @@
context linorder
begin
-lemma total_le: "total {(a \<Colon> 'a, b). a \<le> b}"
+lemma total_le: "total {(a :: 'a, b). a \<le> b}"
unfolding total_on_def by auto
-lemma total_less: "total {(a \<Colon> 'a, b). a < b}"
+lemma total_less: "total {(a :: 'a, b). a < b}"
unfolding total_on_def by auto
lemma linorder_mult: "class.linorder
@@ -205,78 +205,78 @@
lemmas less_multiset\<^sub>H\<^sub>O = mult\<^sub>H\<^sub>O[folded less_multiset_def]
lemma le_multiset\<^sub>H\<^sub>O:
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows "M #\<subseteq># N \<longleftrightarrow> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))"
by (auto simp: le_multiset_def less_multiset\<^sub>H\<^sub>O)
-lemma wf_less_multiset: "wf {(M \<Colon> ('a \<Colon> wellorder) multiset, N). M #\<subset># N}"
+lemma wf_less_multiset: "wf {(M :: ('a :: wellorder) multiset, N). M #\<subset># N}"
unfolding less_multiset_def by (auto intro: wf_mult wf)
lemma order_multiset: "class.order
- (le_multiset :: ('a \<Colon> order) multiset \<Rightarrow> ('a \<Colon> order) multiset \<Rightarrow> bool)
- (less_multiset :: ('a \<Colon> order) multiset \<Rightarrow> ('a \<Colon> order) multiset \<Rightarrow> bool)"
+ (le_multiset :: ('a :: order) multiset \<Rightarrow> ('a :: order) multiset \<Rightarrow> bool)
+ (less_multiset :: ('a :: order) multiset \<Rightarrow> ('a :: order) multiset \<Rightarrow> bool)"
by unfold_locales
lemma linorder_multiset: "class.linorder
- (le_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool)
- (less_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool)"
+ (le_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool)
+ (less_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool)"
by unfold_locales (fastforce simp add: less_multiset\<^sub>H\<^sub>O le_multiset_def not_less_iff_gr_or_eq)
interpretation multiset_linorder: linorder
- "le_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool"
- "less_multiset :: ('a \<Colon> linorder) multiset \<Rightarrow> ('a \<Colon> linorder) multiset \<Rightarrow> bool"
+ "le_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool"
+ "less_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool"
by (rule linorder_multiset)
interpretation multiset_wellorder: wellorder
- "le_multiset :: ('a \<Colon> wellorder) multiset \<Rightarrow> ('a \<Colon> wellorder) multiset \<Rightarrow> bool"
- "less_multiset :: ('a \<Colon> wellorder) multiset \<Rightarrow> ('a \<Colon> wellorder) multiset \<Rightarrow> bool"
+ "le_multiset :: ('a :: wellorder) multiset \<Rightarrow> ('a :: wellorder) multiset \<Rightarrow> bool"
+ "less_multiset :: ('a :: wellorder) multiset \<Rightarrow> ('a :: wellorder) multiset \<Rightarrow> bool"
by unfold_locales (blast intro: wf_less_multiset[unfolded wf_def, rule_format])
lemma le_multiset_total:
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows "\<not> M #\<subseteq># N \<Longrightarrow> N #\<subseteq># M"
by (metis multiset_linorder.le_cases)
lemma less_eq_imp_le_multiset:
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows "M \<le># N \<Longrightarrow> M #\<subseteq># N"
unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O
by (simp add: less_le_not_le subseteq_mset_def)
lemma less_multiset_right_total:
- fixes M :: "('a \<Colon> linorder) multiset"
+ fixes M :: "('a :: linorder) multiset"
shows "M #\<subset># M + {#undefined#}"
unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O by simp
lemma le_multiset_empty_left[simp]:
- fixes M :: "('a \<Colon> linorder) multiset"
+ fixes M :: "('a :: linorder) multiset"
shows "{#} #\<subseteq># M"
by (simp add: less_eq_imp_le_multiset)
lemma le_multiset_empty_right[simp]:
- fixes M :: "('a \<Colon> linorder) multiset"
+ fixes M :: "('a :: linorder) multiset"
shows "M \<noteq> {#} \<Longrightarrow> \<not> M #\<subseteq># {#}"
by (metis le_multiset_empty_left multiset_order.antisym)
lemma less_multiset_empty_left[simp]:
- fixes M :: "('a \<Colon> linorder) multiset"
+ fixes M :: "('a :: linorder) multiset"
shows "M \<noteq> {#} \<Longrightarrow> {#} #\<subset># M"
by (simp add: less_multiset\<^sub>H\<^sub>O)
lemma less_multiset_empty_right[simp]:
- fixes M :: "('a \<Colon> linorder) multiset"
+ fixes M :: "('a :: linorder) multiset"
shows "\<not> M #\<subset># {#}"
using le_empty less_multiset\<^sub>D\<^sub>M by blast
lemma
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows
le_multiset_plus_left[simp]: "N #\<subseteq># (M + N)" and
le_multiset_plus_right[simp]: "M #\<subseteq># (M + N)"
using [[metis_verbose = false]] by (metis less_eq_imp_le_multiset mset_le_add_left add.commute)+
lemma
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows
less_multiset_plus_plus_left_iff[simp]: "M + N #\<subset># M' + N \<longleftrightarrow> M #\<subset># M'" and
less_multiset_plus_plus_right_iff[simp]: "M + N #\<subset># M + N' \<longleftrightarrow> N #\<subset># N'"
@@ -286,7 +286,7 @@
by (metis add.commute add_diff_cancel_right' monoid_add_class.add.left_neutral)
lemma
- fixes M N :: "('a \<Colon> linorder) multiset"
+ fixes M N :: "('a :: linorder) multiset"
shows
less_multiset_plus_left_nonempty[simp]: "M \<noteq> {#} \<Longrightarrow> N #\<subset># M + N" and
less_multiset_plus_right_nonempty[simp]: "N \<noteq> {#} \<Longrightarrow> M #\<subset># M + N"
@@ -294,11 +294,11 @@
by (metis add.right_neutral less_multiset_empty_left less_multiset_plus_plus_right_iff
add.commute)+
-lemma ex_gt_imp_less_multiset: "(\<exists>y \<Colon> 'a \<Colon> linorder. y \<in># N \<and> (\<forall>x. x \<in># M \<longrightarrow> x < y)) \<Longrightarrow> M #\<subset># N"
+lemma ex_gt_imp_less_multiset: "(\<exists>y :: 'a :: linorder. y \<in># N \<and> (\<forall>x. x \<in># M \<longrightarrow> x < y)) \<Longrightarrow> M #\<subset># N"
unfolding less_multiset\<^sub>H\<^sub>O by (metis less_irrefl less_nat_zero_code not_gr0)
lemma ex_gt_count_imp_less_multiset:
- "(\<forall>y \<Colon> 'a \<Colon> linorder. y \<in># M + N \<longrightarrow> y \<le> x) \<Longrightarrow> count M x < count N x \<Longrightarrow> M #\<subset># N"
+ "(\<forall>y :: 'a :: linorder. y \<in># M + N \<longrightarrow> y \<le> x) \<Longrightarrow> count M x < count N x \<Longrightarrow> M #\<subset># N"
unfolding less_multiset\<^sub>H\<^sub>O by (metis add.left_neutral add_lessD1 dual_order.strict_iff_order
less_not_sym mset_leD mset_le_add_left)
--- a/src/HOL/Library/RBT.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/RBT.thy Tue Sep 01 22:32:58 2015 +0200
@@ -10,7 +10,7 @@
subsection \<open>Type definition\<close>
-typedef ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
+typedef ('a, 'b) rbt = "{t :: ('a::linorder, 'b) RBT_Impl.rbt. is_rbt t}"
morphisms impl_of RBT
proof -
have "RBT_Impl.Empty \<in> ?rbt" by simp
@@ -37,32 +37,32 @@
setup_lifting type_definition_rbt
-lift_definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "rbt_lookup" .
+lift_definition lookup :: "('a::linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "rbt_lookup" .
-lift_definition empty :: "('a\<Colon>linorder, 'b) rbt" is RBT_Impl.Empty
+lift_definition empty :: "('a::linorder, 'b) rbt" is RBT_Impl.Empty
by (simp add: empty_def)
-lift_definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_insert"
+lift_definition insert :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_insert"
by simp
-lift_definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_delete"
+lift_definition delete :: "'a::linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_delete"
by simp
-lift_definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" is RBT_Impl.entries .
+lift_definition entries :: "('a::linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" is RBT_Impl.entries .
-lift_definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" is RBT_Impl.keys .
+lift_definition keys :: "('a::linorder, 'b) rbt \<Rightarrow> 'a list" is RBT_Impl.keys .
-lift_definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" is "rbt_bulkload" ..
+lift_definition bulkload :: "('a::linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" is "rbt_bulkload" ..
-lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is rbt_map_entry
+lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a::linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is rbt_map_entry
by simp
-lift_definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'c) rbt" is RBT_Impl.map
+lift_definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::linorder, 'b) rbt \<Rightarrow> ('a, 'c) rbt" is RBT_Impl.map
by simp
-lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" is RBT_Impl.fold .
+lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a::linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" is RBT_Impl.fold .
-lift_definition union :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_union"
+lift_definition union :: "('a::linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" is "rbt_union"
by (simp add: rbt_union_is_rbt)
lift_definition foldi :: "('c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a :: linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c"
@@ -70,7 +70,7 @@
subsection \<open>Derived operations\<close>
-definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
+definition is_empty :: "('a::linorder, 'b) rbt \<Rightarrow> bool" where
[code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"
--- a/src/HOL/Library/RBT_Impl.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/RBT_Impl.thy Tue Sep 01 22:32:58 2015 +0200
@@ -2060,19 +2060,19 @@
(@{const_name rbt_sorted}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> bool"}),
(@{const_name rbt_lookup}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b"}),
(@{const_name is_rbt}, SOME @{typ "('a :: linorder, 'b) rbt \<Rightarrow> bool"}),
- (@{const_name rbt_ins}, SOME @{typ "('a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_insert_with_key}, SOME @{typ "('a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_ins}, SOME @{typ "('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_insert_with_key}, SOME @{typ "('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
(@{const_name rbt_insert_with}, SOME @{typ "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
(@{const_name rbt_insert}, SOME @{typ "('a :: linorder) \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_del_from_left}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_del_from_right}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_del}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_delete}, SOME @{typ "('a\<Colon>linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_union_with_key}, SOME @{typ "('a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_union_with}, SOME @{typ "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_union}, SOME @{typ "('a\<Colon>linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_map_entry}, SOME @{typ "'a\<Colon>linorder \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
- (@{const_name rbt_bulkload}, SOME @{typ "('a \<times> 'b) list \<Rightarrow> ('a\<Colon>linorder,'b) rbt"})]
+ (@{const_name rbt_del_from_left}, SOME @{typ "('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_del_from_right}, SOME @{typ "('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_del}, SOME @{typ "('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_delete}, SOME @{typ "('a::linorder) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_union_with_key}, SOME @{typ "('a::linorder \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_union_with}, SOME @{typ "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a::linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_union}, SOME @{typ "('a::linorder,'b) rbt \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_map_entry}, SOME @{typ "'a::linorder \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) rbt \<Rightarrow> ('a,'b) rbt"}),
+ (@{const_name rbt_bulkload}, SOME @{typ "('a \<times> 'b) list \<Rightarrow> ('a::linorder,'b) rbt"})]
\<close>
hide_const (open) R B Empty entries keys fold gen_keys gen_entries
--- a/src/HOL/Library/RBT_Mapping.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/RBT_Mapping.thy Tue Sep 01 22:32:58 2015 +0200
@@ -12,7 +12,7 @@
subsection \<open>Implementation of mappings\<close>
context includes rbt.lifting begin
-lift_definition Mapping :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" is RBT.lookup .
+lift_definition Mapping :: "('a::linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" is RBT.lookup .
end
code_datatype Mapping
--- a/src/HOL/Library/RBT_Set.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/RBT_Set.thy Tue Sep 01 22:32:58 2015 +0200
@@ -18,10 +18,10 @@
section \<open>Definition of code datatype constructors\<close>
-definition Set :: "('a\<Colon>linorder, unit) rbt \<Rightarrow> 'a set"
+definition Set :: "('a::linorder, unit) rbt \<Rightarrow> 'a set"
where "Set t = {x . RBT.lookup t x = Some ()}"
-definition Coset :: "('a\<Colon>linorder, unit) rbt \<Rightarrow> 'a set"
+definition Coset :: "('a::linorder, unit) rbt \<Rightarrow> 'a set"
where [simp]: "Coset t = - Set t"
--- a/src/HOL/Library/Sublist.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Library/Sublist.thy Tue Sep 01 22:32:58 2015 +0200
@@ -111,7 +111,7 @@
assume a1: "\<exists>zs. ys = xs @ zs"
then obtain sk :: "'a list" where sk: "ys = xs @ sk" by fastforce
assume a2: "length xs < length ys"
- have f1: "\<And>v. ([]\<Colon>'a list) @ v = v" using append_Nil2 by simp
+ have f1: "\<And>v. ([]::'a list) @ v = v" using append_Nil2 by simp
have "[] \<noteq> sk" using a1 a2 sk less_not_refl by force
hence "\<exists>v. xs @ hd sk # v = ys" using sk by (metis hd_Cons_tl)
thus "\<exists>zs. ys = (xs @ [ys ! length xs]) @ zs" using f1 by fastforce
--- a/src/HOL/Limits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Limits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -22,7 +22,7 @@
(auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b])
lemma at_infinity_eq_at_top_bot:
- "(at_infinity \<Colon> real filter) = sup at_top at_bot"
+ "(at_infinity :: real filter) = sup at_top at_bot"
apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity
eventually_at_top_linorder eventually_at_bot_linorder)
apply safe
@@ -872,7 +872,7 @@
using assms unfolding continuous_on_def by (fast intro: tendsto_sgn)
lemma filterlim_at_infinity:
- fixes f :: "_ \<Rightarrow> 'a\<Colon>real_normed_vector"
+ fixes f :: "_ \<Rightarrow> 'a::real_normed_vector"
assumes "0 \<le> c"
shows "(LIM x F. f x :> at_infinity) \<longleftrightarrow> (\<forall>r>c. eventually (\<lambda>x. r \<le> norm (f x)) F)"
unfolding filterlim_iff eventually_at_infinity
@@ -983,7 +983,7 @@
qed
lemma tendsto_inverse_0:
- fixes x :: "_ \<Rightarrow> 'a\<Colon>real_normed_div_algebra"
+ fixes x :: "_ \<Rightarrow> 'a::real_normed_div_algebra"
shows "(inverse ---> (0::'a)) at_infinity"
unfolding tendsto_Zfun_iff diff_0_right Zfun_def eventually_at_infinity
proof safe
@@ -1041,7 +1041,7 @@
unfolding filterlim_def at_top_to_right filtermap_filtermap ..
lemma filterlim_inverse_at_infinity:
- fixes x :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring}"
+ fixes x :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
shows "filterlim inverse at_infinity (at (0::'a))"
unfolding filterlim_at_infinity[OF order_refl]
proof safe
@@ -1053,7 +1053,7 @@
qed
lemma filterlim_inverse_at_iff:
- fixes g :: "'a \<Rightarrow> 'b\<Colon>{real_normed_div_algebra, division_ring}"
+ fixes g :: "'a \<Rightarrow> 'b::{real_normed_div_algebra, division_ring}"
shows "(LIM x F. inverse (g x) :> at 0) \<longleftrightarrow> (LIM x F. g x :> at_infinity)"
unfolding filterlim_def filtermap_filtermap[symmetric]
proof
@@ -1078,7 +1078,7 @@
lemma at_to_infinity:
- fixes x :: "'a \<Colon> {real_normed_field,field}"
+ fixes x :: "'a :: {real_normed_field,field}"
shows "(at (0::'a)) = filtermap inverse at_infinity"
proof (rule antisym)
have "(inverse ---> (0::'a)) at_infinity"
@@ -1226,7 +1226,7 @@
qed
lemma tendsto_divide_0:
- fixes f :: "_ \<Rightarrow> 'a\<Colon>{real_normed_div_algebra, division_ring}"
+ fixes f :: "_ \<Rightarrow> 'a::{real_normed_div_algebra, division_ring}"
assumes f: "(f ---> c) F"
assumes g: "LIM x F. g x :> at_infinity"
shows "((\<lambda>x. f x / g x) ---> 0) F"
--- a/src/HOL/List.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/List.thy Tue Sep 01 22:32:58 2015 +0200
@@ -6280,19 +6280,19 @@
by auto
lemma all_nat_less_eq [code_unfold]:
- "(\<forall>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
+ "(\<forall>m<n::nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
by auto
lemma ex_nat_less_eq [code_unfold]:
- "(\<exists>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
+ "(\<exists>m<n::nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
by auto
lemma all_nat_less [code_unfold]:
- "(\<forall>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
+ "(\<forall>m\<le>n::nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
by auto
lemma ex_nat_less [code_unfold]:
- "(\<exists>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
+ "(\<exists>m\<le>n::nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
by auto
text\<open>Bounded @{text LEAST} operator:\<close>
--- a/src/HOL/MacLaurin.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/MacLaurin.thy Tue Sep 01 22:32:58 2015 +0200
@@ -83,7 +83,7 @@
by (cases n) (simp add: n)
obtain B where f_h: "f h =
- (\<Sum>m<n. diff m (0\<Colon>real) / (fact m) * h ^ m) + B * (h ^ n / (fact n))"
+ (\<Sum>m<n. diff m (0::real) / (fact m) * h ^ m) + B * (h ^ n / (fact n))"
using Maclaurin_lemma [OF h] ..
def g \<equiv> "(\<lambda>t. f t -
@@ -99,8 +99,8 @@
have difg_0: "difg 0 = g"
unfolding difg_def g_def by (simp add: diff_0)
- have difg_Suc: "\<forall>(m\<Colon>nat) t\<Colon>real.
- m < n \<and> (0\<Colon>real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
+ have difg_Suc: "\<forall>(m::nat) t::real.
+ m < n \<and> (0::real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2)
have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
@@ -127,9 +127,9 @@
proof (rule Rolle)
show "0 < h" by fact
show "difg 0 0 = difg 0 h" by (simp add: difg_0 g2)
- show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0\<Colon>nat)) x"
+ show "\<forall>x. 0 \<le> x \<and> x \<le> h \<longrightarrow> isCont (difg (0::nat)) x"
by (simp add: isCont_difg n)
- show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0\<Colon>nat) differentiable (at x)"
+ show "\<forall>x. 0 < x \<and> x < h \<longrightarrow> difg (0::nat) differentiable (at x)"
by (simp add: differentiable_difg n)
qed
next
--- a/src/HOL/Matrix_LP/Matrix.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Matrix_LP/Matrix.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1290,7 +1290,7 @@
le_matrix_def: "A \<le> B \<longleftrightarrow> (\<forall>j i. Rep_matrix A j i \<le> Rep_matrix B j i)"
definition
- less_def: "A < (B\<Colon>'a matrix) \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
+ less_def: "A < (B::'a matrix) \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
instance ..
@@ -1496,7 +1496,7 @@
begin
definition
- abs_matrix_def: "abs (A \<Colon> 'a matrix) = sup A (- A)"
+ abs_matrix_def: "abs (A :: 'a matrix) = sup A (- A)"
instance ..
--- a/src/HOL/Matrix_LP/SparseMatrix.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Matrix_LP/SparseMatrix.thy Tue Sep 01 22:32:58 2015 +0200
@@ -834,14 +834,14 @@
lemma bool1: "(\<not> True) = False" by blast
lemma bool2: "(\<not> False) = True" by blast
-lemma bool3: "((P\<Colon>bool) \<and> True) = P" by blast
-lemma bool4: "(True \<and> (P\<Colon>bool)) = P" by blast
-lemma bool5: "((P\<Colon>bool) \<and> False) = False" by blast
-lemma bool6: "(False \<and> (P\<Colon>bool)) = False" by blast
-lemma bool7: "((P\<Colon>bool) \<or> True) = True" by blast
-lemma bool8: "(True \<or> (P\<Colon>bool)) = True" by blast
-lemma bool9: "((P\<Colon>bool) \<or> False) = P" by blast
-lemma bool10: "(False \<or> (P\<Colon>bool)) = P" by blast
+lemma bool3: "((P::bool) \<and> True) = P" by blast
+lemma bool4: "(True \<and> (P::bool)) = P" by blast
+lemma bool5: "((P::bool) \<and> False) = False" by blast
+lemma bool6: "(False \<and> (P::bool)) = False" by blast
+lemma bool7: "((P::bool) \<or> True) = True" by blast
+lemma bool8: "(True \<or> (P::bool)) = True" by blast
+lemma bool9: "((P::bool) \<or> False) = P" by blast
+lemma bool10: "(False \<or> (P::bool)) = P" by blast
lemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10
lemma if_case_eq: "(if b then x else y) = (case b of True => x | False => y)" by simp
--- a/src/HOL/Meson.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Meson.thy Tue Sep 01 22:32:58 2015 +0200
@@ -182,7 +182,7 @@
definition skolem :: "'a \<Rightarrow> 'a" where
"skolem = (\<lambda>x. x)"
-lemma skolem_COMBK_iff: "P \<longleftrightarrow> skolem (COMBK P (i\<Colon>nat))"
+lemma skolem_COMBK_iff: "P \<longleftrightarrow> skolem (COMBK P (i::nat))"
unfolding skolem_def COMBK_def by (rule refl)
lemmas skolem_COMBK_I = iffD1 [OF skolem_COMBK_iff]
--- a/src/HOL/Metis_Examples/Big_O.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy Tue Sep 01 22:32:58 2015 +0200
@@ -16,11 +16,11 @@
subsection {* Definitions *}
-definition bigo :: "('a => 'b\<Colon>linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where
- "O(f\<Colon>('a => 'b)) == {h. \<exists>c. \<forall>x. abs (h x) <= c * abs (f x)}"
+definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where
+ "O(f::('a => 'b)) == {h. \<exists>c. \<forall>x. abs (h x) <= c * abs (f x)}"
lemma bigo_pos_const:
- "(\<exists>c\<Colon>'a\<Colon>linordered_idom.
+ "(\<exists>c::'a::linordered_idom.
\<forall>x. abs (h x) \<le> c * abs (f x))
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. abs(h x) \<le> c * abs (f x)))"
by (metis (no_types) abs_ge_zero
@@ -32,7 +32,7 @@
sledgehammer_params [isar_proofs, compress = 1]
lemma
- "(\<exists>c\<Colon>'a\<Colon>linordered_idom.
+ "(\<exists>c::'a::linordered_idom.
\<forall>x. abs (h x) \<le> c * abs (f x))
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. abs(h x) \<le> c * abs (f x)))"
apply auto
@@ -42,19 +42,19 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> \<bar>x\<^sub>1\<bar>" by (metis abs_ge_zero)
- have F2: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 0 \<le> \<bar>x\<^sub>1\<bar>" by (metis abs_ge_zero)
+ have F2: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
have F3: "\<forall>x\<^sub>1 x\<^sub>3. x\<^sub>3 \<le> \<bar>h x\<^sub>1\<bar> \<longrightarrow> x\<^sub>3 \<le> c * \<bar>f x\<^sub>1\<bar>" by (metis A1 order_trans)
- have F4: "\<forall>x\<^sub>2 x\<^sub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
+ have F4: "\<forall>x\<^sub>2 x\<^sub>3::'a::linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
by (metis abs_mult)
- have F5: "\<forall>x\<^sub>3 x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1"
+ have F5: "\<forall>x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1"
by (metis abs_mult_pos)
- hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = \<bar>1\<bar> * x\<^sub>1" by (metis F2)
- hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F2 abs_one)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = \<bar>1\<bar> * x\<^sub>1" by (metis F2)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F2 abs_one)
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>h x\<^sub>3\<bar> \<longrightarrow> \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F3)
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F1)
- hence "\<forall>x\<^sub>3. (0\<Colon>'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F5)
- hence "\<forall>x\<^sub>3. (0\<Colon>'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F4)
+ hence "\<forall>x\<^sub>3. (0::'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F5)
+ hence "\<forall>x\<^sub>3. (0::'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F4)
hence "\<forall>x\<^sub>3. c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F1)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F4)
@@ -63,7 +63,7 @@
sledgehammer_params [isar_proofs, compress = 2]
lemma
- "(\<exists>c\<Colon>'a\<Colon>linordered_idom.
+ "(\<exists>c::'a::linordered_idom.
\<forall>x. abs (h x) \<le> c * abs (f x))
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. abs(h x) \<le> c * abs (f x)))"
apply auto
@@ -73,10 +73,10 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
- have F2: "\<forall>x\<^sub>2 x\<^sub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
+ have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F2: "\<forall>x\<^sub>2 x\<^sub>3::'a::linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
by (metis abs_mult)
- have "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_mult_pos abs_one)
+ have "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_mult_pos abs_one)
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis A1 abs_ge_zero order_trans)
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F2 abs_mult_pos)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero)
@@ -86,7 +86,7 @@
sledgehammer_params [isar_proofs, compress = 3]
lemma
- "(\<exists>c\<Colon>'a\<Colon>linordered_idom.
+ "(\<exists>c::'a::linordered_idom.
\<forall>x. abs (h x) \<le> c * abs (f x))
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. abs(h x) \<le> c * abs (f x)))"
apply auto
@@ -96,9 +96,9 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
- have F2: "\<forall>x\<^sub>3 x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1" by (metis abs_mult_pos)
- hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_one)
+ have F1: "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F2: "\<forall>x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1" by (metis abs_mult_pos)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1::'a::linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_one)
hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F2 A1 abs_ge_zero order_trans)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis A1 abs_ge_zero)
qed
@@ -106,7 +106,7 @@
sledgehammer_params [isar_proofs, compress = 4]
lemma
- "(\<exists>c\<Colon>'a\<Colon>linordered_idom.
+ "(\<exists>c::'a::linordered_idom.
\<forall>x. abs (h x) \<le> c * abs (f x))
\<longleftrightarrow> (\<exists>c. 0 < c \<and> (\<forall>x. abs(h x) \<le> c * abs (f x)))"
apply auto
@@ -116,7 +116,7 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have "\<forall>x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>"
by (metis A1 abs_ge_zero order_trans abs_mult_pos abs_one)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero abs_mult_pos abs_mult)
@@ -321,7 +321,7 @@
by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib)
lemma bigo_mult5: "\<forall>x. f x ~= 0 \<Longrightarrow>
- O(f * g) <= (f\<Colon>'a => ('b\<Colon>linordered_field)) *o O(g)"
+ O(f * g) <= (f::'a => ('b::linordered_field)) *o O(g)"
proof -
assume a: "\<forall>x. f x ~= 0"
show "O(f * g) <= f *o O(g)"
@@ -334,31 +334,31 @@
by (rule bigo_mult2)
also have "(\<lambda>x. 1 / f x) * (f * g) = g"
by (simp add: fun_eq_iff a)
- finally have "(\<lambda>x. (1\<Colon>'b) / f x) * h : O(g)".
- then have "f * ((\<lambda>x. (1\<Colon>'b) / f x) * h) : f *o O(g)"
+ finally have "(\<lambda>x. (1::'b) / f x) * h : O(g)".
+ then have "f * ((\<lambda>x. (1::'b) / f x) * h) : f *o O(g)"
by auto
- also have "f * ((\<lambda>x. (1\<Colon>'b) / f x) * h) = h"
+ also have "f * ((\<lambda>x. (1::'b) / f x) * h) = h"
by (simp add: func_times fun_eq_iff a)
finally show "h : f *o O(g)".
qed
qed
lemma bigo_mult6:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = (f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) *o O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = (f::'a \<Rightarrow> ('b::linordered_field)) *o O(g)"
by (metis bigo_mult2 bigo_mult5 order_antisym)
(*proof requires relaxing relevance: 2007-01-25*)
declare bigo_mult6 [simp]
lemma bigo_mult7:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f::'a \<Rightarrow> ('b::linordered_field)) * O(g)"
by (metis bigo_refl bigo_mult6 set_times_mono3)
declare bigo_mult6 [simp del]
declare bigo_mult7 [intro!]
lemma bigo_mult8:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f::'a \<Rightarrow> ('b::linordered_field)) * O(g)"
by (metis bigo_mult bigo_mult7 order_antisym_conv)
lemma bigo_minus [intro]: "f : O(g) \<Longrightarrow> - f : O(g)"
@@ -397,14 +397,14 @@
lemma bigo_const2 [intro]: "O(\<lambda>x. c) \<le> O(\<lambda>x. 1)"
by (metis bigo_const1 bigo_elt_subset)
-lemma bigo_const3: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow> (\<lambda>x. 1) : O(\<lambda>x. c)"
+lemma bigo_const3: "(c::'a::linordered_field) ~= 0 \<Longrightarrow> (\<lambda>x. 1) : O(\<lambda>x. c)"
apply (simp add: bigo_def)
by (metis abs_eq_0 left_inverse order_refl)
-lemma bigo_const4: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow> O(\<lambda>x. 1) <= O(\<lambda>x. c)"
+lemma bigo_const4: "(c::'a::linordered_field) ~= 0 \<Longrightarrow> O(\<lambda>x. 1) <= O(\<lambda>x. c)"
by (metis bigo_elt_subset bigo_const3)
-lemma bigo_const [simp]: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow>
+lemma bigo_const [simp]: "(c::'a::linordered_field) ~= 0 \<Longrightarrow>
O(\<lambda>x. c) = O(\<lambda>x. 1)"
by (metis bigo_const2 bigo_const4 equalityI)
@@ -415,19 +415,19 @@
lemma bigo_const_mult2: "O(\<lambda>x. c * f x) \<le> O(f)"
by (rule bigo_elt_subset, rule bigo_const_mult1)
-lemma bigo_const_mult3: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow> f : O(\<lambda>x. c * f x)"
+lemma bigo_const_mult3: "(c::'a::linordered_field) ~= 0 \<Longrightarrow> f : O(\<lambda>x. c * f x)"
apply (simp add: bigo_def)
by (metis (no_types) abs_mult mult.assoc mult_1 order_refl left_inverse)
lemma bigo_const_mult4:
-"(c\<Colon>'a\<Colon>linordered_field) \<noteq> 0 \<Longrightarrow> O(f) \<le> O(\<lambda>x. c * f x)"
+"(c::'a::linordered_field) \<noteq> 0 \<Longrightarrow> O(f) \<le> O(\<lambda>x. c * f x)"
by (metis bigo_elt_subset bigo_const_mult3)
-lemma bigo_const_mult [simp]: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow>
+lemma bigo_const_mult [simp]: "(c::'a::linordered_field) ~= 0 \<Longrightarrow>
O(\<lambda>x. c * f x) = O(f)"
by (metis equalityI bigo_const_mult2 bigo_const_mult4)
-lemma bigo_const_mult5 [simp]: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow>
+lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) ~= 0 \<Longrightarrow>
(\<lambda>x. c) *o O(f) = O(f)"
apply (auto del: subsetI)
apply (rule order_trans)
@@ -444,20 +444,20 @@
(* couldn't get this proof without the step above *)
proof -
fix g :: "'b \<Rightarrow> 'a" and d :: 'a
- assume A1: "c \<noteq> (0\<Colon>'a)"
- assume A2: "\<forall>x\<Colon>'b. \<bar>g x\<bar> \<le> d * \<bar>f x\<bar>"
+ assume A1: "c \<noteq> (0::'a)"
+ assume A2: "\<forall>x::'b. \<bar>g x\<bar> \<le> d * \<bar>f x\<bar>"
have F1: "inverse \<bar>c\<bar> = \<bar>inverse c\<bar>" using A1 by (metis nonzero_abs_inverse)
- have F2: "(0\<Colon>'a) < \<bar>c\<bar>" using A1 by (metis zero_less_abs_iff)
- have "(0\<Colon>'a) < \<bar>c\<bar> \<longrightarrow> (0\<Colon>'a) < \<bar>inverse c\<bar>" using F1 by (metis positive_imp_inverse_positive)
- hence "(0\<Colon>'a) < \<bar>inverse c\<bar>" using F2 by metis
- hence F3: "(0\<Colon>'a) \<le> \<bar>inverse c\<bar>" by (metis order_le_less)
- have "\<exists>(u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar>"
+ have F2: "(0::'a) < \<bar>c\<bar>" using A1 by (metis zero_less_abs_iff)
+ have "(0::'a) < \<bar>c\<bar> \<longrightarrow> (0::'a) < \<bar>inverse c\<bar>" using F1 by (metis positive_imp_inverse_positive)
+ hence "(0::'a) < \<bar>inverse c\<bar>" using F2 by metis
+ hence F3: "(0::'a) \<le> \<bar>inverse c\<bar>" by (metis order_le_less)
+ have "\<exists>(u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar>"
using A2 by metis
- hence F4: "\<exists>(u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0\<Colon>'a) \<le> \<bar>inverse c\<bar>"
+ hence F4: "\<exists>(u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0::'a) \<le> \<bar>inverse c\<bar>"
using F3 by metis
- hence "\<exists>(v\<Colon>'a) (u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^sub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^sub>7 (u * v))\<bar>)"
+ hence "\<exists>(v::'a) (u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^sub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^sub>7 (u * v))\<bar>)"
by (metis mult_left_mono)
- then show "\<exists>ca\<Colon>'a. \<forall>x\<Colon>'b. inverse \<bar>c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>"
+ then show "\<exists>ca::'a. \<forall>x::'b. inverse \<bar>c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>"
using A2 F4 by (metis F1 `0 < \<bar>inverse c\<bar>` linordered_field_class.sign_simps(23) mult_le_cancel_left_pos)
qed
@@ -582,7 +582,7 @@
apply assumption+
done
-lemma bigo_useful_const_mult: "(c\<Colon>'a\<Colon>linordered_field) ~= 0 \<Longrightarrow>
+lemma bigo_useful_const_mult: "(c::'a::linordered_field) ~= 0 \<Longrightarrow>
(\<lambda>x. c) * f =o O(h) \<Longrightarrow> f =o O(h)"
apply (rule subsetD)
apply (subgoal_tac "(\<lambda>x. 1 / c) *o O(h) <= O(h)")
@@ -594,13 +594,13 @@
apply (simp add: func_times)
done
-lemma bigo_fix: "(\<lambda>x. f ((x\<Colon>nat) + 1)) =o O(\<lambda>x. h(x + 1)) \<Longrightarrow> f 0 = 0 \<Longrightarrow>
+lemma bigo_fix: "(\<lambda>x. f ((x::nat) + 1)) =o O(\<lambda>x. h(x + 1)) \<Longrightarrow> f 0 = 0 \<Longrightarrow>
f =o O(h)"
apply (simp add: bigo_alt_def)
by (metis abs_ge_zero abs_mult abs_of_pos abs_zero not0_implies_Suc)
lemma bigo_fix2:
- "(\<lambda>x. f ((x\<Colon>nat) + 1)) =o (\<lambda>x. g(x + 1)) +o O(\<lambda>x. h(x + 1)) \<Longrightarrow>
+ "(\<lambda>x. f ((x::nat) + 1)) =o (\<lambda>x. g(x + 1)) +o O(\<lambda>x. h(x + 1)) \<Longrightarrow>
f 0 = g 0 \<Longrightarrow> f =o g +o O(h)"
apply (rule set_minus_imp_plus)
apply (rule bigo_fix)
@@ -613,7 +613,7 @@
subsection {* Less than or equal to *}
-definition lesso :: "('a => 'b\<Colon>linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
+definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
"f <o g == (\<lambda>x. max (f x - g x) 0)"
lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. abs (g x) <= abs (f x) \<Longrightarrow>
@@ -691,7 +691,7 @@
by (metis abs_ge_zero linorder_linear max.absorb1 max.commute)
lemma bigo_lesso4:
- "f <o g =o O(k\<Colon>'a=>'b\<Colon>{linordered_field}) \<Longrightarrow>
+ "f <o g =o O(k::'a=>'b::{linordered_field}) \<Longrightarrow>
g =o h +o O(k) \<Longrightarrow> f <o h =o O(k)"
apply (unfold lesso_def)
apply (drule set_plus_imp_minus)
--- a/src/HOL/Metis_Examples/Clausification.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Clausification.thy Tue Sep 01 22:32:58 2015 +0200
@@ -140,7 +140,7 @@
lemma ex_tl: "EX ys. tl ys = xs"
using list.sel(3) by fast
-lemma "(\<exists>ys\<Colon>nat list. tl ys = xs) \<and> (\<exists>bs\<Colon>int list. tl bs = as)"
+lemma "(\<exists>ys::nat list. tl ys = xs) \<and> (\<exists>bs::int list. tl bs = as)"
by (metis ex_tl)
end
--- a/src/HOL/Metis_Examples/Message.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Message.thy Tue Sep 01 22:32:58 2015 +0200
@@ -707,7 +707,7 @@
have F2: "\<forall>x\<^sub>1. parts x\<^sub>1 \<union> synth (analz x\<^sub>1) = parts (synth (analz x\<^sub>1))"
by (metis parts_analz parts_synth)
have F3: "X \<in> synth (analz H)" using A1 by metis
- have "\<forall>x\<^sub>2 x\<^sub>1\<Colon>msg set. x\<^sub>1 \<le> sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3))
+ have "\<forall>x\<^sub>2 x\<^sub>1::msg set. x\<^sub>1 \<le> sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3))
hence F4: "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> analz (synth x\<^sub>1)" by (metis analz_synth)
have F5: "X \<in> synth (analz H)" using F3 by metis
have "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> synth (analz x\<^sub>1)
--- a/src/HOL/Metis_Examples/Proxies.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Proxies.thy Tue Sep 01 22:32:58 2015 +0200
@@ -19,7 +19,7 @@
text {* Extensionality and set constants *}
-lemma plus_1_not_0: "n + (1\<Colon>nat) \<noteq> 0"
+lemma plus_1_not_0: "n + (1::nat) \<noteq> 0"
by simp
definition inc :: "nat \<Rightarrow> nat" where
@@ -40,7 +40,7 @@
sledgehammer [expect = some] (add_swap_def)
by (metis_exhaust add_swap_def)
-definition "A = {xs\<Colon>'a list. True}"
+definition "A = {xs::'a list. True}"
lemma "xs \<in> A"
(* The "add:" argument is unfortunate. *)
--- a/src/HOL/Metis_Examples/Sets.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Sets.thy Tue Sep 01 22:32:58 2015 +0200
@@ -27,47 +27,47 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>1 \<union> x\<^sub>2" by (metis Un_commute Un_upper2)
- have F2a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
- have F2: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>1 \<union> x\<^sub>2" by (metis Un_commute Un_upper2)
+ have F2a: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
+ have F2: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq)
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume AA1: "Y \<union> Z \<noteq> X"
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
moreover
{ assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)
- hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1) }
+ hence "\<exists>x\<^sub>1::'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1) }
moreover
- { assume "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
+ { assume "\<exists>x\<^sub>1::'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
moreover
{ assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)
- hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by blast }
+ hence "\<exists>x\<^sub>1::'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by blast }
moreover
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
- ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
+ ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis
qed
sledgehammer_params [isar_proofs, compress = 2]
@@ -75,36 +75,36 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
- { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
+ have F1: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
+ { assume AA1: "\<exists>x\<^sub>1::'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
- hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1 Un_subset_iff) }
+ hence "\<exists>x\<^sub>1::'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1 Un_subset_iff) }
moreover
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }
moreover
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
moreover
{ assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
{ assume "\<not> Z \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
- hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
- ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }
- ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis
+ hence "\<exists>x\<^sub>1::'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }
+ ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis
qed
sledgehammer_params [isar_proofs, compress = 3]
@@ -112,16 +112,16 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
- have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq)
+ have F1a: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq)
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
moreover
- { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
+ { assume AA1: "\<exists>x\<^sub>1::'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) }
- ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2)
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) }
+ ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2)
qed
sledgehammer_params [isar_proofs, compress = 4]
@@ -129,15 +129,15 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
{ assume "\<not> Y \<subseteq> X"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
moreover
{ assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
- { assume "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z"
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
- hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) }
- ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_subset_iff Un_upper2)
+ { assume "\<exists>x\<^sub>1::'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z"
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
+ hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) }
+ ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V::'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_subset_iff Un_upper2)
qed
sledgehammer_params [isar_proofs, compress = 1]
@@ -154,8 +154,8 @@
lemma (* fixedpoint: *) "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
proof -
- assume "\<exists>!x\<Colon>'a. f (g x) = x"
- thus "\<exists>!y\<Colon>'b. g (f y) = y" by metis
+ assume "\<exists>!x::'a. f (g x) = x"
+ thus "\<exists>!y::'b. g (f y) = y" by metis
qed
lemma (* singleton_example_2: *)
@@ -172,8 +172,8 @@
assume "\<forall>x \<in> S. \<Union>S \<subseteq> x"
hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> \<Union>S \<and> x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis set_eq_subset)
hence "\<forall>x\<^sub>1. x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis Union_upper)
- hence "\<forall>x\<^sub>1\<Colon>('a set) set. \<Union>S \<in> x\<^sub>1 \<longrightarrow> S \<subseteq> x\<^sub>1" by (metis subsetI)
- hence "\<forall>x\<^sub>1\<Colon>('a set) set. S \<subseteq> insert (\<Union>S) x\<^sub>1" by (metis insert_iff)
+ hence "\<forall>x\<^sub>1::('a set) set. \<Union>S \<in> x\<^sub>1 \<longrightarrow> S \<subseteq> x\<^sub>1" by (metis subsetI)
+ hence "\<forall>x\<^sub>1::('a set) set. S \<subseteq> insert (\<Union>S) x\<^sub>1" by (metis insert_iff)
thus "\<exists>z. S \<subseteq> {z}" by metis
qed
--- a/src/HOL/Metis_Examples/Type_Encodings.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Metis_Examples/Type_Encodings.thy Tue Sep 01 22:32:58 2015 +0200
@@ -89,7 +89,7 @@
lemma "P (null xs) \<Longrightarrow> null xs \<Longrightarrow> xs = []"
by (metis_exhaust null_def)
-lemma "(0\<Colon>nat) + 0 = 0"
+lemma "(0::nat) + 0 = 0"
by (metis_exhaust add_0_left)
end
--- a/src/HOL/MicroJava/Comp/LemmasComp.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Tue Sep 01 22:32:58 2015 +0200
@@ -311,10 +311,10 @@
apply (simp add: map_of_map [symmetric])
apply (simp add: split_beta)
apply (simp add: Fun.comp_def split_beta)
- apply (subgoal_tac "(\<lambda>x\<Colon>(vname list \<times> fdecl list \<times> stmt \<times> expr) mdecl.
+ apply (subgoal_tac "(\<lambda>x::(vname list \<times> fdecl list \<times> stmt \<times> expr) mdecl.
(fst x, Object,
snd (compMethod G Object
- (inv (\<lambda>(s\<Colon>sig, m\<Colon>ty \<times> vname list \<times> fdecl list \<times> stmt \<times> expr).
+ (inv (\<lambda>(s::sig, m::ty \<times> vname list \<times> fdecl list \<times> stmt \<times> expr).
(s, Object, m))
(S, Object, snd x)))))
= (\<lambda>x. (fst x, Object, fst (snd x),
--- a/src/HOL/MicroJava/Comp/TypeInf.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/MicroJava/Comp/TypeInf.thy Tue Sep 01 22:32:58 2015 +0200
@@ -108,9 +108,9 @@
(* Uniqueness of types property *)
lemma uniqueness_of_types: "
- (\<forall> (E\<Colon>'a prog \<times> (vname \<Rightarrow> ty option)) T1 T2.
+ (\<forall> (E::'a prog \<times> (vname \<Rightarrow> ty option)) T1 T2.
E\<turnstile>e :: T1 \<longrightarrow> E\<turnstile>e :: T2 \<longrightarrow> T1 = T2) \<and>
- (\<forall> (E\<Colon>'a prog \<times> (vname \<Rightarrow> ty option)) Ts1 Ts2.
+ (\<forall> (E::'a prog \<times> (vname \<Rightarrow> ty option)) Ts1 Ts2.
E\<turnstile>es [::] Ts1 \<longrightarrow> E\<turnstile>es [::] Ts2 \<longrightarrow> Ts1 = Ts2)"
apply (rule compat_expr_expr_list.induct)
(* NewC *)
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -61,7 +61,7 @@
text\<open>The ordering on one-dimensional vectors is linear.\<close>
class cart_one =
- assumes UNIV_one: "card (UNIV \<Colon> 'a set) = Suc 0"
+ assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0"
begin
subclass finite
@@ -690,7 +690,7 @@
{ fix y
have "?P y"
proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR])
- show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
+ show "\<exists>x::real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
by (rule exI[where x=0], simp)
next
fix c y1 y2
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -3706,7 +3706,7 @@
shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
(is "affine hull (insert 0 ?A) = ?B")
proof -
- have *: "\<And>A. op + (0\<Colon>'a) ` A = A" "\<And>A. op + (- (0\<Colon>'a)) ` A = A"
+ have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A"
by auto
show ?thesis
unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * ..
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Tue Sep 01 22:32:58 2015 +0200
@@ -2178,7 +2178,7 @@
apply auto
done
-text \<open>Considering derivative @{typ "real \<Rightarrow> 'b\<Colon>real_normed_vector"} as a vector.\<close>
+text \<open>Considering derivative @{typ "real \<Rightarrow> 'b::real_normed_vector"} as a vector.\<close>
definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)"
--- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Sep 01 22:32:58 2015 +0200
@@ -7948,7 +7948,7 @@
using assms by auto
have *: "a \<le> x"
using assms(5) by auto
- have "((\<lambda>x. 0\<Colon>'a) has_integral f x - f a) {a .. x}"
+ have "((\<lambda>x. 0::'a) has_integral f x - f a) {a .. x}"
apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *])
apply (rule continuous_on_subset[OF assms(2)])
defer
--- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -177,7 +177,7 @@
text\<open>Instantiation for intervals on @{text ordered_euclidean_space}\<close>
lemma
- fixes a :: "'a\<Colon>ordered_euclidean_space"
+ fixes a :: "'a::ordered_euclidean_space"
shows cbox_interval: "cbox a b = {a..b}"
and interval_cbox: "{a..b} = cbox a b"
and eucl_le_atMost: "{x. \<forall>i\<in>Basis. x \<bullet> i <= a \<bullet> i} = {..a}"
@@ -185,17 +185,17 @@
by (auto simp: eucl_le[where 'a='a] eucl_less_def box_def cbox_def)
lemma closed_eucl_atLeastAtMost[simp, intro]:
- fixes a :: "'a\<Colon>ordered_euclidean_space"
+ fixes a :: "'a::ordered_euclidean_space"
shows "closed {a..b}"
by (simp add: cbox_interval[symmetric] closed_cbox)
lemma closed_eucl_atMost[simp, intro]:
- fixes a :: "'a\<Colon>ordered_euclidean_space"
+ fixes a :: "'a::ordered_euclidean_space"
shows "closed {..a}"
by (simp add: eucl_le_atMost[symmetric])
lemma closed_eucl_atLeast[simp, intro]:
- fixes a :: "'a\<Colon>ordered_euclidean_space"
+ fixes a :: "'a::ordered_euclidean_space"
shows "closed {a..}"
by (simp add: eucl_le_atLeast[symmetric])
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -869,7 +869,7 @@
unfolding set_eq_iff and Int_iff and mem_box by auto
lemma rational_boxes:
- fixes x :: "'a\<Colon>euclidean_space"
+ fixes x :: "'a::euclidean_space"
assumes "e > 0"
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat> ) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e"
proof -
@@ -926,7 +926,7 @@
qed
lemma open_UNION_box:
- fixes M :: "'a\<Colon>euclidean_space set"
+ fixes M :: "'a::euclidean_space set"
assumes "open M"
defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)"
defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)"
@@ -6677,25 +6677,25 @@
simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm)
lemma eucl_less_eq_halfspaces:
- fixes a :: "'a\<Colon>euclidean_space"
+ fixes a :: "'a::euclidean_space"
shows "{x. x <e a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
"{x. a <e x} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
by (auto simp: eucl_less_def)
lemma eucl_le_eq_halfspaces:
- fixes a :: "'a\<Colon>euclidean_space"
+ fixes a :: "'a::euclidean_space"
shows "{x. \<forall>i\<in>Basis. x \<bullet> i \<le> a \<bullet> i} = (\<Inter>i\<in>Basis. {x. x \<bullet> i \<le> a \<bullet> i})"
"{x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i} = (\<Inter>i\<in>Basis. {x. a \<bullet> i \<le> x \<bullet> i})"
by auto
lemma open_Collect_eucl_less[simp, intro]:
- fixes a :: "'a\<Colon>euclidean_space"
+ fixes a :: "'a::euclidean_space"
shows "open {x. x <e a}"
"open {x. a <e x}"
by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt)
lemma closed_Collect_eucl_le[simp, intro]:
- fixes a :: "'a\<Colon>euclidean_space"
+ fixes a :: "'a::euclidean_space"
shows "closed {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i}"
"closed {x. \<forall>i\<in>Basis. x \<bullet> i \<le> a \<bullet> i}"
unfolding eucl_le_eq_halfspaces
--- a/src/HOL/NSA/NSA.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/NSA/NSA.thy Tue Sep 01 22:32:58 2015 +0200
@@ -104,11 +104,11 @@
by transfer (rule norm_power)
lemma hnorm_one [simp]:
- "hnorm (1\<Colon>'a::real_normed_div_algebra star) = 1"
+ "hnorm (1::'a::real_normed_div_algebra star) = 1"
by transfer (rule norm_one)
lemma hnorm_zero [simp]:
- "hnorm (0\<Colon>'a::real_normed_vector star) = 0"
+ "hnorm (0::'a::real_normed_vector star) = 0"
by transfer (rule norm_zero)
lemma zero_less_hnorm_iff [simp]:
--- a/src/HOL/NSA/StarDef.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/NSA/StarDef.thy Tue Sep 01 22:32:58 2015 +0200
@@ -927,7 +927,7 @@
then show ?case by simp
next
case (Suc n)
- have "\<And>x::'a star. x * ( *f* (\<lambda>x\<Colon>'a. x ^ n)) x = ( *f* (\<lambda>x\<Colon>'a. x * x ^ n)) x"
+ have "\<And>x::'a star. x * ( *f* (\<lambda>x::'a. x ^ n)) x = ( *f* (\<lambda>x::'a. x * x ^ n)) x"
by transfer simp
with Suc show ?case by simp
qed
--- a/src/HOL/Nat.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nat.thy Tue Sep 01 22:32:58 2015 +0200
@@ -85,10 +85,10 @@
done
free_constructors case_nat for
- "0 \<Colon> nat"
+ "0 :: nat"
| Suc pred
where
- "pred (0 \<Colon> nat) = (0 \<Colon> nat)"
+ "pred (0 :: nat) = (0 :: nat)"
apply atomize_elim
apply (rename_tac n, induct_tac n rule: nat_induct0, auto)
apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject'
@@ -99,7 +99,7 @@
-- \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>
setup \<open>Sign.mandatory_path "old"\<close>
-old_rep_datatype "0 \<Colon> nat" Suc
+old_rep_datatype "0 :: nat" Suc
apply (erule nat_induct0, assumption)
apply (rule nat.inject)
apply (rule nat.distinct(1))
@@ -216,7 +216,7 @@
begin
primrec plus_nat where
- add_0: "0 + n = (n\<Colon>nat)"
+ add_0: "0 + n = (n::nat)"
| add_Suc: "Suc m + n = Suc (m + n)"
lemma add_0_right [simp]: "m + 0 = (m::nat)"
@@ -231,7 +231,7 @@
by simp
primrec minus_nat where
- diff_0 [code]: "m - 0 = (m\<Colon>nat)"
+ diff_0 [code]: "m - 0 = (m::nat)"
| diff_Suc: "m - Suc n = (case m - n of 0 => 0 | Suc k => k)"
declare diff_Suc [simp del]
@@ -263,7 +263,7 @@
One_nat_def [simp]: "1 = Suc 0"
primrec times_nat where
- mult_0: "0 * n = (0\<Colon>nat)"
+ mult_0: "0 * n = (0::nat)"
| mult_Suc: "Suc m * n = n + (m * n)"
lemma mult_0_right [simp]: "(m::nat) * 0 = 0"
@@ -349,7 +349,7 @@
subsubsection \<open>Difference\<close>
-lemma diff_self_eq_0 [simp]: "(m\<Colon>nat) - m = 0"
+lemma diff_self_eq_0 [simp]: "(m::nat) - m = 0"
by (fact diff_cancel)
lemma diff_diff_left: "(i::nat) - j - k = i - (j + k)"
@@ -435,12 +435,12 @@
begin
primrec less_eq_nat where
- "(0\<Colon>nat) \<le> n \<longleftrightarrow> True"
+ "(0::nat) \<le> n \<longleftrightarrow> True"
| "Suc m \<le> n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc n \<Rightarrow> m \<le> n)"
declare less_eq_nat.simps [simp del]
-lemma le0 [iff]: "0 \<le> (n\<Colon>nat)" by (simp add: less_eq_nat.simps)
-lemma [code]: "(0\<Colon>nat) \<le> n \<longleftrightarrow> True" by simp
+lemma le0 [iff]: "0 \<le> (n::nat)" by (simp add: less_eq_nat.simps)
+lemma [code]: "(0::nat) \<le> n \<longleftrightarrow> True" by simp
definition less_nat where
less_eq_Suc_le: "n < m \<longleftrightarrow> Suc n \<le> m"
@@ -451,13 +451,13 @@
lemma Suc_le_eq [code]: "Suc m \<le> n \<longleftrightarrow> m < n"
unfolding less_eq_Suc_le ..
-lemma le_0_eq [iff]: "(n\<Colon>nat) \<le> 0 \<longleftrightarrow> n = 0"
+lemma le_0_eq [iff]: "(n::nat) \<le> 0 \<longleftrightarrow> n = 0"
by (induct n) (simp_all add: less_eq_nat.simps(2))
-lemma not_less0 [iff]: "\<not> n < (0\<Colon>nat)"
+lemma not_less0 [iff]: "\<not> n < (0::nat)"
by (simp add: less_eq_Suc_le)
-lemma less_nat_zero_code [code]: "n < (0\<Colon>nat) \<longleftrightarrow> False"
+lemma less_nat_zero_code [code]: "n < (0::nat) \<longleftrightarrow> False"
by simp
lemma Suc_less_eq [iff]: "Suc m < Suc n \<longleftrightarrow> m < n"
@@ -1290,10 +1290,10 @@
begin
definition
- "(inf \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat) = min"
+ "(inf :: nat \<Rightarrow> nat \<Rightarrow> nat) = min"
definition
- "(sup \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat) = max"
+ "(sup :: nat \<Rightarrow> nat \<Rightarrow> nat) = max"
instance by intro_classes
(auto simp add: inf_nat_def sup_nat_def max_def not_le min_def
@@ -1934,7 +1934,7 @@
text \<open>@{term "op dvd"} is a partial order\<close>
-interpretation dvd: order "op dvd" "\<lambda>n m \<Colon> nat. n dvd m \<and> \<not> m dvd n"
+interpretation dvd: order "op dvd" "\<lambda>n m :: nat. n dvd m \<and> \<not> m dvd n"
proof qed (auto intro: dvd_refl dvd_trans dvd_antisym)
lemma dvd_diff_nat[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)"
@@ -2021,7 +2021,7 @@
begin
definition size_nat where
- [simp, code]: "size (n \<Colon> nat) = n"
+ [simp, code]: "size (n::nat) = n"
instance ..
--- a/src/HOL/Nitpick.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick.thy Tue Sep 01 22:32:58 2015 +0200
@@ -72,7 +72,7 @@
definition card' :: "'a set \<Rightarrow> nat" where
"card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
-definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
+definition setsum' :: "('a \<Rightarrow> 'b::comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
"setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
@@ -193,7 +193,7 @@
definition less_eq_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
[nitpick_simp]: "less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
-definition of_frac :: "'a \<Rightarrow> 'b\<Colon>{inverse,ring_1}" where
+definition of_frac :: "'a \<Rightarrow> 'b::{inverse,ring_1}" where
"of_frac q \<equiv> of_int (num q) / of_int (denom q)"
axiomatization wf_wfrec :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -55,7 +55,7 @@
nitpick [card = 1, expect = genuine]
oops
-lemma "{(a\<Colon>'a\<times>'a, b\<Colon>'b)}^-1 = {(b, a)}"
+lemma "{(a::'a\<times>'a, b::'b)}^-1 = {(b, a)}"
nitpick [card = 1-12, expect = none]
by auto
@@ -67,11 +67,11 @@
nitpick [card = 1-20, expect = none]
by auto
-lemma "(a\<Colon>'a\<Rightarrow>'b, a) \<in> Id\<^sup>*"
+lemma "(a::'a\<Rightarrow>'b, a) \<in> Id\<^sup>*"
nitpick [card = 1-2, expect = none]
by auto
-lemma "(a\<Colon>'a\<times>'a, a) \<in> Id\<^sup>* \<union> {(a, b)}\<^sup>*"
+lemma "(a::'a\<times>'a, a) \<in> Id\<^sup>* \<union> {(a, b)}\<^sup>*"
nitpick [card = 1-4, expect = none]
by auto
@@ -79,11 +79,11 @@
nitpick [card = 1-50, expect = none]
by (auto simp: Id_def)
-lemma "((a\<Colon>'a, b\<Colon>'a), (a, b)) \<in> Id"
+lemma "((a::'a, b::'a), (a, b)) \<in> Id"
nitpick [card = 1-10, expect = none]
by (auto simp: Id_def)
-lemma "(x\<Colon>'a\<times>'a) \<in> UNIV"
+lemma "(x::'a\<times>'a) \<in> UNIV"
nitpick [card = 1-50, expect = none]
sorry
@@ -112,13 +112,13 @@
nitpick [card = 50, expect = genuine]
oops
-lemma "(a\<Colon>'a\<times>'a, a\<Colon>'a\<times>'a) \<in> R"
+lemma "(a::'a\<times>'a, a::'a\<times>'a) \<in> R"
nitpick [card = 1, expect = genuine]
nitpick [card = 10, expect = genuine]
nitpick [card = 5, dont_box, expect = genuine]
oops
-lemma "f (g\<Colon>'a\<Rightarrow>'a) = x"
+lemma "f (g::'a\<Rightarrow>'a) = x"
nitpick [card = 3, dont_box, expect = genuine]
nitpick [card = 8, expect = genuine]
oops
@@ -131,7 +131,7 @@
nitpick [card = 10, expect = genuine]
oops
-lemma "(x\<Colon>'a) = (\<lambda>a. \<lambda>b. \<lambda>c. if c then a else b) x x True"
+lemma "(x::'a) = (\<lambda>a. \<lambda>b. \<lambda>c. if c then a else b) x x True"
nitpick [card = 1-10, expect = none]
by auto
@@ -143,7 +143,7 @@
nitpick [card = 2, expect = genuine]
oops
-lemma "(A\<Colon>'a\<times>'a, B\<Colon>'a\<times>'a) \<in> R \<Longrightarrow> (A, B) \<in> R"
+lemma "(A::'a\<times>'a, B::'a\<times>'a) \<in> R \<Longrightarrow> (A, B) \<in> R"
nitpick [card = 15, expect = none]
by auto
@@ -152,7 +152,7 @@
nitpick [card = 1-25, expect = none]
by auto
-lemma "f = (\<lambda>x\<Colon>'a\<times>'b. x)"
+lemma "f = (\<lambda>x::'a\<times>'b. x)"
nitpick [card = 8, expect = genuine]
oops
@@ -168,16 +168,16 @@
nitpick [card 'a = 100, expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a \<Rightarrow> bool. x = y"
+lemma "\<forall>x::'a \<Rightarrow> bool. x = y"
nitpick [card 'a = 1, expect = genuine]
nitpick [card 'a = 100, expect = genuine]
oops
-lemma "\<exists>x\<Colon>'a \<Rightarrow> bool. x = y"
+lemma "\<exists>x::'a \<Rightarrow> bool. x = y"
nitpick [card 'a = 1-15, expect = none]
by auto
-lemma "\<exists>x y\<Colon>'a \<Rightarrow> bool. x = y"
+lemma "\<exists>x y::'a \<Rightarrow> bool. x = y"
nitpick [card = 1-15, expect = none]
by auto
@@ -208,22 +208,22 @@
nitpick [card = 1-2, expect = genuine]
oops
-lemma "\<forall>u\<Colon>'a \<times> 'b. \<exists>v\<Colon>'c. \<forall>w\<Colon>'d. \<exists>x\<Colon>'e \<times> 'f.
+lemma "\<forall>u::'a \<times> 'b. \<exists>v::'c. \<forall>w::'d. \<exists>x::'e \<times> 'f.
f u v w x = f u (g u) w (h u w)"
nitpick [card = 1-2, expect = none]
sorry
-lemma "\<forall>u\<Colon>'a \<times> 'b. \<exists>v\<Colon>'c. \<forall>w\<Colon>'d. \<exists>x\<Colon>'e \<times> 'f.
+lemma "\<forall>u::'a \<times> 'b. \<exists>v::'c. \<forall>w::'d. \<exists>x::'e \<times> 'f.
f u v w x = f u (g u w) w (h u)"
nitpick [card = 1-2, dont_box, expect = genuine]
oops
-lemma "\<forall>u\<Colon>'a \<Rightarrow> 'b. \<exists>v\<Colon>'c. \<forall>w\<Colon>'d. \<exists>x\<Colon>'e \<Rightarrow> 'f.
+lemma "\<forall>u::'a \<Rightarrow> 'b. \<exists>v::'c. \<forall>w::'d. \<exists>x::'e \<Rightarrow> 'f.
f u v w x = f u (g u) w (h u w)"
nitpick [card = 1-2, dont_box, expect = none]
sorry
-lemma "\<forall>u\<Colon>'a \<Rightarrow> 'b. \<exists>v\<Colon>'c. \<forall>w\<Colon>'d. \<exists>x\<Colon>'e \<Rightarrow> 'f.
+lemma "\<forall>u::'a \<Rightarrow> 'b. \<exists>v::'c. \<forall>w::'d. \<exists>x::'e \<Rightarrow> 'f.
f u v w x = f u (g u w) w (h u)"
nitpick [card = 1-2, dont_box, expect = genuine]
oops
@@ -233,7 +233,7 @@
nitpick [card = 2-5, expect = none]
oops
-lemma "\<forall>x\<Colon>'a\<times>'b. if (\<forall>y. x = y) then False else True"
+lemma "\<forall>x::'a\<times>'b. if (\<forall>y. x = y) then False else True"
nitpick [card = 1, expect = genuine]
nitpick [card = 2, expect = none]
oops
@@ -242,7 +242,7 @@
nitpick [expect = none]
sorry
-lemma "(\<exists>x\<Colon>'a. \<forall>y. P x y) \<or> (\<exists>x\<Colon>'a \<times> 'a. \<forall>y. P y x)"
+lemma "(\<exists>x::'a. \<forall>y. P x y) \<or> (\<exists>x::'a \<times> 'a. \<forall>y. P y x)"
nitpick [card 'a = 1, expect = genuine]
oops
@@ -260,7 +260,7 @@
nitpick [expect = none]
by auto
-lemma "let x = (\<forall>x\<Colon>'a \<times> 'b. P x) in if x then x else \<not> x"
+lemma "let x = (\<forall>x::'a \<times> 'b. P x) in if x then x else \<not> x"
nitpick [expect = none]
by auto
@@ -278,7 +278,7 @@
nitpick [expect = none]
by auto
-schematic_lemma "\<exists>x\<Colon>'a \<Rightarrow> 'b. x = ?x"
+schematic_lemma "\<exists>x::'a \<Rightarrow> 'b. x = ?x"
nitpick [expect = none]
by auto
@@ -491,35 +491,35 @@
nitpick [expect = none]
by (simp add: snd_def)
-lemma "fst (x\<Colon>'a\<Rightarrow>'b, y) = x"
+lemma "fst (x::'a\<Rightarrow>'b, y) = x"
nitpick [expect = none]
by (simp add: fst_def)
-lemma "snd (x\<Colon>'a\<Rightarrow>'b, y) = y"
+lemma "snd (x::'a\<Rightarrow>'b, y) = y"
nitpick [expect = none]
by (simp add: snd_def)
-lemma "fst (x, y\<Colon>'a\<Rightarrow>'b) = x"
+lemma "fst (x, y::'a\<Rightarrow>'b) = x"
nitpick [expect = none]
by (simp add: fst_def)
-lemma "snd (x, y\<Colon>'a\<Rightarrow>'b) = y"
+lemma "snd (x, y::'a\<Rightarrow>'b) = y"
nitpick [expect = none]
by (simp add: snd_def)
-lemma "fst (x\<Colon>'a\<times>'b, y) = x"
+lemma "fst (x::'a\<times>'b, y) = x"
nitpick [expect = none]
by (simp add: fst_def)
-lemma "snd (x\<Colon>'a\<times>'b, y) = y"
+lemma "snd (x::'a\<times>'b, y) = y"
nitpick [expect = none]
by (simp add: snd_def)
-lemma "fst (x, y\<Colon>'a\<times>'b) = x"
+lemma "fst (x, y::'a\<times>'b) = x"
nitpick [expect = none]
by (simp add: fst_def)
-lemma "snd (x, y\<Colon>'a\<times>'b) = y"
+lemma "snd (x, y::'a\<times>'b) = y"
nitpick [expect = none]
by (simp add: snd_def)
@@ -626,7 +626,7 @@
nitpick [card = 1-5, expect = none]
by auto
-lemma "x \<in> ((A\<Colon>'a set) - B) \<longleftrightarrow> x \<in> A \<and> x \<notin> B"
+lemma "x \<in> ((A::'a set) - B) \<longleftrightarrow> x \<in> A \<and> x \<notin> B"
nitpick [card = 1-5, expect = none]
by auto
@@ -650,7 +650,7 @@
nitpick [expect = none]
by auto
-lemma "I = (\<lambda>x\<Colon>'a set. x) \<Longrightarrow> uminus = (\<lambda>x. uminus (I x))"
+lemma "I = (\<lambda>x::'a set. x) \<Longrightarrow> uminus = (\<lambda>x. uminus (I x))"
nitpick [card = 1-7, expect = none]
by auto
@@ -662,7 +662,7 @@
nitpick [expect = none]
by auto
-lemma "A = -(A\<Colon>'a set)"
+lemma "A = -(A::'a set)"
nitpick [card 'a = 10, expect = genuine]
oops
@@ -743,7 +743,7 @@
nitpick [expect = genuine]
oops
-lemma "Eps (\<lambda>x. x \<in> P) \<in> (P\<Colon>nat set)"
+lemma "Eps (\<lambda>x. x \<in> P) \<in> (P::nat set)"
nitpick [expect = genuine]
oops
@@ -751,7 +751,7 @@
nitpick [expect = genuine]
oops
-lemma "\<not> (P \<Colon> nat \<Rightarrow> bool) (Eps P)"
+lemma "\<not> (P :: nat \<Rightarrow> bool) (Eps P)"
nitpick [expect = genuine]
oops
@@ -759,7 +759,7 @@
nitpick [expect = none]
sorry
-lemma "(P \<Colon> nat \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> P (Eps P)"
+lemma "(P :: nat \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> P (Eps P)"
nitpick [expect = none]
sorry
@@ -767,7 +767,7 @@
nitpick [expect = genuine]
oops
-lemma "(P \<Colon> nat \<Rightarrow> bool) (The P)"
+lemma "(P :: nat \<Rightarrow> bool) (The P)"
nitpick [expect = genuine]
oops
@@ -775,7 +775,7 @@
nitpick [expect = genuine]
oops
-lemma "\<not> (P \<Colon> nat \<Rightarrow> bool) (The P)"
+lemma "\<not> (P :: nat \<Rightarrow> bool) (The P)"
nitpick [expect = genuine]
oops
@@ -783,7 +783,7 @@
nitpick [expect = genuine]
oops
-lemma "The P \<noteq> (x\<Colon>nat)"
+lemma "The P \<noteq> (x::nat)"
nitpick [expect = genuine]
oops
@@ -791,7 +791,7 @@
nitpick [expect = genuine]
oops
-lemma "P (x\<Colon>nat) \<Longrightarrow> P (The P)"
+lemma "P (x::nat) \<Longrightarrow> P (The P)"
nitpick [expect = genuine]
oops
@@ -799,7 +799,7 @@
nitpick [expect = none]
oops
-lemma "P = {x\<Colon>nat} \<Longrightarrow> (THE x. x \<in> P) \<in> P"
+lemma "P = {x::nat} \<Longrightarrow> (THE x. x \<in> P) \<in> P"
nitpick [expect = none]
oops
@@ -809,23 +809,23 @@
nitpick [expect = genuine]
oops
-lemma "(Q \<Colon> nat \<Rightarrow> bool) (Eps Q)"
+lemma "(Q :: nat \<Rightarrow> bool) (Eps Q)"
nitpick [expect = none] (* unfortunate *)
oops
-lemma "\<not> (Q \<Colon> nat \<Rightarrow> bool) (Eps Q)"
+lemma "\<not> (Q :: nat \<Rightarrow> bool) (Eps Q)"
nitpick [expect = genuine]
oops
-lemma "\<not> (Q \<Colon> nat \<Rightarrow> bool) (Eps Q)"
+lemma "\<not> (Q :: nat \<Rightarrow> bool) (Eps Q)"
nitpick [expect = genuine]
oops
-lemma "(Q\<Colon>'a \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> (Q\<Colon>'a \<Rightarrow> bool) (Eps Q)"
+lemma "(Q::'a \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> (Q::'a \<Rightarrow> bool) (Eps Q)"
nitpick [expect = none]
sorry
-lemma "(Q\<Colon>nat \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> (Q\<Colon>nat \<Rightarrow> bool) (Eps Q)"
+lemma "(Q::nat \<Rightarrow> bool) \<noteq> bot \<Longrightarrow> (Q::nat \<Rightarrow> bool) (Eps Q)"
nitpick [expect = none]
sorry
@@ -833,7 +833,7 @@
nitpick [expect = genuine]
oops
-lemma "(Q\<Colon>nat \<Rightarrow> bool) (The Q)"
+lemma "(Q::nat \<Rightarrow> bool) (The Q)"
nitpick [expect = genuine]
oops
@@ -841,7 +841,7 @@
nitpick [expect = genuine]
oops
-lemma "\<not> (Q\<Colon>nat \<Rightarrow> bool) (The Q)"
+lemma "\<not> (Q::nat \<Rightarrow> bool) (The Q)"
nitpick [expect = genuine]
oops
@@ -849,7 +849,7 @@
nitpick [expect = genuine]
oops
-lemma "The Q \<noteq> (x\<Colon>nat)"
+lemma "The Q \<noteq> (x::nat)"
nitpick [expect = genuine]
oops
@@ -857,15 +857,15 @@
nitpick [expect = genuine]
oops
-lemma "Q (x\<Colon>nat) \<Longrightarrow> Q (The Q)"
+lemma "Q (x::nat) \<Longrightarrow> Q (The Q)"
nitpick [expect = genuine]
oops
-lemma "Q = (\<lambda>x\<Colon>'a. x = a) \<Longrightarrow> (Q\<Colon>'a \<Rightarrow> bool) (The Q)"
+lemma "Q = (\<lambda>x::'a. x = a) \<Longrightarrow> (Q::'a \<Rightarrow> bool) (The Q)"
nitpick [expect = none]
sorry
-lemma "Q = (\<lambda>x\<Colon>nat. x = a) \<Longrightarrow> (Q\<Colon>nat \<Rightarrow> bool) (The Q)"
+lemma "Q = (\<lambda>x::nat. x = a) \<Longrightarrow> (Q::nat \<Rightarrow> bool) (The Q)"
nitpick [expect = none]
sorry
@@ -921,7 +921,7 @@
subsection {* Destructors and Recursors *}
-lemma "(x\<Colon>'a) = (case True of True \<Rightarrow> x | False \<Rightarrow> x)"
+lemma "(x::'a) = (case True of True \<Rightarrow> x | False \<Rightarrow> x)"
nitpick [card = 2, expect = none]
by auto
--- a/src/HOL/Nitpick_Examples/Integer_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Integer_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -232,7 +232,7 @@
nitpick [expect = potential] (* unfortunate *)
sorry
-lemma "(\<Sum>i \<in> labels (Node x t u). f i\<Colon>nat) = f x + (\<Sum>i \<in> labels t. f i) + (\<Sum>i \<in> labels u. f i)"
+lemma "(\<Sum>i \<in> labels (Node x t u). f i::nat) = f x + (\<Sum>i \<in> labels t. f i) + (\<Sum>i \<in> labels u. f i)"
nitpick [expect = potential] (* unfortunate *)
oops
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -68,7 +68,7 @@
subsection {* 2.5. Natural Numbers and Integers *}
-lemma "\<lbrakk>i \<le> j; n \<le> (m\<Colon>int)\<rbrakk> \<Longrightarrow> i * n + j * m \<le> i * m + j * n"
+lemma "\<lbrakk>i \<le> j; n \<le> (m::int)\<rbrakk> \<Longrightarrow> i * n + j * m \<le> i * m + j * n"
nitpick [expect = genuine]
nitpick [binary_ints, bits = 16, expect = genuine]
oops
@@ -81,7 +81,7 @@
nitpick [expect = none]
oops
-lemma "P (op +\<Colon>nat\<Rightarrow>nat\<Rightarrow>nat)"
+lemma "P (op +::nat\<Rightarrow>nat\<Rightarrow>nat)"
nitpick [card nat = 1, expect = genuine]
nitpick [card nat = 2, expect = none]
oops
@@ -101,7 +101,7 @@
subsection {* 2.7. Typedefs, Records, Rationals, and Reals *}
-definition "three = {0\<Colon>nat, 1, 2}"
+definition "three = {0::nat, 1, 2}"
typedef three = three
unfolding three_def by blast
@@ -120,9 +120,9 @@
by (auto simp add: equivp_def fun_eq_iff)
definition add_raw where
-"add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u\<Colon>nat), y + (v\<Colon>nat))"
+"add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
-quotient_definition "add\<Colon>my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw
+quotient_definition "add::my_int \<Rightarrow> my_int \<Rightarrow> my_int" is add_raw
unfolding add_raw_def by auto
lemma "add x y = add x x"
@@ -148,11 +148,11 @@
Xcoord :: int
Ycoord :: int
-lemma "Xcoord (p\<Colon>point) = Xcoord (q\<Colon>point)"
+lemma "Xcoord (p::point) = Xcoord (q::point)"
nitpick [show_types, expect = genuine]
oops
-lemma "4 * x + 3 * (y\<Colon>real) \<noteq> 1 / 2"
+lemma "4 * x + 3 * (y::real) \<noteq> 1 / 2"
nitpick [show_types, expect = genuine]
oops
@@ -172,7 +172,7 @@
oops
inductive even' where
-"even' (0\<Colon>nat)" |
+"even' (0::nat)" |
"even' 2" |
"\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)"
@@ -185,7 +185,7 @@
oops
coinductive nats where
-"nats (x\<Colon>nat) \<Longrightarrow> nats x"
+"nats (x::nat) \<Longrightarrow> nats x"
lemma "nats = (\<lambda>n. n \<in> {0, 1, 2, 3, 4})"
nitpick [card nat = 10, show_consts, expect = genuine]
@@ -264,7 +264,7 @@
nitpick [verbose, expect = genuine]
oops
-lemma "\<exists>g. \<forall>x\<Colon>'b. g (f x) = x \<Longrightarrow> \<forall>y\<Colon>'a. \<exists>x. y = f x"
+lemma "\<exists>g. \<forall>x::'b. g (f x) = x \<Longrightarrow> \<forall>y::'a. \<exists>x. y = f x"
nitpick [mono, expect = none]
nitpick [expect = genuine]
oops
@@ -273,7 +273,7 @@
subsection {* 2.12. Inductive Properties *}
inductive_set reach where
-"(4\<Colon>nat) \<in> reach" |
+"(4::nat) \<in> reach" |
"n \<in> reach \<Longrightarrow> n < 4 \<Longrightarrow> 3 * n + 1 \<in> reach" |
"n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
@@ -381,7 +381,7 @@
subsection {* 3.2. AA Trees *}
-datatype 'a aa_tree = \<Lambda> | N "'a\<Colon>linorder" nat "'a aa_tree" "'a aa_tree"
+datatype 'a aa_tree = \<Lambda> | N "'a::linorder" nat "'a aa_tree" "'a aa_tree"
primrec data where
"data \<Lambda> = undefined" |
@@ -449,7 +449,7 @@
nitpick [expect = genuine]
oops
-theorem wf_insort\<^sub>1_nat: "wf t \<Longrightarrow> wf (insort\<^sub>1 t (x\<Colon>nat))"
+theorem wf_insort\<^sub>1_nat: "wf t \<Longrightarrow> wf (insort\<^sub>1 t (x::nat))"
nitpick [eval = "insort\<^sub>1 t x", expect = genuine]
oops
--- a/src/HOL/Nitpick_Examples/Mini_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -38,20 +38,20 @@
ML_val {* none 1 @{prop "\<forall>x. x = y"} *}
ML_val {* genuine 2 @{prop "\<forall>x. x = y"} *}
ML_val {* none 2 @{prop "\<exists>x. x = y"} *}
-ML_val {* none 2 @{prop "\<forall>x\<Colon>'a \<times> 'a. x = x"} *}
-ML_val {* none 2 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y"} *}
-ML_val {* genuine 2 @{prop "\<forall>x\<Colon>'a \<times> 'a. x = y"} *}
-ML_val {* none 2 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y"} *}
+ML_val {* none 2 @{prop "\<forall>x::'a \<times> 'a. x = x"} *}
+ML_val {* none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"} *}
+ML_val {* genuine 2 @{prop "\<forall>x::'a \<times> 'a. x = y"} *}
+ML_val {* none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"} *}
ML_val {* none 1 @{prop "All = Ex"} *}
ML_val {* genuine 2 @{prop "All = Ex"} *}
ML_val {* none 1 @{prop "All P = Ex P"} *}
ML_val {* genuine 2 @{prop "All P = Ex P"} *}
ML_val {* none 4 @{prop "x = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 4 @{prop "(x\<Colon>'a \<times> 'a) = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "(x\<Colon>'a \<times> 'a) = y \<longrightarrow> P x y = P y x"} *}
-ML_val {* none 4 @{prop "\<exists>x\<Colon>'a \<times> 'a. x = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "(x\<Colon>'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "\<exists>x\<Colon>'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y"} *}
+ML_val {* none 4 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x = P y"} *}
+ML_val {* none 2 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x y = P y x"} *}
+ML_val {* none 4 @{prop "\<exists>x::'a \<times> 'a. x = y \<longrightarrow> P x = P y"} *}
+ML_val {* none 2 @{prop "(x::'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y"} *}
+ML_val {* none 2 @{prop "\<exists>x::'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y"} *}
ML_val {* genuine 1 @{prop "(op =) X = Ex"} *}
ML_val {* none 2 @{prop "\<forall>x::'a \<Rightarrow> 'a. x = x"} *}
ML_val {* none 1 @{prop "x = y"} *}
@@ -68,8 +68,8 @@
ML_val {* genuine 1 @{prop "{a} \<noteq> {a, b}"} *}
ML_val {* none 4 @{prop "{}\<^sup>+ = {}"} *}
ML_val {* none 4 @{prop "UNIV\<^sup>+ = UNIV"} *}
-ML_val {* none 4 @{prop "(UNIV \<Colon> ('a \<times> 'b) set) - {} = UNIV"} *}
-ML_val {* none 4 @{prop "{} - (UNIV \<Colon> ('a \<times> 'b) set) = {}"} *}
+ML_val {* none 4 @{prop "(UNIV :: ('a \<times> 'b) set) - {} = UNIV"} *}
+ML_val {* none 4 @{prop "{} - (UNIV :: ('a \<times> 'b) set) = {}"} *}
ML_val {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
ML_val {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
ML_val {* none 4 @{prop "a \<noteq> c \<Longrightarrow> {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
@@ -79,12 +79,12 @@
ML_val {* none 4 @{prop "\<exists>a b. (a, b) = (b, a)"} *}
ML_val {* genuine 2 @{prop "(a, b) = (b, a)"} *}
ML_val {* genuine 2 @{prop "(a, b) \<noteq> (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b\<Colon>'a \<times> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a\<Colon>'a \<times> 'a, b) = (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b\<Colon>'a \<times> 'a \<times> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a\<Colon>'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b\<Colon>'a \<Rightarrow> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 1 @{prop "(a\<Colon>'a \<Rightarrow> 'a, b) \<noteq> (b, a)"} *}
+ML_val {* none 4 @{prop "\<exists>a b::'a \<times> 'a. (a, b) = (b, a)"} *}
+ML_val {* genuine 2 @{prop "(a::'a \<times> 'a, b) = (b, a)"} *}
+ML_val {* none 4 @{prop "\<exists>a b::'a \<times> 'a \<times> 'a. (a, b) = (b, a)"} *}
+ML_val {* genuine 2 @{prop "(a::'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"} *}
+ML_val {* none 4 @{prop "\<exists>a b::'a \<Rightarrow> 'a. (a, b) = (b, a)"} *}
+ML_val {* genuine 1 @{prop "(a::'a \<Rightarrow> 'a, b) \<noteq> (b, a)"} *}
ML_val {* none 4 @{prop "fst (a, b) = a"} *}
ML_val {* none 1 @{prop "fst (a, b) = b"} *}
ML_val {* genuine 2 @{prop "fst (a, b) = b"} *}
@@ -104,8 +104,8 @@
ML_val {* none 3 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (y, x)"} *}
ML_val {* genuine 2 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (x, y)"} *}
ML_val {* none 4 @{prop "f = (\<lambda>x. f x)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x. f x\<Colon>'a \<Rightarrow> bool)"} *}
+ML_val {* none 4 @{prop "f = (\<lambda>x. f x::'a \<Rightarrow> bool)"} *}
ML_val {* none 4 @{prop "f = (\<lambda>x y. f x y)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x y. f x y\<Colon>'a \<Rightarrow> bool)"} *}
+ML_val {* none 4 @{prop "f = (\<lambda>x y. f x y::'a \<Rightarrow> bool)"} *}
end
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -66,78 +66,78 @@
ML {* Nitpick_Mono.trace := false *}
-ML_val {* const @{term "A\<Colon>('a\<Rightarrow>'b)"} *}
-ML_val {* const @{term "(A\<Colon>'a set) = A"} *}
-ML_val {* const @{term "(A\<Colon>'a set set) = A"} *}
-ML_val {* const @{term "(\<lambda>x\<Colon>'a set. a \<in> x)"} *}
-ML_val {* const @{term "{{a\<Colon>'a}} = C"} *}
-ML_val {* const @{term "{f\<Colon>'a\<Rightarrow>nat} = {g\<Colon>'a\<Rightarrow>nat}"} *}
-ML_val {* const @{term "A \<union> (B\<Colon>'a set)"} *}
-ML_val {* const @{term "\<lambda>A B x\<Colon>'a. A x \<or> B x"} *}
-ML_val {* const @{term "P (a\<Colon>'a)"} *}
-ML_val {* const @{term "\<lambda>a\<Colon>'a. b (c (d\<Colon>'a)) (e\<Colon>'a) (f\<Colon>'a)"} *}
-ML_val {* const @{term "\<forall>A\<Colon>'a set. a \<in> A"} *}
-ML_val {* const @{term "\<forall>A\<Colon>'a set. P A"} *}
+ML_val {* const @{term "A::('a\<Rightarrow>'b)"} *}
+ML_val {* const @{term "(A::'a set) = A"} *}
+ML_val {* const @{term "(A::'a set set) = A"} *}
+ML_val {* const @{term "(\<lambda>x::'a set. a \<in> x)"} *}
+ML_val {* const @{term "{{a::'a}} = C"} *}
+ML_val {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *}
+ML_val {* const @{term "A \<union> (B::'a set)"} *}
+ML_val {* const @{term "\<lambda>A B x::'a. A x \<or> B x"} *}
+ML_val {* const @{term "P (a::'a)"} *}
+ML_val {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *}
+ML_val {* const @{term "\<forall>A::'a set. a \<in> A"} *}
+ML_val {* const @{term "\<forall>A::'a set. P A"} *}
ML_val {* const @{term "P \<or> Q"} *}
-ML_val {* const @{term "A \<union> B = (C\<Colon>'a set)"} *}
-ML_val {* const @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) A B = C"} *}
-ML_val {* const @{term "(if P then (A\<Colon>'a set) else B) = C"} *}
-ML_val {* const @{term "let A = (C\<Colon>'a set) in A \<union> B"} *}
-ML_val {* const @{term "THE x\<Colon>'b. P x"} *}
-ML_val {* const @{term "(\<lambda>x\<Colon>'a. False)"} *}
-ML_val {* const @{term "(\<lambda>x\<Colon>'a. True)"} *}
-ML_val {* const @{term "(\<lambda>x\<Colon>'a. False) = (\<lambda>x\<Colon>'a. False)"} *}
-ML_val {* const @{term "(\<lambda>x\<Colon>'a. True) = (\<lambda>x\<Colon>'a. True)"} *}
-ML_val {* const @{term "Let (a\<Colon>'a) A"} *}
-ML_val {* const @{term "A (a\<Colon>'a)"} *}
-ML_val {* const @{term "insert (a\<Colon>'a) A = B"} *}
-ML_val {* const @{term "- (A\<Colon>'a set)"} *}
-ML_val {* const @{term "finite (A\<Colon>'a set)"} *}
-ML_val {* const @{term "\<not> finite (A\<Colon>'a set)"} *}
-ML_val {* const @{term "finite (A\<Colon>'a set set)"} *}
-ML_val {* const @{term "\<lambda>a\<Colon>'a. A a \<and> \<not> B a"} *}
-ML_val {* const @{term "A < (B\<Colon>'a set)"} *}
-ML_val {* const @{term "A \<le> (B\<Colon>'a set)"} *}
-ML_val {* const @{term "[a\<Colon>'a]"} *}
-ML_val {* const @{term "[a\<Colon>'a set]"} *}
-ML_val {* const @{term "[A \<union> (B\<Colon>'a set)]"} *}
-ML_val {* const @{term "[A \<union> (B\<Colon>'a set)] = [C]"} *}
-ML_val {* const @{term "{(\<lambda>x\<Colon>'a. x = a)} = C"} *}
-ML_val {* const @{term "(\<lambda>a\<Colon>'a. \<not> A a) = B"} *}
-ML_val {* const @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a"} *}
-ML_val {* const @{term "\<lambda>A B x\<Colon>'a. A x \<and> B x \<and> A = B"} *}
-ML_val {* const @{term "p = (\<lambda>(x\<Colon>'a) (y\<Colon>'a). P x \<or> \<not> Q y)"} *}
-ML_val {* const @{term "p = (\<lambda>(x\<Colon>'a) (y\<Colon>'a). p x y \<Colon> bool)"} *}
+ML_val {* const @{term "A \<union> B = (C::'a set)"} *}
+ML_val {* const @{term "(\<lambda>A B x::'a. A x \<or> B x) A B = C"} *}
+ML_val {* const @{term "(if P then (A::'a set) else B) = C"} *}
+ML_val {* const @{term "let A = (C::'a set) in A \<union> B"} *}
+ML_val {* const @{term "THE x::'b. P x"} *}
+ML_val {* const @{term "(\<lambda>x::'a. False)"} *}
+ML_val {* const @{term "(\<lambda>x::'a. True)"} *}
+ML_val {* const @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. False)"} *}
+ML_val {* const @{term "(\<lambda>x::'a. True) = (\<lambda>x::'a. True)"} *}
+ML_val {* const @{term "Let (a::'a) A"} *}
+ML_val {* const @{term "A (a::'a)"} *}
+ML_val {* const @{term "insert (a::'a) A = B"} *}
+ML_val {* const @{term "- (A::'a set)"} *}
+ML_val {* const @{term "finite (A::'a set)"} *}
+ML_val {* const @{term "\<not> finite (A::'a set)"} *}
+ML_val {* const @{term "finite (A::'a set set)"} *}
+ML_val {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *}
+ML_val {* const @{term "A < (B::'a set)"} *}
+ML_val {* const @{term "A \<le> (B::'a set)"} *}
+ML_val {* const @{term "[a::'a]"} *}
+ML_val {* const @{term "[a::'a set]"} *}
+ML_val {* const @{term "[A \<union> (B::'a set)]"} *}
+ML_val {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
+ML_val {* const @{term "{(\<lambda>x::'a. x = a)} = C"} *}
+ML_val {* const @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
+ML_val {* const @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a"} *}
+ML_val {* const @{term "\<lambda>A B x::'a. A x \<and> B x \<and> A = B"} *}
+ML_val {* const @{term "p = (\<lambda>(x::'a) (y::'a). P x \<or> \<not> Q y)"} *}
+ML_val {* const @{term "p = (\<lambda>(x::'a) (y::'a). p x y :: bool)"} *}
ML_val {* const @{term "p = (\<lambda>A B x. A x \<and> \<not> B x) (\<lambda>x. True) (\<lambda>y. x \<noteq> y)"} *}
ML_val {* const @{term "p = (\<lambda>y. x \<noteq> y)"} *}
-ML_val {* const @{term "(\<lambda>x. (p\<Colon>'a\<Rightarrow>bool\<Rightarrow>bool) x False)"} *}
-ML_val {* const @{term "(\<lambda>x y. (p\<Colon>'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)"} *}
-ML_val {* const @{term "f = (\<lambda>x\<Colon>'a. P x \<longrightarrow> Q x)"} *}
-ML_val {* const @{term "\<forall>a\<Colon>'a. P a"} *}
+ML_val {* const @{term "(\<lambda>x. (p::'a\<Rightarrow>bool\<Rightarrow>bool) x False)"} *}
+ML_val {* const @{term "(\<lambda>x y. (p::'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)"} *}
+ML_val {* const @{term "f = (\<lambda>x::'a. P x \<longrightarrow> Q x)"} *}
+ML_val {* const @{term "\<forall>a::'a. P a"} *}
-ML_val {* nonconst @{term "\<forall>P (a\<Colon>'a). P a"} *}
-ML_val {* nonconst @{term "THE x\<Colon>'a. P x"} *}
-ML_val {* nonconst @{term "SOME x\<Colon>'a. P x"} *}
-ML_val {* nonconst @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) = myunion"} *}
-ML_val {* nonconst @{term "(\<lambda>x\<Colon>'a. False) = (\<lambda>x\<Colon>'a. True)"} *}
-ML_val {* nonconst @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
+ML_val {* nonconst @{term "\<forall>P (a::'a). P a"} *}
+ML_val {* nonconst @{term "THE x::'a. P x"} *}
+ML_val {* nonconst @{term "SOME x::'a. P x"} *}
+ML_val {* nonconst @{term "(\<lambda>A B x::'a. A x \<or> B x) = myunion"} *}
+ML_val {* nonconst @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. True)"} *}
+ML_val {* nonconst @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
-ML_val {* mono @{prop "Q (\<forall>x\<Colon>'a set. P x)"} *}
-ML_val {* mono @{prop "P (a\<Colon>'a)"} *}
-ML_val {* mono @{prop "{a} = {b\<Colon>'a}"} *}
-ML_val {* mono @{prop "(\<lambda>x. x = a) = (\<lambda>y. y = (b\<Colon>'a))"} *}
-ML_val {* mono @{prop "(a\<Colon>'a) \<in> P \<and> P \<union> P = P"} *}
-ML_val {* mono @{prop "\<forall>F\<Colon>'a set set. P"} *}
-ML_val {* mono @{prop "\<not> (\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h)"} *}
-ML_val {* mono @{prop "\<not> Q (\<forall>x\<Colon>'a set. P x)"} *}
-ML_val {* mono @{prop "\<not> (\<forall>x\<Colon>'a. P x)"} *}
-ML_val {* mono @{prop "myall P = (P = (\<lambda>x\<Colon>'a. True))"} *}
-ML_val {* mono @{prop "myall P = (P = (\<lambda>x\<Colon>'a. False))"} *}
-ML_val {* mono @{prop "\<forall>x\<Colon>'a. P x"} *}
-ML_val {* mono @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) \<noteq> myunion"} *}
+ML_val {* mono @{prop "Q (\<forall>x::'a set. P x)"} *}
+ML_val {* mono @{prop "P (a::'a)"} *}
+ML_val {* mono @{prop "{a} = {b::'a}"} *}
+ML_val {* mono @{prop "(\<lambda>x. x = a) = (\<lambda>y. y = (b::'a))"} *}
+ML_val {* mono @{prop "(a::'a) \<in> P \<and> P \<union> P = P"} *}
+ML_val {* mono @{prop "\<forall>F::'a set set. P"} *}
+ML_val {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h)"} *}
+ML_val {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *}
+ML_val {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *}
+ML_val {* mono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *}
+ML_val {* mono @{prop "myall P = (P = (\<lambda>x::'a. False))"} *}
+ML_val {* mono @{prop "\<forall>x::'a. P x"} *}
+ML_val {* mono @{term "(\<lambda>A B x::'a. A x \<or> B x) \<noteq> myunion"} *}
ML_val {* nonmono @{prop "A = (\<lambda>x::'a. True) \<and> A = (\<lambda>x. False)"} *}
-ML_val {* nonmono @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
+ML_val {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
ML {*
val preproc_timeout = seconds 5.0
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -56,7 +56,7 @@
nitpick [expect = genuine]
oops
-lemma "(P\<Colon>bool) = Q"
+lemma "(P::bool) = Q"
nitpick [expect = genuine]
oops
@@ -96,11 +96,11 @@
nitpick [expect = genuine]
oops
-lemma "(f\<Colon>'a\<Rightarrow>'b) = g"
+lemma "(f::'a\<Rightarrow>'b) = g"
nitpick [expect = genuine]
oops
-lemma "(f\<Colon>('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
+lemma "(f::('d\<Rightarrow>'d)\<Rightarrow>('c\<Rightarrow>'d)) = g"
nitpick [expect = genuine]
oops
@@ -202,7 +202,7 @@
text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
-lemma "((P\<Colon>('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
+lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
nitpick [expect = genuine]
oops
@@ -367,11 +367,11 @@
subsubsection {* Sets *}
-lemma "P (A\<Colon>'a set)"
+lemma "P (A::'a set)"
nitpick [expect = genuine]
oops
-lemma "P (A\<Colon>'a set set)"
+lemma "P (A::'a set set)"
nitpick [expect = genuine]
oops
@@ -473,33 +473,33 @@
subsubsection {* Operations on Natural Numbers *}
-lemma "(x\<Colon>nat) + y = 0"
+lemma "(x::nat) + y = 0"
nitpick [expect = genuine]
oops
-lemma "(x\<Colon>nat) = x + x"
+lemma "(x::nat) = x + x"
nitpick [expect = genuine]
oops
-lemma "(x\<Colon>nat) - y + y = x"
+lemma "(x::nat) - y + y = x"
nitpick [expect = genuine]
oops
-lemma "(x\<Colon>nat) = x * x"
+lemma "(x::nat) = x * x"
nitpick [expect = genuine]
oops
-lemma "(x\<Colon>nat) < x + y"
+lemma "(x::nat) < x + y"
nitpick [card = 1, expect = genuine]
oops
text {* \<times> *}
-lemma "P (x\<Colon>'a\<times>'b)"
+lemma "P (x::'a\<times>'b)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a\<times>'b. P x"
+lemma "\<forall>x::'a\<times>'b. P x"
nitpick [expect = genuine]
oops
@@ -532,7 +532,7 @@
typedef 'a myTdef = "myTdef :: 'a set"
unfolding myTdef_def by auto
-lemma "(x\<Colon>'a myTdef) = y"
+lemma "(x::'a myTdef) = y"
nitpick [expect = genuine]
oops
@@ -543,7 +543,7 @@
typedef 'a T_bij = "T_bij :: ('a \<Rightarrow> 'a) set"
unfolding T_bij_def by auto
-lemma "P (f\<Colon>(myTdecl myTdef) T_bij)"
+lemma "P (f::(myTdecl myTdef) T_bij)"
nitpick [expect = genuine]
oops
@@ -551,11 +551,11 @@
text {* unit *}
-lemma "P (x\<Colon>unit)"
+lemma "P (x::unit)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>unit. P x"
+lemma "\<forall>x::unit. P x"
nitpick [expect = genuine]
oops
@@ -569,11 +569,11 @@
text {* option *}
-lemma "P (x\<Colon>'a option)"
+lemma "P (x::'a option)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a option. P x"
+lemma "\<forall>x::'a option. P x"
nitpick [expect = genuine]
oops
@@ -591,11 +591,11 @@
text {* + *}
-lemma "P (x\<Colon>'a+'b)"
+lemma "P (x::'a+'b)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a+'b. P x"
+lemma "\<forall>x::'a+'b. P x"
nitpick [expect = genuine]
oops
@@ -619,11 +619,11 @@
datatype T1 = A | B
-lemma "P (x\<Colon>T1)"
+lemma "P (x::T1)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>T1. P x"
+lemma "\<forall>x::T1. P x"
nitpick [expect = genuine]
oops
@@ -655,11 +655,11 @@
datatype 'a T2 = C T1 | D 'a
-lemma "P (x\<Colon>'a T2)"
+lemma "P (x::'a T2)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a T2. P x"
+lemma "\<forall>x::'a T2. P x"
nitpick [expect = genuine]
oops
@@ -687,11 +687,11 @@
datatype ('a, 'b) T3 = E "'a \<Rightarrow> 'b"
-lemma "P (x\<Colon>('a, 'b) T3)"
+lemma "P (x::('a, 'b) T3)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>('a, 'b) T3. P x"
+lemma "\<forall>x::('a, 'b) T3. P x"
nitpick [expect = genuine]
oops
@@ -716,11 +716,11 @@
text {* nat *}
-lemma "P (x\<Colon>nat)"
+lemma "P (x::nat)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>nat. P x"
+lemma "\<forall>x::nat. P x"
nitpick [expect = genuine]
oops
@@ -752,11 +752,11 @@
text {* 'a list *}
-lemma "P (xs\<Colon>'a list)"
+lemma "P (xs::'a list)"
nitpick [expect = genuine]
oops
-lemma "\<forall>xs\<Colon>'a list. P xs"
+lemma "\<forall>xs::'a list. P xs"
nitpick [expect = genuine]
oops
@@ -782,7 +782,7 @@
nitpick [expect = genuine]
oops
-lemma "(xs\<Colon>'a list) = ys"
+lemma "(xs::'a list) = ys"
nitpick [expect = genuine]
oops
@@ -792,11 +792,11 @@
datatype BitList = BitListNil | Bit0 BitList | Bit1 BitList
-lemma "P (x\<Colon>BitList)"
+lemma "P (x::BitList)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>BitList. P x"
+lemma "\<forall>x::BitList. P x"
nitpick [expect = genuine]
oops
@@ -825,11 +825,11 @@
datatype 'a BinTree = Leaf 'a | Node "'a BinTree" "'a BinTree"
-lemma "P (x\<Colon>'a BinTree)"
+lemma "P (x::'a BinTree)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a BinTree. P x"
+lemma "\<forall>x::'a BinTree. P x"
nitpick [expect = genuine]
oops
@@ -860,11 +860,11 @@
datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp"
and 'a bexp = Equal "'a aexp" "'a aexp"
-lemma "P (x\<Colon>'a aexp)"
+lemma "P (x::'a aexp)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a aexp. P x"
+lemma "\<forall>x::'a aexp. P x"
nitpick [expect = genuine]
oops
@@ -872,11 +872,11 @@
nitpick [expect = genuine]
oops
-lemma "P (x\<Colon>'a bexp)"
+lemma "P (x::'a bexp)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a bexp. P x"
+lemma "\<forall>x::'a bexp. P x"
nitpick [expect = genuine]
oops
@@ -913,11 +913,11 @@
datatype X = A | B X | C Y and Y = D X | E Y | F
-lemma "P (x\<Colon>X)"
+lemma "P (x::X)"
nitpick [expect = genuine]
oops
-lemma "P (y\<Colon>Y)"
+lemma "P (y::Y)"
nitpick [expect = genuine]
oops
@@ -1001,7 +1001,7 @@
datatype XOpt = CX "XOpt option" | DX "bool \<Rightarrow> XOpt option"
-lemma "P (x\<Colon>XOpt)"
+lemma "P (x::XOpt)"
nitpick [expect = genuine]
oops
@@ -1019,7 +1019,7 @@
datatype 'a YOpt = CY "('a \<Rightarrow> 'a YOpt) option"
-lemma "P (x\<Colon>'a YOpt)"
+lemma "P (x::'a YOpt)"
nitpick [expect = genuine]
oops
@@ -1033,11 +1033,11 @@
datatype Trie = TR "Trie list"
-lemma "P (x\<Colon>Trie)"
+lemma "P (x::Trie)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>Trie. P x"
+lemma "\<forall>x::Trie. P x"
nitpick [expect = genuine]
oops
@@ -1047,11 +1047,11 @@
datatype InfTree = Leaf | Node "nat \<Rightarrow> InfTree"
-lemma "P (x\<Colon>InfTree)"
+lemma "P (x::InfTree)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>InfTree. P x"
+lemma "\<forall>x::InfTree. P x"
nitpick [expect = genuine]
oops
@@ -1075,11 +1075,11 @@
datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \<Rightarrow> 'a lambda"
-lemma "P (x\<Colon>'a lambda)"
+lemma "P (x::'a lambda)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'a lambda. P x"
+lemma "\<forall>x::'a lambda. P x"
nitpick [expect = genuine]
oops
@@ -1112,11 +1112,11 @@
datatype (dead 'a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
datatype 'c U = E "('c, 'c U) T"
-lemma "P (x\<Colon>'c U)"
+lemma "P (x::'c U)"
nitpick [expect = genuine]
oops
-lemma "\<forall>x\<Colon>'c U. P x"
+lemma "\<forall>x::'c U. P x"
nitpick [expect = genuine]
oops
@@ -1130,14 +1130,14 @@
xpos :: 'a
ypos :: 'b
-lemma "(x\<Colon>('a, 'b) point) = y"
+lemma "(x::('a, 'b) point) = y"
nitpick [expect = genuine]
oops
record ('a, 'b, 'c) extpoint = "('a, 'b) point" +
ext :: 'c
-lemma "(x\<Colon>('a, 'b, 'c) extpoint) = y"
+lemma "(x::('a, 'b, 'c) extpoint) = y"
nitpick [expect = genuine]
oops
@@ -1218,7 +1218,7 @@
class classA
-lemma "P (x\<Colon>'a\<Colon>classA)"
+lemma "P (x::'a::classA)"
nitpick [expect = genuine]
oops
@@ -1227,11 +1227,11 @@
class classC =
assumes classC_ax: "\<exists>x y. x \<noteq> y"
-lemma "P (x\<Colon>'a\<Colon>classC)"
+lemma "P (x::'a::classC)"
nitpick [expect = genuine]
oops
-lemma "\<exists>x y. (x\<Colon>'a\<Colon>classC) \<noteq> y"
+lemma "\<exists>x y. (x::'a::classC) \<noteq> y"
nitpick [expect = none]
sorry
@@ -1241,7 +1241,7 @@
fixes classD_const :: "'a \<Rightarrow> 'a"
assumes classD_ax: "classD_const (classD_const x) = classD_const x"
-lemma "P (x\<Colon>'a\<Colon>classD)"
+lemma "P (x::'a::classD)"
nitpick [expect = genuine]
oops
@@ -1249,23 +1249,23 @@
class classE = classC + classD
-lemma "P (x\<Colon>'a\<Colon>classE)"
+lemma "P (x::'a::classE)"
nitpick [expect = genuine]
oops
text {* OFCLASS: *}
-lemma "OFCLASS('a\<Colon>type, type_class)"
+lemma "OFCLASS('a::type, type_class)"
nitpick [expect = none]
apply intro_classes
done
-lemma "OFCLASS('a\<Colon>classC, type_class)"
+lemma "OFCLASS('a::classC, type_class)"
nitpick [expect = none]
apply intro_classes
done
-lemma "OFCLASS('a\<Colon>type, classC_class)"
+lemma "OFCLASS('a::type, classC_class)"
nitpick [expect = genuine]
oops
@@ -1274,19 +1274,19 @@
consts inverse :: "'a \<Rightarrow> 'a"
defs (overloaded)
-inverse_bool: "inverse (b\<Colon>bool) \<equiv> \<not> b"
-inverse_set: "inverse (S\<Colon>'a set) \<equiv> -S"
+inverse_bool: "inverse (b::bool) \<equiv> \<not> b"
+inverse_set: "inverse (S::'a set) \<equiv> -S"
inverse_pair: "inverse p \<equiv> (inverse (fst p), inverse (snd p))"
lemma "inverse b"
nitpick [expect = genuine]
oops
-lemma "P (inverse (S\<Colon>'a set))"
+lemma "P (inverse (S::'a set))"
nitpick [expect = genuine]
oops
-lemma "P (inverse (p\<Colon>'a\<times>'b))"
+lemma "P (inverse (p::'a\<times>'b))"
nitpick [expect = genuine]
oops
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -131,7 +131,7 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11::nat).
b\<^sub>1 < b\<^sub>11 \<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 < b\<^sub>11 then a else h b\<^sub>2) x"
nitpick [expect = potential]
nitpick [dont_specialize, expect = none]
@@ -140,7 +140,7 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11::nat).
b\<^sub>1 < b\<^sub>11
\<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 < b\<^sub>11 then
a
@@ -154,7 +154,7 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11::nat).
b\<^sub>1 < b\<^sub>11
\<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 \<ge> b\<^sub>11 then
a
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy Tue Sep 01 22:32:58 2015 +0200
@@ -14,7 +14,7 @@
nitpick_params [verbose, card = 1-4, sat_solver = MiniSat_JNI, max_threads = 1,
timeout = 240]
-definition "three = {0\<Colon>nat, 1, 2}"
+definition "three = {0::nat, 1, 2}"
typedef three = three
unfolding three_def by blast
@@ -22,37 +22,37 @@
definition B :: three where "B \<equiv> Abs_three 1"
definition C :: three where "C \<equiv> Abs_three 2"
-lemma "x = (y\<Colon>three)"
+lemma "x = (y::three)"
nitpick [expect = genuine]
oops
-definition "one_or_two = {undefined False\<Colon>'a, undefined True}"
+definition "one_or_two = {undefined False::'a, undefined True}"
typedef 'a one_or_two = "one_or_two :: 'a set"
unfolding one_or_two_def by auto
-lemma "x = (y\<Colon>unit one_or_two)"
+lemma "x = (y::unit one_or_two)"
nitpick [expect = none]
sorry
-lemma "x = (y\<Colon>bool one_or_two)"
+lemma "x = (y::bool one_or_two)"
nitpick [expect = genuine]
oops
-lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> x = (y\<Colon>bool one_or_two)"
+lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> x = (y::bool one_or_two)"
nitpick [expect = none]
sorry
-lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> \<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
+lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> \<exists>x (y::bool one_or_two). x \<noteq> y"
nitpick [card = 1, expect = potential] (* unfortunate *)
oops
-lemma "\<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
+lemma "\<exists>x (y::bool one_or_two). x \<noteq> y"
nitpick [card = 1, expect = potential] (* unfortunate *)
nitpick [card = 2, expect = none]
oops
-definition "bounded = {n\<Colon>nat. finite (UNIV \<Colon> 'a set) \<longrightarrow> n < card (UNIV \<Colon> 'a set)}"
+definition "bounded = {n::nat. finite (UNIV :: 'a set) \<longrightarrow> n < card (UNIV :: 'a set)}"
typedef 'a bounded = "bounded(TYPE('a))"
unfolding bounded_def
@@ -60,23 +60,23 @@
apply (case_tac "card UNIV = 0")
by auto
-lemma "x = (y\<Colon>unit bounded)"
+lemma "x = (y::unit bounded)"
nitpick [expect = none]
sorry
-lemma "x = (y\<Colon>bool bounded)"
+lemma "x = (y::bool bounded)"
nitpick [expect = genuine]
oops
-lemma "x \<noteq> (y\<Colon>bool bounded) \<Longrightarrow> z = x \<or> z = y"
+lemma "x \<noteq> (y::bool bounded) \<Longrightarrow> z = x \<or> z = y"
nitpick [expect = potential] (* unfortunate *)
sorry
-lemma "x \<noteq> (y\<Colon>(bool \<times> bool) bounded) \<Longrightarrow> z = x \<or> z = y"
+lemma "x \<noteq> (y::(bool \<times> bool) bounded) \<Longrightarrow> z = x \<or> z = y"
nitpick [card = 1-5, expect = genuine]
oops
-lemma "True \<equiv> ((\<lambda>x\<Colon>bool. x) = (\<lambda>x. x))"
+lemma "True \<equiv> ((\<lambda>x::bool. x) = (\<lambda>x. x))"
nitpick [expect = none]
by (rule True_def)
@@ -183,7 +183,7 @@
nitpick [card = 1, expect = none]
by (rule Rep_rat_inverse)
-typedef check = "{x\<Colon>nat. x < 2}" by (rule exI[of _ 0], auto)
+typedef check = "{x::nat. x < 2}" by (rule exI[of _ 0], auto)
lemma "Rep_check (Abs_check n) = n \<Longrightarrow> n < 2"
nitpick [card = 1-3, expect = none]
--- a/src/HOL/Old_Number_Theory/Primes.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Old_Number_Theory/Primes.thy Tue Sep 01 22:32:58 2015 +0200
@@ -108,7 +108,7 @@
declare nat_mult_dvd_cancel_disj[presburger]
lemma nat_mult_dvd_cancel_disj'[presburger]:
- "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult.commute[of m k] mult.commute[of n k] by presburger
+ "(m::nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult.commute[of m k] mult.commute[of n k] by presburger
lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"
by presburger
--- a/src/HOL/Orderings.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Orderings.thy Tue Sep 01 22:32:58 2015 +0200
@@ -1034,21 +1034,21 @@
context order
begin
-definition mono :: "('a \<Rightarrow> 'b\<Colon>order) \<Rightarrow> bool" where
+definition mono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"mono f \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y)"
lemma monoI [intro?]:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
shows "(\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y) \<Longrightarrow> mono f"
unfolding mono_def by iprover
lemma monoD [dest?]:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
shows "mono f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
unfolding mono_def by iprover
lemma monoE:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
assumes "mono f"
assumes "x \<le> y"
obtains "f x \<le> f y"
@@ -1056,21 +1056,21 @@
from assms show "f x \<le> f y" by (simp add: mono_def)
qed
-definition antimono :: "('a \<Rightarrow> 'b\<Colon>order) \<Rightarrow> bool" where
+definition antimono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"antimono f \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> f x \<ge> f y)"
lemma antimonoI [intro?]:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
shows "(\<And>x y. x \<le> y \<Longrightarrow> f x \<ge> f y) \<Longrightarrow> antimono f"
unfolding antimono_def by iprover
lemma antimonoD [dest?]:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
shows "antimono f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<ge> f y"
unfolding antimono_def by iprover
lemma antimonoE:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
assumes "antimono f"
assumes "x \<le> y"
obtains "f x \<ge> f y"
@@ -1078,7 +1078,7 @@
from assms show "f x \<ge> f y" by (simp add: antimono_def)
qed
-definition strict_mono :: "('a \<Rightarrow> 'b\<Colon>order) \<Rightarrow> bool" where
+definition strict_mono :: "('a \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"strict_mono f \<longleftrightarrow> (\<forall>x y. x < y \<longrightarrow> f x < f y)"
lemma strict_monoI [intro?]:
@@ -1112,7 +1112,7 @@
begin
lemma mono_invE:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>order"
+ fixes f :: "'a \<Rightarrow> 'b::order"
assumes "mono f"
assumes "f x < f y"
obtains "x \<le> y"
@@ -1180,10 +1180,10 @@
lemma max_absorb2: "x \<le> y \<Longrightarrow> max x y = y"
by (simp add: max_def)
-lemma min_absorb2: "(y\<Colon>'a\<Colon>order) \<le> x \<Longrightarrow> min x y = y"
+lemma min_absorb2: "(y::'a::order) \<le> x \<Longrightarrow> min x y = y"
by (simp add:min_def)
-lemma max_absorb1: "(y\<Colon>'a\<Colon>order) \<le> x \<Longrightarrow> max x y = x"
+lemma max_absorb1: "(y::'a::order) \<le> x \<Longrightarrow> max x y = x"
by (simp add: max_def)
@@ -1409,7 +1409,7 @@
le_bool_def [simp]: "P \<le> Q \<longleftrightarrow> P \<longrightarrow> Q"
definition
- [simp]: "(P\<Colon>bool) < Q \<longleftrightarrow> \<not> P \<and> Q"
+ [simp]: "(P::bool) < Q \<longleftrightarrow> \<not> P \<and> Q"
definition
[simp]: "\<bottom> \<longleftrightarrow> False"
@@ -1457,7 +1457,7 @@
le_fun_def: "f \<le> g \<longleftrightarrow> (\<forall>x. f x \<le> g x)"
definition
- "(f\<Colon>'a \<Rightarrow> 'b) < g \<longleftrightarrow> f \<le> g \<and> \<not> (g \<le> f)"
+ "(f::'a \<Rightarrow> 'b) < g \<longleftrightarrow> f \<le> g \<and> \<not> (g \<le> f)"
instance ..
@@ -1620,4 +1620,3 @@
lemmas linorder_antisym_conv3 = linorder_class.antisym_conv3
end
-
--- a/src/HOL/Power.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Power.thy Tue Sep 01 22:32:58 2015 +0200
@@ -800,7 +800,7 @@
Premises cannot be weakened: consider the case where @{term "i=0"},
@{term "m=1"} and @{term "n=0"}.\<close>
lemma nat_power_less_imp_less:
- assumes nonneg: "0 < (i\<Colon>nat)"
+ assumes nonneg: "0 < (i::nat)"
assumes less: "i ^ m < i ^ n"
shows "m < n"
proof (cases "i = 1")
--- a/src/HOL/Probability/Borel_Space.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Probability/Borel_Space.thy Tue Sep 01 22:32:58 2015 +0200
@@ -81,7 +81,7 @@
by simp
lemma borel_measurableI:
- fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
+ fixes f :: "'a \<Rightarrow> 'x::topological_space"
assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
shows "f \<in> borel_measurable M"
unfolding borel_def
@@ -452,7 +452,7 @@
by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
lemma [measurable]:
- fixes a b :: "'a\<Colon>linorder_topology"
+ fixes a b :: "'a::linorder_topology"
shows lessThan_borel: "{..< a} \<in> sets borel"
and greaterThan_borel: "{a <..} \<in> sets borel"
and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
@@ -473,7 +473,7 @@
by auto
lemma eucl_ivals[measurable]:
- fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ fixes a b :: "'a::ordered_euclidean_space"
shows "{x. x <e a} \<in> sets borel"
and "{x. a <e x} \<in> sets borel"
and "{..a} \<in> sets borel"
@@ -597,7 +597,7 @@
using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
lemma borel_eq_box:
- "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a \<Colon> euclidean_space set))"
+ "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI1[OF borel_def])
fix M :: "'a set" assume "M \<in> {S. open S}"
@@ -611,13 +611,13 @@
lemma halfspace_gt_in_halfspace:
assumes i: "i \<in> A"
- shows "{x\<Colon>'a. a < x \<bullet> i} \<in>
- sigma_sets UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
+ shows "{x::'a. a < x \<bullet> i} \<in>
+ sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
(is "?set \<in> ?SIGMA")
proof -
interpret sigma_algebra UNIV ?SIGMA
by (intro sigma_algebra_sigma_sets) simp_all
- have *: "?set = (\<Union>n. UNIV - {x\<Colon>'a. x \<bullet> i < a + 1 / real (Suc n)})"
+ have *: "?set = (\<Union>n. UNIV - {x::'a. x \<bullet> i < a + 1 / real (Suc n)})"
proof (safe, simp_all add: not_less)
fix x :: 'a assume "a < x \<bullet> i"
with reals_Archimedean[of "x \<bullet> i - a"]
@@ -673,7 +673,7 @@
qed auto
lemma borel_eq_halfspace_ge:
- "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
fix a :: real and i :: 'a assume i: "(a, i) \<in> UNIV \<times> Basis"
@@ -683,7 +683,7 @@
qed auto
lemma borel_eq_halfspace_greater:
- "borel = sigma UNIV ((\<lambda> (a, i). {x\<Colon>'a\<Colon>euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
@@ -694,7 +694,7 @@
qed auto
lemma borel_eq_atMost:
- "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
+ "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
@@ -713,7 +713,7 @@
qed auto
lemma borel_eq_greaterThan:
- "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {x. a <e x}))"
+ "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. a <e x}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
@@ -740,7 +740,7 @@
qed auto
lemma borel_eq_lessThan:
- "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {x. x <e a}))"
+ "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. x <e a}))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
@@ -766,7 +766,7 @@
qed auto
lemma borel_eq_atLeastAtMost:
- "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
+ "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
(is "_ = ?SIGMA")
proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
fix a::'a
@@ -828,7 +828,7 @@
qed simp_all
lemma borel_measurable_halfspacesI:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a::real. S a i \<in> sets M)"
@@ -843,22 +843,22 @@
qed
lemma borel_measurable_iff_halfspace_le:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
lemma borel_measurable_iff_halfspace_less:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
lemma borel_measurable_iff_halfspace_ge:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
lemma borel_measurable_iff_halfspace_greater:
- fixes f :: "'a \<Rightarrow> 'c\<Colon>euclidean_space"
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
--- a/src/HOL/Product_Type.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Product_Type.thy Tue Sep 01 22:32:58 2015 +0200
@@ -189,7 +189,7 @@
end
lemma [code]:
- "HOL.equal (u\<Colon>unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+
+ "HOL.equal (u::unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+
code_printing
type_constructor unit \<rightharpoonup>
@@ -224,7 +224,7 @@
definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
"Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
-definition "prod = {f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
+definition "prod = {f. \<exists>a b. f = Pair_Rep (a::'a) (b::'b)}"
typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \<Rightarrow> 'b \<Rightarrow> bool) set"
unfolding prod_def by auto
@@ -1285,7 +1285,7 @@
unfolding image_def
proof(rule set_eqI,rule iffI)
fix x :: "'a \<times> 'c"
- assume "x \<in> {y\<Colon>'a \<times> 'c. \<exists>x\<Colon>'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
+ assume "x \<in> {y::'a \<times> 'c. \<exists>x::'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
then obtain y where "y \<in> A \<times> B" and "x = map_prod f g y" by blast
from \<open>image f A = A'\<close> and \<open>y \<in> A \<times> B\<close> have "f (fst y) \<in> A'" by auto
moreover from \<open>image g B = B'\<close> and \<open>y \<in> A \<times> B\<close> have "g (snd y) \<in> B'" by auto
--- a/src/HOL/Proofs/Extraction/Euclid.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Proofs/Extraction/Euclid.thy Tue Sep 01 22:32:58 2015 +0200
@@ -123,7 +123,7 @@
qed
qed
-lemma dvd_prod [iff]: "n dvd (PROD m\<Colon>nat:#mset (n # ns). m)"
+lemma dvd_prod [iff]: "n dvd (PROD m::nat:#mset (n # ns). m)"
by (simp add: msetprod_Un msetprod_singleton)
definition all_prime :: "nat list \<Rightarrow> bool" where
@@ -140,13 +140,13 @@
lemma split_all_prime:
assumes "all_prime ms" and "all_prime ns"
- shows "\<exists>qs. all_prime qs \<and> (PROD m\<Colon>nat:#mset qs. m) =
- (PROD m\<Colon>nat:#mset ms. m) * (PROD m\<Colon>nat:#mset ns. m)" (is "\<exists>qs. ?P qs \<and> ?Q qs")
+ shows "\<exists>qs. all_prime qs \<and> (PROD m::nat:#mset qs. m) =
+ (PROD m::nat:#mset ms. m) * (PROD m::nat:#mset ns. m)" (is "\<exists>qs. ?P qs \<and> ?Q qs")
proof -
from assms have "all_prime (ms @ ns)"
by (simp add: all_prime_append)
- moreover from assms have "(PROD m\<Colon>nat:#mset (ms @ ns). m) =
- (PROD m\<Colon>nat:#mset ms. m) * (PROD m\<Colon>nat:#mset ns. m)"
+ moreover from assms have "(PROD m::nat:#mset (ms @ ns). m) =
+ (PROD m::nat:#mset ms. m) * (PROD m::nat:#mset ns. m)"
by (simp add: msetprod_Un)
ultimately have "?P (ms @ ns) \<and> ?Q (ms @ ns)" ..
then show ?thesis ..
@@ -154,11 +154,11 @@
lemma all_prime_nempty_g_one:
assumes "all_prime ps" and "ps \<noteq> []"
- shows "Suc 0 < (PROD m\<Colon>nat:#mset ps. m)"
+ shows "Suc 0 < (PROD m::nat:#mset ps. m)"
using `ps \<noteq> []` `all_prime ps` unfolding One_nat_def [symmetric] by (induct ps rule: list_nonempty_induct)
(simp_all add: all_prime_simps msetprod_singleton msetprod_Un prime_gt_1_nat less_1_mult del: One_nat_def)
-lemma factor_exists: "Suc 0 < n \<Longrightarrow> (\<exists>ps. all_prime ps \<and> (PROD m\<Colon>nat:#mset ps. m) = n)"
+lemma factor_exists: "Suc 0 < n \<Longrightarrow> (\<exists>ps. all_prime ps \<and> (PROD m::nat:#mset ps. m) = n)"
proof (induct n rule: nat_wf_ind)
case (1 n)
from `Suc 0 < n`
@@ -169,21 +169,21 @@
assume "\<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k"
then obtain m k where m: "Suc 0 < m" and k: "Suc 0 < k" and mn: "m < n"
and kn: "k < n" and nmk: "n = m * k" by iprover
- from mn and m have "\<exists>ps. all_prime ps \<and> (PROD m\<Colon>nat:#mset ps. m) = m" by (rule 1)
- then obtain ps1 where "all_prime ps1" and prod_ps1_m: "(PROD m\<Colon>nat:#mset ps1. m) = m"
+ from mn and m have "\<exists>ps. all_prime ps \<and> (PROD m::nat:#mset ps. m) = m" by (rule 1)
+ then obtain ps1 where "all_prime ps1" and prod_ps1_m: "(PROD m::nat:#mset ps1. m) = m"
by iprover
- from kn and k have "\<exists>ps. all_prime ps \<and> (PROD m\<Colon>nat:#mset ps. m) = k" by (rule 1)
- then obtain ps2 where "all_prime ps2" and prod_ps2_k: "(PROD m\<Colon>nat:#mset ps2. m) = k"
+ from kn and k have "\<exists>ps. all_prime ps \<and> (PROD m::nat:#mset ps. m) = k" by (rule 1)
+ then obtain ps2 where "all_prime ps2" and prod_ps2_k: "(PROD m::nat:#mset ps2. m) = k"
by iprover
from `all_prime ps1` `all_prime ps2`
- have "\<exists>ps. all_prime ps \<and> (PROD m\<Colon>nat:#mset ps. m) =
- (PROD m\<Colon>nat:#mset ps1. m) * (PROD m\<Colon>nat:#mset ps2. m)"
+ have "\<exists>ps. all_prime ps \<and> (PROD m::nat:#mset ps. m) =
+ (PROD m::nat:#mset ps1. m) * (PROD m::nat:#mset ps2. m)"
by (rule split_all_prime)
with prod_ps1_m prod_ps2_k nmk show ?thesis by simp
next
assume "prime n" then have "all_prime [n]" by (simp add: all_prime_simps)
- moreover have "(PROD m\<Colon>nat:#mset [n]. m) = n" by (simp add: msetprod_singleton)
- ultimately have "all_prime [n] \<and> (PROD m\<Colon>nat:#mset [n]. m) = n" ..
+ moreover have "(PROD m::nat:#mset [n]. m) = n" by (simp add: msetprod_singleton)
+ ultimately have "all_prime [n] \<and> (PROD m::nat:#mset [n]. m) = n" ..
then show ?thesis ..
qed
qed
@@ -193,7 +193,7 @@
shows "\<exists>p. prime p \<and> p dvd n"
proof -
from N obtain ps where "all_prime ps"
- and prod_ps: "n = (PROD m\<Colon>nat:#mset ps. m)" using factor_exists
+ and prod_ps: "n = (PROD m::nat:#mset ps. m)" using factor_exists
by simp iprover
with N have "ps \<noteq> []"
by (auto simp add: all_prime_nempty_g_one msetprod_empty)
--- a/src/HOL/Quickcheck_Random.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Quickcheck_Random.thy Tue Sep 01 22:32:58 2015 +0200
@@ -131,7 +131,7 @@
\<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
definition random_fun_lift :: "(Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
- \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+ \<Rightarrow> Random.seed \<Rightarrow> (('a::term_of \<Rightarrow> 'b::typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
where
"random_fun_lift f =
random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
--- a/src/HOL/Quotient_Examples/Quotient_Int.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_Int.thy Tue Sep 01 22:32:58 2015 +0200
@@ -22,10 +22,10 @@
begin
quotient_definition
- "0 \<Colon> int" is "(0\<Colon>nat, 0\<Colon>nat)" done
+ "0 :: int" is "(0::nat, 0::nat)" done
quotient_definition
- "1 \<Colon> int" is "(1\<Colon>nat, 0\<Colon>nat)" done
+ "1 :: int" is "(1::nat, 0::nat)" done
fun
plus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
@@ -33,7 +33,7 @@
"plus_int_raw (x, y) (u, v) = (x + u, y + v)"
quotient_definition
- "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)" is "plus_int_raw" by auto
+ "(op +) :: (int \<Rightarrow> int \<Rightarrow> int)" is "plus_int_raw" by auto
fun
uminus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
@@ -41,10 +41,10 @@
"uminus_int_raw (x, y) = (y, x)"
quotient_definition
- "(uminus \<Colon> (int \<Rightarrow> int))" is "uminus_int_raw" by auto
+ "(uminus :: (int \<Rightarrow> int))" is "uminus_int_raw" by auto
definition
- minus_int_def: "z - w = z + (-w\<Colon>int)"
+ minus_int_def: "z - w = z + (-w::int)"
fun
times_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
@@ -95,13 +95,13 @@
le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" is "le_int_raw" by auto
definition
- less_int_def: "(z\<Colon>int) < w = (z \<le> w \<and> z \<noteq> w)"
+ less_int_def: "(z::int) < w = (z \<le> w \<and> z \<noteq> w)"
definition
- zabs_def: "\<bar>i\<Colon>int\<bar> = (if i < 0 then - i else i)"
+ zabs_def: "\<bar>i::int\<bar> = (if i < 0 then - i else i)"
definition
- zsgn_def: "sgn (i\<Colon>int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
+ zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
instance ..
@@ -182,10 +182,10 @@
begin
definition
- "(inf \<Colon> int \<Rightarrow> int \<Rightarrow> int) = min"
+ "(inf :: int \<Rightarrow> int \<Rightarrow> int) = min"
definition
- "(sup \<Colon> int \<Rightarrow> int \<Rightarrow> int) = max"
+ "(sup :: int \<Rightarrow> int \<Rightarrow> int) = max"
instance
by default
@@ -245,7 +245,7 @@
by (rule zmult_zless_mono2)
show "\<bar>i\<bar> = (if i < 0 then -i else i)"
by (simp only: zabs_def)
- show "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
+ show "sgn (i::int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
by (simp only: zsgn_def)
qed
--- a/src/HOL/Quotient_Examples/Quotient_Rat.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_Rat.thy Tue Sep 01 22:32:58 2015 +0200
@@ -32,10 +32,10 @@
begin
quotient_definition
- "0 \<Colon> rat" is "(0\<Colon>int, 1\<Colon>int)" by simp
+ "0 :: rat" is "(0::int, 1::int)" by simp
quotient_definition
- "1 \<Colon> rat" is "(1\<Colon>int, 1\<Colon>int)" by simp
+ "1 :: rat" is "(1::int, 1::int)" by simp
fun times_rat_raw where
"times_rat_raw (a :: int, b :: int) (c, d) = (a * c, b * d)"
@@ -54,10 +54,10 @@
"uminus_rat_raw (a :: int, b :: int) = (-a, b)"
quotient_definition
- "(uminus \<Colon> (rat \<Rightarrow> rat))" is "uminus_rat_raw" by fastforce
+ "(uminus :: (rat \<Rightarrow> rat))" is "uminus_rat_raw" by fastforce
definition
- minus_rat_def: "a - b = a + (-b\<Colon>rat)"
+ minus_rat_def: "a - b = a + (-b::rat)"
fun le_rat_raw where
"le_rat_raw (a :: int, b) (c, d) \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
@@ -92,13 +92,13 @@
qed
definition
- less_rat_def: "(z\<Colon>rat) < w = (z \<le> w \<and> z \<noteq> w)"
+ less_rat_def: "(z::rat) < w = (z \<le> w \<and> z \<noteq> w)"
definition
- rabs_rat_def: "\<bar>i\<Colon>rat\<bar> = (if i < 0 then - i else i)"
+ rabs_rat_def: "\<bar>i::rat\<bar> = (if i < 0 then - i else i)"
definition
- sgn_rat_def: "sgn (i\<Colon>rat) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
+ sgn_rat_def: "sgn (i::rat) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
instance by intro_classes
(auto simp add: rabs_rat_def sgn_rat_def)
@@ -259,7 +259,7 @@
assume "b \<noteq> 0"
then have "a * b \<le> (a div b + 1) * b * b"
by (metis mult.commute div_mult_self1_is_id less_int_def linorder_le_cases zdiv_mono1 zdiv_mono1_neg zle_add1_eq_le)
- then show "\<exists>z\<Colon>int. a * b \<le> z * b * b" by auto
+ then show "\<exists>z::int. a * b \<le> z * b * b" by auto
qed
qed
*)
--- a/src/HOL/Real.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Real.thy Tue Sep 01 22:32:58 2015 +0200
@@ -491,7 +491,7 @@
by transfer (simp add: ac_simps realrel_def)
show "(a + b) * c = a * c + b * c"
by transfer (simp add: distrib_right realrel_def)
- show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
+ show "(0::real) \<noteq> (1::real)"
by transfer (simp add: realrel_def)
show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
apply transfer
@@ -1975,7 +1975,7 @@
instantiation real :: equal
begin
-definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
+definition "HOL.equal (x::real) y \<longleftrightarrow> x - y = 0"
instance proof
qed (simp add: equal_real_def)
--- a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Tue Sep 01 22:32:58 2015 +0200
@@ -427,7 +427,7 @@
h4 = ce_init\<rparr>"
have steps_to_steps':
"steps
- (\<lambda>n\<Colon>nat. word_of_int (x (int n)))
+ (\<lambda>n::nat. word_of_int (x (int n)))
(from_chain ?INIT_CHAIN, from_chain ?INIT_CHAIN)
80 =
from_chain_pair (
--- a/src/HOL/String.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/String.thy Tue Sep 01 22:32:58 2015 +0200
@@ -374,7 +374,7 @@
definition size_literal :: "literal \<Rightarrow> nat"
where
- [code]: "size_literal (s\<Colon>literal) = 0"
+ [code]: "size_literal (s::literal) = 0"
instance ..
--- a/src/HOL/Sum_Type.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Sum_Type.thy Tue Sep 01 22:32:58 2015 +0200
@@ -98,7 +98,7 @@
proof -
fix P
fix s :: "'a + 'b"
- assume x: "\<And>x\<Colon>'a. P (Inl x)" and y: "\<And>y\<Colon>'b. P (Inr y)"
+ assume x: "\<And>x::'a. P (Inl x)" and y: "\<And>y::'b. P (Inr y)"
then show "P s" by (auto intro: sumE [of s])
qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)
@@ -155,7 +155,7 @@
lemma surjective_sum: "case_sum (\<lambda>x::'a. f (Inl x)) (\<lambda>y::'b. f (Inr y)) = f"
proof
fix s :: "'a + 'b"
- show "(case s of Inl (x\<Colon>'a) \<Rightarrow> f (Inl x) | Inr (y\<Colon>'b) \<Rightarrow> f (Inr y)) = f s"
+ show "(case s of Inl (x::'a) \<Rightarrow> f (Inl x) | Inr (y::'b) \<Rightarrow> f (Inr y)) = f s"
by (cases s) simp_all
qed
@@ -186,7 +186,7 @@
assumes "Suml f = Suml g" shows "f = g"
proof
fix x :: 'a
- let ?s = "Inl x \<Colon> 'a + 'b"
+ let ?s = "Inl x :: 'a + 'b"
from assms have "Suml f ?s = Suml g ?s" by simp
then show "f x = g x" by simp
qed
@@ -195,7 +195,7 @@
assumes "Sumr f = Sumr g" shows "f = g"
proof
fix x :: 'b
- let ?s = "Inr x \<Colon> 'a + 'b"
+ let ?s = "Inr x :: 'a + 'b"
from assms have "Sumr f ?s = Sumr g ?s" by simp
then show "f x = g x" by simp
qed
--- a/src/HOL/TPTP/THF_Arith.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/TPTP/THF_Arith.thy Tue Sep 01 22:32:58 2015 +0200
@@ -16,17 +16,17 @@
overloading rat_is_int \<equiv> "is_int :: rat \<Rightarrow> bool"
begin
- definition "rat_is_int (q\<Colon>rat) \<longleftrightarrow> (\<exists>n\<Colon>int. q = of_int n)"
+ definition "rat_is_int (q::rat) \<longleftrightarrow> (\<exists>n::int. q = of_int n)"
end
overloading real_is_int \<equiv> "is_int :: real \<Rightarrow> bool"
begin
- definition "real_is_int (x\<Colon>real) \<longleftrightarrow> x \<in> \<int>"
+ definition "real_is_int (x::real) \<longleftrightarrow> x \<in> \<int>"
end
overloading real_is_rat \<equiv> "is_rat :: real \<Rightarrow> bool"
begin
- definition "real_is_rat (x\<Colon>real) \<longleftrightarrow> x \<in> \<rat>"
+ definition "real_is_rat (x::real) \<longleftrightarrow> x \<in> \<rat>"
end
consts
@@ -36,32 +36,32 @@
overloading rat_to_int \<equiv> "to_int :: rat \<Rightarrow> int"
begin
- definition "rat_to_int (q\<Colon>rat) = floor q"
+ definition "rat_to_int (q::rat) = floor q"
end
overloading real_to_int \<equiv> "to_int :: real \<Rightarrow> int"
begin
- definition "real_to_int (x\<Colon>real) = floor x"
+ definition "real_to_int (x::real) = floor x"
end
overloading int_to_rat \<equiv> "to_rat :: int \<Rightarrow> rat"
begin
- definition "int_to_rat (n\<Colon>int) = (of_int n\<Colon>rat)"
+ definition "int_to_rat (n::int) = (of_int n::rat)"
end
overloading real_to_rat \<equiv> "to_rat :: real \<Rightarrow> rat"
begin
- definition "real_to_rat (x\<Colon>real) = (inv of_rat x\<Colon>rat)"
+ definition "real_to_rat (x::real) = (inv of_rat x::rat)"
end
overloading int_to_real \<equiv> "to_real :: int \<Rightarrow> real"
begin
- definition "int_to_real (n\<Colon>int) = real n"
+ definition "int_to_real (n::int) = real n"
end
overloading rat_to_real \<equiv> "to_real :: rat \<Rightarrow> real"
begin
- definition "rat_to_real (x\<Colon>rat) = (of_rat x\<Colon>real)"
+ definition "rat_to_real (x::rat) = (of_rat x::real)"
end
declare
@@ -75,16 +75,16 @@
int_to_real_def [simp]
rat_to_real_def [simp]
-lemma to_rat_is_int [intro, simp]: "is_int (to_rat (n\<Colon>int))"
+lemma to_rat_is_int [intro, simp]: "is_int (to_rat (n::int))"
by (metis int_to_rat_def rat_is_int_def)
-lemma to_real_is_int [intro, simp]: "is_int (to_real (n\<Colon>int))"
+lemma to_real_is_int [intro, simp]: "is_int (to_real (n::int))"
by (metis Ints_real_of_int int_to_real_def real_is_int_def)
-lemma to_real_is_rat [intro, simp]: "is_rat (to_real (q\<Colon>rat))"
+lemma to_real_is_rat [intro, simp]: "is_rat (to_real (q::rat))"
by (metis Rats_of_rat rat_to_real_def real_is_rat_def)
-lemma inj_of_rat [intro, simp]: "inj (of_rat\<Colon>rat\<Rightarrow>real)"
+lemma inj_of_rat [intro, simp]: "inj (of_rat::rat\<Rightarrow>real)"
by (metis injI of_rat_eq_iff)
end
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy Tue Sep 01 22:32:58 2015 +0200
@@ -348,7 +348,7 @@
(*FIXME move these examples elsewhere*)
(*
-lemma "\<forall>(Xj\<Colon>TPTP_Interpret.ind) Xk\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>(Xj::TPTP_Interpret.ind) Xk::TPTP_Interpret.ind.
bnd_cCKB6_BLACK Xj Xk \<longrightarrow>
bnd_cCKB6_BLACK (bnd_s (bnd_s (bnd_s Xj))) (bnd_s Xk)"
apply (tactic {*nth (nth just_the_tacs 0) 0*})
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy Tue Sep 01 22:32:58 2015 +0200
@@ -27,8 +27,8 @@
(*
(* SEU581^2.p_nux *)
(* (Annotated_step ("inode1", "bind"), *)
-lemma "\<forall>(SV5\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- SV6\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>(SV5::TPTP_Interpret.ind \<Rightarrow> bool)
+ SV6::TPTP_Interpret.ind.
(bnd_in (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
(bnd_powerset bnd_sK1_A) =
bnd_in (bnd_dsetconstr SV6 SV5)
@@ -66,7 +66,7 @@
done
(* (Annotated_step ("inode2", "bind"), *)
-lemma "\<forall>(SV7\<Colon>TPTP_Interpret.ind) SV8\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>(SV7::TPTP_Interpret.ind) SV8::TPTP_Interpret.ind.
(bnd_subset SV8 SV7 =
bnd_subset (bnd_dsetconstr bnd_sK1_A bnd_sK2_SY15)
bnd_sK1_A) =
@@ -303,12 +303,12 @@
(*SEU882^5*)
(*
lemma
- "\<forall>(SV2\<Colon>TPTP_Interpret.ind)
- SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind.
+ "\<forall>(SV2::TPTP_Interpret.ind)
+ SV1::TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind.
(SV1 SV2 = bnd_sK1_Xy) =
False
\<Longrightarrow>
- \<forall>SV2\<Colon>TPTP_Interpret.ind.
+ \<forall>SV2::TPTP_Interpret.ind.
(bnd_sK1_Xy = bnd_sK1_Xy) =
False"
ML_prf {*
@@ -462,12 +462,12 @@
SEU602_2_bnd_in :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
(* (Annotated_step ("113", "extuni_func"), *)
-lemma "\<forall>SV49\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+lemma "\<forall>SV49::TPTP_Interpret.ind \<Rightarrow> bool.
(SV49 =
- (\<lambda>SY23\<Colon>TPTP_Interpret.ind.
+ (\<lambda>SY23::TPTP_Interpret.ind.
\<not> SEU602_2_bnd_in SY23 SEU602_2_bnd_sK2_SY17)) =
False \<Longrightarrow>
- \<forall>SV49\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ \<forall>SV49::TPTP_Interpret.ind \<Rightarrow> bool.
(SV49 (SEU602_2_bnd_sK7_E SV49) =
(\<not> SEU602_2_bnd_in (SEU602_2_bnd_sK7_E SV49) SEU602_2_bnd_sK2_SY17)) =
False"
@@ -478,12 +478,12 @@
SEV405_5_bnd_sK1_SY2 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
SEV405_5_bnd_cA :: bool
-lemma "\<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- (\<forall>SY2\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<forall>SY2::TPTP_Interpret.ind.
\<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or>
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) =
False \<Longrightarrow>
- \<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ \<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool.
(\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or>
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
False"
@@ -1046,28 +1046,28 @@
\<Rightarrow> TPTP_Interpret.ind
\<Rightarrow> TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
-lemma "\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
- (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
- (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind)
+ (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind)
+ (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow>
-\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
- (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
- (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+\<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind)
+ (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind)
+ (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 = SV3) = True \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
by (tactic {*nonfull_extcnf_combined_tac @{context} [Not_neg]*})
lemma "
-\<forall>(SV8\<Colon>TPTP_Interpret.ind) (SV6\<Colon>TPTP_Interpret.ind)
- (SV4\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
- (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+\<forall>(SV8::TPTP_Interpret.ind) (SV6::TPTP_Interpret.ind)
+ (SV4::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind)
+ (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 \<noteq> SV3 \<or> SV2 \<noteq> SV4) = False \<or>
PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \<Longrightarrow>
-\<forall>(SV4\<Colon>TPTP_Interpret.ind) (SV8\<Colon>TPTP_Interpret.ind)
- (SV6\<Colon>TPTP_Interpret.ind) (SV2\<Colon>TPTP_Interpret.ind)
- (SV3\<Colon>TPTP_Interpret.ind) SV1\<Colon>TPTP_Interpret.ind.
+\<forall>(SV4::TPTP_Interpret.ind) (SV8::TPTP_Interpret.ind)
+ (SV6::TPTP_Interpret.ind) (SV2::TPTP_Interpret.ind)
+ (SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
by (tactic {*nonfull_extcnf_combined_tac @{context} [Or_neg]*})
@@ -1081,70 +1081,70 @@
NUM016_5_bnd_less :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind \<Rightarrow> bool"
(* (Annotated_step ("6", "unfold_def"), *)
-lemma "((((((((((((\<forall>X\<Colon>TPTP_Interpret.ind. \<not> NUM016_5_bnd_less X X) \<and>
- (\<forall>(X\<Colon>TPTP_Interpret.ind)
- Y\<Colon>TPTP_Interpret.ind.
+lemma "((((((((((((\<forall>X::TPTP_Interpret.ind. \<not> NUM016_5_bnd_less X X) \<and>
+ (\<forall>(X::TPTP_Interpret.ind)
+ Y::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_less X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \<and>
- (\<forall>(X\<Colon>TPTP_Interpret.ind)
- (Y\<Colon>TPTP_Interpret.ind)
- Z\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \<and>
+ (\<forall>(X::TPTP_Interpret.ind)
+ (Y::TPTP_Interpret.ind)
+ Z::TPTP_Interpret.ind.
(\<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_divides Y Z) \<or>
NUM016_5_bnd_divides X Z)) \<and>
- (\<forall>(X\<Colon>TPTP_Interpret.ind) Y\<Colon>TPTP_Interpret.ind.
+ (\<forall>(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_divides X Y \<or> \<not> NUM016_5_bnd_less Y X)) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind.
NUM016_5_bnd_less X (NUM016_5_bnd_factorial_plus_one X))) \<and>
- (\<forall>(X\<Colon>TPTP_Interpret.ind) Y\<Colon>TPTP_Interpret.ind.
+ (\<forall>(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_divides X (NUM016_5_bnd_factorial_plus_one Y) \<or>
NUM016_5_bnd_less Y X)) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind.
NUM016_5_bnd_prime X \<or>
NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor X) X)) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind.
NUM016_5_bnd_prime X \<or>
NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor X))) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind.
NUM016_5_bnd_prime X \<or>
NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor X) X)) \<and>
NUM016_5_bnd_prime NUM016_5_bnd_a) \<and>
- (\<forall>X\<Colon>TPTP_Interpret.ind.
+ (\<forall>X::TPTP_Interpret.ind.
(\<not> NUM016_5_bnd_prime X \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a X) \<or>
NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a) X)) =
True \<Longrightarrow>
- (\<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ (\<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> \<not> (\<not> (\<forall>SX0::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_less SX0 SX0) \<or>
- \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
- SX1\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>(SX0::TPTP_Interpret.ind)
+ SX1::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_less SX0 SX1 \<or> \<not> NUM016_5_bnd_less SX1 SX0)) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
NUM016_5_bnd_divides SX0 SX0)) \<or>
- \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
- (SX1\<Colon>TPTP_Interpret.ind)
- SX2\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>(SX0::TPTP_Interpret.ind)
+ (SX1::TPTP_Interpret.ind)
+ SX2::TPTP_Interpret.ind.
(\<not> NUM016_5_bnd_divides SX0 SX1 \<or>
\<not> NUM016_5_bnd_divides SX1 SX2) \<or>
NUM016_5_bnd_divides SX0 SX2)) \<or>
- \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind)
- SX1\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>(SX0::TPTP_Interpret.ind)
+ SX1::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_divides SX0 SX1 \<or>
\<not> NUM016_5_bnd_less SX1 SX0)) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
NUM016_5_bnd_less SX0 (NUM016_5_bnd_factorial_plus_one SX0))) \<or>
- \<not> (\<forall>(SX0\<Colon>TPTP_Interpret.ind) SX1\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>(SX0::TPTP_Interpret.ind) SX1::TPTP_Interpret.ind.
\<not> NUM016_5_bnd_divides SX0 (NUM016_5_bnd_factorial_plus_one SX1) \<or>
NUM016_5_bnd_less SX1 SX0)) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
NUM016_5_bnd_prime SX0 \<or>
NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor SX0) SX0)) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
NUM016_5_bnd_prime SX0 \<or> NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor SX0))) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
NUM016_5_bnd_prime SX0 \<or>
NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor SX0)
SX0)) \<or>
\<not> NUM016_5_bnd_prime NUM016_5_bnd_a) \<or>
- \<not> (\<forall>SX0\<Colon>TPTP_Interpret.ind.
+ \<not> (\<forall>SX0::TPTP_Interpret.ind.
(\<not> NUM016_5_bnd_prime SX0 \<or> \<not> NUM016_5_bnd_less NUM016_5_bnd_a SX0) \<or>
NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a)
SX0))) =
@@ -1248,7 +1248,7 @@
(*test that nullary skolem terms are OK*)
(* (Annotated_step ("79", "extcnf_forall_neg"), *)
-lemma "(\<forall>SX0\<Colon>TPTP_Interpret.ind.
+lemma "(\<forall>SX0::TPTP_Interpret.ind.
AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola SX0) =
False \<Longrightarrow>
AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola AGT037_2_bnd_sK1_SX0 =
@@ -1256,27 +1256,27 @@
by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
(* (Annotated_step ("202", "extcnf_forall_neg"), *)
-lemma "\<forall>(SV13\<Colon>TPTP_Interpret.ind) (SV39\<Colon>AGT037_2_bnd_mu) (SV29\<Colon>AGT037_2_bnd_mu)
- SV45\<Colon>TPTP_Interpret.ind.
- ((((\<forall>SY68\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>(SV13::TPTP_Interpret.ind) (SV39::AGT037_2_bnd_mu) (SV29::AGT037_2_bnd_mu)
+ SV45::TPTP_Interpret.ind.
+ ((((\<forall>SY68::TPTP_Interpret.ind.
\<not> AGT037_2_bnd_a1 SV45 SY68 \<or>
AGT037_2_bnd_likes SV29 SV39 SY68) =
False \<or>
- (\<not> (\<forall>SY69\<Colon>TPTP_Interpret.ind.
+ (\<not> (\<forall>SY69::TPTP_Interpret.ind.
\<not> AGT037_2_bnd_a2 SV45 SY69 \<or>
AGT037_2_bnd_likes SV29 SV39 SY69)) =
True) \<or>
AGT037_2_bnd_likes SV29 SV39 SV45 = False) \<or>
AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \<or>
AGT037_2_bnd_a3 SV13 SV45 = False \<Longrightarrow>
- \<forall>(SV29\<Colon>AGT037_2_bnd_mu) (SV39\<Colon>AGT037_2_bnd_mu) (SV13\<Colon>TPTP_Interpret.ind)
- SV45\<Colon>TPTP_Interpret.ind.
+ \<forall>(SV29::AGT037_2_bnd_mu) (SV39::AGT037_2_bnd_mu) (SV13::TPTP_Interpret.ind)
+ SV45::TPTP_Interpret.ind.
((((\<not> AGT037_2_bnd_a1 SV45
(AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45) \<or>
AGT037_2_bnd_likes SV29 SV39
(AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45)) =
False \<or>
- (\<not> (\<forall>SY69\<Colon>TPTP_Interpret.ind.
+ (\<not> (\<forall>SY69::TPTP_Interpret.ind.
\<not> AGT037_2_bnd_a2 SV45 SY69 \<or>
AGT037_2_bnd_likes SV29 SV39 SY69)) =
True) \<or>
@@ -1555,10 +1555,10 @@
bnd_addition bnd_sK2_X2 bnd_sK1_X1) =
True \<Longrightarrow>
(bnd_sup
- (\<lambda>SX0\<Colon>TPTP_Interpret.ind.
+ (\<lambda>SX0::TPTP_Interpret.ind.
SX0 = bnd_sK1_X1 \<or> SX0 = bnd_sK2_X2) \<noteq>
bnd_sup
- (\<lambda>SX0\<Colon>TPTP_Interpret.ind.
+ (\<lambda>SX0::TPTP_Interpret.ind.
SX0 = bnd_sK2_X2 \<or> SX0 = bnd_sK1_X1)) =
True"
by (tactic {*rtac (leo2_tac @{context} (hd prob_names) "20") 1*})
@@ -2011,60 +2011,60 @@
(* (Annotated_step ("12", "unfold_def"), *)
lemma "bnd_mor =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool)
+ (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
X U \<or> Y U) \<and>
bnd_mnot =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
\<not> X U) \<and>
bnd_mimplies =
- (\<lambda>U\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
+ (\<lambda>U::TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
bnd_mbox_s4 =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) X\<Colon>TPTP_Interpret.ind.
- \<forall>Y\<Colon>TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) X::TPTP_Interpret.ind.
+ \<forall>Y::TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
bnd_mand =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool)
+ (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
X U \<and> Y U) \<and>
bnd_ixor =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_inot (bnd_iequiv P Q)) \<and>
bnd_ivalid = All \<and>
- bnd_itrue = (\<lambda>W\<Colon>TPTP_Interpret.ind. True) \<and>
+ bnd_itrue = (\<lambda>W::TPTP_Interpret.ind. True) \<and>
bnd_isatisfiable = Ex \<and>
bnd_ior =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
bnd_inot =
- (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mnot (bnd_mbox_s4 P)) \<and>
bnd_iinvalid =
- (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- \<forall>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+ (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool.
+ \<forall>W::TPTP_Interpret.ind. \<not> Phi W) \<and>
bnd_iimplies =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
bnd_iimplied =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
bnd_ifalse = bnd_inot bnd_itrue \<and>
bnd_iequiv =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and>
bnd_icountersatisfiable =
- (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- \<exists>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
- bnd_iatom = (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
+ (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool.
+ \<exists>W::TPTP_Interpret.ind. \<not> Phi W) \<and>
+ bnd_iatom = (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
bnd_iand = bnd_mand \<and>
- (\<forall>(X\<Colon>TPTP_Interpret.ind) (Y\<Colon>TPTP_Interpret.ind)
- Z\<Colon>TPTP_Interpret.ind.
+ (\<forall>(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind)
+ Z::TPTP_Interpret.ind.
bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) \<Longrightarrow>
- (\<forall>(X\<Colon>TPTP_Interpret.ind) (Y\<Colon>TPTP_Interpret.ind)
- Z\<Colon>TPTP_Interpret.ind.
+ (\<forall>(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind)
+ Z::TPTP_Interpret.ind.
bnd_irel X Y \<and> bnd_irel Y Z \<longrightarrow> bnd_irel X Z) =
True"
(* by (tactic {*tectoc @{context}*}) *)
@@ -2072,61 +2072,61 @@
(* (Annotated_step ("11", "unfold_def"), *)
lemma "bnd_mor =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool)
+ (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
X U \<or> Y U) \<and>
bnd_mnot =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
\<not> X U) \<and>
bnd_mimplies =
- (\<lambda>U\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
+ (\<lambda>U::TPTP_Interpret.ind \<Rightarrow> bool. bnd_mor (bnd_mnot U)) \<and>
bnd_mbox_s4 =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) X\<Colon>TPTP_Interpret.ind.
- \<forall>Y\<Colon>TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool) X::TPTP_Interpret.ind.
+ \<forall>Y::TPTP_Interpret.ind. bnd_irel X Y \<longrightarrow> P Y) \<and>
bnd_mand =
- (\<lambda>(X\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- (Y\<Colon>TPTP_Interpret.ind \<Rightarrow> bool) U\<Colon>TPTP_Interpret.ind.
+ (\<lambda>(X::TPTP_Interpret.ind \<Rightarrow> bool)
+ (Y::TPTP_Interpret.ind \<Rightarrow> bool) U::TPTP_Interpret.ind.
X U \<and> Y U) \<and>
bnd_ixor =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_inot (bnd_iequiv P Q)) \<and>
bnd_ivalid = All \<and>
- bnd_itrue = (\<lambda>W\<Colon>TPTP_Interpret.ind. True) \<and>
+ bnd_itrue = (\<lambda>W::TPTP_Interpret.ind. True) \<and>
bnd_isatisfiable = Ex \<and>
bnd_ior =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
bnd_inot =
- (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mnot (bnd_mbox_s4 P)) \<and>
bnd_iinvalid =
- (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- \<forall>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
+ (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool.
+ \<forall>W::TPTP_Interpret.ind. \<not> Phi W) \<and>
bnd_iimplies =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \<and>
bnd_iimplied =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool. bnd_iimplies Q P) \<and>
bnd_ifalse = bnd_inot bnd_itrue \<and>
bnd_iequiv =
- (\<lambda>(P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool)
- Q\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<lambda>(P::TPTP_Interpret.ind \<Rightarrow> bool)
+ Q::TPTP_Interpret.ind \<Rightarrow> bool.
bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \<and>
bnd_icountersatisfiable =
- (\<lambda>Phi\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- \<exists>W\<Colon>TPTP_Interpret.ind. \<not> Phi W) \<and>
- bnd_iatom = (\<lambda>P\<Colon>TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
+ (\<lambda>Phi::TPTP_Interpret.ind \<Rightarrow> bool.
+ \<exists>W::TPTP_Interpret.ind. \<not> Phi W) \<and>
+ bnd_iatom = (\<lambda>P::TPTP_Interpret.ind \<Rightarrow> bool. P) \<and>
bnd_iand = bnd_mand \<and>
bnd_ivalid
(bnd_iimplies (bnd_iatom bnd_q) (bnd_iatom bnd_r)) \<Longrightarrow>
- (\<forall>SY161\<Colon>TPTP_Interpret.ind.
- \<not> (\<forall>SY162\<Colon>TPTP_Interpret.ind.
+ (\<forall>SY161::TPTP_Interpret.ind.
+ \<not> (\<forall>SY162::TPTP_Interpret.ind.
bnd_irel SY161 SY162 \<longrightarrow> bnd_q SY162) \<or>
- (\<forall>SY163\<Colon>TPTP_Interpret.ind.
+ (\<forall>SY163::TPTP_Interpret.ind.
bnd_irel SY161 SY163 \<longrightarrow> bnd_r SY163)) =
True"
(* by (tactic {*tectoc @{context}*}) *)
@@ -2147,12 +2147,12 @@
*)
(* (Annotated_step ("11", "extcnf_forall_neg"), *)
-lemma "\<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
- (\<forall>SY2\<Colon>TPTP_Interpret.ind.
+lemma "\<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool.
+ (\<forall>SY2::TPTP_Interpret.ind.
\<not> (\<not> (\<not> SV1 SY2 \<or> SEV405_5_bnd_cA) \<or>
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 SY2))) =
False \<Longrightarrow>
- \<forall>SV1\<Colon>TPTP_Interpret.ind \<Rightarrow> bool.
+ \<forall>SV1::TPTP_Interpret.ind \<Rightarrow> bool.
(\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or>
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
False"
@@ -2184,7 +2184,7 @@
(*from SYO198^5.p.out*)
(* [[(Annotated_step ("11", "extcnf_forall_special_pos"), *)
-lemma "(\<forall>SX0\<Colon>bool \<Rightarrow> bool.
+lemma "(\<forall>SX0::bool \<Rightarrow> bool.
\<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
True \<Longrightarrow>
(\<not> \<not> (\<not> True \<or> \<not> True)) = True"
@@ -2192,7 +2192,7 @@
done
(* (Annotated_step ("13", "extcnf_forall_special_pos"), *)
-lemma "(\<forall>SX0\<Colon>bool \<Rightarrow> bool.
+lemma "(\<forall>SX0::bool \<Rightarrow> bool.
\<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
True \<Longrightarrow>
(\<not> \<not> (\<not> bnd_sK1_Xx \<or> \<not> bnd_sK2_Xy)) = True"
@@ -2200,9 +2200,9 @@
done
(* [[(Annotated_step ("8", "polarity_switch"), *)
-lemma "(\<forall>(Xx\<Colon>bool) (Xy\<Colon>bool) Xz\<Colon>bool. True \<and> True \<longrightarrow> True) =
+lemma "(\<forall>(Xx::bool) (Xy::bool) Xz::bool. True \<and> True \<longrightarrow> True) =
False \<Longrightarrow>
- (\<not> (\<forall>(Xx\<Colon>bool) (Xy\<Colon>bool) Xz\<Colon>bool.
+ (\<not> (\<forall>(Xx::bool) (Xy::bool) Xz::bool.
True \<and> True \<longrightarrow> True)) =
True"
apply (tactic {*nonfull_extcnf_combined_tac @{context} [Polarity_switch]*})
--- a/src/HOL/Transcendental.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Transcendental.thy Tue Sep 01 22:32:58 2015 +0200
@@ -171,13 +171,13 @@
using powser_times_n_limit_0 [of "inverse x"]
by (simp add: norm_divide divide_simps)
-lemma lim_1_over_n: "((\<lambda>n. 1 / of_nat n) ---> (0::'a\<Colon>real_normed_field)) sequentially"
+lemma lim_1_over_n: "((\<lambda>n. 1 / of_nat n) ---> (0::'a::real_normed_field)) sequentially"
apply (clarsimp simp: lim_sequentially norm_divide dist_norm divide_simps)
apply (auto simp: mult_ac dest!: ex_less_of_nat_mult [of _ 1])
by (metis le_eq_less_or_eq less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono
of_nat_less_0_iff of_nat_less_iff zero_less_mult_iff zero_less_one)
-lemma lim_inverse_n: "((\<lambda>n. inverse(of_nat n)) ---> (0::'a\<Colon>real_normed_field)) sequentially"
+lemma lim_inverse_n: "((\<lambda>n. inverse(of_nat n)) ---> (0::'a::real_normed_field)) sequentially"
using lim_1_over_n
by (simp add: inverse_eq_divide)
--- a/src/HOL/UNITY/Follows.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/UNITY/Follows.thy Tue Sep 01 22:32:58 2015 +0200
@@ -172,7 +172,7 @@
instantiation multiset :: (order) ordered_ab_semigroup_add
begin
-definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
+definition less_multiset :: "'a::order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
"M' < M \<longleftrightarrow> M' #<# M"
definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
--- a/src/HOL/Word/Word.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/Word/Word.thy Tue Sep 01 22:32:58 2015 +0200
@@ -43,7 +43,7 @@
"uint a = uint b \<Longrightarrow> a = b"
by (simp add: word_uint_eq_iff)
-definition word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word"
+definition word_of_int :: "int \<Rightarrow> 'a::len0 word"
where
-- {* representation of words using unsigned or signed bins,
only difference in these is the type class *}
@@ -2707,7 +2707,7 @@
by (auto simp add: word_ubin.eq_norm nth_bintr nth_2p_bin)
lemma nth_w2p:
- "((2\<Colon>'a\<Colon>len word) ^ n) !! m \<longleftrightarrow> m = n \<and> m < len_of TYPE('a\<Colon>len)"
+ "((2::'a::len word) ^ n) !! m \<longleftrightarrow> m = n \<and> m < len_of TYPE('a::len)"
unfolding test_bit_2p [symmetric] word_of_int [symmetric]
by (simp add: of_int_power)
@@ -3736,7 +3736,7 @@
lemma test_bit_split:
"word_split c = (a, b) \<Longrightarrow>
- (\<forall>n\<Colon>nat. b !! n \<longleftrightarrow> n < size b \<and> c !! n) \<and> (\<forall>m\<Colon>nat. a !! m \<longleftrightarrow> m < size a \<and> c !! (m + size b))"
+ (\<forall>n::nat. b !! n \<longleftrightarrow> n < size b \<and> c !! n) \<and> (\<forall>m::nat. a !! m \<longleftrightarrow> m < size a \<and> c !! (m + size b))"
by (simp add: test_bit_split')
lemma test_bit_split_eq: "word_split c = (a, b) <->
--- a/src/HOL/ZF/HOLZF.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ZF/HOLZF.thy Tue Sep 01 22:32:58 2015 +0200
@@ -583,8 +583,8 @@
ultimately have "False" using u by arith
}
note lemma_nat2Nat = this
- have th:"\<And>x y. \<not> (x < y \<and> (\<forall>(m\<Colon>nat). y \<noteq> x + m))" by presburger
- have th': "\<And>x y. \<not> (x \<noteq> y \<and> (\<not> x < y) \<and> (\<forall>(m\<Colon>nat). x \<noteq> y + m))" by presburger
+ have th:"\<And>x y. \<not> (x < y \<and> (\<forall>(m::nat). y \<noteq> x + m))" by presburger
+ have th': "\<And>x y. \<not> (x \<noteq> y \<and> (\<not> x < y) \<and> (\<forall>(m::nat). x \<noteq> y + m))" by presburger
show ?thesis
apply (auto simp add: inj_on_def)
apply (case_tac "x = y")
--- a/src/HOL/ex/Dedekind_Real.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/Dedekind_Real.thy Tue Sep 01 22:32:58 2015 +0200
@@ -102,7 +102,7 @@
preal_inverse_def:
"inverse R == Abs_preal (inverse_set (Rep_preal R))"
-definition "R div S = R * inverse (S\<Colon>preal)"
+definition "R div S = R * inverse (S::preal)"
definition
preal_one_def:
@@ -220,10 +220,10 @@
begin
definition
- "(inf \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = min"
+ "(inf :: preal \<Rightarrow> preal \<Rightarrow> preal) = min"
definition
- "(sup \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = max"
+ "(sup :: preal \<Rightarrow> preal \<Rightarrow> preal) = max"
instance
by intro_classes
@@ -1229,7 +1229,7 @@
(\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)"
definition
- real_less_def: "x < (y\<Colon>real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
+ real_less_def: "x < (y::real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
definition
real_abs_def: "abs (r::real) = (if r < 0 then - r else r)"
@@ -1561,10 +1561,10 @@
begin
definition
- "(inf \<Colon> real \<Rightarrow> real \<Rightarrow> real) = min"
+ "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
definition
- "(sup \<Colon> real \<Rightarrow> real \<Rightarrow> real) = max"
+ "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
instance
by default (auto simp add: inf_real_def sup_real_def max_min_distrib2)
--- a/src/HOL/ex/Eval_Examples.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/Eval_Examples.thy Tue Sep 01 22:32:58 2015 +0200
@@ -43,7 +43,7 @@
text {* a fancy datatype *}
datatype ('a, 'b) foo =
- Foo "'a\<Colon>order" 'b
+ Foo "'a::order" 'b
| Bla "('a, 'b) bar"
| Dummy nat
and ('a, 'b) bar =
--- a/src/HOL/ex/FinFunPred.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/FinFunPred.thy Tue Sep 01 22:32:58 2015 +0200
@@ -15,7 +15,7 @@
definition le_finfun_def [code del]: "f \<le> g \<longleftrightarrow> (\<forall>x. f $ x \<le> g $ x)"
-definition [code del]: "(f\<Colon>'a \<Rightarrow>f 'b) < g \<longleftrightarrow> f \<le> g \<and> \<not> g \<le> f"
+definition [code del]: "(f::'a \<Rightarrow>f 'b) < g \<longleftrightarrow> f \<le> g \<and> \<not> g \<le> f"
instance ..
--- a/src/HOL/ex/Normalization_by_Evaluation.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/Normalization_by_Evaluation.thy Tue Sep 01 22:32:58 2015 +0200
@@ -99,8 +99,8 @@
lemma "(2::int) < 3" by normalization
lemma "(2::int) <= 3" by normalization
lemma "abs ((-4::int) + 2 * 1) = 2" by normalization
-lemma "4 - 42 * abs (3 + (-7\<Colon>int)) = -164" by normalization
-lemma "(if (0\<Colon>nat) \<le> (x\<Colon>nat) then 0\<Colon>nat else x) = 0" by normalization
+lemma "4 - 42 * abs (3 + (-7::int)) = -164" by normalization
+lemma "(if (0::nat) \<le> (x::nat) then 0::nat else x) = 0" by normalization
lemma "4 = Suc (Suc (Suc (Suc 0)))" by normalization
lemma "nat 4 = Suc (Suc (Suc (Suc 0)))" by normalization
lemma "[Suc 0, 0] = [Suc 0, 0]" by normalization
@@ -127,8 +127,8 @@
lemma "map f [x, y] = [f x, f y]" by normalization
lemma "(map f [x, y], w) = ([f x, f y], w)" by normalization
-lemma "map f [x, y] = [f x \<Colon> 'a\<Colon>semigroup_add, f y]" by normalization
-lemma "map f [x \<Colon> 'a\<Colon>semigroup_add, y] = [f x, f y]" by normalization
-lemma "(map f [x \<Colon> 'a\<Colon>semigroup_add, y], w \<Colon> 'b\<Colon>finite) = ([f x, f y], w)" by normalization
+lemma "map f [x, y] = [f x :: 'a::semigroup_add, f y]" by normalization
+lemma "map f [x :: 'a::semigroup_add, y] = [f x, f y]" by normalization
+lemma "(map f [x :: 'a::semigroup_add, y], w :: 'b::finite) = ([f x, f y], w)" by normalization
end
--- a/src/HOL/ex/Transfer_Ex.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/Transfer_Ex.thy Tue Sep 01 22:32:58 2015 +0200
@@ -8,58 +8,58 @@
lemma ex1: "(x::nat) + y = y + x"
by auto
-lemma "0 \<le> (y\<Colon>int) \<Longrightarrow> 0 \<le> (x\<Colon>int) \<Longrightarrow> x + y = y + x"
+lemma "0 \<le> (y::int) \<Longrightarrow> 0 \<le> (x::int) \<Longrightarrow> x + y = y + x"
by (fact ex1 [transferred])
(* Using new transfer package *)
-lemma "0 \<le> (x\<Colon>int) \<Longrightarrow> 0 \<le> (y\<Colon>int) \<Longrightarrow> x + y = y + x"
+lemma "0 \<le> (x::int) \<Longrightarrow> 0 \<le> (y::int) \<Longrightarrow> x + y = y + x"
by (fact ex1 [untransferred])
lemma ex2: "(a::nat) div b * b + a mod b = a"
by (rule mod_div_equality)
-lemma "0 \<le> (b\<Colon>int) \<Longrightarrow> 0 \<le> (a\<Colon>int) \<Longrightarrow> a div b * b + a mod b = a"
+lemma "0 \<le> (b::int) \<Longrightarrow> 0 \<le> (a::int) \<Longrightarrow> a div b * b + a mod b = a"
by (fact ex2 [transferred])
(* Using new transfer package *)
-lemma "0 \<le> (a\<Colon>int) \<Longrightarrow> 0 \<le> (b\<Colon>int) \<Longrightarrow> a div b * b + a mod b = a"
+lemma "0 \<le> (a::int) \<Longrightarrow> 0 \<le> (b::int) \<Longrightarrow> a div b * b + a mod b = a"
by (fact ex2 [untransferred])
lemma ex3: "ALL (x::nat). ALL y. EX z. z >= x + y"
by auto
-lemma "\<forall>x\<ge>0\<Colon>int. \<forall>y\<ge>0. \<exists>z\<ge>0. x + y \<le> z"
+lemma "\<forall>x\<ge>0::int. \<forall>y\<ge>0. \<exists>z\<ge>0. x + y \<le> z"
by (fact ex3 [transferred nat_int])
(* Using new transfer package *)
-lemma "\<forall>x\<Colon>int\<in>{0..}. \<forall>y\<in>{0..}. \<exists>z\<in>{0..}. x + y \<le> z"
+lemma "\<forall>x::int\<in>{0..}. \<forall>y\<in>{0..}. \<exists>z\<in>{0..}. x + y \<le> z"
by (fact ex3 [untransferred])
lemma ex4: "(x::nat) >= y \<Longrightarrow> (x - y) + y = x"
by auto
-lemma "0 \<le> (x\<Colon>int) \<Longrightarrow> 0 \<le> (y\<Colon>int) \<Longrightarrow> y \<le> x \<Longrightarrow> tsub x y + y = x"
+lemma "0 \<le> (x::int) \<Longrightarrow> 0 \<le> (y::int) \<Longrightarrow> y \<le> x \<Longrightarrow> tsub x y + y = x"
by (fact ex4 [transferred])
(* Using new transfer package *)
-lemma "0 \<le> (y\<Colon>int) \<Longrightarrow> 0 \<le> (x\<Colon>int) \<Longrightarrow> y \<le> x \<Longrightarrow> tsub x y + y = x"
+lemma "0 \<le> (y::int) \<Longrightarrow> 0 \<le> (x::int) \<Longrightarrow> y \<le> x \<Longrightarrow> tsub x y + y = x"
by (fact ex4 [untransferred])
lemma ex5: "(2::nat) * \<Sum>{..n} = n * (n + 1)"
by (induct n rule: nat_induct, auto)
-lemma "0 \<le> (n\<Colon>int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
+lemma "0 \<le> (n::int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
by (fact ex5 [transferred])
(* Using new transfer package *)
-lemma "0 \<le> (n\<Colon>int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
+lemma "0 \<le> (n::int) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
by (fact ex5 [untransferred])
-lemma "0 \<le> (n\<Colon>nat) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
+lemma "0 \<le> (n::nat) \<Longrightarrow> 2 * \<Sum>{0..n} = n * (n + 1)"
by (fact ex5 [transferred, transferred])
(* Using new transfer package *)
-lemma "0 \<le> (n\<Colon>nat) \<Longrightarrow> 2 * \<Sum>{..n} = n * (n + 1)"
+lemma "0 \<le> (n::nat) \<Longrightarrow> 2 * \<Sum>{..n} = n * (n + 1)"
by (fact ex5 [untransferred, Transfer.transferred])
end
--- a/src/HOL/ex/Transfer_Int_Nat.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/HOL/ex/Transfer_Int_Nat.thy Tue Sep 01 22:32:58 2015 +0200
@@ -200,7 +200,7 @@
involved are bi-unique. *}
lemma
- assumes "\<And>xs\<Colon>int list. \<lbrakk>list_all (\<lambda>x. x \<ge> 0) xs; xs \<noteq> []\<rbrakk> \<Longrightarrow>
+ assumes "\<And>xs::int list. \<lbrakk>list_all (\<lambda>x. x \<ge> 0) xs; xs \<noteq> []\<rbrakk> \<Longrightarrow>
listsum xs < listsum (map (\<lambda>x. x + 1) xs)"
shows "xs \<noteq> [] \<Longrightarrow> listsum xs < listsum (map Suc xs)"
apply transfer
--- a/src/Tools/Code_Generator.thy Tue Sep 01 17:25:36 2015 +0200
+++ b/src/Tools/Code_Generator.thy Tue Sep 01 22:32:58 2015 +0200
@@ -26,7 +26,7 @@
ML_file "~~/src/Tools/Code/code_haskell.ML"
ML_file "~~/src/Tools/Code/code_scala.ML"
-code_datatype "TYPE('a\<Colon>{})"
+code_datatype "TYPE('a::{})"
definition holds :: "prop" where
"holds \<equiv> ((\<lambda>x::prop. x) \<equiv> (\<lambda>x. x))"
@@ -59,4 +59,3 @@
hide_const (open) holds
end
-