src/HOL/Nitpick_Examples/Typedef_Nits.thy
author blanchet
Tue, 09 Feb 2010 16:07:51 +0100
changeset 35075 6fd1052fe463
parent 35073 cc19e2aef17e
child 35284 9edc2bd6d2bd
permissions -rw-r--r--
optimization to quantifiers in Nitpick's handling of simp rules + renamed some SAT solvers
     1 (*  Title:      HOL/Nitpick_Examples/Typedef_Nits.thy
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2009, 2010
     4 
     5 Examples featuring Nitpick applied to typedefs.
     6 *)
     7 
     8 header {* Examples Featuring Nitpick Applied to Typedefs *}
     9 
    10 theory Typedef_Nits
    11 imports Main Rational
    12 begin
    13 
    14 nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
    15                 card = 1\<midarrow>4]
    16 
    17 typedef three = "{0\<Colon>nat, 1, 2}"
    18 by blast
    19 
    20 definition A :: three where "A \<equiv> Abs_three 0"
    21 definition B :: three where "B \<equiv> Abs_three 1"
    22 definition C :: three where "C \<equiv> Abs_three 2"
    23 
    24 lemma "x = (y\<Colon>three)"
    25 nitpick [expect = genuine]
    26 oops
    27 
    28 typedef 'a one_or_two = "{undefined False\<Colon>'a, undefined True}"
    29 by auto
    30 
    31 lemma "x = (y\<Colon>unit one_or_two)"
    32 nitpick [expect = none]
    33 sorry
    34 
    35 lemma "x = (y\<Colon>bool one_or_two)"
    36 nitpick [expect = genuine]
    37 oops
    38 
    39 lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> x = (y\<Colon>bool one_or_two)"
    40 nitpick [expect = none]
    41 sorry
    42 
    43 lemma "undefined False \<longleftrightarrow> undefined True \<Longrightarrow> \<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
    44 nitpick [card = 1, expect = potential] (* unfortunate *)
    45 oops
    46 
    47 lemma "\<exists>x (y\<Colon>bool one_or_two). x \<noteq> y"
    48 nitpick [card = 1, expect = potential] (* unfortunate *)
    49 nitpick [card = 2, expect = none]
    50 oops
    51 
    52 typedef 'a bounded =
    53         "{n\<Colon>nat. finite (UNIV\<Colon>'a \<Rightarrow> bool) \<longrightarrow> n < card (UNIV\<Colon>'a \<Rightarrow> bool)}"
    54 apply (rule_tac x = 0 in exI)
    55 apply (case_tac "card UNIV = 0")
    56 by auto
    57 
    58 lemma "x = (y\<Colon>unit bounded)"
    59 nitpick [expect = none]
    60 sorry
    61 
    62 lemma "x = (y\<Colon>bool bounded)"
    63 nitpick [expect = genuine]
    64 oops
    65 
    66 lemma "x \<noteq> (y\<Colon>bool bounded) \<Longrightarrow> z = x \<or> z = y"
    67 nitpick [expect = none]
    68 sorry
    69 
    70 lemma "x \<noteq> (y\<Colon>(bool \<times> bool) bounded) \<Longrightarrow> z = x \<or> z = y"
    71 nitpick [card = 1\<midarrow>5, expect = genuine]
    72 oops
    73 
    74 lemma "True \<equiv> ((\<lambda>x\<Colon>bool. x) = (\<lambda>x. x))"
    75 nitpick [expect = none]
    76 by (rule True_def)
    77 
    78 lemma "False \<equiv> \<forall>P. P"
    79 nitpick [expect = none]
    80 by (rule False_def)
    81 
    82 lemma "() = Abs_unit True"
    83 nitpick [expect = none]
    84 by (rule Unity_def)
    85 
    86 lemma "() = Abs_unit False"
    87 nitpick [expect = none]
    88 by simp
    89 
    90 lemma "Rep_unit () = True"
    91 nitpick [expect = none]
    92 by (insert Rep_unit) (simp add: unit_def)
    93 
    94 lemma "Rep_unit () = False"
    95 nitpick [expect = genuine]
    96 oops
    97 
    98 lemma "Pair a b \<equiv> Abs_Prod (Pair_Rep a b)"
    99 nitpick [card = 1\<midarrow>2, expect = none]
   100 by (rule Pair_def)
   101 
   102 lemma "Pair a b \<equiv> Abs_Prod (Pair_Rep b a)"
   103 nitpick [card = 1\<midarrow>2, expect = none]
   104 nitpick [dont_box, expect = genuine]
   105 oops
   106 
   107 lemma "fst (Abs_Prod (Pair_Rep a b)) = a"
   108 nitpick [card = 2, expect = none]
   109 by (simp add: Pair_def [THEN symmetric])
   110 
   111 lemma "fst (Abs_Prod (Pair_Rep a b)) = b"
   112 nitpick [card = 1\<midarrow>2, expect = none]
   113 nitpick [dont_box, expect = genuine]
   114 oops
   115 
   116 lemma "a \<noteq> a' \<Longrightarrow> Pair_Rep a b \<noteq> Pair_Rep a' b"
   117 nitpick [expect = none]
   118 apply (rule ccontr)
   119 apply simp
   120 apply (drule subst [where P = "\<lambda>r. Abs_Prod r = Abs_Prod (Pair_Rep a b)"])
   121  apply (rule refl)
   122 by (simp add: Pair_def [THEN symmetric])
   123 
   124 lemma "Abs_Prod (Rep_Prod a) = a"
   125 nitpick [card = 2, expect = none]
   126 by (rule Rep_Prod_inverse)
   127 
   128 lemma "Inl \<equiv> \<lambda>a. Abs_Sum (Inl_Rep a)"
   129 nitpick [card = 1, expect = none]
   130 by (simp only: Inl_def o_def)
   131 
   132 lemma "Inl \<equiv> \<lambda>a. Abs_Sum (Inr_Rep a)"
   133 nitpick [card = 1, card "'a + 'a" = 2, expect = genuine]
   134 oops
   135 
   136 lemma "Inl_Rep a \<noteq> Inr_Rep a"
   137 nitpick [expect = none]
   138 by (rule Inl_Rep_not_Inr_Rep)
   139 
   140 lemma "Abs_Sum (Rep_Sum a) = a"
   141 nitpick [card = 1\<midarrow>2, timeout = 60 s, expect = none]
   142 by (rule Rep_Sum_inverse)
   143 
   144 lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
   145 (* nitpick [expect = none] FIXME *)
   146 by (rule Zero_nat_def_raw)
   147 
   148 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Rep_Nat n))"
   149 (* nitpick [expect = none] FIXME *)
   150 by (rule Suc_def)
   151 
   152 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
   153 nitpick [expect = genuine]
   154 oops
   155 
   156 lemma "Abs_Nat (Rep_Nat a) = a"
   157 nitpick [expect = none]
   158 by (rule Rep_Nat_inverse)
   159 
   160 lemma "0 \<equiv> Abs_Integ (intrel `` {(0, 0)})"
   161 nitpick [card = 1, unary_ints, max_potential = 0, expect = none]
   162 by (rule Zero_int_def_raw)
   163 
   164 lemma "Abs_Integ (Rep_Integ a) = a"
   165 nitpick [card = 1, unary_ints, max_potential = 0, expect = none]
   166 by (rule Rep_Integ_inverse)
   167 
   168 lemma "Abs_list (Rep_list a) = a"
   169 nitpick [card = 1\<midarrow>2, expect = none]
   170 by (rule Rep_list_inverse)
   171 
   172 record point =
   173   Xcoord :: int
   174   Ycoord :: int
   175 
   176 lemma "Abs_point_ext_type (Rep_point_ext_type a) = a"
   177 nitpick [expect = none]
   178 by (rule Rep_point_ext_type_inverse)
   179 
   180 lemma "Fract a b = of_int a / of_int b"
   181 nitpick [card = 1, expect = none]
   182 by (rule Fract_of_int_quotient)
   183 
   184 lemma "Abs_Rat (Rep_Rat a) = a"
   185 nitpick [card = 1, expect = none]
   186 by (rule Rep_Rat_inverse)
   187 
   188 end