src/HOL/Quickcheck_Narrowing.thy
author bulwahn
Thu, 09 Jun 2011 08:32:21 +0200
changeset 44186 893de45ac28d
parent 44185 a9090cabca14
child 44187 3e274608f06b
permissions -rw-r--r--
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
     1 (* Author: Lukas Bulwahn, TU Muenchen *)
     2 
     3 header {* Counterexample generator preforming narrowing-based testing *}
     4 
     5 theory Quickcheck_Narrowing
     6 imports Quickcheck_Exhaustive
     7 uses
     8   ("~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs")
     9   ("~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs")
    10   ("~~/src/HOL/Tools/Quickcheck/narrowing_generators.ML")
    11 begin
    12 
    13 subsection {* Counterexample generator *}
    14 
    15 text {* We create a new target for the necessary code generation setup. *}
    16 
    17 setup {* Code_Target.extend_target ("Haskell_Quickcheck", (Code_Haskell.target, K I)) *}
    18 
    19 subsubsection {* Code generation setup *}
    20 
    21 code_type typerep
    22   (Haskell_Quickcheck "Typerep")
    23 
    24 code_const Typerep.Typerep
    25   (Haskell_Quickcheck "Typerep")
    26 
    27 code_reserved Haskell_Quickcheck Typerep
    28 
    29 code_type char
    30   (Haskell_Quickcheck "Char")
    31 
    32 setup {*
    33   fold String_Code.add_literal_char ["Haskell_Quickcheck"] 
    34   #> String_Code.add_literal_list_string "Haskell_Quickcheck"
    35 *}
    36 
    37 code_instance char :: equal
    38   (Haskell_Quickcheck -)
    39 
    40 code_const "HOL.equal \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
    41   (Haskell_Quickcheck infix 4 "==")
    42 
    43 subsubsection {* Type @{text "code_int"} for Haskell_Quickcheck's Int type *}
    44 
    45 typedef (open) code_int = "UNIV \<Colon> int set"
    46   morphisms int_of of_int by rule
    47 
    48 lemma of_int_int_of [simp]:
    49   "of_int (int_of k) = k"
    50   by (rule int_of_inverse)
    51 
    52 lemma int_of_of_int [simp]:
    53   "int_of (of_int n) = n"
    54   by (rule of_int_inverse) (rule UNIV_I)
    55 
    56 lemma code_int:
    57   "(\<And>n\<Colon>code_int. PROP P n) \<equiv> (\<And>n\<Colon>int. PROP P (of_int n))"
    58 proof
    59   fix n :: int
    60   assume "\<And>n\<Colon>code_int. PROP P n"
    61   then show "PROP P (of_int n)" .
    62 next
    63   fix n :: code_int
    64   assume "\<And>n\<Colon>int. PROP P (of_int n)"
    65   then have "PROP P (of_int (int_of n))" .
    66   then show "PROP P n" by simp
    67 qed
    68 
    69 
    70 lemma int_of_inject [simp]:
    71   "int_of k = int_of l \<longleftrightarrow> k = l"
    72   by (rule int_of_inject)
    73 
    74 lemma of_int_inject [simp]:
    75   "of_int n = of_int m \<longleftrightarrow> n = m"
    76   by (rule of_int_inject) (rule UNIV_I)+
    77 
    78 instantiation code_int :: equal
    79 begin
    80 
    81 definition
    82   "HOL.equal k l \<longleftrightarrow> HOL.equal (int_of k) (int_of l)"
    83 
    84 instance proof
    85 qed (auto simp add: equal_code_int_def equal_int_def eq_int_refl)
    86 
    87 end
    88 
    89 instantiation code_int :: number
    90 begin
    91 
    92 definition
    93   "number_of = of_int"
    94 
    95 instance ..
    96 
    97 end
    98 
    99 lemma int_of_number [simp]:
   100   "int_of (number_of k) = number_of k"
   101   by (simp add: number_of_code_int_def number_of_is_id)
   102 
   103 
   104 definition nat_of :: "code_int => nat"
   105 where
   106   "nat_of i = nat (int_of i)"
   107 
   108 
   109 code_datatype "number_of \<Colon> int \<Rightarrow> code_int"
   110   
   111   
   112 instantiation code_int :: "{minus, linordered_semidom, semiring_div, linorder}"
   113 begin
   114 
   115 definition [simp, code del]:
   116   "0 = of_int 0"
   117 
   118 definition [simp, code del]:
   119   "1 = of_int 1"
   120 
   121 definition [simp, code del]:
   122   "n + m = of_int (int_of n + int_of m)"
   123 
   124 definition [simp, code del]:
   125   "n - m = of_int (int_of n - int_of m)"
   126 
   127 definition [simp, code del]:
   128   "n * m = of_int (int_of n * int_of m)"
   129 
   130 definition [simp, code del]:
   131   "n div m = of_int (int_of n div int_of m)"
   132 
   133 definition [simp, code del]:
   134   "n mod m = of_int (int_of n mod int_of m)"
   135 
   136 definition [simp, code del]:
   137   "n \<le> m \<longleftrightarrow> int_of n \<le> int_of m"
   138 
   139 definition [simp, code del]:
   140   "n < m \<longleftrightarrow> int_of n < int_of m"
   141 
   142 
   143 instance proof
   144 qed (auto simp add: code_int left_distrib zmult_zless_mono2)
   145 
   146 end
   147 
   148 lemma zero_code_int_code [code, code_unfold]:
   149   "(0\<Colon>code_int) = Numeral0"
   150   by (simp add: number_of_code_int_def Pls_def)
   151 lemma [code_post]: "Numeral0 = (0\<Colon>code_int)"
   152   using zero_code_int_code ..
   153 
   154 lemma one_code_int_code [code, code_unfold]:
   155   "(1\<Colon>code_int) = Numeral1"
   156   by (simp add: number_of_code_int_def Pls_def Bit1_def)
   157 lemma [code_post]: "Numeral1 = (1\<Colon>code_int)"
   158   using one_code_int_code ..
   159 
   160 
   161 definition div_mod_code_int :: "code_int \<Rightarrow> code_int \<Rightarrow> code_int \<times> code_int" where
   162   [code del]: "div_mod_code_int n m = (n div m, n mod m)"
   163 
   164 lemma [code]:
   165   "div_mod_code_int n m = (if m = 0 then (0, n) else (n div m, n mod m))"
   166   unfolding div_mod_code_int_def by auto
   167 
   168 lemma [code]:
   169   "n div m = fst (div_mod_code_int n m)"
   170   unfolding div_mod_code_int_def by simp
   171 
   172 lemma [code]:
   173   "n mod m = snd (div_mod_code_int n m)"
   174   unfolding div_mod_code_int_def by simp
   175 
   176 lemma int_of_code [code]:
   177   "int_of k = (if k = 0 then 0
   178     else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
   179 proof -
   180   have 1: "(int_of k div 2) * 2 + int_of k mod 2 = int_of k" 
   181     by (rule mod_div_equality)
   182   have "int_of k mod 2 = 0 \<or> int_of k mod 2 = 1" by auto
   183   from this show ?thesis
   184     apply auto
   185     apply (insert 1) by (auto simp add: mult_ac)
   186 qed
   187 
   188 
   189 code_instance code_numeral :: equal
   190   (Haskell_Quickcheck -)
   191 
   192 setup {* fold (Numeral.add_code @{const_name number_code_int_inst.number_of_code_int}
   193   false Code_Printer.literal_numeral) ["Haskell_Quickcheck"]  *}
   194 
   195 code_const "0 \<Colon> code_int"
   196   (Haskell_Quickcheck "0")
   197 
   198 code_const "1 \<Colon> code_int"
   199   (Haskell_Quickcheck "1")
   200 
   201 code_const "minus \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> code_int"
   202   (Haskell_Quickcheck "(_/ -/ _)")
   203 
   204 code_const div_mod_code_int
   205   (Haskell_Quickcheck "divMod")
   206 
   207 code_const "HOL.equal \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
   208   (Haskell_Quickcheck infix 4 "==")
   209 
   210 code_const "op \<le> \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
   211   (Haskell_Quickcheck infix 4 "<=")
   212 
   213 code_const "op < \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"
   214   (Haskell_Quickcheck infix 4 "<")
   215 
   216 code_type code_int
   217   (Haskell_Quickcheck "Int")
   218 
   219 code_abort of_int
   220 
   221 subsubsection {* Narrowing's deep representation of types and terms *}
   222 
   223 datatype narrowing_type = SumOfProd "narrowing_type list list"
   224 
   225 datatype narrowing_term = Var "code_int list" narrowing_type | Ctr code_int "narrowing_term list"
   226 datatype 'a cons = C narrowing_type "(narrowing_term list => 'a) list"
   227 
   228 subsubsection {* From narrowing's deep representation of terms to Code_Evaluation's terms *}
   229 
   230 class partial_term_of = typerep +
   231   fixes partial_term_of :: "'a itself => narrowing_term => Code_Evaluation.term"
   232 
   233 lemma partial_term_of_anything: "partial_term_of x nt \<equiv> t"
   234   by (rule eq_reflection) (cases "partial_term_of x nt", cases t, simp)
   235 
   236 
   237 subsubsection {* Auxilary functions for Narrowing *}
   238 
   239 consts nth :: "'a list => code_int => 'a"
   240 
   241 code_const nth (Haskell_Quickcheck infixl 9  "!!")
   242 
   243 consts error :: "char list => 'a"
   244 
   245 code_const error (Haskell_Quickcheck "error")
   246 
   247 consts toEnum :: "code_int => char"
   248 
   249 code_const toEnum (Haskell_Quickcheck "toEnum")
   250 
   251 consts map_index :: "(code_int * 'a => 'b) => 'a list => 'b list"  
   252 
   253 consts split_At :: "code_int => 'a list => 'a list * 'a list"
   254  
   255 subsubsection {* Narrowing's basic operations *}
   256 
   257 type_synonym 'a narrowing = "code_int => 'a cons"
   258 
   259 definition empty :: "'a narrowing"
   260 where
   261   "empty d = C (SumOfProd []) []"
   262   
   263 definition cons :: "'a => 'a narrowing"
   264 where
   265   "cons a d = (C (SumOfProd [[]]) [(%_. a)])"
   266 
   267 fun conv :: "(narrowing_term list => 'a) list => narrowing_term => 'a"
   268 where
   269   "conv cs (Var p _) = error (Char Nibble0 Nibble0 # map toEnum p)"
   270 | "conv cs (Ctr i xs) = (nth cs i) xs"
   271 
   272 fun nonEmpty :: "narrowing_type => bool"
   273 where
   274   "nonEmpty (SumOfProd ps) = (\<not> (List.null ps))"
   275 
   276 definition "apply" :: "('a => 'b) narrowing => 'a narrowing => 'b narrowing"
   277 where
   278   "apply f a d =
   279      (case f d of C (SumOfProd ps) cfs =>
   280        case a (d - 1) of C ta cas =>
   281        let
   282          shallow = (d > 0 \<and> nonEmpty ta);
   283          cs = [(%xs'. (case xs' of [] => undefined | x # xs => cf xs (conv cas x))). shallow, cf <- cfs]
   284        in C (SumOfProd [ta # p. shallow, p <- ps]) cs)"
   285 
   286 definition sum :: "'a narrowing => 'a narrowing => 'a narrowing"
   287 where
   288   "sum a b d =
   289     (case a d of C (SumOfProd ssa) ca => 
   290       case b d of C (SumOfProd ssb) cb =>
   291       C (SumOfProd (ssa @ ssb)) (ca @ cb))"
   292 
   293 lemma [fundef_cong]:
   294   assumes "a d = a' d" "b d = b' d" "d = d'"
   295   shows "sum a b d = sum a' b' d'"
   296 using assms unfolding sum_def by (auto split: cons.split narrowing_type.split)
   297 
   298 lemma [fundef_cong]:
   299   assumes "f d = f' d" "(\<And>d'. 0 <= d' & d' < d ==> a d' = a' d')"
   300   assumes "d = d'"
   301   shows "apply f a d = apply f' a' d'"
   302 proof -
   303   note assms moreover
   304   have "int_of (of_int 0) < int_of d' ==> int_of (of_int 0) <= int_of (of_int (int_of d' - int_of (of_int 1)))"
   305     by (simp add: of_int_inverse)
   306   moreover
   307   have "int_of (of_int (int_of d' - int_of (of_int 1))) < int_of d'"
   308     by (simp add: of_int_inverse)
   309   ultimately show ?thesis
   310     unfolding apply_def by (auto split: cons.split narrowing_type.split simp add: Let_def)
   311 qed
   312 
   313 subsubsection {* Narrowing generator type class *}
   314 
   315 class narrowing =
   316   fixes narrowing :: "code_int => 'a cons"
   317 
   318 definition drawn_from :: "'a list => 'a cons"
   319 where "drawn_from xs = C (SumOfProd (map (%_. []) xs)) (map (%x y. x) xs)"
   320 
   321 instantiation int :: narrowing
   322 begin
   323 
   324 definition
   325   "narrowing_int d = (let i = Quickcheck_Narrowing.int_of d in drawn_from [-i .. i])"
   326 
   327 instance ..
   328 
   329 end
   330 
   331 datatype property = Universal narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Existential narrowing_type "(narrowing_term => property)" "narrowing_term => Code_Evaluation.term" | Property bool
   332 
   333 (* FIXME: hard-wired maximal depth of 100 here *)
   334 definition exists :: "('a :: {narrowing, partial_term_of} => property) => property"
   335 where
   336   "exists f = (case narrowing (100 :: code_int) of C ty cs => Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
   337 
   338 definition "all" :: "('a :: {narrowing, partial_term_of} => property) => property"
   339 where
   340   "all f = (case narrowing (100 :: code_int) of C ty cs => Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
   341 
   342 subsubsection {* class @{text is_testable} *}
   343 
   344 text {* The class @{text is_testable} ensures that all necessary type instances are generated. *}
   345 
   346 class is_testable
   347 
   348 instance bool :: is_testable ..
   349 
   350 instance "fun" :: ("{term_of, narrowing, partial_term_of}", is_testable) is_testable ..
   351 
   352 definition ensure_testable :: "'a :: is_testable => 'a :: is_testable"
   353 where
   354   "ensure_testable f = f"
   355 
   356 
   357 subsubsection {* Defining a simple datatype to represent functions in an incomplete and redundant way *}
   358 
   359 datatype ('a, 'b) ffun = Constant 'b | Update 'a 'b "('a, 'b) ffun"
   360 
   361 primrec eval_ffun :: "('a, 'b) ffun => 'a => 'b"
   362 where
   363   "eval_ffun (Constant c) x = c"
   364 | "eval_ffun (Update x' y f) x = (if x = x' then y else eval_ffun f x)"
   365 
   366 hide_type (open) ffun
   367 hide_const (open) Constant Update eval_ffun
   368 
   369 datatype 'b cfun = Constant 'b
   370 
   371 primrec eval_cfun :: "'b cfun => 'a => 'b"
   372 where
   373   "eval_cfun (Constant c) y = c"
   374 
   375 hide_type (open) cfun
   376 hide_const (open) Constant eval_cfun
   377 
   378 subsubsection {* Setting up the counterexample generator *}
   379 
   380 setup {* Thy_Load.provide_file (Path.explode ("~~/src/HOL/Tools/Quickcheck/PNF_Narrowing_Engine.hs")) *}
   381 setup {* Thy_Load.provide_file (Path.explode ("~~/src/HOL/Tools/Quickcheck/Narrowing_Engine.hs")) *}
   382 use "~~/src/HOL/Tools/Quickcheck/narrowing_generators.ML"
   383 
   384 setup {* Narrowing_Generators.setup *}
   385 
   386 hide_type code_int narrowing_type narrowing_term cons property
   387 hide_const int_of of_int nth error toEnum map_index split_At empty
   388   C cons conv nonEmpty "apply" sum ensure_testable all exists 
   389 hide_fact empty_def cons_def conv.simps nonEmpty.simps apply_def sum_def ensure_testable_def all_def exists_def
   390 
   391 end