# HG changeset patch # User wenzelm # Date 1441139578 -7200 # Node ID bdc1e2f0a86a94e422877ca475f1fa0065ac6e63 # Parent f6b0d827240ebbc30cb079ed65ae1f048adaba82 eliminated \; diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Classes/Classes.thy --- 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 \ \ \ \ \ bool"} which is overloaded on + equality function @{text "eq :: \ \ \ \ bool"} which is overloaded on different types for @{text "\"}, 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 \ \ \ \ \ bool"} + \hspace*{2ex}@{text "eq :: \ \ \ \ bool"} - \medskip\noindent@{text "instance nat \ 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 (\\eq, \\eq) pair \ eq where"} \\ + \medskip\noindent@{text "instance (\::eq, \::eq) pair :: eq where"} \\ \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} \medskip\noindent@{text "class ord extends eq where"} \\ - \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ - \hspace*{2ex}@{text "less \ \ \ \ \ bool"} + \hspace*{2ex}@{text "less_eq :: \ \ \ \ bool"} \\ + \hspace*{2ex}@{text "less :: \ \ \ \ bool"} \end{quote} @@ -57,7 +57,7 @@ \begin{quote} \noindent@{text "class eq where"} \\ - \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ + \hspace*{2ex}@{text "eq :: \ \ \ \ bool"} \\ @{text "satisfying"} \\ \hspace*{2ex}@{text "refl: eq x x"} \\ \hspace*{2ex}@{text "sym: eq x y \ 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 \ \\semigroup \ \ \ \"} and the - global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y z \ - \\semigroup. (x \ y) \ z = x \ (y \ z)"}. + parameter @{term [source] "mult :: \::semigroup \ \ \ \"} and the + global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y z :: + \::semigroup. (x \ y) \ z = x \ (y \ z)"}. *} @@ -130,7 +130,7 @@ begin definition %quote - mult_int_def: "i \ j = i + (j\int)" + mult_int_def: "i \ 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\nat) \ n = n" + "(0::nat) \ n = n" | "Suc m \ n = Suc (m \ n)" instance %quote proof @@ -197,7 +197,7 @@ mult_prod_def: "p\<^sub>1 \ p\<^sub>2 = (fst p\<^sub>1 \ fst p\<^sub>2, snd p\<^sub>1 \ snd p\<^sub>2)" instance %quote proof - fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\\semigroup \ \\semigroup" + fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\::semigroup \ \::semigroup" show "p\<^sub>1 \ p\<^sub>2 \ p\<^sub>3 = p\<^sub>1 \ (p\<^sub>2 \ p\<^sub>3)" unfolding mult_prod_def by (simp add: assoc) qed @@ -237,10 +237,10 @@ begin definition %quote - neutral_nat_def: "\ = (0\nat)" + neutral_nat_def: "\ = (0::nat)" definition %quote - neutral_int_def: "\ = (0\int)" + neutral_int_def: "\ = (0::int)" instance %quote proof fix n :: nat @@ -261,7 +261,7 @@ neutral_prod_def: "\ = (\, \)" instance %quote proof - fix p :: "\\monoidl \ \\monoidl" + fix p :: "\::monoidl \ \::monoidl" show "\ \ p = p" unfolding neutral_prod_def mult_prod_def by (simp add: neutl) qed @@ -295,7 +295,7 @@ begin instance %quote proof - fix p :: "\\monoid \ \\monoid" + fix p :: "\::monoid \ \::monoid" show "p \ \ = p" unfolding neutral_prod_def mult_prod_def by (simp add: neutr) qed @@ -315,7 +315,7 @@ begin definition %quote - inverse_int_def: "i\
= - (i\int)" + inverse_int_def: "i\
= - (i::int)" instance %quote proof fix i :: int @@ -361,7 +361,7 @@ *} interpretation %quote idem_class: - idem "f \ (\\idem) \ \" + idem "f :: (\::idem) \ \" (*<*)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] "\x y z \ - \\group. x \ y = x \ z \ y = z"}. Since type @{text "int"} has been + global one @{fact "group.left_cancel:"} @{prop [source] "\x y z :: + \::group. x \ y = x \ z \ 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] "\x y z \ int. x \ y = x \ z \ y = + fact as well: @{prop [source] "\x y z :: int. x \ y = x \ z \ 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 \ nat \ \\monoid \ \\monoid"} with corresponding theorems + "pow_nat :: nat \ \::monoid \ \::monoid"} with corresponding theorems @{thm pow_nat.simps [no_vars]}. @@ -542,8 +542,8 @@ else (pow_nat (nat (- k)) x)\
)" text {* - \noindent yields the global definition of @{term [source] "pow_int \ - int \ \\group \ \\group"} with the corresponding theorem @{thm + \noindent yields the global definition of @{term [source] "pow_int :: + int \ \::group \ \::group"} with the corresponding theorem @{thm pow_int_def [no_vars]}. *} @@ -559,7 +559,7 @@ begin term %quote "x \ y" -- {* example 1 *} -term %quote "(x\nat) \ y" -- {* example 2 *} +term %quote "(x::nat) \ y" -- {* example 2 *} end %quote @@ -570,7 +570,7 @@ operation @{text "mult [\]"}, 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 [?\ \ semigroup]"}. + polymorphic global class operation @{text "mult [?\ :: semigroup]"}. *} section {* Further issues *} diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Classes/Setup.thy --- 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_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + "_alpha_ofsort" :: "sort \ type" ("\()::_" [0] 1000) "_beta" :: "type" ("\") - "_beta_ofsort" :: "sort \ type" ("\()\_" [0] 1000) + "_beta_ofsort" :: "sort \ type" ("\()::_" [0] 1000) parse_ast_translation {* let diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Codegen/Adaptation.thy --- 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\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\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\bar) y \ x = y" +definition %quote "HOL.equal (x::bar) y \ x = y" instance %quote by default (simp add: equal_bar_def) diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Codegen/Foundations.thy --- 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\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. diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Codegen/Introduction.thy --- 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 \ n = (0\nat)" + "0 \ n = (0::nat)" | "Suc m \ n = n + m \ n" definition %quote neutral_nat where diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Datatypes/Datatypes.thy --- 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 \ 'a option"} and -@{text "Some \ 'a \ 'a option"}. +The constructors are @{text "None :: 'a option"} and +@{text "Some :: 'a \ 'a option"}. The next example has three type parameters: *} @@ -203,7 +203,7 @@ text {* \noindent The constructor is -@{text "Triple \ 'a \ 'b \ 'c \ ('a, 'b, 'c) triple"}. +@{text "Triple :: 'a \ 'b \ 'c \ ('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 \ 'a list"} \\ + @{text "Nil :: 'a list"} \\ & - @{text "Cons \ 'a \ 'a list \ 'a list"} \\ + @{text "Cons :: 'a \ 'a list \ 'a list"} \\ Discriminator: & - @{text "null \ 'a list \ bool"} \\ + @{text "null :: 'a list \ bool"} \\ Selectors: & - @{text "hd \ 'a list \ 'a"} \\ + @{text "hd :: 'a list \ 'a"} \\ & - @{text "tl \ 'a list \ 'a list"} \\ + @{text "tl :: 'a list \ 'a list"} \\ Set function: & - @{text "set \ 'a list \ 'a set"} \\ + @{text "set :: 'a list \ 'a set"} \\ Map function: & - @{text "map \ ('a \ 'b) \ 'a list \ 'b list"} \\ + @{text "map :: ('a \ 'b) \ 'a list \ 'b list"} \\ Relator: & - @{text "list_all2 \ ('a \ 'b \ bool) \ 'a list \ 'b list \ bool"} + @{text "list_all2 :: ('a \ 'b \ bool) \ 'a list \ 'b list \ bool"} \end{tabular} \medskip @@ -1327,7 +1327,7 @@ map function @{const map_option}: *} - primrec (*<*)(in early) (*>*)sum_btree :: "('a\{zero,plus}) btree \ 'a" where + primrec (*<*)(in early) (*>*)sum_btree :: "('a::{zero,plus}) btree \ '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\{zero,plus}) btree \ 'a" and + sum_btree :: "('a::{zero,plus}) btree \ 'a" and sum_btree_option :: "'a btree option \ '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 \ '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 \ "termi\<^sub>0 \ ('a, 'b) tlist \ 'b" + termi\<^sub>0 \ "termi\<^sub>0 :: ('a, 'b) tlist \ 'b" begin primrec termi\<^sub>0 :: "('a, 'b) tlist \ '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\int) \ 0"}, -@{term "n mod (4\int) \ 1"}, and -@{term "n mod (4\int) \ 2"}. +without having first to discharge @{term "n mod (4::int) \ 0"}, +@{term "n mod (4::int) \ 1"}, and +@{term "n mod (4::int) \ 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\int) = i, n mod (4\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 \ +Its relator @{text "llist_all2 :: ('a \ 'b \ bool) \ 'a llist \ 'b llist \ bool"} extends binary predicates over elements to binary predicates over parallel @@ -2680,7 +2680,7 @@ set function, and relator. *} - typedef ('d, 'a) fn = "UNIV \ ('d \ 'a) set" + typedef ('d, 'a) fn = "UNIV :: ('d \ 'a) set" by simp text {* \blankline *} @@ -2723,22 +2723,22 @@ show "set_fn \ map_fn f = op ` f \ set_fn" by transfer (auto simp add: comp_def) next - show "card_order (natLeq +c |UNIV \ '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 \ '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| \o |UNIV \ 'd set|" (is "_ \o ?U") + have "|set_fn F| \o |UNIV :: 'd set|" (is "_ \o ?U") by transfer (rule card_of_image) also have "?U \o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order) - finally show "|set_fn F| \o natLeq +c |UNIV \ 'd set|" . + finally show "|set_fn F| \o natLeq +c |UNIV :: 'd set|" . next fix R :: "'a \ 'b \ bool" and S :: "'b \ 'c \ bool" show "rel_fn R OO rel_fn S \ 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 \ t \ nat"} belonging to the @{text size} type +@{text "size :: t \ nat"} belonging to the @{text size} type class. The \keyw{fun} command relies on @{const size} to prove termination of recursive functions on datatypes. diff -r f6b0d827240e -r bdc1e2f0a86a src/Doc/Tutorial/Types/Axioms.thy --- 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] "\x y z \ 'a\semigroup. (x \ y) \ z = x \ (y \ z)"}. +[source] "\x y z :: 'a::semigroup. (x \ y) \ z = x \ (y \ 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\semigroup" + fixes x y z :: "'a::semigroup" shows "x \ (y \ z) = (x \ y) \ z" using assoc by (rule sym) @@ -63,7 +63,7 @@ begin instance proof - fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a\semigroup \ 'b\semigroup" + fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a::semigroup \ 'b::semigroup" show "p\<^sub>1 \ p\<^sub>2 \ p\<^sub>3 = p\<^sub>1 \ (p\<^sub>2 \ 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: "\ = (0\nat)" + neutral_nat_def: "\ = (0::nat)" instance proof fix n :: nat @@ -119,7 +119,7 @@ neutral_prod_def: "\ = (\, \)" instance proof - fix p :: "'a\monoidl \ 'b\monoidl" + fix p :: "'a::monoidl \ 'b::monoidl" show "\ \ 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\group" + fixes x y z :: "'a::group" shows "x \ y = x \ z \ y = z" proof assume "x \ y = x \ 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 \ 'a\finite +which specifies a type to be finite: @{lemma [source] "finite (UNIV :: 'a::finite set)" by (fact finite_UNIV)}. *} (*<*)end(*>*) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Bali/AxCompl.thy --- 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\{=:n} \c1\\<^sub>s\ {G\}" and mgf_c2: "G,A\{=:n} \c2\\<^sub>s\ {G\}" - shows "G,(A\state triple set)\{=:n} \c1 Finally c2\\<^sub>s\ {G\}" + shows "G,(A::state triple set)\{=:n} \c1 Finally c2\\<^sub>s\ {G\}" proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp) fix T L accC C assume wt: "\prg=G,cls=accC,lcl=L\\In1r (c1 Finally c2)\T" @@ -956,7 +956,7 @@ assumes wf: "wf_prog G" and mgf_init: "G,A\{=:n} \Init D\\<^sub>s\ {G\}" and mgf_c: "G,A\{=:n} \c\\<^sub>s\ {G\}" - shows "G,(A\state triple set)\{=:n} \Body D c\\<^sub>e\ {G\}" + shows "G,(A::state triple set)\{=:n} \Body D c\\<^sub>e\ {G\}" proof (rule MGFn_free_wt_da_NormalConformI [rule_format],clarsimp) fix T L accC E assume wt: "\prg=G, cls=accC,lcl=L\\\Body D c\\<^sub>e\T" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Basic_BNF_LFPs.thy --- 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\bool) = 0" +lemma size_bool[code]: "size (b::bool) = 0" by (cases b) auto declare prod.size[no_atp] diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Binomial.thy --- 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 \ {1 .. n}" by auto have eq: "insert 0 {1 .. n} = {0..n}" by auto - have **: "(\n\{1\nat..n}. a + of_nat n) = (\n\{0\nat..n - 1}. a + 1 + of_nat n)" + have **: "(\n\{1::nat..n}. a + of_nat n) = (\n\{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\'a)) ^ n = setprod (\i. - 1) {0 .. n - 1}" + have eq: "(- (1::'a)) ^ n = setprod (\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 * (\i = 0..h. a - of_nat i) + 2 * (\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)) + (\i\{0\nat..Suc h}. a - of_nat i)" + ((a gchoose Suc h) * (fact (Suc h)) * of_nat (Suc k)) + (\i\{0::nat..Suc h}. a - of_nat i)" unfolding gbinomial_mult_fact' by (simp add: comm_semiring_class.distrib field_simps Suc) also have "\ = (\i\{0..h}. a - of_nat i) * (a + 1)" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Code_Evaluation.thy --- 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 \ 'a \ 'b) = + "term_of (f :: 'a \ '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 \ integer \ term" \ (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT" -| constant "term_of \ String.literal \ term" \ (Eval) "HOLogic.mk'_literal" + constant "term_of :: integer \ term" \ (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT" +| constant "term_of :: String.literal \ term" \ (Eval) "HOLogic.mk'_literal" declare [[code drop: "term_of :: integer \ _"]] diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Code_Numeral.thy --- 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 \Type of target language integers\ -typedef integer = "UNIV \ int set" +typedef integer = "UNIV :: int set" morphisms int_of_integer integer_of_int .. setup_lifting type_definition_integer @@ -615,7 +615,7 @@ subsection \Type of target language naturals\ -typedef natural = "UNIV \ nat set" +typedef natural = "UNIV :: nat set" morphisms nat_of_natural natural_of_nat .. setup_lifting type_definition_natural diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Complex.thy --- 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\complex) = x * inverse y" +definition "x div (y::complex) = x * inverse y" instance by intro_classes diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Datatype_Examples/Misc_Primcorec.thy --- 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 \ ord, 'b \ ord, 'c, 'd) some_passive \ ('d, 'a, 'b, 'c) some_passive" where +primcorec shuffle_sp :: "('a :: ord, 'b :: ord, 'c, 'd) some_passive \ ('d, 'a, 'b, 'c) some_passive" where "shuffle_sp sp = (case sp of SP1 sp' \ SP1 (shuffle_sp sp') @@ -48,7 +48,7 @@ | Let SL l \ Let (fimage (map_prod f (rename_lam f)) SL) (rename_lam f l))" primcorec - j1_sum :: "('a\{zero,one,plus}) \ 'a J1" and + j1_sum :: "('a::{zero,one,plus}) \ 'a J1" and j2_sum :: "'a \ 'a J2" where "n = 0 \ is_J11 (j1_sum n)" | diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Datatype_Examples/Misc_Primrec.thy --- 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 \ ord, 'b \ ord, 'c, 'd) some_passive \ ('d, 'a, 'b, 'c) some_passive" where +primrec shuffle_sp :: "('a :: ord, 'b :: ord, 'c, 'd) some_passive \ ('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\{zero,plus}) I1 \ 'a" and + sum_i1 :: "('a::{zero,plus}) I1 \ 'a" and sum_i2 :: "'a I2 \ 'a" where "sum_i1 (I11 n i) = n + sum_i1 i" | diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Decision_Procs/MIR.thy --- 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 \ real \ bool" (infixl "rdvd" 50) - where "x rdvd y \ (\k\int. y = x * real k)" + where "x rdvd y \ (\k::int. y = x * real k)" lemma int_rdvd_real: "real (i::int) rdvd x = (i dvd (floor x) \ real (floor x) = x)" (is "?l = ?r") @@ -60,7 +60,7 @@ hence "\ 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\int) dvd \x\real\" .. + assume "?r" hence "(i::int) dvd \x::real\" .. hence "\ k. real (floor x) = real (i*k)" by (simp only: real_of_int_inject) (simp add: dvd_def) thus ?l using \?r\ 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\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 "\ = (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 \ (0\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 "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (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\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 "\ = (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 \ (0\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 "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (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\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 "\ = (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 \ (0\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 "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (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 \ j" and jn: "j \ n'" - from 5 pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ + from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ numbound0 s' \ 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' \ j" and jn: "j \ 0" - from 5 pns have H:"(Ifm ((x\real) # (bs\real list)) p' \ + from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ numbound0 s' \ isrlfm p'" by blast hence nb: "numbound0 s'" by simp diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Divides.thy --- 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 \ We define @{const divide} and @{const mod} on @{typ nat} by means of a characteristic relation with two input arguments - @{term "m\nat"}, @{term "n\nat"} and two output arguments - @{term "q\nat"}(uotient) and @{term "r\nat"}(emainder). + @{term "m::nat"}, @{term "n::nat"} and two output arguments + @{term "q::nat"}(uotient) and @{term "r::nat"}(emainder). \ definition divmod_nat_rel :: "nat \ nat \ nat \ nat \ bool" where @@ -835,7 +835,7 @@ have "\q r. m = q * n + r \ r < n" proof (induct m) case 0 with \n \ 0\ - have "(0\nat) = 0 * n + 0 \ 0 < n" by simp + have "(0::nat) = 0 * n + 0 \ 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: "\q r q' r'. q' * n + r' = q * n + r \ r < n \ q' \ (q\nat)" + have aux: "\q r q' r'. q' * n + r' = q * n + r \ r < n \ q' \ (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 \ m" by auto qed -lemma mod_geq: "\ m < (n\nat) \ m mod n = (m - n) mod n" +lemma mod_geq: "\ m < (n::nat) \ m mod n = (m - n) mod n" by (simp add: le_mod_geq linorder_not_less) -lemma mod_if: "m mod (n\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 \ m \ m < n * Suc q \ q = ((m\nat) div n)" (is "?lhs \ ?rhs") + shows "n * q \ m \ m < n * Suc q \ q = ((m::nat) div n)" (is "?lhs \ ?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\nat) mod 2 \ m mod 2 = 1" +lemma mod2_gr_0 [simp]: "0 < (m::nat) mod 2 \ m mod 2 = 1" proof - { fix n :: nat have "(n::nat) < 2 \ n = 0 \ 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\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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Enum.thy --- 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\enum \ 'b\order" + fixes f g :: "'a::enum \ 'b::order" shows "f \ g \ enum_all (\x. f x \ g x)" and "f < g \ f \ g \ enum_ex (\x. f x \ g x)" by (simp_all add: fun_eq_iff le_fun_def order_less_le) @@ -244,32 +244,32 @@ subsection \Default instances for @{class enum}\ lemma map_of_zip_enum_is_Some: - assumes "length ys = length (enum \ 'a\enum list)" - shows "\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y" + assumes "length ys = length (enum :: 'a::enum list)" + shows "\y. map_of (zip (enum :: 'a::enum list) ys) x = Some y" proof - - from assms have "x \ set (enum \ 'a\enum list) \ - (\y. map_of (zip (enum \ 'a\enum list) ys) x = Some y)" + from assms have "x \ set (enum :: 'a::enum list) \ + (\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\enum list" - assumes length: "length xs = length (enum \ 'a\enum list)" - "length ys = length (enum \ 'a\enum list)" - and map_of: "the \ map_of (zip (enum \ 'a\enum list) xs) = the \ map_of (zip (enum \ 'a\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 \ map_of (zip (enum :: 'a::enum list) xs) = the \ map_of (zip (enum :: 'a::enum list) ys)" shows "xs = ys" proof - - have "map_of (zip (enum \ 'a list) xs) = map_of (zip (enum \ '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 \ 'a list) xs) x = Some y1" - and "map_of (zip (enum \ '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 \ 'a\enum list) xs) x) = the (map_of (zip (enum \ 'a\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 \ 'a\enum list) xs) x = map_of (zip (enum \ 'a\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 (\ys. the o map_of (zip (enum\'a list) ys)) (List.n_lists (length (enum\'a\enum list)) enum)" + "enum = map (\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 (\bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" @@ -306,17 +306,17 @@ "enum_ex P = ex_n_lists (\bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" instance proof - show "UNIV = set (enum \ ('a \ 'b) list)" + show "UNIV = set (enum :: ('a \ 'b) list)" proof (rule UNIV_eq_I) fix f :: "'a \ 'b" - have "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" + have "f = the \ 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 \ 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 \ ('a \ 'b) list)" + show "distinct (enum :: ('a \ '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 \ 'b" - have f: "f = the \ map_of (zip (enum \ 'a\enum list) (map f enum))" + have f: "f = the \ 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 \enum_all P\ have "P (the \ 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 \ map_of (zip (enum \ 'a\enum list) (map f enum))" + have f: "f = the \ 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 \P f\ this have "P (the \ map_of (zip (enum \ 'a\enum list) (map f enum)))" + from \P f\ this have "P (the \ 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 \ 'a\{enum, equal} list) +lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list) in map (\ys. the o map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))" by (simp add: enum_fun_def Let_def) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Fields.thy --- 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 \0 < x\] by simp next assume "\0 < x" hence "x \ 0" by simp - obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1\'a"] by auto + obtain s::'a where s: "0 < s" "s < 1" using dense[of 0 "1::'a"] by auto hence "x \ s * x" using mult_le_cancel_right[of 1 x s] \x \ 0\ by auto also note *[OF s] finally show ?thesis . diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Finite_Set.thy --- 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 \Class @{text finite}\ class finite = - assumes finite_UNIV: "finite (UNIV \ 'a set)" + assumes finite_UNIV: "finite (UNIV :: 'a set)" begin -lemma finite [simp]: "finite (A \ 'a set)" +lemma finite [simp]: "finite (A :: 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ -lemma finite_code [code]: "finite (A \ 'a set) \ True" +lemma finite_code [code]: "finite (A :: 'a set) \ True" by simp end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Groups.thy --- 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 \Tools setup\ lemma add_mono_thms_linordered_semiring: - fixes i j k :: "'a\ordered_ab_semigroup_add" + fixes i j k :: "'a::ordered_ab_semigroup_add" shows "i \ j \ k \ l \ i + k \ j + l" and "i = j \ k \ l \ i + k \ j + l" and "i \ j \ k = l \ i + k \ j + l" @@ -1400,7 +1400,7 @@ by (rule add_mono, clarify+)+ lemma add_mono_thms_linordered_field: - fixes i j k :: "'a\ordered_cancel_ab_semigroup_add" + fixes i j k :: "'a::ordered_cancel_ab_semigroup_add" shows "i < j \ k = l \ i + k < j + l" and "i = j \ k < l \ i + k < j + l" and "i < j \ k \ l \ i + k < j + l" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/HOL.thy --- 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) \ True" by (simp add: equal) -setup \Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\type \ 'a \ bool"})\ +setup \Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a::type \ 'a \ bool"})\ lemma equal_alias_cert: "OFCLASS('a, equal_class) \ ((op = :: 'a \ 'a \ bool) \ equal)" (is "?ofclass \ ?equal") proof @@ -1806,7 +1806,7 @@ qed (simp add: \PROP ?equal\) qed -setup \Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\equal \ 'a \ bool"})\ +setup \Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a::equal \ 'a \ bool"})\ setup \Nbe.add_const_alias @{thm equal_alias_cert}\ diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/HOLCF/Lift.thy --- 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 "\\'a lift" Def +old_rep_datatype "\::'a lift" Def by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject inst_lift_pcpo) text {* @{term bottom} and @{term Def} *} diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Hahn_Banach/Linearform.thy --- 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 @@ \ locale linearform = - fixes V :: "'a\{minus, plus, zero, uminus} set" and f + fixes V :: "'a::{minus, plus, zero, uminus} set" and f assumes add [iff]: "x \ V \ y \ V \ f (x + y) = f x + f y" and mult [iff]: "x \ V \ f (a \ x) = a * f x" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Hahn_Banach/Normed_Space.thy --- 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 @@ \ locale seminorm = - fixes V :: "'a\{minus, plus, zero, uminus} set" + fixes V :: "'a::{minus, plus, zero, uminus} set" fixes norm :: "'a \ real" ("\_\") assumes ge_zero [iff?]: "x \ V \ 0 \ \x\" and abs_homogenous [iff?]: "x \ V \ \a \ x\ = \a\ * \x\" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Hahn_Banach/Subspace.thy --- 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 @@ \ locale subspace = - fixes U :: "'a\{minus, plus, zero, uminus} set" and V + fixes U :: "'a::{minus, plus, zero, uminus} set" and V assumes non_empty [iff, intro]: "U \ {}" and subset [iff]: "U \ V" and add_closed [iff]: "x \ U \ y \ U \ x + y \ U" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Hilbert_Choice.thy --- 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 (\y. if y = x then b1 else b2) b1" by (simp add: inv_into_def) - thus "x \ range (\f\'a \ 'b. inv f b1)" by blast + thus "x \ range (\f::'a \ 'b. inv f b1)" by blast qed ultimately show "finite (UNIV :: 'a set)" by simp qed diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy --- 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 \i=0.. {(c, c \\<^sub>c c') |c c'\ivl st option acom. +lemma measure_m_c: "finite X \ {(c, c \\<^sub>c c') |c c'::ivl st option acom. strip c' = strip c \ c : Com X \ c' : Com X \ \ c' \ c}\ \ measure(m_c(m_o (m_st m_ivl) (2*card(X))))" apply(auto simp: m_c_def Let_def Com_def) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/Array.thy --- 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 \ 'a\heap array \ bool" where +definition present :: "heap \ 'a::heap array \ bool" where "present h a \ addr_of_array a < lim h" -definition get :: "heap \ 'a\heap array \ 'a list" where +definition get :: "heap \ 'a::heap array \ 'a list" where "get h a = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" -definition set :: "'a\heap array \ 'a list \ heap \ heap" where +definition set :: "'a::heap array \ 'a list \ heap \ heap" where "set a x = arrays_update (\h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))" -definition alloc :: "'a list \ heap \ 'a\heap array \ heap" where +definition alloc :: "'a list \ heap \ 'a::heap array \ heap" where "alloc xs h = (let l = lim h; r = Array l; h'' = set r xs (h\lim := l + 1\) in (r, h''))" -definition length :: "heap \ 'a\heap array \ nat" where +definition length :: "heap \ 'a::heap array \ nat" where "length h a = List.length (get h a)" -definition update :: "'a\heap array \ nat \ 'a \ heap \ heap" where +definition update :: "'a::heap array \ nat \ 'a \ heap \ heap" where "update a i x h = set a ((get h a)[i:=x]) h" -definition noteq :: "'a\heap array \ 'b\heap array \ bool" (infix "=!!=" 70) where +definition noteq :: "'a::heap array \ 'b::heap array \ bool" (infix "=!!=" 70) where "r =!!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_array r \ addr_of_array s" subsection {* Monad operations *} -definition new :: "nat \ 'a\heap \ 'a array Heap" where +definition new :: "nat \ 'a::heap \ 'a array Heap" where [code del]: "new n x = Heap_Monad.heap (alloc (replicate n x))" -definition of_list :: "'a\heap list \ 'a array Heap" where +definition of_list :: "'a::heap list \ 'a array Heap" where [code del]: "of_list xs = Heap_Monad.heap (alloc xs)" -definition make :: "nat \ (nat \ 'a\heap) \ 'a array Heap" where +definition make :: "nat \ (nat \ 'a::heap) \ 'a array Heap" where [code del]: "make n f = Heap_Monad.heap (alloc (map f [0 ..< n]))" -definition len :: "'a\heap array \ nat Heap" where +definition len :: "'a::heap array \ nat Heap" where [code del]: "len a = Heap_Monad.tap (\h. length h a)" -definition nth :: "'a\heap array \ nat \ 'a Heap" where +definition nth :: "'a::heap array \ nat \ 'a Heap" where [code del]: "nth a i = Heap_Monad.guard (\h. i < length h a) (\h. (get h a ! i, h))" -definition upd :: "nat \ 'a \ 'a\heap array \ 'a\heap array Heap" where +definition upd :: "nat \ 'a \ 'a::heap array \ 'a::heap array Heap" where [code del]: "upd i x a = Heap_Monad.guard (\h. i < length h a) (\h. (a, update a i x h))" -definition map_entry :: "nat \ ('a\heap \ 'a) \ 'a array \ 'a array Heap" where +definition map_entry :: "nat \ ('a::heap \ 'a) \ 'a array \ 'a array Heap" where [code del]: "map_entry i f a = Heap_Monad.guard (\h. i < length h a) (\h. (a, update a i (f (get h a ! i)) h))" -definition swap :: "nat \ 'a \ 'a\heap array \ 'a Heap" where +definition swap :: "nat \ 'a \ 'a::heap array \ 'a Heap" where [code del]: "swap i x a = Heap_Monad.guard (\h. i < length h a) (\h. (get h a ! i, update a i x h))" -definition freeze :: "'a\heap array \ 'a list Heap" where +definition freeze :: "'a::heap array \ 'a list Heap" where [code del]: "freeze a = Heap_Monad.tap (\h. get h a)" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/Heap.thy --- 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 \ 'a\heap array"}) - #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a\heap ref"}) - #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\heap array \ nat"}) - #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\heap ref \ nat"}) + Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \ 'a::heap array"}) + #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a::heap ref"}) + #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a::heap array \ nat"}) + #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a::heap ref \ nat"}) *} hide_const (open) empty diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/Ref.thy --- 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 \ 'a\heap ref \ bool" where +definition present :: "heap \ 'a::heap ref \ bool" where "present h r \ addr_of_ref r < lim h" -definition get :: "heap \ 'a\heap ref \ 'a" where +definition get :: "heap \ 'a::heap ref \ 'a" where "get h = from_nat \ refs h TYPEREP('a) \ addr_of_ref" -definition set :: "'a\heap ref \ 'a \ heap \ heap" where +definition set :: "'a::heap ref \ 'a \ heap \ heap" where "set r x = refs_update (\h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))" -definition alloc :: "'a \ heap \ 'a\heap ref \ heap" where +definition alloc :: "'a \ heap \ 'a::heap ref \ heap" where "alloc x h = (let l = lim h; r = Ref l in (r, set r x (h\lim := l + 1\)))" -definition noteq :: "'a\heap ref \ 'b\heap ref \ bool" (infix "=!=" 70) where +definition noteq :: "'a::heap ref \ 'b::heap ref \ bool" (infix "=!=" 70) where "r =!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_ref r \ addr_of_ref s" subsection {* Monad operations *} -definition ref :: "'a\heap \ 'a ref Heap" where +definition ref :: "'a::heap \ 'a ref Heap" where [code del]: "ref v = Heap_Monad.heap (alloc v)" -definition lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where +definition lookup :: "'a::heap ref \ 'a Heap" ("!_" 61) where [code del]: "lookup r = Heap_Monad.tap (\h. get h r)" -definition update :: "'a ref \ 'a\heap \ unit Heap" ("_ := _" 62) where +definition update :: "'a ref \ 'a::heap \ unit Heap" ("_ := _" 62) where [code del]: "update r v = Heap_Monad.heap (\h. ((), set r v h))" -definition change :: "('a\heap \ 'a) \ 'a ref \ 'a Heap" where +definition change :: "('a::heap \ 'a) \ 'a ref \ 'a Heap" where "change f r = do { x \ ! r; let y = f x; diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy --- 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\heap array \ nat \ nat \ unit Heap" where +fun swap :: "'a::heap array \ nat \ nat \ unit Heap" where "swap a i j = do { x \ Array.nth a i; y \ Array.nth a j; @@ -17,7 +17,7 @@ return () }" -fun rev :: "'a\heap array \ nat \ nat \ unit Heap" where +fun rev :: "'a::heap array \ nat \ nat \ unit Heap" where "rev a i j = (if (i < j) then do { swap a i j; rev a (i + 1) (j - 1) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- 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 \ 'a\type ref"}) *} +setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a::type ref"}) *} datatype 'a node = Empty | Node 'a "'a node ref" primrec - node_encode :: "'a\countable node \ nat" + node_encode :: "'a::countable node \ 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\countable node" + fix x y :: "'a::countable node" show "node_encode x = node_encode y \ x = y" by (induct x, auto, induct y, auto, induct y, auto) qed instance node :: (heap) heap .. -primrec make_llist :: "'a\heap list \ 'a node Heap" +primrec make_llist :: "'a::heap list \ 'a node Heap" where [simp del]: "make_llist [] = return Empty" | "make_llist (x#xs) = do { tl \ make_llist xs; @@ -37,7 +37,7 @@ }" -partial_function (heap) traverse :: "'a\heap node \ 'a list Heap" +partial_function (heap) traverse :: "'a::heap node \ 'a list Heap" where [code del]: "traverse l = (case l of Empty \ return [] diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Imperative_HOL/ex/SatChecker.thy --- 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\heap) option array \ heap \ 'a set" where + array_ran :: "('a::heap) option array \ heap \ 'a set" where "array_ran a h = {e. Some e \ set (Array.get h a)}" -- {*FIXME*} diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Import/HOL_Light_Maps.thy --- 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 \) = (\p q. (\f. f p q \ bool) = (\f. f True True))" + "(op \) = (\p q. (\f. f p q :: bool) = (\f. f True True))" by metis lemma [import_const "==>"]: - "(op \) = (\(p\bool) q\bool. (p \ q) = p)" + "(op \) = (\(p::bool) q::bool. (p \ q) = p)" by auto lemma [import_const "!"]: - "All = (\P\'A \ bool. P = (\x\'A. True))" + "All = (\P::'A \ bool. P = (\x::'A. True))" by auto lemma [import_const "?"]: - "Ex = (\P\'A \ bool. All (\q\bool. (All (\x\'A\type. (P x) \ q)) \ q))" + "Ex = (\P::'A \ bool. All (\q::bool. (All (\x::'A::type. (P x) \ q)) \ q))" by auto lemma [import_const "\\/"]: @@ -44,16 +44,16 @@ by simp lemma [import_const "?!"]: - "Ex1 = (\P\'A \ bool. Ex P \ (\x y. P x \ P y \ x = y))" + "Ex1 = (\P::'A \ bool. Ex P \ (\x y. P x \ P y \ x = y))" by auto lemma [import_const "_FALSITY_"]: "False = False" by simp -lemma hl_ax1: "\t\'A \ 'B. (\x. t x) = t" +lemma hl_ax1: "\t::'A \ 'B. (\x. t x) = t" by metis -lemma hl_ax2: "\P x\'A. P x \ P (Eps P)" +lemma hl_ax2: "\P x::'A. P x \ 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 \) = (\(f\'B \ 'C) g x\'A. f (g x))" + "(op \) = (\(f::'B \ 'C) g x::'A. f (g x))" unfolding fun_eq_iff by simp -lemma [import_const I]: "id = (\x\'A. x)" +lemma [import_const I]: "id = (\x::'A. x)" by auto lemma [import_type 1 one_ABS one_REP]: "type_definition Rep_unit Abs_unit (Collect (\b. b))" by (metis (full_types) Collect_cong singleton_conv2 type_definition_unit) -lemma [import_const one]: "() = (SOME x\unit. True)" +lemma [import_const one]: "() = (SOME x::unit. True)" by auto lemma [import_const mk_pair]: - "Pair_Rep = (\(x\'A) (y\'B) (a\'A) b\'B. a = x \ b = y)" + "Pair_Rep = (\(x::'A) (y::'B) (a::'A) b::'B. a = x \ 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 (\x\'A \ 'B \ bool. \a b. x = Pair_Rep a b))" + "type_definition Rep_prod Abs_prod (Collect (\x::'A \ 'B \ bool. \a b. x = Pair_Rep a b))" using type_definition_prod[unfolded Product_Type.prod_def] by simp -lemma [import_const ","]: "Pair = (\(x\'A) y\'B. Abs_prod (Pair_Rep x y))" +lemma [import_const ","]: "Pair = (\(x::'A) y::'B. Abs_prod (Pair_Rep x y))" by (metis Pair_def) lemma [import_const FST]: - "fst = (\p\'A \ 'B. SOME x\'A. \y\'B. p = (x, y))" + "fst = (\p::'A \ 'B. SOME x::'A. \y::'B. p = (x, y))" by auto lemma [import_const SND]: - "snd = (\p\'A \ 'B. SOME y\'B. \x\'A. p = (x, y))" + "snd = (\p::'A \ 'B. SOME y::'B. \x::'A. p = (x, y))" by auto lemma CURRY_DEF[import_const CURRY]: - "curry = (\(f\'A \ 'B \ 'C) x y. f (x, y))" + "curry = (\(f::'A \ 'B \ 'C) x y. f (x, y))" using curry_def . lemma [import_const ONE_ONE : inj]: - "inj = (\(f\'A \ 'B). \x1 x2. f x1 = f x2 \ x1 = x2)" + "inj = (\(f::'A \ 'B). \x1 x2. f x1 = f x2 \ x1 = x2)" by (auto simp add: fun_eq_iff inj_on_def) lemma [import_const ONTO : surj]: - "surj = (\(f\'A \ 'B). \y. \x. y = f x)" + "surj = (\(f::'A \ 'B). \y. \x. y = f x)" by (auto simp add: fun_eq_iff) -lemma hl_ax3: "\f\ind \ ind. inj f \ \ surj f" +lemma hl_ax3: "\f::ind \ ind. inj f \ \ 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\nat)) = id (0\nat) \ (\n\nat. pred (Suc n) = n)" + "pred (id (0::nat)) = id (0::nat) \ (\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\nat) = True \ (\n. even (Suc n) = (\ even n))" + "even (id 0::nat) = True \ (\n. even (Suc n) = (\ even n))" by (simp add: even_def) lemma SUB[import_const "-" : minus]: - "(\m\nat. m - (id 0) = m) \ (\m n. m - (Suc n) = pred (m - n))" + "(\m::nat. m - (id 0) = m) \ (\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: - "\m n\nat. if n = id 0 then m div n = id 0 \ m mod n = m else m = m div n * n + m mod n \ m mod n < n" + "\m n::nat. if n = id 0 then m div n = id 0 \ m mod n = m else m = m div n * n + m mod n \ 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: - "\P\'A list \ bool. P [] \ (\a0 a1. P a1 \ P (a0 # a1)) \ (\x. P x)" + "\P::'A list \ bool. P [] \ (\a0 a1. P a1 \ P (a0 # a1)) \ (\x. P x)" using list.induct by auto lemma list_RECURSION: - "\nil' cons'. \fn\'A list \ 'Z. fn [] = nil' \ (\(a0\'A) a1\'A list. fn (a0 # a1) = cons' a0 a1 (fn a1))" + "\nil' cons'. \fn::'A list \ 'Z. fn [] = nil' \ (\(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\'A) # t) = h" + "hd ((h::'A) # t) = h" by simp lemma TL[import_const TL : tl]: - "tl ((h\'A) # t) = t" + "tl ((h::'A) # t) = t" by simp lemma APPEND[import_const APPEND : append]: - "(\l\'A list. [] @ l = l) \ (\(h\'A) t l. (h # t) @ l = h # t @ l)" + "(\l::'A list. [] @ l = l) \ (\(h::'A) t l. (h # t) @ l = h # t @ l)" by simp lemma REVERSE[import_const REVERSE : rev]: - "rev [] = ([] :: 'A list) \ rev ((x\'A) # l) = rev l @ [x]" + "rev [] = ([] :: 'A list) \ rev ((x::'A) # l) = rev l @ [x]" by simp lemma LENGTH[import_const LENGTH : length]: - "length ([] :: 'A list) = id 0 \ (\(h\'A) t. length (h # t) = Suc (length t))" + "length ([] :: 'A list) = id 0 \ (\(h::'A) t. length (h # t) = Suc (length t))" by simp lemma MAP[import_const MAP : map]: - "(\f\'A \ 'B. map f [] = []) \ - (\(f\'A \ 'B) h t. map f (h # t) = f h # map f t)" + "(\f::'A \ 'B. map f [] = []) \ + (\(f::'A \ 'B) h t. map f (h # t) = f h # map f t)" by simp lemma LAST[import_const LAST : last]: - "last ((h\'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\nat)) (x\'t18358) = [] \ + "replicate (id (0::nat)) (x::'t18358) = [] \ replicate (Suc n) x = x # replicate n x" by simp lemma NULL[import_const NULL : List.null]: - "List.null ([]\'t18373 list) = True \ List.null ((h\'t18373) # t) = False" + "List.null ([]::'t18373 list) = True \ List.null ((h::'t18373) # t) = False" unfolding null_def by simp lemma ALL[import_const ALL : list_all]: - "list_all (P\'t18393 \ bool) [] = True \ + "list_all (P::'t18393 \ bool) [] = True \ list_all P (h # t) = (P h \ list_all P t)" by simp lemma EX[import_const EX : list_ex]: - "list_ex (P\'t18414 \ bool) [] = False \ + "list_ex (P::'t18414 \ bool) [] = False \ list_ex P (h # t) = (P h \ list_ex P t)" by simp lemma ITLIST[import_const ITLIST : foldr]: - "foldr (f\'t18437 \ 't18436 \ 't18436) [] b = b \ + "foldr (f::'t18437 \ 't18436 \ 't18436) [] b = b \ foldr f (h # t) b = f h (foldr f t b)" by simp lemma ALL2_DEF[import_const ALL2 : list_all2]: - "list_all2 (P\'t18495 \ 't18502 \ bool) [] (l2\'t18502 list) = (l2 = []) \ - list_all2 P ((h1\'t18495) # (t1\'t18495 list)) l2 = + "list_all2 (P::'t18495 \ 't18502 \ bool) [] (l2::'t18502 list) = (l2 = []) \ + list_all2 P ((h1::'t18495) # (t1::'t18495 list)) l2 = (if l2 = [] then False else P h1 (hd l2) \ list_all2 P t1 (tl l2))" by simp (induct_tac l2, simp_all) lemma FILTER[import_const FILTER : filter]: - "filter (P\'t18680 \ bool) [] = [] \ - filter P ((h\'t18680) # t) = (if P h then h # filter P t else filter P t)" + "filter (P::'t18680 \ bool) [] = [] \ + 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 \ 't18825) list) \ - zip ((h1\'t18849) # t1) ((h2\'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]: diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Inductive.thy --- 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 \General induction rules for least fixed points\ lemma lfp_ordinal_induct[case_names mono step union]: - fixes f :: "'a\complete_lattice \ 'a" + fixes f :: "'a::complete_lattice \ 'a" assumes mono: "mono f" and P_f: "\S. P S \ S \ lfp f \ P (f S)" and P_Union: "\M. \S\M. P S \ 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\complete_lattice \ 'a" + fixes f :: "'a::complete_lattice \ 'a" assumes mono: "mono f" and P_f: "\S. P S \ gfp f \ S \ P (f S)" and P_Union: "\M. \S\M. P S \ P (Inf M)" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Int.thy --- 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 \ int \ int \ int) = min" + "(inf :: int \ int \ int) = min" definition - "(sup \ int \ int \ int) = max" + "(sup :: int \ int \ int) = max" instance by intro_classes @@ -161,10 +161,10 @@ begin definition - zabs_def: "\i\int\ = (if i < 0 then - i else i)" + zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition - zsgn_def: "sgn (i\int) = (if i=0 then 0 else if 0i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) - show "sgn (i\int) = (if i=0 then 0 else if 0 w + (1\int) \ z" +lemma zless_imp_add1_zle: "w < z \ w + (1::int) \ z" by transfer clarsimp lemma zless_iff_Suc_zadd: - "(w \ int) < z \ (\n. z = w + int (Suc n))" + "(w :: int) < z \ (\n. z = w + int (Suc n))" apply transfer apply auto apply (rename_tac a b c d) @@ -438,7 +438,7 @@ subsection\Lemmas about the Function @{term of_nat} and Orderings\ -lemma negative_zless_0: "- (int (Suc n)) < (0 \ 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: "\ finite (UNIV::int set)" proof assume "finite (UNIV::int set)" - moreover have "inj (\i\int. 2 * i)" + moreover have "inj (\i::int. 2 * i)" by (rule injI) simp - ultimately have "surj (\i\int. 2 * i)" + ultimately have "surj (\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) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Lattice/Orders.thy --- 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 "\ \ 'a \ 'a \ bool"}. For a + 'a} with relation @{text "\ :: 'a \ 'a \ 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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Lattices.thy --- 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 \ 'b\semilattice_inf" + fixes f :: "'a \ 'b::semilattice_inf" shows "mono f \ f (A \ B) \ f A \ 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 \ 'b\semilattice_sup" + fixes f :: "'a \ 'b::semilattice_sup" shows "mono f \ f A \ f B \ f (A \ 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 \ 'b\linorder" + fixes f :: "'a \ 'b::linorder" shows "mono f \ 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 \ 'b\linorder" + fixes f :: "'a \ 'b::linorder" shows "mono f \ 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 \ 'a\{semilattice_inf, linorder} \ 'a \ 'a)" +lemma inf_min: "inf = (min :: 'a::{semilattice_inf,linorder} \ 'a \ 'a)" by (auto intro: antisym simp add: min_def fun_eq_iff) -lemma sup_max: "sup = (max \ 'a\{semilattice_sup, linorder} \ 'a \ 'a)" +lemma sup_max: "sup = (max :: 'a::{semilattice_sup,linorder} \ 'a \ 'a)" by (auto intro: antisym simp add: max_def fun_eq_iff) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Cardinality.thy --- 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 \ 't set)" +translations "CARD('t)" => "CONST card (CONST UNIV :: 't set)" print_translation \ let diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Char_ord.thy --- 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 \ nibble \ _) = min" -definition "(sup \ nibble \ _) = max" +definition "(inf :: nibble \ _) = min" +definition "(sup :: nibble \ _) = 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 \ char \ _) = min" -definition "(sup \ char \ _) = max" +definition "(inf :: char \ _) = min" +definition "(sup :: char \ _) = max" instance by standard (auto simp add: inf_char_def sup_char_def max_min_distrib2) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Code_Binary_Nat.thy --- 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\nat) = Numeral1" + "(1::nat) = Numeral1" by simp -lemma [code_abbrev]: "Numeral1 = (1\nat)" +lemma [code_abbrev]: "Numeral1 = (1::nat)" by simp lemma [code]: diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Code_Char.thy --- 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 \ char \ term" \ +| constant "Code_Evaluation.term_of :: char \ term" \ (Eval) "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))" code_reserved SML diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Countable.thy --- 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 \The class of countable types\ class countable = - assumes ex_inj: "\to_nat \ 'a \ nat. inj to_nat" + assumes ex_inj: "\to_nat :: 'a \ nat. inj to_nat" lemma countable_classI: fixes f :: "'a \ nat" @@ -27,11 +27,11 @@ subsection \Conversion functions\ -definition to_nat :: "'a\countable \ nat" where +definition to_nat :: "'a::countable \ nat" where "to_nat = (SOME f. inj f)" -definition from_nat :: "nat \ 'a\countable" where - "from_nat = inv (to_nat \ 'a \ nat)" +definition from_nat :: "nat \ 'a::countable" where + "from_nat = inv (to_nat :: 'a \ 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\'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 \ '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 "\to_nat \ 'a \ nat. inj to_nat" by (rule exI[of inj]) + then show "\to_nat :: 'a \ nat. inj to_nat" by (rule exI[of inj]) qed diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Extended_Nat.thy --- 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]: - "\ 0 = (1\enat)" - "\ 1 = (0\enat)" + "\ 0 = (1::enat)" + "\ 1 = (0::enat)" unfolding zero_enat_def one_enat_def by simp_all lemma infinity_ne_i1 [simp]: "(\::enat) \ 1" @@ -380,14 +380,14 @@ a generalize linordered_semidom to a new class that includes enat? *) lemma enat_ord_number [simp]: - "(numeral m \ enat) \ numeral n \ (numeral m \ nat) \ numeral n" - "(numeral m \ enat) < numeral n \ (numeral m \ nat) < numeral n" + "(numeral m :: enat) \ numeral n \ (numeral m :: nat) \ numeral n" + "(numeral m :: enat) < numeral n \ (numeral m :: nat) < numeral n" by (simp_all add: numeral_eq_enat) -lemma i0_lb [simp]: "(0\enat) \ n" +lemma i0_lb [simp]: "(0::enat) \ n" by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) -lemma ile0_eq [simp]: "n \ (0\enat) \ n = 0" +lemma ile0_eq [simp]: "n \ (0::enat) \ n = 0" by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) lemma infinity_ileE [elim!]: "\ \ enat m \ R" @@ -396,10 +396,10 @@ lemma infinity_ilessE [elim!]: "\ < enat m \ R" by simp -lemma not_iless0 [simp]: "\ n < (0\enat)" +lemma not_iless0 [simp]: "\ n < (0::enat)" by (simp add: zero_enat_def less_enat_def split: enat.splits) -lemma i0_less [simp]: "(0\enat) < n \ n \ 0" +lemma i0_less [simp]: "(0::enat) < n \ n \ 0" by (simp add: zero_enat_def less_enat_def split: enat.splits) lemma eSuc_ile_mono [simp]: "eSuc n \ eSuc m \ n \ m" @@ -623,7 +623,7 @@ instance enat :: wellorder proof fix P and n - assume hyp: "(\n\enat. (\m\enat. m < n \ P m) \ P n)" + assume hyp: "(\n::enat. (\m::enat. m < n \ P m) \ P n)" show "P n" by (blast intro: enat_less_induct hyp) qed diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Fraction_Field.thy --- 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 \ 'a fract \ 'a fract \ 'a fract) = min" + "(inf :: 'a fract \ 'a fract \ 'a fract) = min" definition sup_fract_def: - "(sup \ 'a fract \ 'a fract \ 'a fract) = max" + "(sup :: 'a fract \ 'a fract \ 'a fract) = max" instance by intro_classes diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/List_lexord.thy --- 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 \ 'a list \ _) = min" +definition "(inf :: 'a list \ _) = min" -definition "(sup \ 'a list \ _) = max" +definition "(sup :: 'a list \ _) = 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 < ([]\'a\{equal, order} list) \ False" - "[] < (x\'a\{equal, order}) # xs \ True" - "(x\'a\{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" + "xs < ([]::'a::{equal, order} list) \ False" + "[] < (x::'a::{equal, order}) # xs \ True" + "(x::'a::{equal, order}) # xs < y # ys \ x < y \ x = y \ xs < ys" by simp_all lemma less_eq_list_code [code]: - "x # xs \ ([]\'a\{equal, order} list) \ False" - "[] \ (xs\'a\{equal, order} list) \ True" - "(x\'a\{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" + "x # xs \ ([]::'a::{equal, order} list) \ False" + "[] \ (xs::'a::{equal, order} list) \ True" + "(x::'a::{equal, order}) # xs \ y # ys \ x < y \ x = y \ xs \ ys" by simp_all end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Mapping.thy --- 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 \Derived operations\ -definition ordered_keys :: "('a\linorder, 'b) mapping \ 'a list" +definition ordered_keys :: "('a::linorder, 'b) mapping \ 'a list" where "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Multiset.thy --- 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 + \ 'a multiset \ _ \ _)" +lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + :: 'a multiset \ _ \ _)" 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 \ ordered_cancel_comm_monoid_diff) multiset" + fixes M N :: "('a :: ordered_cancel_comm_monoid_diff) multiset" shows "N \# M \ msetsum (M - N) = msetsum M - msetsum N" by (metis add_diff_cancel_right' msetsum.union subset_mset.diff_add) @@ -1756,10 +1756,10 @@ subsubsection \Partial-order properties\ -definition less_multiset :: "'a\order multiset \ 'a multiset \ bool" (infix "#<#" 50) where +definition less_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<#" 50) where "M' #<# M \ (M', M) \ mult {(x', x). x' < x}" -definition le_multiset :: "'a\order multiset \ 'a multiset \ bool" (infix "#<=#" 50) where +definition le_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<=#" 50) where "M' #<=# M \ M' #<# M \ M' = M" notation (xsymbols) less_multiset (infix "#\#" 50) @@ -2210,7 +2210,7 @@ text \Quickcheck generators\ definition (in term_syntax) - msetify :: "'a\typerep list \ (unit \ Code_Evaluation.term) + msetify :: "'a::typerep list \ (unit \ Code_Evaluation.term) \ 'a multiset \ (unit \ Code_Evaluation.term)" where [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\} xs" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Multiset_Order.thy --- 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 \ 'a, b). a \ b}" +lemma total_le: "total {(a :: 'a, b). a \ b}" unfolding total_on_def by auto -lemma total_less: "total {(a \ '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 \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows "M #\# N \ (\y. count N y < count M y \ (\x. y < x \ count M x < count N x))" by (auto simp: le_multiset_def less_multiset\<^sub>H\<^sub>O) -lemma wf_less_multiset: "wf {(M \ ('a \ wellorder) multiset, N). M #\# N}" +lemma wf_less_multiset: "wf {(M :: ('a :: wellorder) multiset, N). M #\# N}" unfolding less_multiset_def by (auto intro: wf_mult wf) lemma order_multiset: "class.order - (le_multiset :: ('a \ order) multiset \ ('a \ order) multiset \ bool) - (less_multiset :: ('a \ order) multiset \ ('a \ order) multiset \ bool)" + (le_multiset :: ('a :: order) multiset \ ('a :: order) multiset \ bool) + (less_multiset :: ('a :: order) multiset \ ('a :: order) multiset \ bool)" by unfold_locales lemma linorder_multiset: "class.linorder - (le_multiset :: ('a \ linorder) multiset \ ('a \ linorder) multiset \ bool) - (less_multiset :: ('a \ linorder) multiset \ ('a \ linorder) multiset \ bool)" + (le_multiset :: ('a :: linorder) multiset \ ('a :: linorder) multiset \ bool) + (less_multiset :: ('a :: linorder) multiset \ ('a :: linorder) multiset \ 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 \ linorder) multiset \ ('a \ linorder) multiset \ bool" - "less_multiset :: ('a \ linorder) multiset \ ('a \ linorder) multiset \ bool" + "le_multiset :: ('a :: linorder) multiset \ ('a :: linorder) multiset \ bool" + "less_multiset :: ('a :: linorder) multiset \ ('a :: linorder) multiset \ bool" by (rule linorder_multiset) interpretation multiset_wellorder: wellorder - "le_multiset :: ('a \ wellorder) multiset \ ('a \ wellorder) multiset \ bool" - "less_multiset :: ('a \ wellorder) multiset \ ('a \ wellorder) multiset \ bool" + "le_multiset :: ('a :: wellorder) multiset \ ('a :: wellorder) multiset \ bool" + "less_multiset :: ('a :: wellorder) multiset \ ('a :: wellorder) multiset \ bool" by unfold_locales (blast intro: wf_less_multiset[unfolded wf_def, rule_format]) lemma le_multiset_total: - fixes M N :: "('a \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows "\ M #\# N \ N #\# M" by (metis multiset_linorder.le_cases) lemma less_eq_imp_le_multiset: - fixes M N :: "('a \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows "M \# N \ M #\# 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 \ linorder) multiset" + fixes M :: "('a :: linorder) multiset" shows "M #\# M + {#undefined#}" unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O by simp lemma le_multiset_empty_left[simp]: - fixes M :: "('a \ linorder) multiset" + fixes M :: "('a :: linorder) multiset" shows "{#} #\# M" by (simp add: less_eq_imp_le_multiset) lemma le_multiset_empty_right[simp]: - fixes M :: "('a \ linorder) multiset" + fixes M :: "('a :: linorder) multiset" shows "M \ {#} \ \ M #\# {#}" by (metis le_multiset_empty_left multiset_order.antisym) lemma less_multiset_empty_left[simp]: - fixes M :: "('a \ linorder) multiset" + fixes M :: "('a :: linorder) multiset" shows "M \ {#} \ {#} #\# M" by (simp add: less_multiset\<^sub>H\<^sub>O) lemma less_multiset_empty_right[simp]: - fixes M :: "('a \ linorder) multiset" + fixes M :: "('a :: linorder) multiset" shows "\ M #\# {#}" using le_empty less_multiset\<^sub>D\<^sub>M by blast lemma - fixes M N :: "('a \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows le_multiset_plus_left[simp]: "N #\# (M + N)" and le_multiset_plus_right[simp]: "M #\# (M + N)" using [[metis_verbose = false]] by (metis less_eq_imp_le_multiset mset_le_add_left add.commute)+ lemma - fixes M N :: "('a \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows less_multiset_plus_plus_left_iff[simp]: "M + N #\# M' + N \ M #\# M'" and less_multiset_plus_plus_right_iff[simp]: "M + N #\# M + N' \ N #\# N'" @@ -286,7 +286,7 @@ by (metis add.commute add_diff_cancel_right' monoid_add_class.add.left_neutral) lemma - fixes M N :: "('a \ linorder) multiset" + fixes M N :: "('a :: linorder) multiset" shows less_multiset_plus_left_nonempty[simp]: "M \ {#} \ N #\# M + N" and less_multiset_plus_right_nonempty[simp]: "N \ {#} \ M #\# 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: "(\y \ 'a \ linorder. y \# N \ (\x. x \# M \ x < y)) \ M #\# N" +lemma ex_gt_imp_less_multiset: "(\y :: 'a :: linorder. y \# N \ (\x. x \# M \ x < y)) \ M #\# 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: - "(\y \ 'a \ linorder. y \# M + N \ y \ x) \ count M x < count N x \ M #\# N" + "(\y :: 'a :: linorder. y \# M + N \ y \ x) \ count M x < count N x \ M #\# 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) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/RBT.thy --- 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 \Type definition\ -typedef ('a, 'b) rbt = "{t :: ('a\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 \ ?rbt" by simp @@ -37,32 +37,32 @@ setup_lifting type_definition_rbt -lift_definition lookup :: "('a\linorder, 'b) rbt \ 'a \ 'b" is "rbt_lookup" . +lift_definition lookup :: "('a::linorder, 'b) rbt \ 'a \ 'b" is "rbt_lookup" . -lift_definition empty :: "('a\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\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_insert" +lift_definition insert :: "'a::linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_insert" by simp -lift_definition delete :: "'a\linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_delete" +lift_definition delete :: "'a::linorder \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_delete" by simp -lift_definition entries :: "('a\linorder, 'b) rbt \ ('a \ 'b) list" is RBT_Impl.entries . +lift_definition entries :: "('a::linorder, 'b) rbt \ ('a \ 'b) list" is RBT_Impl.entries . -lift_definition keys :: "('a\linorder, 'b) rbt \ 'a list" is RBT_Impl.keys . +lift_definition keys :: "('a::linorder, 'b) rbt \ 'a list" is RBT_Impl.keys . -lift_definition bulkload :: "('a\linorder \ 'b) list \ ('a, 'b) rbt" is "rbt_bulkload" .. +lift_definition bulkload :: "('a::linorder \ 'b) list \ ('a, 'b) rbt" is "rbt_bulkload" .. -lift_definition map_entry :: "'a \ ('b \ 'b) \ ('a\linorder, 'b) rbt \ ('a, 'b) rbt" is rbt_map_entry +lift_definition map_entry :: "'a \ ('b \ 'b) \ ('a::linorder, 'b) rbt \ ('a, 'b) rbt" is rbt_map_entry by simp -lift_definition map :: "('a \ 'b \ 'c) \ ('a\linorder, 'b) rbt \ ('a, 'c) rbt" is RBT_Impl.map +lift_definition map :: "('a \ 'b \ 'c) \ ('a::linorder, 'b) rbt \ ('a, 'c) rbt" is RBT_Impl.map by simp -lift_definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a\linorder, 'b) rbt \ 'c \ 'c" is RBT_Impl.fold . +lift_definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a::linorder, 'b) rbt \ 'c \ 'c" is RBT_Impl.fold . -lift_definition union :: "('a\linorder, 'b) rbt \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_union" +lift_definition union :: "('a::linorder, 'b) rbt \ ('a, 'b) rbt \ ('a, 'b) rbt" is "rbt_union" by (simp add: rbt_union_is_rbt) lift_definition foldi :: "('c \ bool) \ ('a \ 'b \ 'c \ 'c) \ ('a :: linorder, 'b) rbt \ 'c \ 'c" @@ -70,7 +70,7 @@ subsection \Derived operations\ -definition is_empty :: "('a\linorder, 'b) rbt \ bool" where +definition is_empty :: "('a::linorder, 'b) rbt \ bool" where [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \ True | _ \ False)" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/RBT_Impl.thy --- 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 \ bool"}), (@{const_name rbt_lookup}, SOME @{typ "('a :: linorder, 'b) rbt \ 'a \ 'b"}), (@{const_name is_rbt}, SOME @{typ "('a :: linorder, 'b) rbt \ bool"}), - (@{const_name rbt_ins}, SOME @{typ "('a\linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_insert_with_key}, SOME @{typ "('a\linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_ins}, SOME @{typ "('a::linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_insert_with_key}, SOME @{typ "('a::linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), (@{const_name rbt_insert_with}, SOME @{typ "('b \ 'b \ 'b) \ ('a :: linorder) \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), (@{const_name rbt_insert}, SOME @{typ "('a :: linorder) \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_del_from_left}, SOME @{typ "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_del_from_right}, SOME @{typ "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_del}, SOME @{typ "('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_delete}, SOME @{typ "('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_union_with_key}, SOME @{typ "('a\linorder \ 'b \ 'b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_union_with}, SOME @{typ "('b \ 'b \ 'b) \ ('a\linorder,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_union}, SOME @{typ "('a\linorder,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_map_entry}, SOME @{typ "'a\linorder \ ('b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt"}), - (@{const_name rbt_bulkload}, SOME @{typ "('a \ 'b) list \ ('a\linorder,'b) rbt"})] + (@{const_name rbt_del_from_left}, SOME @{typ "('a::linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_del_from_right}, SOME @{typ "('a::linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_del}, SOME @{typ "('a::linorder) \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_delete}, SOME @{typ "('a::linorder) \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_union_with_key}, SOME @{typ "('a::linorder \ 'b \ 'b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_union_with}, SOME @{typ "('b \ 'b \ 'b) \ ('a::linorder,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_union}, SOME @{typ "('a::linorder,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_map_entry}, SOME @{typ "'a::linorder \ ('b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt"}), + (@{const_name rbt_bulkload}, SOME @{typ "('a \ 'b) list \ ('a::linorder,'b) rbt"})] \ hide_const (open) R B Empty entries keys fold gen_keys gen_entries diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/RBT_Mapping.thy --- 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 \Implementation of mappings\ context includes rbt.lifting begin -lift_definition Mapping :: "('a\linorder, 'b) rbt \ ('a, 'b) mapping" is RBT.lookup . +lift_definition Mapping :: "('a::linorder, 'b) rbt \ ('a, 'b) mapping" is RBT.lookup . end code_datatype Mapping diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/RBT_Set.thy --- 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 \Definition of code datatype constructors\ -definition Set :: "('a\linorder, unit) rbt \ 'a set" +definition Set :: "('a::linorder, unit) rbt \ 'a set" where "Set t = {x . RBT.lookup t x = Some ()}" -definition Coset :: "('a\linorder, unit) rbt \ 'a set" +definition Coset :: "('a::linorder, unit) rbt \ 'a set" where [simp]: "Coset t = - Set t" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Library/Sublist.thy --- 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: "\zs. ys = xs @ zs" then obtain sk :: "'a list" where sk: "ys = xs @ sk" by fastforce assume a2: "length xs < length ys" - have f1: "\v. ([]\'a list) @ v = v" using append_Nil2 by simp + have f1: "\v. ([]::'a list) @ v = v" using append_Nil2 by simp have "[] \ sk" using a1 a2 sk less_not_refl by force hence "\v. xs @ hd sk # v = ys" using sk by (metis hd_Cons_tl) thus "\zs. ys = (xs @ [ys ! length xs]) @ zs" using f1 by fastforce diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Limits.thy --- 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 \ 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 :: "_ \ 'a\real_normed_vector" + fixes f :: "_ \ 'a::real_normed_vector" assumes "0 \ c" shows "(LIM x F. f x :> at_infinity) \ (\r>c. eventually (\x. r \ norm (f x)) F)" unfolding filterlim_iff eventually_at_infinity @@ -983,7 +983,7 @@ qed lemma tendsto_inverse_0: - fixes x :: "_ \ 'a\real_normed_div_algebra" + fixes x :: "_ \ '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 :: "_ \ 'a\{real_normed_div_algebra, division_ring}" + fixes x :: "_ \ '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 \ 'b\{real_normed_div_algebra, division_ring}" + fixes g :: "'a \ 'b::{real_normed_div_algebra, division_ring}" shows "(LIM x F. inverse (g x) :> at 0) \ (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 \ {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 :: "_ \ 'a\{real_normed_div_algebra, division_ring}" + fixes f :: "_ \ 'a::{real_normed_div_algebra, division_ring}" assumes f: "(f ---> c) F" assumes g: "LIM x F. g x :> at_infinity" shows "((\x. f x / g x) ---> 0) F" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/List.thy --- 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]: - "(\mnat. P m) \ (\m \ {0..m (\m \ {0..mnat. P m) \ (\m \ {0..m (\m \ {0..m\n\nat. P m) \ (\m \ {0..n}. P m)" + "(\m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto lemma ex_nat_less [code_unfold]: - "(\m\n\nat. P m) \ (\m \ {0..n}. P m)" + "(\m\n::nat. P m) \ (\m \ {0..n}. P m)" by auto text\Bounded @{text LEAST} operator:\ diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/MacLaurin.thy --- 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 = - (\mreal) / (fact m) * h ^ m) + B * (h ^ n / (fact n))" + (\m "(\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: "\(m\nat) t\real. - m < n \ (0\real) \ t \ t \ h \ DERIV (difg m) t :> difg (Suc m) t" + have difg_Suc: "\(m::nat) t::real. + m < n \ (0::real) \ t \ t \ h \ DERIV (difg m) t :> difg (Suc m) t" using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2) have difg_eq_0: "\mx. 0 \ x \ x \ h \ isCont (difg (0\nat)) x" + show "\x. 0 \ x \ x \ h \ isCont (difg (0::nat)) x" by (simp add: isCont_difg n) - show "\x. 0 < x \ x < h \ difg (0\nat) differentiable (at x)" + show "\x. 0 < x \ x < h \ difg (0::nat) differentiable (at x)" by (simp add: differentiable_difg n) qed next diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Matrix_LP/Matrix.thy --- 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 \ B \ (\j i. Rep_matrix A j i \ Rep_matrix B j i)" definition - less_def: "A < (B\'a matrix) \ A \ B \ \ B \ A" + less_def: "A < (B::'a matrix) \ A \ B \ \ B \ A" instance .. @@ -1496,7 +1496,7 @@ begin definition - abs_matrix_def: "abs (A \ 'a matrix) = sup A (- A)" + abs_matrix_def: "abs (A :: 'a matrix) = sup A (- A)" instance .. diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Matrix_LP/SparseMatrix.thy --- 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: "(\ True) = False" by blast lemma bool2: "(\ False) = True" by blast -lemma bool3: "((P\bool) \ True) = P" by blast -lemma bool4: "(True \ (P\bool)) = P" by blast -lemma bool5: "((P\bool) \ False) = False" by blast -lemma bool6: "(False \ (P\bool)) = False" by blast -lemma bool7: "((P\bool) \ True) = True" by blast -lemma bool8: "(True \ (P\bool)) = True" by blast -lemma bool9: "((P\bool) \ False) = P" by blast -lemma bool10: "(False \ (P\bool)) = P" by blast +lemma bool3: "((P::bool) \ True) = P" by blast +lemma bool4: "(True \ (P::bool)) = P" by blast +lemma bool5: "((P::bool) \ False) = False" by blast +lemma bool6: "(False \ (P::bool)) = False" by blast +lemma bool7: "((P::bool) \ True) = True" by blast +lemma bool8: "(True \ (P::bool)) = True" by blast +lemma bool9: "((P::bool) \ False) = P" by blast +lemma bool10: "(False \ (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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Meson.thy --- 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 \ 'a" where "skolem = (\x. x)" -lemma skolem_COMBK_iff: "P \ skolem (COMBK P (i\nat))" +lemma skolem_COMBK_iff: "P \ skolem (COMBK P (i::nat))" unfolding skolem_def COMBK_def by (rule refl) lemmas skolem_COMBK_I = iffD1 [OF skolem_COMBK_iff] diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Metis_Examples/Big_O.thy --- 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\linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where - "O(f\('a => 'b)) == {h. \c. \x. abs (h x) <= c * abs (f x)}" +definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where + "O(f::('a => 'b)) == {h. \c. \x. abs (h x) <= c * abs (f x)}" lemma bigo_pos_const: - "(\c\'a\linordered_idom. + "(\c::'a::linordered_idom. \x. abs (h x) \ c * abs (f x)) \ (\c. 0 < c \ (\x. abs(h x) \ c * abs (f x)))" by (metis (no_types) abs_ge_zero @@ -32,7 +32,7 @@ sledgehammer_params [isar_proofs, compress = 1] lemma - "(\c\'a\linordered_idom. + "(\c::'a::linordered_idom. \x. abs (h x) \ c * abs (f x)) \ (\c. 0 < c \ (\x. abs(h x) \ c * abs (f x)))" apply auto @@ -42,19 +42,19 @@ proof - fix c :: 'a and x :: 'b assume A1: "\x. \h x\ \ c * \f x\" - have F1: "\x\<^sub>1\'a\linordered_idom. 0 \ \x\<^sub>1\" by (metis abs_ge_zero) - have F2: "\x\<^sub>1\'a\linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) + have F1: "\x\<^sub>1::'a::linordered_idom. 0 \ \x\<^sub>1\" by (metis abs_ge_zero) + have F2: "\x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) have F3: "\x\<^sub>1 x\<^sub>3. x\<^sub>3 \ \h x\<^sub>1\ \ x\<^sub>3 \ c * \f x\<^sub>1\" by (metis A1 order_trans) - have F4: "\x\<^sub>2 x\<^sub>3\'a\linordered_idom. \x\<^sub>3\ * \x\<^sub>2\ = \x\<^sub>3 * x\<^sub>2\" + have F4: "\x\<^sub>2 x\<^sub>3::'a::linordered_idom. \x\<^sub>3\ * \x\<^sub>2\ = \x\<^sub>3 * x\<^sub>2\" by (metis abs_mult) - have F5: "\x\<^sub>3 x\<^sub>1\'a\linordered_idom. 0 \ x\<^sub>1 \ \x\<^sub>3 * x\<^sub>1\ = \x\<^sub>3\ * x\<^sub>1" + have F5: "\x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \ x\<^sub>1 \ \x\<^sub>3 * x\<^sub>1\ = \x\<^sub>3\ * x\<^sub>1" by (metis abs_mult_pos) - hence "\x\<^sub>1\0. \x\<^sub>1\'a\linordered_idom\ = \1\ * x\<^sub>1" by (metis F2) - hence "\x\<^sub>1\0. \x\<^sub>1\'a\linordered_idom\ = x\<^sub>1" by (metis F2 abs_one) + hence "\x\<^sub>1\0. \x\<^sub>1::'a::linordered_idom\ = \1\ * x\<^sub>1" by (metis F2) + hence "\x\<^sub>1\0. \x\<^sub>1::'a::linordered_idom\ = x\<^sub>1" by (metis F2 abs_one) hence "\x\<^sub>3. 0 \ \h x\<^sub>3\ \ \c * \f x\<^sub>3\\ = c * \f x\<^sub>3\" by (metis F3) hence "\x\<^sub>3. \c * \f x\<^sub>3\\ = c * \f x\<^sub>3\" by (metis F1) - hence "\x\<^sub>3. (0\'a) \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c\ * \f x\<^sub>3\" by (metis F5) - hence "\x\<^sub>3. (0\'a) \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c * f x\<^sub>3\" by (metis F4) + hence "\x\<^sub>3. (0::'a) \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c\ * \f x\<^sub>3\" by (metis F5) + hence "\x\<^sub>3. (0::'a) \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c * f x\<^sub>3\" by (metis F4) hence "\x\<^sub>3. c * \f x\<^sub>3\ = \c * f x\<^sub>3\" by (metis F1) hence "\h x\ \ \c * f x\" by (metis A1) thus "\h x\ \ \c\ * \f x\" by (metis F4) @@ -63,7 +63,7 @@ sledgehammer_params [isar_proofs, compress = 2] lemma - "(\c\'a\linordered_idom. + "(\c::'a::linordered_idom. \x. abs (h x) \ c * abs (f x)) \ (\c. 0 < c \ (\x. abs(h x) \ c * abs (f x)))" apply auto @@ -73,10 +73,10 @@ proof - fix c :: 'a and x :: 'b assume A1: "\x. \h x\ \ c * \f x\" - have F1: "\x\<^sub>1\'a\linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) - have F2: "\x\<^sub>2 x\<^sub>3\'a\linordered_idom. \x\<^sub>3\ * \x\<^sub>2\ = \x\<^sub>3 * x\<^sub>2\" + have F1: "\x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) + have F2: "\x\<^sub>2 x\<^sub>3::'a::linordered_idom. \x\<^sub>3\ * \x\<^sub>2\ = \x\<^sub>3 * x\<^sub>2\" by (metis abs_mult) - have "\x\<^sub>1\0. \x\<^sub>1\'a\linordered_idom\ = x\<^sub>1" by (metis F1 abs_mult_pos abs_one) + have "\x\<^sub>1\0. \x\<^sub>1::'a::linordered_idom\ = x\<^sub>1" by (metis F1 abs_mult_pos abs_one) hence "\x\<^sub>3. \c * \f x\<^sub>3\\ = c * \f x\<^sub>3\" by (metis A1 abs_ge_zero order_trans) hence "\x\<^sub>3. 0 \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c * f x\<^sub>3\" by (metis F2 abs_mult_pos) hence "\h x\ \ \c * f x\" by (metis A1 abs_ge_zero) @@ -86,7 +86,7 @@ sledgehammer_params [isar_proofs, compress = 3] lemma - "(\c\'a\linordered_idom. + "(\c::'a::linordered_idom. \x. abs (h x) \ c * abs (f x)) \ (\c. 0 < c \ (\x. abs(h x) \ c * abs (f x)))" apply auto @@ -96,9 +96,9 @@ proof - fix c :: 'a and x :: 'b assume A1: "\x. \h x\ \ c * \f x\" - have F1: "\x\<^sub>1\'a\linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) - have F2: "\x\<^sub>3 x\<^sub>1\'a\linordered_idom. 0 \ x\<^sub>1 \ \x\<^sub>3 * x\<^sub>1\ = \x\<^sub>3\ * x\<^sub>1" by (metis abs_mult_pos) - hence "\x\<^sub>1\0. \x\<^sub>1\'a\linordered_idom\ = x\<^sub>1" by (metis F1 abs_one) + have F1: "\x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) + have F2: "\x\<^sub>3 x\<^sub>1::'a::linordered_idom. 0 \ x\<^sub>1 \ \x\<^sub>3 * x\<^sub>1\ = \x\<^sub>3\ * x\<^sub>1" by (metis abs_mult_pos) + hence "\x\<^sub>1\0. \x\<^sub>1::'a::linordered_idom\ = x\<^sub>1" by (metis F1 abs_one) hence "\x\<^sub>3. 0 \ \f x\<^sub>3\ \ c * \f x\<^sub>3\ = \c\ * \f x\<^sub>3\" by (metis F2 A1 abs_ge_zero order_trans) thus "\h x\ \ \c\ * \f x\" by (metis A1 abs_ge_zero) qed @@ -106,7 +106,7 @@ sledgehammer_params [isar_proofs, compress = 4] lemma - "(\c\'a\linordered_idom. + "(\c::'a::linordered_idom. \x. abs (h x) \ c * abs (f x)) \ (\c. 0 < c \ (\x. abs(h x) \ c * abs (f x)))" apply auto @@ -116,7 +116,7 @@ proof - fix c :: 'a and x :: 'b assume A1: "\x. \h x\ \ c * \f x\" - have "\x\<^sub>1\'a\linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) + have "\x\<^sub>1::'a::linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1) hence "\x\<^sub>3. \c * \f x\<^sub>3\\ = c * \f x\<^sub>3\" by (metis A1 abs_ge_zero order_trans abs_mult_pos abs_one) hence "\h x\ \ \c * f x\" 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: "\x. f x ~= 0 \ - O(f * g) <= (f\'a => ('b\linordered_field)) *o O(g)" + O(f * g) <= (f::'a => ('b::linordered_field)) *o O(g)" proof - assume a: "\x. f x ~= 0" show "O(f * g) <= f *o O(g)" @@ -334,31 +334,31 @@ by (rule bigo_mult2) also have "(\x. 1 / f x) * (f * g) = g" by (simp add: fun_eq_iff a) - finally have "(\x. (1\'b) / f x) * h : O(g)". - then have "f * ((\x. (1\'b) / f x) * h) : f *o O(g)" + finally have "(\x. (1::'b) / f x) * h : O(g)". + then have "f * ((\x. (1::'b) / f x) * h) : f *o O(g)" by auto - also have "f * ((\x. (1\'b) / f x) * h) = h" + also have "f * ((\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: -"\x. f x \ 0 \ O(f * g) = (f\'a \ ('b\linordered_field)) *o O(g)" +"\x. f x \ 0 \ O(f * g) = (f::'a \ ('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: -"\x. f x \ 0 \ O(f * g) \ O(f\'a \ ('b\linordered_field)) * O(g)" +"\x. f x \ 0 \ O(f * g) \ O(f::'a \ ('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: -"\x. f x \ 0 \ O(f * g) = O(f\'a \ ('b\linordered_field)) * O(g)" +"\x. f x \ 0 \ O(f * g) = O(f::'a \ ('b::linordered_field)) * O(g)" by (metis bigo_mult bigo_mult7 order_antisym_conv) lemma bigo_minus [intro]: "f : O(g) \ - f : O(g)" @@ -397,14 +397,14 @@ lemma bigo_const2 [intro]: "O(\x. c) \ O(\x. 1)" by (metis bigo_const1 bigo_elt_subset) -lemma bigo_const3: "(c\'a\linordered_field) ~= 0 \ (\x. 1) : O(\x. c)" +lemma bigo_const3: "(c::'a::linordered_field) ~= 0 \ (\x. 1) : O(\x. c)" apply (simp add: bigo_def) by (metis abs_eq_0 left_inverse order_refl) -lemma bigo_const4: "(c\'a\linordered_field) ~= 0 \ O(\x. 1) <= O(\x. c)" +lemma bigo_const4: "(c::'a::linordered_field) ~= 0 \ O(\x. 1) <= O(\x. c)" by (metis bigo_elt_subset bigo_const3) -lemma bigo_const [simp]: "(c\'a\linordered_field) ~= 0 \ +lemma bigo_const [simp]: "(c::'a::linordered_field) ~= 0 \ O(\x. c) = O(\x. 1)" by (metis bigo_const2 bigo_const4 equalityI) @@ -415,19 +415,19 @@ lemma bigo_const_mult2: "O(\x. c * f x) \ O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) -lemma bigo_const_mult3: "(c\'a\linordered_field) ~= 0 \ f : O(\x. c * f x)" +lemma bigo_const_mult3: "(c::'a::linordered_field) ~= 0 \ f : O(\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\'a\linordered_field) \ 0 \ O(f) \ O(\x. c * f x)" +"(c::'a::linordered_field) \ 0 \ O(f) \ O(\x. c * f x)" by (metis bigo_elt_subset bigo_const_mult3) -lemma bigo_const_mult [simp]: "(c\'a\linordered_field) ~= 0 \ +lemma bigo_const_mult [simp]: "(c::'a::linordered_field) ~= 0 \ O(\x. c * f x) = O(f)" by (metis equalityI bigo_const_mult2 bigo_const_mult4) -lemma bigo_const_mult5 [simp]: "(c\'a\linordered_field) ~= 0 \ +lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) ~= 0 \ (\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 \ 'a" and d :: 'a - assume A1: "c \ (0\'a)" - assume A2: "\x\'b. \g x\ \ d * \f x\" + assume A1: "c \ (0::'a)" + assume A2: "\x::'b. \g x\ \ d * \f x\" have F1: "inverse \c\ = \inverse c\" using A1 by (metis nonzero_abs_inverse) - have F2: "(0\'a) < \c\" using A1 by (metis zero_less_abs_iff) - have "(0\'a) < \c\ \ (0\'a) < \inverse c\" using F1 by (metis positive_imp_inverse_positive) - hence "(0\'a) < \inverse c\" using F2 by metis - hence F3: "(0\'a) \ \inverse c\" by (metis order_le_less) - have "\(u\'a) SKF\<^sub>7\'a \ 'b. \g (SKF\<^sub>7 (\inverse c\ * u))\ \ u * \f (SKF\<^sub>7 (\inverse c\ * u))\" + have F2: "(0::'a) < \c\" using A1 by (metis zero_less_abs_iff) + have "(0::'a) < \c\ \ (0::'a) < \inverse c\" using F1 by (metis positive_imp_inverse_positive) + hence "(0::'a) < \inverse c\" using F2 by metis + hence F3: "(0::'a) \ \inverse c\" by (metis order_le_less) + have "\(u::'a) SKF\<^sub>7::'a \ 'b. \g (SKF\<^sub>7 (\inverse c\ * u))\ \ u * \f (SKF\<^sub>7 (\inverse c\ * u))\" using A2 by metis - hence F4: "\(u\'a) SKF\<^sub>7\'a \ 'b. \g (SKF\<^sub>7 (\inverse c\ * u))\ \ u * \f (SKF\<^sub>7 (\inverse c\ * u))\ \ (0\'a) \ \inverse c\" + hence F4: "\(u::'a) SKF\<^sub>7::'a \ 'b. \g (SKF\<^sub>7 (\inverse c\ * u))\ \ u * \f (SKF\<^sub>7 (\inverse c\ * u))\ \ (0::'a) \ \inverse c\" using F3 by metis - hence "\(v\'a) (u\'a) SKF\<^sub>7\'a \ 'b. \inverse c\ * \g (SKF\<^sub>7 (u * v))\ \ u * (v * \f (SKF\<^sub>7 (u * v))\)" + hence "\(v::'a) (u::'a) SKF\<^sub>7::'a \ 'b. \inverse c\ * \g (SKF\<^sub>7 (u * v))\ \ u * (v * \f (SKF\<^sub>7 (u * v))\)" by (metis mult_left_mono) - then show "\ca\'a. \x\'b. inverse \c\ * \g x\ \ ca * \f x\" + then show "\ca::'a. \x::'b. inverse \c\ * \g x\ \ ca * \f x\" using A2 F4 by (metis F1 `0 < \inverse c\` 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\'a\linordered_field) ~= 0 \ +lemma bigo_useful_const_mult: "(c::'a::linordered_field) ~= 0 \ (\x. c) * f =o O(h) \ f =o O(h)" apply (rule subsetD) apply (subgoal_tac "(\x. 1 / c) *o O(h) <= O(h)") @@ -594,13 +594,13 @@ apply (simp add: func_times) done -lemma bigo_fix: "(\x. f ((x\nat) + 1)) =o O(\x. h(x + 1)) \ f 0 = 0 \ +lemma bigo_fix: "(\x. f ((x::nat) + 1)) =o O(\x. h(x + 1)) \ f 0 = 0 \ 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: - "(\x. f ((x\nat) + 1)) =o (\x. g(x + 1)) +o O(\x. h(x + 1)) \ + "(\x. f ((x::nat) + 1)) =o (\x. g(x + 1)) +o O(\x. h(x + 1)) \ f 0 = g 0 \ 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\linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl " 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) \ \x. abs (g x) <= abs (f x) \ @@ -691,7 +691,7 @@ by (metis abs_ge_zero linorder_linear max.absorb1 max.commute) lemma bigo_lesso4: - "f 'a=>'b\{linordered_field}) \ + "f 'b::{linordered_field}) \ g =o h +o O(k) \ f ys\nat list. tl ys = xs) \ (\bs\int list. tl bs = as)" +lemma "(\ys::nat list. tl ys = xs) \ (\bs::int list. tl bs = as)" by (metis ex_tl) end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Metis_Examples/Message.thy --- 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: "\x\<^sub>1. parts x\<^sub>1 \ synth (analz x\<^sub>1) = parts (synth (analz x\<^sub>1))" by (metis parts_analz parts_synth) have F3: "X \ synth (analz H)" using A1 by metis - have "\x\<^sub>2 x\<^sub>1\msg set. x\<^sub>1 \ sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3)) + have "\x\<^sub>2 x\<^sub>1::msg set. x\<^sub>1 \ sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3)) hence F4: "\x\<^sub>1. analz x\<^sub>1 \ analz (synth x\<^sub>1)" by (metis analz_synth) have F5: "X \ synth (analz H)" using F3 by metis have "\x\<^sub>1. analz x\<^sub>1 \ synth (analz x\<^sub>1) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Metis_Examples/Proxies.thy --- 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\nat) \ 0" +lemma plus_1_not_0: "n + (1::nat) \ 0" by simp definition inc :: "nat \ nat" where @@ -40,7 +40,7 @@ sledgehammer [expect = some] (add_swap_def) by (metis_exhaust add_swap_def) -definition "A = {xs\'a list. True}" +definition "A = {xs::'a list. True}" lemma "xs \ A" (* The "add:" argument is unfortunate. *) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Metis_Examples/Sets.thy --- 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 \ Z) = (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" proof - - have F1: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>1 \ x\<^sub>2" by (metis Un_commute Un_upper2) - have F2a: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 = x\<^sub>2 \ x\<^sub>1" by (metis Un_commute subset_Un_eq) - have F2: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq) + have F1: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>1 \ x\<^sub>2" by (metis Un_commute Un_upper2) + have F2a: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 = x\<^sub>2 \ x\<^sub>1" by (metis Un_commute subset_Un_eq) + have F2: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq) { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume AA1: "Y \ Z \ X" { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } moreover { assume AAA1: "Y \ X \ Y \ Z \ X" { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume "(Z \ X \ Y \ X) \ Y \ Z \ X" hence "Y \ Z \ X \ X \ Y \ Z" by (metis Un_subset_iff) hence "Y \ Z \ X \ \ X \ Y \ Z" by (metis F2) - hence "\x\<^sub>1\'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1) - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1) } + hence "\x\<^sub>1::'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1) + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1) } moreover - { assume "\x\<^sub>1\'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" + { assume "\x\<^sub>1::'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } moreover { assume AAA1: "Y \ X \ Y \ Z \ X" { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume "(Z \ X \ Y \ X) \ Y \ Z \ X" hence "Y \ Z \ X \ X \ Y \ Z" by (metis Un_subset_iff) hence "Y \ Z \ X \ \ X \ Y \ Z" by (metis F2) - hence "\x\<^sub>1\'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1) - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by blast } + hence "\x\<^sub>1::'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1) + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by blast } moreover { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } - ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by metis + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1) } + ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by metis qed sledgehammer_params [isar_proofs, compress = 2] @@ -75,36 +75,36 @@ lemma (*equal_union: *) "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" proof - - have F1: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) - { assume AA1: "\x\<^sub>1\'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" + have F1: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) + { assume AA1: "\x\<^sub>1::'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" { assume AAA1: "Y \ X \ Y \ Z \ X" { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume "Y \ Z \ X \ X \ Y \ Z" - hence "\x\<^sub>1\'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1 Un_commute Un_upper2) - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1 Un_subset_iff) } + hence "\x\<^sub>1::'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1 Un_commute Un_upper2) + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AAA1 Un_subset_iff) } moreover { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_subset_iff) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_subset_iff) } moreover { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } moreover { assume AA1: "Y \ X \ Y \ Z \ X" { assume "\ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } moreover { assume "Y \ Z \ X \ X \ Y \ Z" - hence "\x\<^sub>1\'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1 Un_commute Un_upper2) - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } - ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_subset_iff) } - ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by metis + hence "\x\<^sub>1::'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" by (metis F1 Un_commute Un_upper2) + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + ultimately have "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_subset_iff) } + ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by metis qed sledgehammer_params [isar_proofs, compress = 3] @@ -112,16 +112,16 @@ lemma (*equal_union: *) "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" proof - - have F1a: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 = x\<^sub>2 \ x\<^sub>1" by (metis Un_commute subset_Un_eq) - have F1: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq) + have F1a: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 = x\<^sub>2 \ x\<^sub>1" by (metis Un_commute subset_Un_eq) + have F1: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq) { assume "(Z \ X \ Y \ X) \ Y \ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } moreover - { assume AA1: "\x\<^sub>1\'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" + { assume AA1: "\x\<^sub>1::'a set. (Z \ x\<^sub>1 \ Y \ x\<^sub>1) \ \ X \ x\<^sub>1" { assume "(Z \ X \ Y \ X) \ Y \ Z \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) } - ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) } + ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) qed sledgehammer_params [isar_proofs, compress = 4] @@ -129,15 +129,15 @@ lemma (*equal_union: *) "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" proof - - have F1: "\(x\<^sub>2\'b set) x\<^sub>1\'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) + have F1: "\(x\<^sub>2::'b set) x\<^sub>1::'b set. x\<^sub>1 \ x\<^sub>2 \ x\<^sub>2 \ x\<^sub>1 \ x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) { assume "\ Y \ X" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_commute Un_upper2) } moreover { assume AA1: "Y \ X \ Y \ Z \ X" - { assume "\x\<^sub>1\'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } - hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) } - ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V\'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_subset_iff Un_upper2) + { assume "\x\<^sub>1::'a set. Y \ x\<^sub>1 \ Z \ Y \ Z \ X \ \ X \ x\<^sub>1 \ Z" + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_upper2) } + hence "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) } + ultimately show "(X = Y \ Z) = (Y \ X \ Z \ X \ (\V::'a set. Y \ V \ Z \ V \ X \ V))" by (metis Un_subset_iff Un_upper2) qed sledgehammer_params [isar_proofs, compress = 1] @@ -154,8 +154,8 @@ lemma (* fixedpoint: *) "\!x. f (g x) = x \ \!y. g (f y) = y" proof - - assume "\!x\'a. f (g x) = x" - thus "\!y\'b. g (f y) = y" by metis + assume "\!x::'a. f (g x) = x" + thus "\!y::'b. g (f y) = y" by metis qed lemma (* singleton_example_2: *) @@ -172,8 +172,8 @@ assume "\x \ S. \S \ x" hence "\x\<^sub>1. x\<^sub>1 \ \S \ x\<^sub>1 \ S \ x\<^sub>1 = \S" by (metis set_eq_subset) hence "\x\<^sub>1. x\<^sub>1 \ S \ x\<^sub>1 = \S" by (metis Union_upper) - hence "\x\<^sub>1\('a set) set. \S \ x\<^sub>1 \ S \ x\<^sub>1" by (metis subsetI) - hence "\x\<^sub>1\('a set) set. S \ insert (\S) x\<^sub>1" by (metis insert_iff) + hence "\x\<^sub>1::('a set) set. \S \ x\<^sub>1 \ S \ x\<^sub>1" by (metis subsetI) + hence "\x\<^sub>1::('a set) set. S \ insert (\S) x\<^sub>1" by (metis insert_iff) thus "\z. S \ {z}" by metis qed diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Metis_Examples/Type_Encodings.thy --- 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) \ null xs \ xs = []" by (metis_exhaust null_def) -lemma "(0\nat) + 0 = 0" +lemma "(0::nat) + 0 = 0" by (metis_exhaust add_0_left) end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/MicroJava/Comp/LemmasComp.thy --- 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 "(\x\(vname list \ fdecl list \ stmt \ expr) mdecl. + apply (subgoal_tac "(\x::(vname list \ fdecl list \ stmt \ expr) mdecl. (fst x, Object, snd (compMethod G Object - (inv (\(s\sig, m\ty \ vname list \ fdecl list \ stmt \ expr). + (inv (\(s::sig, m::ty \ vname list \ fdecl list \ stmt \ expr). (s, Object, m)) (S, Object, snd x))))) = (\x. (fst x, Object, fst (snd x), diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/MicroJava/Comp/TypeInf.thy --- 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: " - (\ (E\'a prog \ (vname \ ty option)) T1 T2. + (\ (E::'a prog \ (vname \ ty option)) T1 T2. E\e :: T1 \ E\e :: T2 \ T1 = T2) \ - (\ (E\'a prog \ (vname \ ty option)) Ts1 Ts2. + (\ (E::'a prog \ (vname \ ty option)) Ts1 Ts2. E\es [::] Ts1 \ E\es [::] Ts2 \ Ts1 = Ts2)" apply (rule compat_expr_expr_list.induct) (* NewC *) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- 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\The ordering on one-dimensional vectors is linear.\ class cart_one = - assumes UNIV_one: "card (UNIV \ '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 "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" + show "\x::real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" by (rule exI[where x=0], simp) next fix c y1 y2 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- 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. \i\Basis. i \ d \ x\i = 0}" (is "affine hull (insert 0 ?A) = ?B") proof - - have *: "\A. op + (0\'a) ` A = A" "\A. op + (- (0\'a)) ` A = A" + have *: "\A. op + (0::'a) ` A = A" "\A. op + (- (0::'a)) ` A = A" by auto show ?thesis unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Derivative.thy --- 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 \Considering derivative @{typ "real \ 'b\real_normed_vector"} as a vector.\ +text \Considering derivative @{typ "real \ 'b::real_normed_vector"} as a vector.\ definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Integration.thy --- 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 \ x" using assms(5) by auto - have "((\x. 0\'a) has_integral f x - f a) {a .. x}" + have "((\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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy --- 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\Instantiation for intervals on @{text ordered_euclidean_space}\ lemma - fixes a :: "'a\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. \i\Basis. x \ i <= a \ 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\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\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\ordered_euclidean_space" + fixes a :: "'a::ordered_euclidean_space" shows "closed {a..}" by (simp add: eucl_le_atLeast[symmetric]) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- 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\euclidean_space" + fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \ ) \ x \ box a b \ box a b \ ball x e" proof - @@ -926,7 +926,7 @@ qed lemma open_UNION_box: - fixes M :: "'a\euclidean_space set" + fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f :: 'a \ real \ real. (\(i::'a)\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\euclidean_space" + fixes a :: "'a::euclidean_space" shows "{x. x i\Basis. {x. x \ i < a \ i})" "{x. a i\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less_def) lemma eucl_le_eq_halfspaces: - fixes a :: "'a\euclidean_space" + fixes a :: "'a::euclidean_space" shows "{x. \i\Basis. x \ i \ a \ i} = (\i\Basis. {x. x \ i \ a \ i})" "{x. \i\Basis. a \ i \ x \ i} = (\i\Basis. {x. a \ i \ x \ i})" by auto lemma open_Collect_eucl_less[simp, intro]: - fixes a :: "'a\euclidean_space" + fixes a :: "'a::euclidean_space" shows "open {x. x euclidean_space" + fixes a :: "'a::euclidean_space" shows "closed {x. \i\Basis. a \ i \ x \ i}" "closed {x. \i\Basis. x \ i \ a \ i}" unfolding eucl_le_eq_halfspaces diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/NSA/NSA.thy --- 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\'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\'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]: diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/NSA/StarDef.thy --- 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 "\x::'a star. x * ( *f* (\x\'a. x ^ n)) x = ( *f* (\x\'a. x * x ^ n)) x" + have "\x::'a star. x * ( *f* (\x::'a. x ^ n)) x = ( *f* (\x::'a. x * x ^ n)) x" by transfer simp with Suc show ?case by simp qed diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nat.thy --- 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 \ nat" + "0 :: nat" | Suc pred where - "pred (0 \ nat) = (0 \ 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 @@ -- \Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\ setup \Sign.mandatory_path "old"\ -old_rep_datatype "0 \ 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\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\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\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 \Difference\ -lemma diff_self_eq_0 [simp]: "(m\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\nat) \ n \ True" + "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] -lemma le0 [iff]: "0 \ (n\nat)" by (simp add: less_eq_nat.simps) -lemma [code]: "(0\nat) \ n \ True" by simp +lemma le0 [iff]: "0 \ (n::nat)" by (simp add: less_eq_nat.simps) +lemma [code]: "(0::nat) \ n \ True" by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" @@ -451,13 +451,13 @@ lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. -lemma le_0_eq [iff]: "(n\nat) \ 0 \ n = 0" +lemma le_0_eq [iff]: "(n::nat) \ 0 \ n = 0" by (induct n) (simp_all add: less_eq_nat.simps(2)) -lemma not_less0 [iff]: "\ n < (0\nat)" +lemma not_less0 [iff]: "\ n < (0::nat)" by (simp add: less_eq_Suc_le) -lemma less_nat_zero_code [code]: "n < (0\nat) \ False" +lemma less_nat_zero_code [code]: "n < (0::nat) \ False" by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" @@ -1290,10 +1290,10 @@ begin definition - "(inf \ nat \ nat \ nat) = min" + "(inf :: nat \ nat \ nat) = min" definition - "(sup \ nat \ nat \ nat) = max" + "(sup :: nat \ nat \ 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 \@{term "op dvd"} is a partial order\ -interpretation dvd: order "op dvd" "\n m \ nat. n dvd m \ \ m dvd n" +interpretation dvd: order "op dvd" "\n m :: nat. n dvd m \ \ 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 \ nat) = n" + [simp, code]: "size (n::nat) = n" instance .. diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick.thy --- 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 \ nat" where "card' A \ if finite A then length (SOME xs. set xs = A \ distinct xs) else 0" -definition setsum' :: "('a \ 'b\comm_monoid_add) \ 'a set \ 'b" where +definition setsum' :: "('a \ 'b::comm_monoid_add) \ 'a set \ 'b" where "setsum' f A \ if finite A then listsum (map f (SOME xs. set xs = A \ distinct xs)) else 0" inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" where @@ -193,7 +193,7 @@ definition less_eq_frac :: "'a \ 'a \ bool" where [nitpick_simp]: "less_eq_frac q r \ num (plus_frac q (uminus_frac r)) \ 0" -definition of_frac :: "'a \ 'b\{inverse,ring_1}" where +definition of_frac :: "'a \ 'b::{inverse,ring_1}" where "of_frac q \ of_int (num q) / of_int (denom q)" axiomatization wf_wfrec :: "('a \ 'a) set \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Core_Nits.thy --- 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\'a\'a, b\'b)}^-1 = {(b, a)}" +lemma "{(a::'a\'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\'a\'b, a) \ Id\<^sup>*" +lemma "(a::'a\'b, a) \ Id\<^sup>*" nitpick [card = 1-2, expect = none] by auto -lemma "(a\'a\'a, a) \ Id\<^sup>* \ {(a, b)}\<^sup>*" +lemma "(a::'a\'a, a) \ Id\<^sup>* \ {(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\'a, b\'a), (a, b)) \ Id" +lemma "((a::'a, b::'a), (a, b)) \ Id" nitpick [card = 1-10, expect = none] by (auto simp: Id_def) -lemma "(x\'a\'a) \ UNIV" +lemma "(x::'a\'a) \ UNIV" nitpick [card = 1-50, expect = none] sorry @@ -112,13 +112,13 @@ nitpick [card = 50, expect = genuine] oops -lemma "(a\'a\'a, a\'a\'a) \ R" +lemma "(a::'a\'a, a::'a\'a) \ R" nitpick [card = 1, expect = genuine] nitpick [card = 10, expect = genuine] nitpick [card = 5, dont_box, expect = genuine] oops -lemma "f (g\'a\'a) = x" +lemma "f (g::'a\'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\'a) = (\a. \b. \c. if c then a else b) x x True" +lemma "(x::'a) = (\a. \b. \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\'a\'a, B\'a\'a) \ R \ (A, B) \ R" +lemma "(A::'a\'a, B::'a\'a) \ R \ (A, B) \ R" nitpick [card = 15, expect = none] by auto @@ -152,7 +152,7 @@ nitpick [card = 1-25, expect = none] by auto -lemma "f = (\x\'a\'b. x)" +lemma "f = (\x::'a\'b. x)" nitpick [card = 8, expect = genuine] oops @@ -168,16 +168,16 @@ nitpick [card 'a = 100, expect = genuine] oops -lemma "\x\'a \ bool. x = y" +lemma "\x::'a \ bool. x = y" nitpick [card 'a = 1, expect = genuine] nitpick [card 'a = 100, expect = genuine] oops -lemma "\x\'a \ bool. x = y" +lemma "\x::'a \ bool. x = y" nitpick [card 'a = 1-15, expect = none] by auto -lemma "\x y\'a \ bool. x = y" +lemma "\x y::'a \ bool. x = y" nitpick [card = 1-15, expect = none] by auto @@ -208,22 +208,22 @@ nitpick [card = 1-2, expect = genuine] oops -lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. +lemma "\u::'a \ 'b. \v::'c. \w::'d. \x::'e \ 'f. f u v w x = f u (g u) w (h u w)" nitpick [card = 1-2, expect = none] sorry -lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. +lemma "\u::'a \ 'b. \v::'c. \w::'d. \x::'e \ 'f. f u v w x = f u (g u w) w (h u)" nitpick [card = 1-2, dont_box, expect = genuine] oops -lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. +lemma "\u::'a \ 'b. \v::'c. \w::'d. \x::'e \ 'f. f u v w x = f u (g u) w (h u w)" nitpick [card = 1-2, dont_box, expect = none] sorry -lemma "\u\'a \ 'b. \v\'c. \w\'d. \x\'e \ 'f. +lemma "\u::'a \ 'b. \v::'c. \w::'d. \x::'e \ '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 "\x\'a\'b. if (\y. x = y) then False else True" +lemma "\x::'a\'b. if (\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 "(\x\'a. \y. P x y) \ (\x\'a \ 'a. \y. P y x)" +lemma "(\x::'a. \y. P x y) \ (\x::'a \ 'a. \y. P y x)" nitpick [card 'a = 1, expect = genuine] oops @@ -260,7 +260,7 @@ nitpick [expect = none] by auto -lemma "let x = (\x\'a \ 'b. P x) in if x then x else \ x" +lemma "let x = (\x::'a \ 'b. P x) in if x then x else \ x" nitpick [expect = none] by auto @@ -278,7 +278,7 @@ nitpick [expect = none] by auto -schematic_lemma "\x\'a \ 'b. x = ?x" +schematic_lemma "\x::'a \ 'b. x = ?x" nitpick [expect = none] by auto @@ -491,35 +491,35 @@ nitpick [expect = none] by (simp add: snd_def) -lemma "fst (x\'a\'b, y) = x" +lemma "fst (x::'a\'b, y) = x" nitpick [expect = none] by (simp add: fst_def) -lemma "snd (x\'a\'b, y) = y" +lemma "snd (x::'a\'b, y) = y" nitpick [expect = none] by (simp add: snd_def) -lemma "fst (x, y\'a\'b) = x" +lemma "fst (x, y::'a\'b) = x" nitpick [expect = none] by (simp add: fst_def) -lemma "snd (x, y\'a\'b) = y" +lemma "snd (x, y::'a\'b) = y" nitpick [expect = none] by (simp add: snd_def) -lemma "fst (x\'a\'b, y) = x" +lemma "fst (x::'a\'b, y) = x" nitpick [expect = none] by (simp add: fst_def) -lemma "snd (x\'a\'b, y) = y" +lemma "snd (x::'a\'b, y) = y" nitpick [expect = none] by (simp add: snd_def) -lemma "fst (x, y\'a\'b) = x" +lemma "fst (x, y::'a\'b) = x" nitpick [expect = none] by (simp add: fst_def) -lemma "snd (x, y\'a\'b) = y" +lemma "snd (x, y::'a\'b) = y" nitpick [expect = none] by (simp add: snd_def) @@ -626,7 +626,7 @@ nitpick [card = 1-5, expect = none] by auto -lemma "x \ ((A\'a set) - B) \ x \ A \ x \ B" +lemma "x \ ((A::'a set) - B) \ x \ A \ x \ B" nitpick [card = 1-5, expect = none] by auto @@ -650,7 +650,7 @@ nitpick [expect = none] by auto -lemma "I = (\x\'a set. x) \ uminus = (\x. uminus (I x))" +lemma "I = (\x::'a set. x) \ uminus = (\x. uminus (I x))" nitpick [card = 1-7, expect = none] by auto @@ -662,7 +662,7 @@ nitpick [expect = none] by auto -lemma "A = -(A\'a set)" +lemma "A = -(A::'a set)" nitpick [card 'a = 10, expect = genuine] oops @@ -743,7 +743,7 @@ nitpick [expect = genuine] oops -lemma "Eps (\x. x \ P) \ (P\nat set)" +lemma "Eps (\x. x \ P) \ (P::nat set)" nitpick [expect = genuine] oops @@ -751,7 +751,7 @@ nitpick [expect = genuine] oops -lemma "\ (P \ nat \ bool) (Eps P)" +lemma "\ (P :: nat \ bool) (Eps P)" nitpick [expect = genuine] oops @@ -759,7 +759,7 @@ nitpick [expect = none] sorry -lemma "(P \ nat \ bool) \ bot \ P (Eps P)" +lemma "(P :: nat \ bool) \ bot \ P (Eps P)" nitpick [expect = none] sorry @@ -767,7 +767,7 @@ nitpick [expect = genuine] oops -lemma "(P \ nat \ bool) (The P)" +lemma "(P :: nat \ bool) (The P)" nitpick [expect = genuine] oops @@ -775,7 +775,7 @@ nitpick [expect = genuine] oops -lemma "\ (P \ nat \ bool) (The P)" +lemma "\ (P :: nat \ bool) (The P)" nitpick [expect = genuine] oops @@ -783,7 +783,7 @@ nitpick [expect = genuine] oops -lemma "The P \ (x\nat)" +lemma "The P \ (x::nat)" nitpick [expect = genuine] oops @@ -791,7 +791,7 @@ nitpick [expect = genuine] oops -lemma "P (x\nat) \ P (The P)" +lemma "P (x::nat) \ P (The P)" nitpick [expect = genuine] oops @@ -799,7 +799,7 @@ nitpick [expect = none] oops -lemma "P = {x\nat} \ (THE x. x \ P) \ P" +lemma "P = {x::nat} \ (THE x. x \ P) \ P" nitpick [expect = none] oops @@ -809,23 +809,23 @@ nitpick [expect = genuine] oops -lemma "(Q \ nat \ bool) (Eps Q)" +lemma "(Q :: nat \ bool) (Eps Q)" nitpick [expect = none] (* unfortunate *) oops -lemma "\ (Q \ nat \ bool) (Eps Q)" +lemma "\ (Q :: nat \ bool) (Eps Q)" nitpick [expect = genuine] oops -lemma "\ (Q \ nat \ bool) (Eps Q)" +lemma "\ (Q :: nat \ bool) (Eps Q)" nitpick [expect = genuine] oops -lemma "(Q\'a \ bool) \ bot \ (Q\'a \ bool) (Eps Q)" +lemma "(Q::'a \ bool) \ bot \ (Q::'a \ bool) (Eps Q)" nitpick [expect = none] sorry -lemma "(Q\nat \ bool) \ bot \ (Q\nat \ bool) (Eps Q)" +lemma "(Q::nat \ bool) \ bot \ (Q::nat \ bool) (Eps Q)" nitpick [expect = none] sorry @@ -833,7 +833,7 @@ nitpick [expect = genuine] oops -lemma "(Q\nat \ bool) (The Q)" +lemma "(Q::nat \ bool) (The Q)" nitpick [expect = genuine] oops @@ -841,7 +841,7 @@ nitpick [expect = genuine] oops -lemma "\ (Q\nat \ bool) (The Q)" +lemma "\ (Q::nat \ bool) (The Q)" nitpick [expect = genuine] oops @@ -849,7 +849,7 @@ nitpick [expect = genuine] oops -lemma "The Q \ (x\nat)" +lemma "The Q \ (x::nat)" nitpick [expect = genuine] oops @@ -857,15 +857,15 @@ nitpick [expect = genuine] oops -lemma "Q (x\nat) \ Q (The Q)" +lemma "Q (x::nat) \ Q (The Q)" nitpick [expect = genuine] oops -lemma "Q = (\x\'a. x = a) \ (Q\'a \ bool) (The Q)" +lemma "Q = (\x::'a. x = a) \ (Q::'a \ bool) (The Q)" nitpick [expect = none] sorry -lemma "Q = (\x\nat. x = a) \ (Q\nat \ bool) (The Q)" +lemma "Q = (\x::nat. x = a) \ (Q::nat \ bool) (The Q)" nitpick [expect = none] sorry @@ -921,7 +921,7 @@ subsection {* Destructors and Recursors *} -lemma "(x\'a) = (case True of True \ x | False \ x)" +lemma "(x::'a) = (case True of True \ x | False \ x)" nitpick [card = 2, expect = none] by auto diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Integer_Nits.thy --- 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 "(\i \ labels (Node x t u). f i\nat) = f x + (\i \ labels t. f i) + (\i \ labels u. f i)" +lemma "(\i \ labels (Node x t u). f i::nat) = f x + (\i \ labels t. f i) + (\i \ labels u. f i)" nitpick [expect = potential] (* unfortunate *) oops diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Manual_Nits.thy --- 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 "\i \ j; n \ (m\int)\ \ i * n + j * m \ i * m + j * n" +lemma "\i \ j; n \ (m::int)\ \ i * n + j * m \ 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 +\nat\nat\nat)" +lemma "P (op +::nat\nat\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\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 \ \(x, y) (u, v). (x + (u\nat), y + (v\nat))" +"add_raw \ \(x, y) (u, v). (x + (u::nat), y + (v::nat))" -quotient_definition "add\my_int \ my_int \ my_int" is add_raw +quotient_definition "add::my_int \ my_int \ 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\point) = Xcoord (q\point)" +lemma "Xcoord (p::point) = Xcoord (q::point)" nitpick [show_types, expect = genuine] oops -lemma "4 * x + 3 * (y\real) \ 1 / 2" +lemma "4 * x + 3 * (y::real) \ 1 / 2" nitpick [show_types, expect = genuine] oops @@ -172,7 +172,7 @@ oops inductive even' where -"even' (0\nat)" | +"even' (0::nat)" | "even' 2" | "\even' m; even' n\ \ even' (m + n)" @@ -185,7 +185,7 @@ oops coinductive nats where -"nats (x\nat) \ nats x" +"nats (x::nat) \ nats x" lemma "nats = (\n. n \ {0, 1, 2, 3, 4})" nitpick [card nat = 10, show_consts, expect = genuine] @@ -264,7 +264,7 @@ nitpick [verbose, expect = genuine] oops -lemma "\g. \x\'b. g (f x) = x \ \y\'a. \x. y = f x" +lemma "\g. \x::'b. g (f x) = x \ \y::'a. \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\nat) \ reach" | +"(4::nat) \ reach" | "n \ reach \ n < 4 \ 3 * n + 1 \ reach" | "n \ reach \ n + 2 \ reach" @@ -381,7 +381,7 @@ subsection {* 3.2. AA Trees *} -datatype 'a aa_tree = \ | N "'a\linorder" nat "'a aa_tree" "'a aa_tree" +datatype 'a aa_tree = \ | N "'a::linorder" nat "'a aa_tree" "'a aa_tree" primrec data where "data \ = undefined" | @@ -449,7 +449,7 @@ nitpick [expect = genuine] oops -theorem wf_insort\<^sub>1_nat: "wf t \ wf (insort\<^sub>1 t (x\nat))" +theorem wf_insort\<^sub>1_nat: "wf t \ wf (insort\<^sub>1 t (x::nat))" nitpick [eval = "insort\<^sub>1 t x", expect = genuine] oops diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Mini_Nits.thy --- 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 "\x. x = y"} *} ML_val {* genuine 2 @{prop "\x. x = y"} *} ML_val {* none 2 @{prop "\x. x = y"} *} -ML_val {* none 2 @{prop "\x\'a \ 'a. x = x"} *} -ML_val {* none 2 @{prop "\x\'a \ 'a. x = y"} *} -ML_val {* genuine 2 @{prop "\x\'a \ 'a. x = y"} *} -ML_val {* none 2 @{prop "\x\'a \ 'a. x = y"} *} +ML_val {* none 2 @{prop "\x::'a \ 'a. x = x"} *} +ML_val {* none 2 @{prop "\x::'a \ 'a. x = y"} *} +ML_val {* genuine 2 @{prop "\x::'a \ 'a. x = y"} *} +ML_val {* none 2 @{prop "\x::'a \ '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 \ P x = P y"} *} -ML_val {* none 4 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} -ML_val {* none 2 @{prop "(x\'a \ 'a) = y \ P x y = P y x"} *} -ML_val {* none 4 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} -ML_val {* none 2 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} -ML_val {* none 2 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} +ML_val {* none 4 @{prop "(x::'a \ 'a) = y \ P x = P y"} *} +ML_val {* none 2 @{prop "(x::'a \ 'a) = y \ P x y = P y x"} *} +ML_val {* none 4 @{prop "\x::'a \ 'a. x = y \ P x = P y"} *} +ML_val {* none 2 @{prop "(x::'a \ 'a) = y \ P x = P y"} *} +ML_val {* none 2 @{prop "\x::'a \ 'a. x = y \ P x = P y"} *} ML_val {* genuine 1 @{prop "(op =) X = Ex"} *} ML_val {* none 2 @{prop "\x::'a \ 'a. x = x"} *} ML_val {* none 1 @{prop "x = y"} *} @@ -68,8 +68,8 @@ ML_val {* genuine 1 @{prop "{a} \ {a, b}"} *} ML_val {* none 4 @{prop "{}\<^sup>+ = {}"} *} ML_val {* none 4 @{prop "UNIV\<^sup>+ = UNIV"} *} -ML_val {* none 4 @{prop "(UNIV \ ('a \ 'b) set) - {} = UNIV"} *} -ML_val {* none 4 @{prop "{} - (UNIV \ ('a \ 'b) set) = {}"} *} +ML_val {* none 4 @{prop "(UNIV :: ('a \ 'b) set) - {} = UNIV"} *} +ML_val {* none 4 @{prop "{} - (UNIV :: ('a \ '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 \ c \ {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} @@ -79,12 +79,12 @@ ML_val {* none 4 @{prop "\a b. (a, b) = (b, a)"} *} ML_val {* genuine 2 @{prop "(a, b) = (b, a)"} *} ML_val {* genuine 2 @{prop "(a, b) \ (b, a)"} *} -ML_val {* none 4 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} -ML_val {* genuine 2 @{prop "(a\'a \ 'a, b) = (b, a)"} *} -ML_val {* none 4 @{prop "\a b\'a \ 'a \ 'a. (a, b) = (b, a)"} *} -ML_val {* genuine 2 @{prop "(a\'a \ 'a \ 'a, b) \ (b, a)"} *} -ML_val {* none 4 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} -ML_val {* genuine 1 @{prop "(a\'a \ 'a, b) \ (b, a)"} *} +ML_val {* none 4 @{prop "\a b::'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a::'a \ 'a, b) = (b, a)"} *} +ML_val {* none 4 @{prop "\a b::'a \ 'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a::'a \ 'a \ 'a, b) \ (b, a)"} *} +ML_val {* none 4 @{prop "\a b::'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 1 @{prop "(a::'a \ 'a, b) \ (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 = (\a b. (b, a)) \ f x y = (y, x)"} *} ML_val {* genuine 2 @{prop "f = (\a b. (b, a)) \ f x y = (x, y)"} *} ML_val {* none 4 @{prop "f = (\x. f x)"} *} -ML_val {* none 4 @{prop "f = (\x. f x\'a \ bool)"} *} +ML_val {* none 4 @{prop "f = (\x. f x::'a \ bool)"} *} ML_val {* none 4 @{prop "f = (\x y. f x y)"} *} -ML_val {* none 4 @{prop "f = (\x y. f x y\'a \ bool)"} *} +ML_val {* none 4 @{prop "f = (\x y. f x y::'a \ bool)"} *} end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Mono_Nits.thy --- 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\('a\'b)"} *} -ML_val {* const @{term "(A\'a set) = A"} *} -ML_val {* const @{term "(A\'a set set) = A"} *} -ML_val {* const @{term "(\x\'a set. a \ x)"} *} -ML_val {* const @{term "{{a\'a}} = C"} *} -ML_val {* const @{term "{f\'a\nat} = {g\'a\nat}"} *} -ML_val {* const @{term "A \ (B\'a set)"} *} -ML_val {* const @{term "\A B x\'a. A x \ B x"} *} -ML_val {* const @{term "P (a\'a)"} *} -ML_val {* const @{term "\a\'a. b (c (d\'a)) (e\'a) (f\'a)"} *} -ML_val {* const @{term "\A\'a set. a \ A"} *} -ML_val {* const @{term "\A\'a set. P A"} *} +ML_val {* const @{term "A::('a\'b)"} *} +ML_val {* const @{term "(A::'a set) = A"} *} +ML_val {* const @{term "(A::'a set set) = A"} *} +ML_val {* const @{term "(\x::'a set. a \ x)"} *} +ML_val {* const @{term "{{a::'a}} = C"} *} +ML_val {* const @{term "{f::'a\nat} = {g::'a\nat}"} *} +ML_val {* const @{term "A \ (B::'a set)"} *} +ML_val {* const @{term "\A B x::'a. A x \ B x"} *} +ML_val {* const @{term "P (a::'a)"} *} +ML_val {* const @{term "\a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *} +ML_val {* const @{term "\A::'a set. a \ A"} *} +ML_val {* const @{term "\A::'a set. P A"} *} ML_val {* const @{term "P \ Q"} *} -ML_val {* const @{term "A \ B = (C\'a set)"} *} -ML_val {* const @{term "(\A B x\'a. A x \ 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 \ B"} *} -ML_val {* const @{term "THE x\'b. P x"} *} -ML_val {* const @{term "(\x\'a. False)"} *} -ML_val {* const @{term "(\x\'a. True)"} *} -ML_val {* const @{term "(\x\'a. False) = (\x\'a. False)"} *} -ML_val {* const @{term "(\x\'a. True) = (\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 "\ finite (A\'a set)"} *} -ML_val {* const @{term "finite (A\'a set set)"} *} -ML_val {* const @{term "\a\'a. A a \ \ B a"} *} -ML_val {* const @{term "A < (B\'a set)"} *} -ML_val {* const @{term "A \ (B\'a set)"} *} -ML_val {* const @{term "[a\'a]"} *} -ML_val {* const @{term "[a\'a set]"} *} -ML_val {* const @{term "[A \ (B\'a set)]"} *} -ML_val {* const @{term "[A \ (B\'a set)] = [C]"} *} -ML_val {* const @{term "{(\x\'a. x = a)} = C"} *} -ML_val {* const @{term "(\a\'a. \ A a) = B"} *} -ML_val {* const @{prop "\F f g (h\'a set). F f \ F g \ \ f a \ g a \ \ f a"} *} -ML_val {* const @{term "\A B x\'a. A x \ B x \ A = B"} *} -ML_val {* const @{term "p = (\(x\'a) (y\'a). P x \ \ Q y)"} *} -ML_val {* const @{term "p = (\(x\'a) (y\'a). p x y \ bool)"} *} +ML_val {* const @{term "A \ B = (C::'a set)"} *} +ML_val {* const @{term "(\A B x::'a. A x \ 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 \ B"} *} +ML_val {* const @{term "THE x::'b. P x"} *} +ML_val {* const @{term "(\x::'a. False)"} *} +ML_val {* const @{term "(\x::'a. True)"} *} +ML_val {* const @{term "(\x::'a. False) = (\x::'a. False)"} *} +ML_val {* const @{term "(\x::'a. True) = (\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 "\ finite (A::'a set)"} *} +ML_val {* const @{term "finite (A::'a set set)"} *} +ML_val {* const @{term "\a::'a. A a \ \ B a"} *} +ML_val {* const @{term "A < (B::'a set)"} *} +ML_val {* const @{term "A \ (B::'a set)"} *} +ML_val {* const @{term "[a::'a]"} *} +ML_val {* const @{term "[a::'a set]"} *} +ML_val {* const @{term "[A \ (B::'a set)]"} *} +ML_val {* const @{term "[A \ (B::'a set)] = [C]"} *} +ML_val {* const @{term "{(\x::'a. x = a)} = C"} *} +ML_val {* const @{term "(\a::'a. \ A a) = B"} *} +ML_val {* const @{prop "\F f g (h::'a set). F f \ F g \ \ f a \ g a \ \ f a"} *} +ML_val {* const @{term "\A B x::'a. A x \ B x \ A = B"} *} +ML_val {* const @{term "p = (\(x::'a) (y::'a). P x \ \ Q y)"} *} +ML_val {* const @{term "p = (\(x::'a) (y::'a). p x y :: bool)"} *} ML_val {* const @{term "p = (\A B x. A x \ \ B x) (\x. True) (\y. x \ y)"} *} ML_val {* const @{term "p = (\y. x \ y)"} *} -ML_val {* const @{term "(\x. (p\'a\bool\bool) x False)"} *} -ML_val {* const @{term "(\x y. (p\'a\'a\bool\bool) x y False)"} *} -ML_val {* const @{term "f = (\x\'a. P x \ Q x)"} *} -ML_val {* const @{term "\a\'a. P a"} *} +ML_val {* const @{term "(\x. (p::'a\bool\bool) x False)"} *} +ML_val {* const @{term "(\x y. (p::'a\'a\bool\bool) x y False)"} *} +ML_val {* const @{term "f = (\x::'a. P x \ Q x)"} *} +ML_val {* const @{term "\a::'a. P a"} *} -ML_val {* nonconst @{term "\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 "(\A B x\'a. A x \ B x) = myunion"} *} -ML_val {* nonconst @{term "(\x\'a. False) = (\x\'a. True)"} *} -ML_val {* nonconst @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} +ML_val {* nonconst @{term "\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 "(\A B x::'a. A x \ B x) = myunion"} *} +ML_val {* nonconst @{term "(\x::'a. False) = (\x::'a. True)"} *} +ML_val {* nonconst @{prop "\F f g (h::'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} -ML_val {* mono @{prop "Q (\x\'a set. P x)"} *} -ML_val {* mono @{prop "P (a\'a)"} *} -ML_val {* mono @{prop "{a} = {b\'a}"} *} -ML_val {* mono @{prop "(\x. x = a) = (\y. y = (b\'a))"} *} -ML_val {* mono @{prop "(a\'a) \ P \ P \ P = P"} *} -ML_val {* mono @{prop "\F\'a set set. P"} *} -ML_val {* mono @{prop "\ (\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h)"} *} -ML_val {* mono @{prop "\ Q (\x\'a set. P x)"} *} -ML_val {* mono @{prop "\ (\x\'a. P x)"} *} -ML_val {* mono @{prop "myall P = (P = (\x\'a. True))"} *} -ML_val {* mono @{prop "myall P = (P = (\x\'a. False))"} *} -ML_val {* mono @{prop "\x\'a. P x"} *} -ML_val {* mono @{term "(\A B x\'a. A x \ B x) \ myunion"} *} +ML_val {* mono @{prop "Q (\x::'a set. P x)"} *} +ML_val {* mono @{prop "P (a::'a)"} *} +ML_val {* mono @{prop "{a} = {b::'a}"} *} +ML_val {* mono @{prop "(\x. x = a) = (\y. y = (b::'a))"} *} +ML_val {* mono @{prop "(a::'a) \ P \ P \ P = P"} *} +ML_val {* mono @{prop "\F::'a set set. P"} *} +ML_val {* mono @{prop "\ (\F f g (h::'a set). F f \ F g \ \ a \ f \ a \ g \ F h)"} *} +ML_val {* mono @{prop "\ Q (\x::'a set. P x)"} *} +ML_val {* mono @{prop "\ (\x::'a. P x)"} *} +ML_val {* mono @{prop "myall P = (P = (\x::'a. True))"} *} +ML_val {* mono @{prop "myall P = (P = (\x::'a. False))"} *} +ML_val {* mono @{prop "\x::'a. P x"} *} +ML_val {* mono @{term "(\A B x::'a. A x \ B x) \ myunion"} *} ML_val {* nonmono @{prop "A = (\x::'a. True) \ A = (\x. False)"} *} -ML_val {* nonmono @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} +ML_val {* nonmono @{prop "\F f g (h::'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} ML {* val preproc_timeout = seconds 5.0 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Refute_Nits.thy --- 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\bool) = Q" +lemma "(P::bool) = Q" nitpick [expect = genuine] oops @@ -96,11 +96,11 @@ nitpick [expect = genuine] oops -lemma "(f\'a\'b) = g" +lemma "(f::'a\'b) = g" nitpick [expect = genuine] oops -lemma "(f\('d\'d)\('c\'d)) = g" +lemma "(f::('d\'d)\('c\'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\('a\'b)\bool) f = P g) \ (f x = g x)" +lemma "((P::('a\'b)\bool) f = P g) \ (f x = g x)" nitpick [expect = genuine] oops @@ -367,11 +367,11 @@ subsubsection {* Sets *} -lemma "P (A\'a set)" +lemma "P (A::'a set)" nitpick [expect = genuine] oops -lemma "P (A\'a set set)" +lemma "P (A::'a set set)" nitpick [expect = genuine] oops @@ -473,33 +473,33 @@ subsubsection {* Operations on Natural Numbers *} -lemma "(x\nat) + y = 0" +lemma "(x::nat) + y = 0" nitpick [expect = genuine] oops -lemma "(x\nat) = x + x" +lemma "(x::nat) = x + x" nitpick [expect = genuine] oops -lemma "(x\nat) - y + y = x" +lemma "(x::nat) - y + y = x" nitpick [expect = genuine] oops -lemma "(x\nat) = x * x" +lemma "(x::nat) = x * x" nitpick [expect = genuine] oops -lemma "(x\nat) < x + y" +lemma "(x::nat) < x + y" nitpick [card = 1, expect = genuine] oops text {* \ *} -lemma "P (x\'a\'b)" +lemma "P (x::'a\'b)" nitpick [expect = genuine] oops -lemma "\x\'a\'b. P x" +lemma "\x::'a\'b. P x" nitpick [expect = genuine] oops @@ -532,7 +532,7 @@ typedef 'a myTdef = "myTdef :: 'a set" unfolding myTdef_def by auto -lemma "(x\'a myTdef) = y" +lemma "(x::'a myTdef) = y" nitpick [expect = genuine] oops @@ -543,7 +543,7 @@ typedef 'a T_bij = "T_bij :: ('a \ 'a) set" unfolding T_bij_def by auto -lemma "P (f\(myTdecl myTdef) T_bij)" +lemma "P (f::(myTdecl myTdef) T_bij)" nitpick [expect = genuine] oops @@ -551,11 +551,11 @@ text {* unit *} -lemma "P (x\unit)" +lemma "P (x::unit)" nitpick [expect = genuine] oops -lemma "\x\unit. P x" +lemma "\x::unit. P x" nitpick [expect = genuine] oops @@ -569,11 +569,11 @@ text {* option *} -lemma "P (x\'a option)" +lemma "P (x::'a option)" nitpick [expect = genuine] oops -lemma "\x\'a option. P x" +lemma "\x::'a option. P x" nitpick [expect = genuine] oops @@ -591,11 +591,11 @@ text {* + *} -lemma "P (x\'a+'b)" +lemma "P (x::'a+'b)" nitpick [expect = genuine] oops -lemma "\x\'a+'b. P x" +lemma "\x::'a+'b. P x" nitpick [expect = genuine] oops @@ -619,11 +619,11 @@ datatype T1 = A | B -lemma "P (x\T1)" +lemma "P (x::T1)" nitpick [expect = genuine] oops -lemma "\x\T1. P x" +lemma "\x::T1. P x" nitpick [expect = genuine] oops @@ -655,11 +655,11 @@ datatype 'a T2 = C T1 | D 'a -lemma "P (x\'a T2)" +lemma "P (x::'a T2)" nitpick [expect = genuine] oops -lemma "\x\'a T2. P x" +lemma "\x::'a T2. P x" nitpick [expect = genuine] oops @@ -687,11 +687,11 @@ datatype ('a, 'b) T3 = E "'a \ 'b" -lemma "P (x\('a, 'b) T3)" +lemma "P (x::('a, 'b) T3)" nitpick [expect = genuine] oops -lemma "\x\('a, 'b) T3. P x" +lemma "\x::('a, 'b) T3. P x" nitpick [expect = genuine] oops @@ -716,11 +716,11 @@ text {* nat *} -lemma "P (x\nat)" +lemma "P (x::nat)" nitpick [expect = genuine] oops -lemma "\x\nat. P x" +lemma "\x::nat. P x" nitpick [expect = genuine] oops @@ -752,11 +752,11 @@ text {* 'a list *} -lemma "P (xs\'a list)" +lemma "P (xs::'a list)" nitpick [expect = genuine] oops -lemma "\xs\'a list. P xs" +lemma "\xs::'a list. P xs" nitpick [expect = genuine] oops @@ -782,7 +782,7 @@ nitpick [expect = genuine] oops -lemma "(xs\'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\BitList)" +lemma "P (x::BitList)" nitpick [expect = genuine] oops -lemma "\x\BitList. P x" +lemma "\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\'a BinTree)" +lemma "P (x::'a BinTree)" nitpick [expect = genuine] oops -lemma "\x\'a BinTree. P x" +lemma "\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\'a aexp)" +lemma "P (x::'a aexp)" nitpick [expect = genuine] oops -lemma "\x\'a aexp. P x" +lemma "\x::'a aexp. P x" nitpick [expect = genuine] oops @@ -872,11 +872,11 @@ nitpick [expect = genuine] oops -lemma "P (x\'a bexp)" +lemma "P (x::'a bexp)" nitpick [expect = genuine] oops -lemma "\x\'a bexp. P x" +lemma "\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\X)" +lemma "P (x::X)" nitpick [expect = genuine] oops -lemma "P (y\Y)" +lemma "P (y::Y)" nitpick [expect = genuine] oops @@ -1001,7 +1001,7 @@ datatype XOpt = CX "XOpt option" | DX "bool \ XOpt option" -lemma "P (x\XOpt)" +lemma "P (x::XOpt)" nitpick [expect = genuine] oops @@ -1019,7 +1019,7 @@ datatype 'a YOpt = CY "('a \ 'a YOpt) option" -lemma "P (x\'a YOpt)" +lemma "P (x::'a YOpt)" nitpick [expect = genuine] oops @@ -1033,11 +1033,11 @@ datatype Trie = TR "Trie list" -lemma "P (x\Trie)" +lemma "P (x::Trie)" nitpick [expect = genuine] oops -lemma "\x\Trie. P x" +lemma "\x::Trie. P x" nitpick [expect = genuine] oops @@ -1047,11 +1047,11 @@ datatype InfTree = Leaf | Node "nat \ InfTree" -lemma "P (x\InfTree)" +lemma "P (x::InfTree)" nitpick [expect = genuine] oops -lemma "\x\InfTree. P x" +lemma "\x::InfTree. P x" nitpick [expect = genuine] oops @@ -1075,11 +1075,11 @@ datatype 'a lambda = Var 'a | App "'a lambda" "'a lambda" | Lam "'a \ 'a lambda" -lemma "P (x\'a lambda)" +lemma "P (x::'a lambda)" nitpick [expect = genuine] oops -lemma "\x\'a lambda. P x" +lemma "\x::'a lambda. P x" nitpick [expect = genuine] oops @@ -1112,11 +1112,11 @@ datatype (dead 'a, 'b) T = C "'a \ bool" | D "'b list" datatype 'c U = E "('c, 'c U) T" -lemma "P (x\'c U)" +lemma "P (x::'c U)" nitpick [expect = genuine] oops -lemma "\x\'c U. P x" +lemma "\x::'c U. P x" nitpick [expect = genuine] oops @@ -1130,14 +1130,14 @@ xpos :: 'a ypos :: 'b -lemma "(x\('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\('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\'a\classA)" +lemma "P (x::'a::classA)" nitpick [expect = genuine] oops @@ -1227,11 +1227,11 @@ class classC = assumes classC_ax: "\x y. x \ y" -lemma "P (x\'a\classC)" +lemma "P (x::'a::classC)" nitpick [expect = genuine] oops -lemma "\x y. (x\'a\classC) \ y" +lemma "\x y. (x::'a::classC) \ y" nitpick [expect = none] sorry @@ -1241,7 +1241,7 @@ fixes classD_const :: "'a \ 'a" assumes classD_ax: "classD_const (classD_const x) = classD_const x" -lemma "P (x\'a\classD)" +lemma "P (x::'a::classD)" nitpick [expect = genuine] oops @@ -1249,23 +1249,23 @@ class classE = classC + classD -lemma "P (x\'a\classE)" +lemma "P (x::'a::classE)" nitpick [expect = genuine] oops text {* OFCLASS: *} -lemma "OFCLASS('a\type, type_class)" +lemma "OFCLASS('a::type, type_class)" nitpick [expect = none] apply intro_classes done -lemma "OFCLASS('a\classC, type_class)" +lemma "OFCLASS('a::classC, type_class)" nitpick [expect = none] apply intro_classes done -lemma "OFCLASS('a\type, classC_class)" +lemma "OFCLASS('a::type, classC_class)" nitpick [expect = genuine] oops @@ -1274,19 +1274,19 @@ consts inverse :: "'a \ 'a" defs (overloaded) -inverse_bool: "inverse (b\bool) \ \ b" -inverse_set: "inverse (S\'a set) \ -S" +inverse_bool: "inverse (b::bool) \ \ b" +inverse_set: "inverse (S::'a set) \ -S" inverse_pair: "inverse p \ (inverse (fst p), inverse (snd p))" lemma "inverse b" nitpick [expect = genuine] oops -lemma "P (inverse (S\'a set))" +lemma "P (inverse (S::'a set))" nitpick [expect = genuine] oops -lemma "P (inverse (p\'a\'b))" +lemma "P (inverse (p::'a\'b))" nitpick [expect = genuine] oops diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Special_Nits.thy --- 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 "\a. g a = a - \ \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>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 \ f5 g x = f5 (\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 "\a. g a = a - \ \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>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 \ f5 g x = f5 (\a. if b\<^sub>1 < b\<^sub>11 then a @@ -154,7 +154,7 @@ sorry lemma "\a. g a = a - \ \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>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 \ f5 g x = f5 (\a. if b\<^sub>1 \ b\<^sub>11 then a diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Nitpick_Examples/Typedef_Nits.thy --- 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\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 \ Abs_three 1" definition C :: three where "C \ Abs_three 2" -lemma "x = (y\three)" +lemma "x = (y::three)" nitpick [expect = genuine] oops -definition "one_or_two = {undefined False\'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\unit one_or_two)" +lemma "x = (y::unit one_or_two)" nitpick [expect = none] sorry -lemma "x = (y\bool one_or_two)" +lemma "x = (y::bool one_or_two)" nitpick [expect = genuine] oops -lemma "undefined False \ undefined True \ x = (y\bool one_or_two)" +lemma "undefined False \ undefined True \ x = (y::bool one_or_two)" nitpick [expect = none] sorry -lemma "undefined False \ undefined True \ \x (y\bool one_or_two). x \ y" +lemma "undefined False \ undefined True \ \x (y::bool one_or_two). x \ y" nitpick [card = 1, expect = potential] (* unfortunate *) oops -lemma "\x (y\bool one_or_two). x \ y" +lemma "\x (y::bool one_or_two). x \ y" nitpick [card = 1, expect = potential] (* unfortunate *) nitpick [card = 2, expect = none] oops -definition "bounded = {n\nat. finite (UNIV \ 'a set) \ n < card (UNIV \ 'a set)}" +definition "bounded = {n::nat. finite (UNIV :: 'a set) \ 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\unit bounded)" +lemma "x = (y::unit bounded)" nitpick [expect = none] sorry -lemma "x = (y\bool bounded)" +lemma "x = (y::bool bounded)" nitpick [expect = genuine] oops -lemma "x \ (y\bool bounded) \ z = x \ z = y" +lemma "x \ (y::bool bounded) \ z = x \ z = y" nitpick [expect = potential] (* unfortunate *) sorry -lemma "x \ (y\(bool \ bool) bounded) \ z = x \ z = y" +lemma "x \ (y::(bool \ bool) bounded) \ z = x \ z = y" nitpick [card = 1-5, expect = genuine] oops -lemma "True \ ((\x\bool. x) = (\x. x))" +lemma "True \ ((\x::bool. x) = (\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\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 \ n < 2" nitpick [card = 1-3, expect = none] diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Old_Number_Theory/Primes.thy --- 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\nat)*k dvd n*k \ k = 0 \ m dvd n" unfolding mult.commute[of m k] mult.commute[of n k] by presburger + "(m::nat)*k dvd n*k \ k = 0 \ 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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Orderings.thy --- 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 \ 'b\order) \ bool" where +definition mono :: "('a \ 'b::order) \ bool" where "mono f \ (\x y. x \ y \ f x \ f y)" lemma monoI [intro?]: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ mono f" unfolding mono_def by iprover lemma monoD [dest?]: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" shows "mono f \ x \ y \ f x \ f y" unfolding mono_def by iprover lemma monoE: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" assumes "mono f" assumes "x \ y" obtains "f x \ f y" @@ -1056,21 +1056,21 @@ from assms show "f x \ f y" by (simp add: mono_def) qed -definition antimono :: "('a \ 'b\order) \ bool" where +definition antimono :: "('a \ 'b::order) \ bool" where "antimono f \ (\x y. x \ y \ f x \ f y)" lemma antimonoI [intro?]: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ antimono f" unfolding antimono_def by iprover lemma antimonoD [dest?]: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" shows "antimono f \ x \ y \ f x \ f y" unfolding antimono_def by iprover lemma antimonoE: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" @@ -1078,7 +1078,7 @@ from assms show "f x \ f y" by (simp add: antimono_def) qed -definition strict_mono :: "('a \ 'b\order) \ bool" where +definition strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono f \ (\x y. x < y \ f x < f y)" lemma strict_monoI [intro?]: @@ -1112,7 +1112,7 @@ begin lemma mono_invE: - fixes f :: "'a \ 'b\order" + fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" @@ -1180,10 +1180,10 @@ lemma max_absorb2: "x \ y \ max x y = y" by (simp add: max_def) -lemma min_absorb2: "(y\'a\order) \ x \ min x y = y" +lemma min_absorb2: "(y::'a::order) \ x \ min x y = y" by (simp add:min_def) -lemma max_absorb1: "(y\'a\order) \ x \ max x y = x" +lemma max_absorb1: "(y::'a::order) \ x \ max x y = x" by (simp add: max_def) @@ -1409,7 +1409,7 @@ le_bool_def [simp]: "P \ Q \ P \ Q" definition - [simp]: "(P\bool) < Q \ \ P \ Q" + [simp]: "(P::bool) < Q \ \ P \ Q" definition [simp]: "\ \ False" @@ -1457,7 +1457,7 @@ le_fun_def: "f \ g \ (\x. f x \ g x)" definition - "(f\'a \ 'b) < g \ f \ g \ \ (g \ f)" + "(f::'a \ 'b) < g \ f \ g \ \ (g \ f)" instance .. @@ -1620,4 +1620,3 @@ lemmas linorder_antisym_conv3 = linorder_class.antisym_conv3 end - diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Power.thy --- 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"}.\ lemma nat_power_less_imp_less: - assumes nonneg: "0 < (i\nat)" + assumes nonneg: "0 < (i::nat)" assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Probability/Borel_Space.thy --- 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 \ 'x\topological_space" + fixes f :: "'a \ 'x::topological_space" assumes "\S. open S \ f -` S \ space M \ sets M" shows "f \ 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\linorder_topology" + fixes a b :: "'a::linorder_topology" shows lessThan_borel: "{..< a} \ sets borel" and greaterThan_borel: "{a <..} \ sets borel" and greaterThanLessThan_borel: "{a<.. sets borel" @@ -473,7 +473,7 @@ by auto lemma eucl_ivals[measurable]: - fixes a b :: "'a\ordered_euclidean_space" + fixes a b :: "'a::ordered_euclidean_space" shows "{x. x sets borel" and "{x. a sets borel" and "{..a} \ sets borel" @@ -597,7 +597,7 @@ using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto lemma borel_eq_box: - "borel = sigma UNIV (range (\ (a, b). box a b :: 'a \ euclidean_space set))" + "borel = sigma UNIV (range (\ (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 \ {S. open S}" @@ -611,13 +611,13 @@ lemma halfspace_gt_in_halfspace: assumes i: "i \ A" - shows "{x\'a. a < x \ i} \ - sigma_sets UNIV ((\ (a, i). {x\'a\euclidean_space. x \ i < a}) ` (UNIV \ A))" + shows "{x::'a. a < x \ i} \ + sigma_sets UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ A))" (is "?set \ ?SIGMA") proof - interpret sigma_algebra UNIV ?SIGMA by (intro sigma_algebra_sigma_sets) simp_all - have *: "?set = (\n. UNIV - {x\'a. x \ i < a + 1 / real (Suc n)})" + have *: "?set = (\n. UNIV - {x::'a. x \ i < a + 1 / real (Suc n)})" proof (safe, simp_all add: not_less) fix x :: 'a assume "a < x \ i" with reals_Archimedean[of "x \ i - a"] @@ -673,7 +673,7 @@ qed auto lemma borel_eq_halfspace_ge: - "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" + "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" @@ -683,7 +683,7 @@ qed auto lemma borel_eq_halfspace_greater: - "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a < x \ i}) ` (UNIV \ Basis))" + "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a < x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) fix a :: real and i :: 'a assume "(a, i) \ (UNIV \ Basis)" @@ -694,7 +694,7 @@ qed auto lemma borel_eq_atMost: - "borel = sigma UNIV (range (\a. {..a\'a\ordered_euclidean_space}))" + "borel = sigma UNIV (range (\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) \ UNIV \ Basis" @@ -713,7 +713,7 @@ qed auto lemma borel_eq_greaterThan: - "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {x. a a::'a::ordered_euclidean_space. {x. a UNIV \ Basis" @@ -740,7 +740,7 @@ qed auto lemma borel_eq_lessThan: - "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {x. x a::'a::ordered_euclidean_space. {x. x UNIV \ Basis" @@ -766,7 +766,7 @@ qed auto lemma borel_eq_atLeastAtMost: - "borel = sigma UNIV (range (\(a,b). {a..b} \'a\ordered_euclidean_space set))" + "borel = sigma UNIV (range (\(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 \ 'c\euclidean_space" + fixes f :: "'a \ 'c::euclidean_space" assumes F: "borel = sigma UNIV (F ` (UNIV \ Basis))" and S_eq: "\a i. S a i = f -` F (a,i) \ space M" shows "f \ borel_measurable M = (\i\Basis. \a::real. S a i \ sets M)" @@ -843,22 +843,22 @@ qed lemma borel_measurable_iff_halfspace_le: - fixes f :: "'a \ 'c\euclidean_space" + fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. f w \ i \ a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto lemma borel_measurable_iff_halfspace_less: - fixes f :: "'a \ 'c\euclidean_space" + fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. f w \ i < a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto lemma borel_measurable_iff_halfspace_ge: - fixes f :: "'a \ 'c\euclidean_space" + fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. a \ f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto lemma borel_measurable_iff_halfspace_greater: - fixes f :: "'a \ 'c\euclidean_space" + fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. a < f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Product_Type.thy --- 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\unit) v \ True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+ + "HOL.equal (u::unit) v \ True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+ code_printing type_constructor unit \ @@ -224,7 +224,7 @@ definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" where "Pair_Rep a b = (\x y. x = a \ y = b)" -definition "prod = {f. \a b. f = Pair_Rep (a\'a) (b\'b)}" +definition "prod = {f. \a b. f = Pair_Rep (a::'a) (b::'b)}" typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \ 'b \ bool) set" unfolding prod_def by auto @@ -1285,7 +1285,7 @@ unfolding image_def proof(rule set_eqI,rule iffI) fix x :: "'a \ 'c" - assume "x \ {y\'a \ 'c. \x\'b \ 'd\A \ B. y = map_prod f g x}" + assume "x \ {y::'a \ 'c. \x::'b \ 'd\A \ B. y = map_prod f g x}" then obtain y where "y \ A \ B" and "x = map_prod f g y" by blast from \image f A = A'\ and \y \ A \ B\ have "f (fst y) \ A'" by auto moreover from \image g B = B'\ and \y \ A \ B\ have "g (snd y) \ B'" by auto diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Proofs/Extraction/Euclid.thy --- 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\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 \ bool" where @@ -140,13 +140,13 @@ lemma split_all_prime: assumes "all_prime ms" and "all_prime ns" - shows "\qs. all_prime qs \ (PROD m\nat:#mset qs. m) = - (PROD m\nat:#mset ms. m) * (PROD m\nat:#mset ns. m)" (is "\qs. ?P qs \ ?Q qs") + shows "\qs. all_prime qs \ (PROD m::nat:#mset qs. m) = + (PROD m::nat:#mset ms. m) * (PROD m::nat:#mset ns. m)" (is "\qs. ?P qs \ ?Q qs") proof - from assms have "all_prime (ms @ ns)" by (simp add: all_prime_append) - moreover from assms have "(PROD m\nat:#mset (ms @ ns). m) = - (PROD m\nat:#mset ms. m) * (PROD m\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) \ ?Q (ms @ ns)" .. then show ?thesis .. @@ -154,11 +154,11 @@ lemma all_prime_nempty_g_one: assumes "all_prime ps" and "ps \ []" - shows "Suc 0 < (PROD m\nat:#mset ps. m)" + shows "Suc 0 < (PROD m::nat:#mset ps. m)" using `ps \ []` `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 \ (\ps. all_prime ps \ (PROD m\nat:#mset ps. m) = n)" +lemma factor_exists: "Suc 0 < n \ (\ps. all_prime ps \ (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 "\m k. Suc 0 < m \ Suc 0 < k \ m < n \ k < n \ 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 "\ps. all_prime ps \ (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" + from mn and m have "\ps. all_prime ps \ (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 "\ps. all_prime ps \ (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" + from kn and k have "\ps. all_prime ps \ (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 "\ps. all_prime ps \ (PROD m\nat:#mset ps. m) = - (PROD m\nat:#mset ps1. m) * (PROD m\nat:#mset ps2. m)" + have "\ps. all_prime ps \ (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\nat:#mset [n]. m) = n" by (simp add: msetprod_singleton) - ultimately have "all_prime [n] \ (PROD m\nat:#mset [n]. m) = n" .. + moreover have "(PROD m::nat:#mset [n]. m) = n" by (simp add: msetprod_singleton) + ultimately have "all_prime [n] \ (PROD m::nat:#mset [n]. m) = n" .. then show ?thesis .. qed qed @@ -193,7 +193,7 @@ shows "\p. prime p \ p dvd n" proof - from N obtain ps where "all_prime ps" - and prod_ps: "n = (PROD m\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 \ []" by (auto simp add: all_prime_nempty_g_one msetprod_empty) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Quickcheck_Random.thy --- 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 @@ \ Random.seed \ (('a \ 'b) \ (unit \ term)) \ Random.seed" definition random_fun_lift :: "(Random.seed \ ('b \ (unit \ term)) \ Random.seed) - \ Random.seed \ (('a\term_of \ 'b\typerep) \ (unit \ term)) \ Random.seed" + \ Random.seed \ (('a::term_of \ 'b::typerep) \ (unit \ term)) \ Random.seed" where "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed" diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Quotient_Examples/Quotient_Int.thy --- 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 \ int" is "(0\nat, 0\nat)" done + "0 :: int" is "(0::nat, 0::nat)" done quotient_definition - "1 \ int" is "(1\nat, 0\nat)" done + "1 :: int" is "(1::nat, 0::nat)" done fun plus_int_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" @@ -33,7 +33,7 @@ "plus_int_raw (x, y) (u, v) = (x + u, y + v)" quotient_definition - "(op +) \ (int \ int \ int)" is "plus_int_raw" by auto + "(op +) :: (int \ int \ int)" is "plus_int_raw" by auto fun uminus_int_raw :: "(nat \ nat) \ (nat \ nat)" @@ -41,10 +41,10 @@ "uminus_int_raw (x, y) = (y, x)" quotient_definition - "(uminus \ (int \ int))" is "uminus_int_raw" by auto + "(uminus :: (int \ int))" is "uminus_int_raw" by auto definition - minus_int_def: "z - w = z + (-w\int)" + minus_int_def: "z - w = z + (-w::int)" fun times_int_raw :: "(nat \ nat) \ (nat \ nat) \ (nat \ nat)" @@ -95,13 +95,13 @@ le_int_def: "(op \) :: int \ int \ bool" is "le_int_raw" by auto definition - less_int_def: "(z\int) < w = (z \ w \ z \ w)" + less_int_def: "(z::int) < w = (z \ w \ z \ w)" definition - zabs_def: "\i\int\ = (if i < 0 then - i else i)" + zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition - zsgn_def: "sgn (i\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 \ int \ int \ int) = min" + "(inf :: int \ int \ int) = min" definition - "(sup \ int \ int \ int) = max" + "(sup :: int \ int \ int) = max" instance by default @@ -245,7 +245,7 @@ by (rule zmult_zless_mono2) show "\i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) - show "sgn (i\int) = (if i=0 then 0 else if 0 rat" is "(0\int, 1\int)" by simp + "0 :: rat" is "(0::int, 1::int)" by simp quotient_definition - "1 \ rat" is "(1\int, 1\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 \ (rat \ rat))" is "uminus_rat_raw" by fastforce + "(uminus :: (rat \ rat))" is "uminus_rat_raw" by fastforce definition - minus_rat_def: "a - b = a + (-b\rat)" + minus_rat_def: "a - b = a + (-b::rat)" fun le_rat_raw where "le_rat_raw (a :: int, b) (c, d) \ (a * d) * (b * d) \ (c * b) * (b * d)" @@ -92,13 +92,13 @@ qed definition - less_rat_def: "(z\rat) < w = (z \ w \ z \ w)" + less_rat_def: "(z::rat) < w = (z \ w \ z \ w)" definition - rabs_rat_def: "\i\rat\ = (if i < 0 then - i else i)" + rabs_rat_def: "\i::rat\ = (if i < 0 then - i else i)" definition - sgn_rat_def: "sgn (i\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 \ 0" then have "a * b \ (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 "\z\int. a * b \ z * b * b" by auto + then show "\z::int. a * b \ z * b * b" by auto qed qed *) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Real.thy --- 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\real) \ (1\real)" + show "(0::real) \ (1::real)" by transfer (simp add: realrel_def) show "a \ 0 \ inverse a * a = 1" apply transfer @@ -1975,7 +1975,7 @@ instantiation real :: equal begin -definition "HOL.equal (x\real) y \ x - y = 0" +definition "HOL.equal (x::real) y \ x - y = 0" instance proof qed (simp add: equal_real_def) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/SPARK/Examples/RIPEMD-160/Round.thy --- 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\" have steps_to_steps': "steps - (\n\nat. word_of_int (x (int n))) + (\n::nat. word_of_int (x (int n))) (from_chain ?INIT_CHAIN, from_chain ?INIT_CHAIN) 80 = from_chain_pair ( diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/String.thy --- 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 \ nat" where - [code]: "size_literal (s\literal) = 0" + [code]: "size_literal (s::literal) = 0" instance .. diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Sum_Type.thy --- 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: "\x\'a. P (Inl x)" and y: "\y\'b. P (Inr y)" + assume x: "\x::'a. P (Inl x)" and y: "\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 (\x::'a. f (Inl x)) (\y::'b. f (Inr y)) = f" proof fix s :: "'a + 'b" - show "(case s of Inl (x\'a) \ f (Inl x) | Inr (y\'b) \ f (Inr y)) = f s" + show "(case s of Inl (x::'a) \ f (Inl x) | Inr (y::'b) \ 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 \ '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 \ '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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/TPTP/THF_Arith.thy --- 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 \ "is_int :: rat \ bool" begin - definition "rat_is_int (q\rat) \ (\n\int. q = of_int n)" + definition "rat_is_int (q::rat) \ (\n::int. q = of_int n)" end overloading real_is_int \ "is_int :: real \ bool" begin - definition "real_is_int (x\real) \ x \ \" + definition "real_is_int (x::real) \ x \ \" end overloading real_is_rat \ "is_rat :: real \ bool" begin - definition "real_is_rat (x\real) \ x \ \" + definition "real_is_rat (x::real) \ x \ \" end consts @@ -36,32 +36,32 @@ overloading rat_to_int \ "to_int :: rat \ int" begin - definition "rat_to_int (q\rat) = floor q" + definition "rat_to_int (q::rat) = floor q" end overloading real_to_int \ "to_int :: real \ int" begin - definition "real_to_int (x\real) = floor x" + definition "real_to_int (x::real) = floor x" end overloading int_to_rat \ "to_rat :: int \ rat" begin - definition "int_to_rat (n\int) = (of_int n\rat)" + definition "int_to_rat (n::int) = (of_int n::rat)" end overloading real_to_rat \ "to_rat :: real \ rat" begin - definition "real_to_rat (x\real) = (inv of_rat x\rat)" + definition "real_to_rat (x::real) = (inv of_rat x::rat)" end overloading int_to_real \ "to_real :: int \ real" begin - definition "int_to_real (n\int) = real n" + definition "int_to_real (n::int) = real n" end overloading rat_to_real \ "to_real :: rat \ real" begin - definition "rat_to_real (x\rat) = (of_rat x\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\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\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\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\rat\real)" +lemma inj_of_rat [intro, simp]: "inj (of_rat::rat\real)" by (metis injI of_rat_eq_iff) end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy --- 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 "\(Xj\TPTP_Interpret.ind) Xk\TPTP_Interpret.ind. +lemma "\(Xj::TPTP_Interpret.ind) Xk::TPTP_Interpret.ind. bnd_cCKB6_BLACK Xj Xk \ bnd_cCKB6_BLACK (bnd_s (bnd_s (bnd_s Xj))) (bnd_s Xk)" apply (tactic {*nth (nth just_the_tacs 0) 0*}) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy --- 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 "\(SV5\TPTP_Interpret.ind \ bool) - SV6\TPTP_Interpret.ind. +lemma "\(SV5::TPTP_Interpret.ind \ 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 "\(SV7\TPTP_Interpret.ind) SV8\TPTP_Interpret.ind. +lemma "\(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 - "\(SV2\TPTP_Interpret.ind) - SV1\TPTP_Interpret.ind \ TPTP_Interpret.ind. + "\(SV2::TPTP_Interpret.ind) + SV1::TPTP_Interpret.ind \ TPTP_Interpret.ind. (SV1 SV2 = bnd_sK1_Xy) = False \ - \SV2\TPTP_Interpret.ind. + \SV2::TPTP_Interpret.ind. (bnd_sK1_Xy = bnd_sK1_Xy) = False" ML_prf {* @@ -462,12 +462,12 @@ SEU602_2_bnd_in :: "TPTP_Interpret.ind \ TPTP_Interpret.ind \ bool" (* (Annotated_step ("113", "extuni_func"), *) -lemma "\SV49\TPTP_Interpret.ind \ bool. +lemma "\SV49::TPTP_Interpret.ind \ bool. (SV49 = - (\SY23\TPTP_Interpret.ind. + (\SY23::TPTP_Interpret.ind. \ SEU602_2_bnd_in SY23 SEU602_2_bnd_sK2_SY17)) = False \ - \SV49\TPTP_Interpret.ind \ bool. + \SV49::TPTP_Interpret.ind \ bool. (SV49 (SEU602_2_bnd_sK7_E SV49) = (\ 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 \ bool) \ TPTP_Interpret.ind" SEV405_5_bnd_cA :: bool -lemma "\SV1\TPTP_Interpret.ind \ bool. - (\SY2\TPTP_Interpret.ind. +lemma "\SV1::TPTP_Interpret.ind \ bool. + (\SY2::TPTP_Interpret.ind. \ (\ (\ SV1 SY2 \ SEV405_5_bnd_cA) \ \ (\ SEV405_5_bnd_cA \ SV1 SY2))) = False \ - \SV1\TPTP_Interpret.ind \ bool. + \SV1::TPTP_Interpret.ind \ bool. (\ (\ (\ SV1 (SEV405_5_bnd_sK1_SY2 SV1) \ SEV405_5_bnd_cA) \ \ (\ SEV405_5_bnd_cA \ SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) = False" @@ -1046,28 +1046,28 @@ \ TPTP_Interpret.ind \ TPTP_Interpret.ind \ TPTP_Interpret.ind \ bool" -lemma "\(SV4\TPTP_Interpret.ind) (SV8\TPTP_Interpret.ind) - (SV6\TPTP_Interpret.ind) (SV2\TPTP_Interpret.ind) - (SV3\TPTP_Interpret.ind) SV1\TPTP_Interpret.ind. +lemma "\(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) = False \ PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \ PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \ -\(SV4\TPTP_Interpret.ind) (SV8\TPTP_Interpret.ind) - (SV6\TPTP_Interpret.ind) (SV2\TPTP_Interpret.ind) - (SV3\TPTP_Interpret.ind) SV1\TPTP_Interpret.ind. +\(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 \ PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \ PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False" by (tactic {*nonfull_extcnf_combined_tac @{context} [Not_neg]*}) lemma " -\(SV8\TPTP_Interpret.ind) (SV6\TPTP_Interpret.ind) - (SV4\TPTP_Interpret.ind) (SV2\TPTP_Interpret.ind) - (SV3\TPTP_Interpret.ind) SV1\TPTP_Interpret.ind. +\(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 \ SV3 \ SV2 \ SV4) = False \ PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \ PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False \ -\(SV4\TPTP_Interpret.ind) (SV8\TPTP_Interpret.ind) - (SV6\TPTP_Interpret.ind) (SV2\TPTP_Interpret.ind) - (SV3\TPTP_Interpret.ind) SV1\TPTP_Interpret.ind. +\(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) = False \ PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \ 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 \ TPTP_Interpret.ind \ bool" (* (Annotated_step ("6", "unfold_def"), *) -lemma "((((((((((((\X\TPTP_Interpret.ind. \ NUM016_5_bnd_less X X) \ - (\(X\TPTP_Interpret.ind) - Y\TPTP_Interpret.ind. +lemma "((((((((((((\X::TPTP_Interpret.ind. \ NUM016_5_bnd_less X X) \ + (\(X::TPTP_Interpret.ind) + Y::TPTP_Interpret.ind. \ NUM016_5_bnd_less X Y \ \ NUM016_5_bnd_less Y X)) \ - (\X\TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \ - (\(X\TPTP_Interpret.ind) - (Y\TPTP_Interpret.ind) - Z\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. NUM016_5_bnd_divides X X)) \ + (\(X::TPTP_Interpret.ind) + (Y::TPTP_Interpret.ind) + Z::TPTP_Interpret.ind. (\ NUM016_5_bnd_divides X Y \ \ NUM016_5_bnd_divides Y Z) \ NUM016_5_bnd_divides X Z)) \ - (\(X\TPTP_Interpret.ind) Y\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind. \ NUM016_5_bnd_divides X Y \ \ NUM016_5_bnd_less Y X)) \ - (\X\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. NUM016_5_bnd_less X (NUM016_5_bnd_factorial_plus_one X))) \ - (\(X\TPTP_Interpret.ind) Y\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind) Y::TPTP_Interpret.ind. \ NUM016_5_bnd_divides X (NUM016_5_bnd_factorial_plus_one Y) \ NUM016_5_bnd_less Y X)) \ - (\X\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. NUM016_5_bnd_prime X \ NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor X) X)) \ - (\X\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. NUM016_5_bnd_prime X \ NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor X))) \ - (\X\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. NUM016_5_bnd_prime X \ NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor X) X)) \ NUM016_5_bnd_prime NUM016_5_bnd_a) \ - (\X\TPTP_Interpret.ind. + (\X::TPTP_Interpret.ind. (\ NUM016_5_bnd_prime X \ \ NUM016_5_bnd_less NUM016_5_bnd_a X) \ NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a) X)) = True \ - (\ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ (\SX0\TPTP_Interpret.ind. + (\ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ \ (\ (\SX0::TPTP_Interpret.ind. \ NUM016_5_bnd_less SX0 SX0) \ - \ (\(SX0\TPTP_Interpret.ind) - SX1\TPTP_Interpret.ind. + \ (\(SX0::TPTP_Interpret.ind) + SX1::TPTP_Interpret.ind. \ NUM016_5_bnd_less SX0 SX1 \ \ NUM016_5_bnd_less SX1 SX0)) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. NUM016_5_bnd_divides SX0 SX0)) \ - \ (\(SX0\TPTP_Interpret.ind) - (SX1\TPTP_Interpret.ind) - SX2\TPTP_Interpret.ind. + \ (\(SX0::TPTP_Interpret.ind) + (SX1::TPTP_Interpret.ind) + SX2::TPTP_Interpret.ind. (\ NUM016_5_bnd_divides SX0 SX1 \ \ NUM016_5_bnd_divides SX1 SX2) \ NUM016_5_bnd_divides SX0 SX2)) \ - \ (\(SX0\TPTP_Interpret.ind) - SX1\TPTP_Interpret.ind. + \ (\(SX0::TPTP_Interpret.ind) + SX1::TPTP_Interpret.ind. \ NUM016_5_bnd_divides SX0 SX1 \ \ NUM016_5_bnd_less SX1 SX0)) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. NUM016_5_bnd_less SX0 (NUM016_5_bnd_factorial_plus_one SX0))) \ - \ (\(SX0\TPTP_Interpret.ind) SX1\TPTP_Interpret.ind. + \ (\(SX0::TPTP_Interpret.ind) SX1::TPTP_Interpret.ind. \ NUM016_5_bnd_divides SX0 (NUM016_5_bnd_factorial_plus_one SX1) \ NUM016_5_bnd_less SX1 SX0)) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. NUM016_5_bnd_prime SX0 \ NUM016_5_bnd_divides (NUM016_5_bnd_prime_divisor SX0) SX0)) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. NUM016_5_bnd_prime SX0 \ NUM016_5_bnd_prime (NUM016_5_bnd_prime_divisor SX0))) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. NUM016_5_bnd_prime SX0 \ NUM016_5_bnd_less (NUM016_5_bnd_prime_divisor SX0) SX0)) \ \ NUM016_5_bnd_prime NUM016_5_bnd_a) \ - \ (\SX0\TPTP_Interpret.ind. + \ (\SX0::TPTP_Interpret.ind. (\ NUM016_5_bnd_prime SX0 \ \ NUM016_5_bnd_less NUM016_5_bnd_a SX0) \ 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 "(\SX0\TPTP_Interpret.ind. +lemma "(\SX0::TPTP_Interpret.ind. AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola SX0) = False \ 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 "\(SV13\TPTP_Interpret.ind) (SV39\AGT037_2_bnd_mu) (SV29\AGT037_2_bnd_mu) - SV45\TPTP_Interpret.ind. - ((((\SY68\TPTP_Interpret.ind. +lemma "\(SV13::TPTP_Interpret.ind) (SV39::AGT037_2_bnd_mu) (SV29::AGT037_2_bnd_mu) + SV45::TPTP_Interpret.ind. + ((((\SY68::TPTP_Interpret.ind. \ AGT037_2_bnd_a1 SV45 SY68 \ AGT037_2_bnd_likes SV29 SV39 SY68) = False \ - (\ (\SY69\TPTP_Interpret.ind. + (\ (\SY69::TPTP_Interpret.ind. \ AGT037_2_bnd_a2 SV45 SY69 \ AGT037_2_bnd_likes SV29 SV39 SY69)) = True) \ AGT037_2_bnd_likes SV29 SV39 SV45 = False) \ AGT037_2_bnd_very_much_likes SV29 SV39 SV45 = True) \ AGT037_2_bnd_a3 SV13 SV45 = False \ - \(SV29\AGT037_2_bnd_mu) (SV39\AGT037_2_bnd_mu) (SV13\TPTP_Interpret.ind) - SV45\TPTP_Interpret.ind. + \(SV29::AGT037_2_bnd_mu) (SV39::AGT037_2_bnd_mu) (SV13::TPTP_Interpret.ind) + SV45::TPTP_Interpret.ind. ((((\ AGT037_2_bnd_a1 SV45 (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45) \ AGT037_2_bnd_likes SV29 SV39 (AGT037_2_bnd_sK5_SY68 SV13 SV39 SV29 SV45)) = False \ - (\ (\SY69\TPTP_Interpret.ind. + (\ (\SY69::TPTP_Interpret.ind. \ AGT037_2_bnd_a2 SV45 SY69 \ AGT037_2_bnd_likes SV29 SV39 SY69)) = True) \ @@ -1555,10 +1555,10 @@ bnd_addition bnd_sK2_X2 bnd_sK1_X1) = True \ (bnd_sup - (\SX0\TPTP_Interpret.ind. + (\SX0::TPTP_Interpret.ind. SX0 = bnd_sK1_X1 \ SX0 = bnd_sK2_X2) \ bnd_sup - (\SX0\TPTP_Interpret.ind. + (\SX0::TPTP_Interpret.ind. SX0 = bnd_sK2_X2 \ 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 = - (\(X\TPTP_Interpret.ind \ bool) - (Y\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) + (Y::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. X U \ Y U) \ bnd_mnot = - (\(X\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. \ X U) \ bnd_mimplies = - (\U\TPTP_Interpret.ind \ bool. bnd_mor (bnd_mnot U)) \ + (\U::TPTP_Interpret.ind \ bool. bnd_mor (bnd_mnot U)) \ bnd_mbox_s4 = - (\(P\TPTP_Interpret.ind \ bool) X\TPTP_Interpret.ind. - \Y\TPTP_Interpret.ind. bnd_irel X Y \ P Y) \ + (\(P::TPTP_Interpret.ind \ bool) X::TPTP_Interpret.ind. + \Y::TPTP_Interpret.ind. bnd_irel X Y \ P Y) \ bnd_mand = - (\(X\TPTP_Interpret.ind \ bool) - (Y\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) + (Y::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. X U \ Y U) \ bnd_ixor = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_inot (bnd_iequiv P Q)) \ bnd_ivalid = All \ - bnd_itrue = (\W\TPTP_Interpret.ind. True) \ + bnd_itrue = (\W::TPTP_Interpret.ind. True) \ bnd_isatisfiable = Ex \ bnd_ior = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \ bnd_inot = - (\P\TPTP_Interpret.ind \ bool. + (\P::TPTP_Interpret.ind \ bool. bnd_mnot (bnd_mbox_s4 P)) \ bnd_iinvalid = - (\Phi\TPTP_Interpret.ind \ bool. - \W\TPTP_Interpret.ind. \ Phi W) \ + (\Phi::TPTP_Interpret.ind \ bool. + \W::TPTP_Interpret.ind. \ Phi W) \ bnd_iimplies = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \ bnd_iimplied = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. bnd_iimplies Q P) \ + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_iimplies Q P) \ bnd_ifalse = bnd_inot bnd_itrue \ bnd_iequiv = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \ bnd_icountersatisfiable = - (\Phi\TPTP_Interpret.ind \ bool. - \W\TPTP_Interpret.ind. \ Phi W) \ - bnd_iatom = (\P\TPTP_Interpret.ind \ bool. P) \ + (\Phi::TPTP_Interpret.ind \ bool. + \W::TPTP_Interpret.ind. \ Phi W) \ + bnd_iatom = (\P::TPTP_Interpret.ind \ bool. P) \ bnd_iand = bnd_mand \ - (\(X\TPTP_Interpret.ind) (Y\TPTP_Interpret.ind) - Z\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind) + Z::TPTP_Interpret.ind. bnd_irel X Y \ bnd_irel Y Z \ bnd_irel X Z) \ - (\(X\TPTP_Interpret.ind) (Y\TPTP_Interpret.ind) - Z\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind) (Y::TPTP_Interpret.ind) + Z::TPTP_Interpret.ind. bnd_irel X Y \ bnd_irel Y Z \ bnd_irel X Z) = True" (* by (tactic {*tectoc @{context}*}) *) @@ -2072,61 +2072,61 @@ (* (Annotated_step ("11", "unfold_def"), *) lemma "bnd_mor = - (\(X\TPTP_Interpret.ind \ bool) - (Y\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) + (Y::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. X U \ Y U) \ bnd_mnot = - (\(X\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. \ X U) \ bnd_mimplies = - (\U\TPTP_Interpret.ind \ bool. bnd_mor (bnd_mnot U)) \ + (\U::TPTP_Interpret.ind \ bool. bnd_mor (bnd_mnot U)) \ bnd_mbox_s4 = - (\(P\TPTP_Interpret.ind \ bool) X\TPTP_Interpret.ind. - \Y\TPTP_Interpret.ind. bnd_irel X Y \ P Y) \ + (\(P::TPTP_Interpret.ind \ bool) X::TPTP_Interpret.ind. + \Y::TPTP_Interpret.ind. bnd_irel X Y \ P Y) \ bnd_mand = - (\(X\TPTP_Interpret.ind \ bool) - (Y\TPTP_Interpret.ind \ bool) U\TPTP_Interpret.ind. + (\(X::TPTP_Interpret.ind \ bool) + (Y::TPTP_Interpret.ind \ bool) U::TPTP_Interpret.ind. X U \ Y U) \ bnd_ixor = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_inot (bnd_iequiv P Q)) \ bnd_ivalid = All \ - bnd_itrue = (\W\TPTP_Interpret.ind. True) \ + bnd_itrue = (\W::TPTP_Interpret.ind. True) \ bnd_isatisfiable = Ex \ bnd_ior = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_mor (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \ bnd_inot = - (\P\TPTP_Interpret.ind \ bool. + (\P::TPTP_Interpret.ind \ bool. bnd_mnot (bnd_mbox_s4 P)) \ bnd_iinvalid = - (\Phi\TPTP_Interpret.ind \ bool. - \W\TPTP_Interpret.ind. \ Phi W) \ + (\Phi::TPTP_Interpret.ind \ bool. + \W::TPTP_Interpret.ind. \ Phi W) \ bnd_iimplies = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_mimplies (bnd_mbox_s4 P) (bnd_mbox_s4 Q)) \ bnd_iimplied = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. bnd_iimplies Q P) \ + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_iimplies Q P) \ bnd_ifalse = bnd_inot bnd_itrue \ bnd_iequiv = - (\(P\TPTP_Interpret.ind \ bool) - Q\TPTP_Interpret.ind \ bool. + (\(P::TPTP_Interpret.ind \ bool) + Q::TPTP_Interpret.ind \ bool. bnd_iand (bnd_iimplies P Q) (bnd_iimplies Q P)) \ bnd_icountersatisfiable = - (\Phi\TPTP_Interpret.ind \ bool. - \W\TPTP_Interpret.ind. \ Phi W) \ - bnd_iatom = (\P\TPTP_Interpret.ind \ bool. P) \ + (\Phi::TPTP_Interpret.ind \ bool. + \W::TPTP_Interpret.ind. \ Phi W) \ + bnd_iatom = (\P::TPTP_Interpret.ind \ bool. P) \ bnd_iand = bnd_mand \ bnd_ivalid (bnd_iimplies (bnd_iatom bnd_q) (bnd_iatom bnd_r)) \ - (\SY161\TPTP_Interpret.ind. - \ (\SY162\TPTP_Interpret.ind. + (\SY161::TPTP_Interpret.ind. + \ (\SY162::TPTP_Interpret.ind. bnd_irel SY161 SY162 \ bnd_q SY162) \ - (\SY163\TPTP_Interpret.ind. + (\SY163::TPTP_Interpret.ind. bnd_irel SY161 SY163 \ bnd_r SY163)) = True" (* by (tactic {*tectoc @{context}*}) *) @@ -2147,12 +2147,12 @@ *) (* (Annotated_step ("11", "extcnf_forall_neg"), *) -lemma "\SV1\TPTP_Interpret.ind \ bool. - (\SY2\TPTP_Interpret.ind. +lemma "\SV1::TPTP_Interpret.ind \ bool. + (\SY2::TPTP_Interpret.ind. \ (\ (\ SV1 SY2 \ SEV405_5_bnd_cA) \ \ (\ SEV405_5_bnd_cA \ SV1 SY2))) = False \ - \SV1\TPTP_Interpret.ind \ bool. + \SV1::TPTP_Interpret.ind \ bool. (\ (\ (\ SV1 (SEV405_5_bnd_sK1_SY2 SV1) \ SEV405_5_bnd_cA) \ \ (\ SEV405_5_bnd_cA \ 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 "(\SX0\bool \ bool. +lemma "(\SX0::bool \ bool. \ \ (\ SX0 bnd_sK1_Xx \ \ SX0 bnd_sK2_Xy)) = True \ (\ \ (\ True \ \ True)) = True" @@ -2192,7 +2192,7 @@ done (* (Annotated_step ("13", "extcnf_forall_special_pos"), *) -lemma "(\SX0\bool \ bool. +lemma "(\SX0::bool \ bool. \ \ (\ SX0 bnd_sK1_Xx \ \ SX0 bnd_sK2_Xy)) = True \ (\ \ (\ bnd_sK1_Xx \ \ bnd_sK2_Xy)) = True" @@ -2200,9 +2200,9 @@ done (* [[(Annotated_step ("8", "polarity_switch"), *) -lemma "(\(Xx\bool) (Xy\bool) Xz\bool. True \ True \ True) = +lemma "(\(Xx::bool) (Xy::bool) Xz::bool. True \ True \ True) = False \ - (\ (\(Xx\bool) (Xy\bool) Xz\bool. + (\ (\(Xx::bool) (Xy::bool) Xz::bool. True \ True \ True)) = True" apply (tactic {*nonfull_extcnf_combined_tac @{context} [Polarity_switch]*}) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Transcendental.thy --- 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: "((\n. 1 / of_nat n) ---> (0::'a\real_normed_field)) sequentially" +lemma lim_1_over_n: "((\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: "((\n. inverse(of_nat n)) ---> (0::'a\real_normed_field)) sequentially" +lemma lim_inverse_n: "((\n. inverse(of_nat n)) ---> (0::'a::real_normed_field)) sequentially" using lim_1_over_n by (simp add: inverse_eq_divide) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/UNITY/Follows.thy --- 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\order multiset \ 'a multiset \ bool" where +definition less_multiset :: "'a::order multiset \ 'a multiset \ bool" where "M' < M \ M' #<# M" definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" where diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/Word/Word.thy --- 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 \ a = b" by (simp add: word_uint_eq_iff) -definition word_of_int :: "int \ 'a\len0 word" +definition word_of_int :: "int \ '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\'a\len word) ^ n) !! m \ m = n \ m < len_of TYPE('a\len)" + "((2::'a::len word) ^ n) !! m \ m = n \ 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) \ - (\n\nat. b !! n \ n < size b \ c !! n) \ (\m\nat. a !! m \ m < size a \ c !! (m + size b))" + (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) <-> diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ZF/HOLZF.thy --- 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:"\x y. \ (x < y \ (\(m\nat). y \ x + m))" by presburger - have th': "\x y. \ (x \ y \ (\ x < y) \ (\(m\nat). x \ y + m))" by presburger + have th:"\x y. \ (x < y \ (\(m::nat). y \ x + m))" by presburger + have th': "\x y. \ (x \ y \ (\ x < y) \ (\(m::nat). x \ y + m))" by presburger show ?thesis apply (auto simp add: inj_on_def) apply (case_tac "x = y") diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/Dedekind_Real.thy --- 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\preal)" +definition "R div S = R * inverse (S::preal)" definition preal_one_def: @@ -220,10 +220,10 @@ begin definition - "(inf \ preal \ preal \ preal) = min" + "(inf :: preal \ preal \ preal) = min" definition - "(sup \ preal \ preal \ preal) = max" + "(sup :: preal \ preal \ preal) = max" instance by intro_classes @@ -1229,7 +1229,7 @@ (\x y u v. x+v \ u+y & (x,y) \ Rep_Real z & (u,v) \ Rep_Real w)" definition - real_less_def: "x < (y\real) \ x \ y \ x \ y" + real_less_def: "x < (y::real) \ x \ y \ x \ y" definition real_abs_def: "abs (r::real) = (if r < 0 then - r else r)" @@ -1561,10 +1561,10 @@ begin definition - "(inf \ real \ real \ real) = min" + "(inf :: real \ real \ real) = min" definition - "(sup \ real \ real \ real) = max" + "(sup :: real \ real \ real) = max" instance by default (auto simp add: inf_real_def sup_real_def max_min_distrib2) diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/Eval_Examples.thy --- 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\order" 'b + Foo "'a::order" 'b | Bla "('a, 'b) bar" | Dummy nat and ('a, 'b) bar = diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/FinFunPred.thy --- 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 \ g \ (\x. f $ x \ g $ x)" -definition [code del]: "(f\'a \f 'b) < g \ f \ g \ \ g \ f" +definition [code del]: "(f::'a \f 'b) < g \ f \ g \ \ g \ f" instance .. diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/Normalization_by_Evaluation.thy --- 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\int)) = -164" by normalization -lemma "(if (0\nat) \ (x\nat) then 0\nat else x) = 0" by normalization +lemma "4 - 42 * abs (3 + (-7::int)) = -164" by normalization +lemma "(if (0::nat) \ (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 \ '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 +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 diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/Transfer_Ex.thy --- 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 \ (y\int) \ 0 \ (x\int) \ x + y = y + x" +lemma "0 \ (y::int) \ 0 \ (x::int) \ x + y = y + x" by (fact ex1 [transferred]) (* Using new transfer package *) -lemma "0 \ (x\int) \ 0 \ (y\int) \ x + y = y + x" +lemma "0 \ (x::int) \ 0 \ (y::int) \ 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 \ (b\int) \ 0 \ (a\int) \ a div b * b + a mod b = a" +lemma "0 \ (b::int) \ 0 \ (a::int) \ a div b * b + a mod b = a" by (fact ex2 [transferred]) (* Using new transfer package *) -lemma "0 \ (a\int) \ 0 \ (b\int) \ a div b * b + a mod b = a" +lemma "0 \ (a::int) \ 0 \ (b::int) \ 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 "\x\0\int. \y\0. \z\0. x + y \ z" +lemma "\x\0::int. \y\0. \z\0. x + y \ z" by (fact ex3 [transferred nat_int]) (* Using new transfer package *) -lemma "\x\int\{0..}. \y\{0..}. \z\{0..}. x + y \ z" +lemma "\x::int\{0..}. \y\{0..}. \z\{0..}. x + y \ z" by (fact ex3 [untransferred]) lemma ex4: "(x::nat) >= y \ (x - y) + y = x" by auto -lemma "0 \ (x\int) \ 0 \ (y\int) \ y \ x \ tsub x y + y = x" +lemma "0 \ (x::int) \ 0 \ (y::int) \ y \ x \ tsub x y + y = x" by (fact ex4 [transferred]) (* Using new transfer package *) -lemma "0 \ (y\int) \ 0 \ (x\int) \ y \ x \ tsub x y + y = x" +lemma "0 \ (y::int) \ 0 \ (x::int) \ y \ x \ tsub x y + y = x" by (fact ex4 [untransferred]) lemma ex5: "(2::nat) * \{..n} = n * (n + 1)" by (induct n rule: nat_induct, auto) -lemma "0 \ (n\int) \ 2 * \{0..n} = n * (n + 1)" +lemma "0 \ (n::int) \ 2 * \{0..n} = n * (n + 1)" by (fact ex5 [transferred]) (* Using new transfer package *) -lemma "0 \ (n\int) \ 2 * \{0..n} = n * (n + 1)" +lemma "0 \ (n::int) \ 2 * \{0..n} = n * (n + 1)" by (fact ex5 [untransferred]) -lemma "0 \ (n\nat) \ 2 * \{0..n} = n * (n + 1)" +lemma "0 \ (n::nat) \ 2 * \{0..n} = n * (n + 1)" by (fact ex5 [transferred, transferred]) (* Using new transfer package *) -lemma "0 \ (n\nat) \ 2 * \{..n} = n * (n + 1)" +lemma "0 \ (n::nat) \ 2 * \{..n} = n * (n + 1)" by (fact ex5 [untransferred, Transfer.transferred]) end diff -r f6b0d827240e -r bdc1e2f0a86a src/HOL/ex/Transfer_Int_Nat.thy --- 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 "\xs\int list. \list_all (\x. x \ 0) xs; xs \ []\ \ + assumes "\xs::int list. \list_all (\x. x \ 0) xs; xs \ []\ \ listsum xs < listsum (map (\x. x + 1) xs)" shows "xs \ [] \ listsum xs < listsum (map Suc xs)" apply transfer diff -r f6b0d827240e -r bdc1e2f0a86a src/Tools/Code_Generator.thy --- 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\{})" +code_datatype "TYPE('a::{})" definition holds :: "prop" where "holds \ ((\x::prop. x) \ (\x. x))" @@ -59,4 +59,3 @@ hide_const (open) holds end -