src/HOL/Nitpick_Examples/Datatype_Nits.thy
changeset 33197 de6285ebcc05
child 33199 6c9b2a94a69c
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy	Fri Oct 23 18:59:24 2009 +0200
     1.3 @@ -0,0 +1,104 @@
     1.4 +(*  Title:      HOL/Nitpick_Examples/Datatype_Nits.thy
     1.5 +    Author:     Jasmin Blanchette, TU Muenchen
     1.6 +    Copyright   2009
     1.7 +
     1.8 +Examples featuring Nitpick applied to datatypes.
     1.9 +*)
    1.10 +
    1.11 +header {* Examples Featuring Nitpick Applied to Datatypes *}
    1.12 +
    1.13 +theory Datatype_Nits
    1.14 +imports Main
    1.15 +begin
    1.16 +
    1.17 +primrec rot where
    1.18 +"rot Nibble0 = Nibble1" |
    1.19 +"rot Nibble1 = Nibble2" |
    1.20 +"rot Nibble2 = Nibble3" |
    1.21 +"rot Nibble3 = Nibble4" |
    1.22 +"rot Nibble4 = Nibble5" |
    1.23 +"rot Nibble5 = Nibble6" |
    1.24 +"rot Nibble6 = Nibble7" |
    1.25 +"rot Nibble7 = Nibble8" |
    1.26 +"rot Nibble8 = Nibble9" |
    1.27 +"rot Nibble9 = NibbleA" |
    1.28 +"rot NibbleA = NibbleB" |
    1.29 +"rot NibbleB = NibbleC" |
    1.30 +"rot NibbleC = NibbleD" |
    1.31 +"rot NibbleD = NibbleE" |
    1.32 +"rot NibbleE = NibbleF" |
    1.33 +"rot NibbleF = Nibble0"
    1.34 +
    1.35 +lemma "rot n \<noteq> n"
    1.36 +nitpick [card = 1\<midarrow>16, expect = none]
    1.37 +sorry
    1.38 +
    1.39 +lemma "rot Nibble2 \<noteq> Nibble3"
    1.40 +nitpick [card = 1, expect = none]
    1.41 +nitpick [card = 2, expect = genuine]
    1.42 +nitpick [card = 2, max Nibble2 = 0, expect = none]
    1.43 +nitpick [card = 2, max Nibble3 = 0, expect = none]
    1.44 +oops
    1.45 +
    1.46 +lemma "fun_pow 15 rot n \<noteq> n"
    1.47 +nitpick [card = 17, expect = none]
    1.48 +sorry
    1.49 +
    1.50 +lemma "fun_pow 15 rot n = n"
    1.51 +nitpick [card = 17, expect = genuine]
    1.52 +oops
    1.53 +
    1.54 +lemma "fun_pow 16 rot n = n"
    1.55 +nitpick [card = 17, expect = none]
    1.56 +oops
    1.57 +
    1.58 +datatype ('a, 'b) pd = Pd "'a \<times> 'b"
    1.59 +
    1.60 +fun fs where
    1.61 +"fs (Pd (a, _)) = a"
    1.62 +
    1.63 +fun sn where
    1.64 +"sn (Pd (_, b)) = b"
    1.65 +
    1.66 +lemma "fs (Pd p) = fst p"
    1.67 +nitpick [card = 20, expect = none]
    1.68 +sorry
    1.69 +
    1.70 +lemma "fs (Pd p) = snd p"
    1.71 +nitpick [expect = genuine]
    1.72 +oops
    1.73 +
    1.74 +lemma "sn (Pd p) = snd p"
    1.75 +nitpick [card = 20, expect = none]
    1.76 +sorry
    1.77 +
    1.78 +lemma "sn (Pd p) = fst p"
    1.79 +nitpick [expect = genuine]
    1.80 +oops
    1.81 +
    1.82 +lemma "fs (Pd ((a, b), (c, d))) = (a, b)"
    1.83 +nitpick [card = 1\<midarrow>12, expect = none]
    1.84 +sorry
    1.85 +
    1.86 +lemma "fs (Pd ((a, b), (c, d))) = (c, d)"
    1.87 +nitpick [expect = genuine]
    1.88 +oops
    1.89 +
    1.90 +datatype ('a, 'b) fn = Fn "'a \<Rightarrow> 'b"
    1.91 +
    1.92 +fun app where
    1.93 +"app (Fn f) x = f x"
    1.94 +
    1.95 +lemma "app (Fn g) y = g y"
    1.96 +nitpick [card = 1\<midarrow>12, expect = none]
    1.97 +sorry
    1.98 +
    1.99 +lemma "app (Fn g) y = g' y"
   1.100 +nitpick [expect = genuine]
   1.101 +oops
   1.102 +
   1.103 +lemma "app (Fn g) y = g y'"
   1.104 +nitpick [expect = genuine]
   1.105 +oops
   1.106 +
   1.107 +end