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