src/HOL/Nitpick_Examples/Typedef_Nits.thy
author haftmann
Sat, 24 Dec 2011 15:53:10 +0100
changeset 45970 b6d0cff57d96
parent 45694 4a8743618257
child 46082 1c436a920730
permissions -rw-r--r--
adjusted to set/pred distinction by means of type constructor `set`
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Nitpick_Examples/Typedef_Nits.thy
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
45035
60d2c03d5c70 reintroduced Minipick as Nitpick example
blanchet
parents: 44328
diff changeset
     3
    Copyright   2009-2011
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     4
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     5
Examples featuring Nitpick applied to typedefs.
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     6
*)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     7
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     8
header {* Examples Featuring Nitpick Applied to Typedefs *}
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
     9
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    10
theory Typedef_Nits
35372
ca158c7b1144 renamed theory Rational to Rat
haftmann
parents: 35312
diff changeset
    11
imports Complex_Main
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    12
begin
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    13
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
    14
nitpick_params [verbose, card = 1\<emdash>4, sat_solver = MiniSat_JNI, max_threads = 1,
42208
02513eb26eb7 raised timeouts further, for SML/NJ -- because of variations in machines/compilers, fixed timeouts can merely prevent non-termination, not enforce particular performance characteristics.
krauss
parents: 42187
diff changeset
    15
                timeout = 240]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    16
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    17
definition "three = {0\<Colon>nat, 1, 2}"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    18
typedef (open) three = three
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    19
unfolding three_def by blast
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    20
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    21
definition A :: three where "A \<equiv> Abs_three 0"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    22
definition B :: three where "B \<equiv> Abs_three 1"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    23
definition C :: three where "C \<equiv> Abs_three 2"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    24
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    25
lemma "x = (y\<Colon>three)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    26
nitpick [expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    27
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    28
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    29
definition "one_or_two = {undefined False\<Colon>'a, undefined True}"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    30
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    31
typedef (open) 'a one_or_two = "one_or_two :: 'a set"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    32
unfolding one_or_two_def by auto
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    33
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    34
lemma "x = (y\<Colon>unit one_or_two)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    35
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    36
sorry
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    37
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    38
lemma "x = (y\<Colon>bool one_or_two)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    39
nitpick [expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    40
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    41
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    42
lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> x = (y\<Colon>bool one_or_two)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    43
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    44
sorry
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    45
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    46
lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> \<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    47
nitpick [card = 1, expect = potential] (* unfortunate *)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    48
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    49
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    50
lemma "\<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    51
nitpick [card = 1, expect = potential] (* unfortunate *)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    52
nitpick [card = 2, expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    53
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    54
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45694
diff changeset
    55
definition "bounded = {n\<Colon>nat. finite (UNIV \<Colon> 'a set) \<longrightarrow> n < card (UNIV \<Colon> 'a set)}"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45694
diff changeset
    56
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    57
typedef (open) 'a bounded = "bounded(TYPE('a))"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45035
diff changeset
    58
unfolding bounded_def
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    59
apply (rule_tac x = 0 in exI)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    60
apply (case_tac "card UNIV = 0")
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    61
by auto
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    62
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    63
lemma "x = (y\<Colon>unit bounded)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    64
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    65
sorry
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    66
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    67
lemma "x = (y\<Colon>bool bounded)"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    68
nitpick [expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    69
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    70
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    71
lemma "x \<noteq> (y\<Colon>bool bounded) \<Longrightarrow> z = x \<or> z = y"
39362
ee65900bfced adapt examples to latest Nitpick changes + speed them up a little bit
blanchet
parents: 38542
diff changeset
    72
nitpick [expect = potential] (* unfortunate *)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    73
sorry
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    74
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    75
lemma "x \<noteq> (y\<Colon>(bool \<times> bool) bounded) \<Longrightarrow> z = x \<or> z = y"
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
    76
nitpick [card = 1\<emdash>5, expect = genuine]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    77
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    78
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    79
lemma "True \<equiv> ((\<lambda>x\<Colon>bool. x) = (\<lambda>x. x))"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    80
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    81
by (rule True_def)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    82
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    83
lemma "False \<equiv> \<forall>P. P"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    84
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    85
by (rule False_def)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    86
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    87
lemma "() = Abs_unit True"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    88
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    89
by (rule Unity_def)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    90
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    91
lemma "() = Abs_unit False"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    92
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    93
by simp
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    94
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    95
lemma "Rep_unit () = True"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    96
nitpick [expect = none]
40590
b994d855dbd2 typedef (open) unit
huffman
parents: 40341
diff changeset
    97
by (insert Rep_unit) simp
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    98
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
    99
lemma "Rep_unit () = False"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   100
nitpick [expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   101
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   102
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   103
lemma "Pair a b = Abs_prod (Pair_Rep a b)"
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
   104
nitpick [card = 1\<emdash>2, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   105
by (rule Pair_def)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   106
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   107
lemma "Pair a b = Abs_prod (Pair_Rep b a)"
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
   108
nitpick [card = 1\<emdash>2, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   109
nitpick [dont_box, expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   110
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   111
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   112
lemma "fst (Abs_prod (Pair_Rep a b)) = a"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   113
nitpick [card = 2, expect = none]
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   114
by (simp add: Pair_def [THEN sym])
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   115
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   116
lemma "fst (Abs_prod (Pair_Rep a b)) = b"
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
   117
nitpick [card = 1\<emdash>2, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   118
nitpick [dont_box, expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   119
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   120
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   121
lemma "a \<noteq> a' \<Longrightarrow> Pair_Rep a b \<noteq> Pair_Rep a' b"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   122
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   123
apply (rule ccontr)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   124
apply simp
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   125
apply (drule subst [where P = "\<lambda>r. Abs_prod r = Abs_prod (Pair_Rep a b)"])
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   126
 apply (rule refl)
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   127
by (simp add: Pair_def [THEN sym])
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   128
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   129
lemma "Abs_prod (Rep_prod a) = a"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   130
nitpick [card = 2, expect = none]
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   131
by (rule Rep_prod_inverse)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   132
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   133
lemma "Inl \<equiv> \<lambda>a. Abs_sum (Inl_Rep a)"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   134
nitpick [card = 1, expect = none]
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   135
by (simp add: Inl_def o_def)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   136
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   137
lemma "Inl \<equiv> \<lambda>a. Abs_sum (Inr_Rep a)"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   138
nitpick [card = 1, card "'a + 'a" = 2, expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   139
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   140
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   141
lemma "Inl_Rep a \<noteq> Inr_Rep a"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   142
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   143
by (rule Inl_Rep_not_Inr_Rep)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   144
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   145
lemma "Abs_sum (Rep_sum a) = a"
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35078
diff changeset
   146
nitpick [card = 1, expect = none]
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35078
diff changeset
   147
nitpick [card = 2, expect = none]
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   148
by (rule Rep_sum_inverse)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   149
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   150
lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
35312
99cd1f96b400 improved precision of small sets in Nitpick
blanchet
parents: 35284
diff changeset
   151
nitpick [expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   152
by (rule Zero_nat_def_raw)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   153
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   154
lemma "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
35312
99cd1f96b400 improved precision of small sets in Nitpick
blanchet
parents: 35284
diff changeset
   155
nitpick [expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   156
by (rule Suc_def)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   157
37400
cf5e06d5ecaf adjust Nitpick example to follow latest wave of renamings
blanchet
parents: 37255
diff changeset
   158
lemma "Suc n = Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   159
nitpick [expect = genuine]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   160
oops
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   161
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   162
lemma "Abs_Nat (Rep_Nat a) = a"
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   163
nitpick [expect = none]
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   164
by (rule Rep_Nat_inverse)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   165
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   166
lemma "0 \<equiv> Abs_Integ (intrel `` {(0, 0)})"
44328
cbc6187710c9 deactivated »unknown« nitpick example
haftmann
parents: 42959
diff changeset
   167
(*nitpick [card = 1, unary_ints, max_potential = 0, expect = none] (?)*)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   168
by (rule Zero_int_def_raw)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   169
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   170
lemma "Abs_list (Rep_list a) = a"
42959
ee829022381d use \<emdash> rather than \<midarrow>
blanchet
parents: 42208
diff changeset
   171
nitpick [card = 1\<emdash>2, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   172
by (rule Rep_list_inverse)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   173
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   174
record point =
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   175
  Xcoord :: int
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   176
  Ycoord :: int
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   177
38542
c9599ce8cbfc adjusted to restored naming convention of logical record types
haftmann
parents: 37481
diff changeset
   178
lemma "Abs_point_ext (Rep_point_ext a) = a"
33571
3655e51f9958 minor cleanup in Nitpick
blanchet
parents: 33199
diff changeset
   179
nitpick [expect = none]
38542
c9599ce8cbfc adjusted to restored naming convention of logical record types
haftmann
parents: 37481
diff changeset
   180
by (fact Rep_point_ext_inverse)
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   181
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   182
lemma "Fract a b = of_int a / of_int b"
33737
e441fede163d fixed "expect" of Nitpick examples to reflect latest changes in Nitpick
blanchet
parents: 33571
diff changeset
   183
nitpick [card = 1, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   184
by (rule Fract_of_int_quotient)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   185
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   186
lemma "Abs_Rat (Rep_Rat a) = a"
33737
e441fede163d fixed "expect" of Nitpick examples to reflect latest changes in Nitpick
blanchet
parents: 33571
diff changeset
   187
nitpick [card = 1, expect = none]
33197
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   188
by (rule Rep_Rat_inverse)
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   189
de6285ebcc05 continuation of Nitpick's integration into Isabelle;
blanchet
parents:
diff changeset
   190
end