| author | wenzelm | 
| Fri, 08 Jul 2011 14:37:19 +0200 | |
| changeset 43702 | 24fb44c1086a | 
| parent 43562 | 2c55eac2e5a9 | 
| child 43887 | 442aceb54969 | 
| permissions | -rw-r--r-- | 
| 41905 | 1  | 
(* Author: Lukas Bulwahn, TU Muenchen *)  | 
2  | 
||
| 43356 | 3  | 
header {* Counterexample generator performing narrowing-based testing *}
 | 
| 41905 | 4  | 
|
| 
41930
 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 
bulwahn 
parents: 
41929 
diff
changeset
 | 
5  | 
theory Quickcheck_Narrowing  | 
| 
43312
 
7a31f9064f99
adapting Quickcheck_Narrowing: adding setup for characters; correcting import statement
 
bulwahn 
parents: 
43309 
diff
changeset
 | 
6  | 
imports Quickcheck_Exhaustive  | 
| 
41962
 
27a61a3266d8
correcting names in Narrowing_Engine and example theory for Quickcheck_Narrowing
 
bulwahn 
parents: 
41961 
diff
changeset
 | 
7  | 
uses  | 
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43562 
diff
changeset
 | 
8  | 
  ("Tools/Quickcheck/PNF_Narrowing_Engine.hs")
 | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43562 
diff
changeset
 | 
9  | 
  ("Tools/Quickcheck/Narrowing_Engine.hs")
 | 
| 
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43562 
diff
changeset
 | 
10  | 
  ("Tools/Quickcheck/narrowing_generators.ML")
 | 
| 41905 | 11  | 
begin  | 
12  | 
||
13  | 
subsection {* Counterexample generator *}
 | 
|
14  | 
||
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
15  | 
text {* We create a new target for the necessary code generation setup. *}
 | 
| 
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
16  | 
|
| 
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
17  | 
setup {* Code_Target.extend_target ("Haskell_Quickcheck", (Code_Haskell.target, K I)) *}
 | 
| 
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
18  | 
|
| 
41909
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
19  | 
subsubsection {* Code generation setup *}
 | 
| 
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
20  | 
|
| 
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
21  | 
code_type typerep  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
22  | 
(Haskell_Quickcheck "Typerep")  | 
| 
41909
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
23  | 
|
| 
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
24  | 
code_const Typerep.Typerep  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
25  | 
(Haskell_Quickcheck "Typerep")  | 
| 
41909
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
26  | 
|
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
27  | 
code_reserved Haskell_Quickcheck Typerep  | 
| 
41909
 
383bbdad1650
replacing strings in generated Code resolves the changing names of Typerep in lazysmallcheck prototype
 
bulwahn 
parents: 
41908 
diff
changeset
 | 
28  | 
|
| 43341 | 29  | 
subsubsection {* Type @{text "code_int"} for Haskell Quickcheck's Int type *}
 | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
30  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
31  | 
typedef (open) code_int = "UNIV \<Colon> int set"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
32  | 
morphisms int_of of_int by rule  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
33  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
34  | 
lemma of_int_int_of [simp]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
35  | 
"of_int (int_of k) = k"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
36  | 
by (rule int_of_inverse)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
37  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
38  | 
lemma int_of_of_int [simp]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
39  | 
"int_of (of_int n) = n"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
40  | 
by (rule of_int_inverse) (rule UNIV_I)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
41  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
42  | 
lemma code_int:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
43  | 
"(\<And>n\<Colon>code_int. PROP P n) \<equiv> (\<And>n\<Colon>int. PROP P (of_int n))"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
44  | 
proof  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
45  | 
fix n :: int  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
46  | 
assume "\<And>n\<Colon>code_int. PROP P n"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
47  | 
then show "PROP P (of_int n)" .  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
48  | 
next  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
49  | 
fix n :: code_int  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
50  | 
assume "\<And>n\<Colon>int. PROP P (of_int n)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
51  | 
then have "PROP P (of_int (int_of n))" .  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
52  | 
then show "PROP P n" by simp  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
53  | 
qed  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
54  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
55  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
56  | 
lemma int_of_inject [simp]:  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
57  | 
"int_of k = int_of l \<longleftrightarrow> k = l"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
58  | 
by (rule int_of_inject)  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
59  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
60  | 
lemma of_int_inject [simp]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
61  | 
"of_int n = of_int m \<longleftrightarrow> n = m"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
62  | 
by (rule of_int_inject) (rule UNIV_I)+  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
63  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
64  | 
instantiation code_int :: equal  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
65  | 
begin  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
66  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
67  | 
definition  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
68  | 
"HOL.equal k l \<longleftrightarrow> HOL.equal (int_of k) (int_of l)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
69  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
70  | 
instance proof  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
71  | 
qed (auto simp add: equal_code_int_def equal_int_def eq_int_refl)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
72  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
73  | 
end  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
74  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
75  | 
instantiation code_int :: number  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
76  | 
begin  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
77  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
78  | 
definition  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
79  | 
"number_of = of_int"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
80  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
81  | 
instance ..  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
82  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
83  | 
end  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
84  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
85  | 
lemma int_of_number [simp]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
86  | 
"int_of (number_of k) = number_of k"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
87  | 
by (simp add: number_of_code_int_def number_of_is_id)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
88  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
89  | 
|
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
90  | 
definition nat_of :: "code_int => nat"  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
91  | 
where  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
92  | 
"nat_of i = nat (int_of i)"  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
93  | 
|
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
94  | 
|
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
95  | 
code_datatype "number_of \<Colon> int \<Rightarrow> code_int"  | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
96  | 
|
| 
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
97  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
98  | 
instantiation code_int :: "{minus, linordered_semidom, semiring_div, linorder}"
 | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
99  | 
begin  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
100  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
101  | 
definition [simp, code del]:  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
102  | 
"0 = of_int 0"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
103  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
104  | 
definition [simp, code del]:  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
105  | 
"1 = of_int 1"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
106  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
107  | 
definition [simp, code del]:  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
108  | 
"n + m = of_int (int_of n + int_of m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
109  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
110  | 
definition [simp, code del]:  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
111  | 
"n - m = of_int (int_of n - int_of m)"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
112  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
113  | 
definition [simp, code del]:  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
114  | 
"n * m = of_int (int_of n * int_of m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
115  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
116  | 
definition [simp, code del]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
117  | 
"n div m = of_int (int_of n div int_of m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
118  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
119  | 
definition [simp, code del]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
120  | 
"n mod m = of_int (int_of n mod int_of m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
121  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
122  | 
definition [simp, code del]:  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
123  | 
"n \<le> m \<longleftrightarrow> int_of n \<le> int_of m"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
124  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
125  | 
definition [simp, code del]:  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
126  | 
"n < m \<longleftrightarrow> int_of n < int_of m"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
127  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
128  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
129  | 
instance proof  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
130  | 
qed (auto simp add: code_int left_distrib zmult_zless_mono2)  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
131  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
132  | 
end  | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
133  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
134  | 
lemma zero_code_int_code [code, code_unfold]:  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
135  | 
"(0\<Colon>code_int) = Numeral0"  | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
136  | 
by (simp add: number_of_code_int_def Pls_def)  | 
| 
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
137  | 
lemma [code_post]: "Numeral0 = (0\<Colon>code_int)"  | 
| 
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
138  | 
using zero_code_int_code ..  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
139  | 
|
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
140  | 
lemma one_code_int_code [code, code_unfold]:  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
141  | 
"(1\<Colon>code_int) = Numeral1"  | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
142  | 
by (simp add: number_of_code_int_def Pls_def Bit1_def)  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
143  | 
lemma [code_post]: "Numeral1 = (1\<Colon>code_int)"  | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
144  | 
using one_code_int_code ..  | 
| 
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
145  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
146  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
147  | 
definition div_mod_code_int :: "code_int \<Rightarrow> code_int \<Rightarrow> code_int \<times> code_int" where  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
148  | 
[code del]: "div_mod_code_int n m = (n div m, n mod m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
149  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
150  | 
lemma [code]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
151  | 
"div_mod_code_int n m = (if m = 0 then (0, n) else (n div m, n mod m))"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
152  | 
unfolding div_mod_code_int_def by auto  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
153  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
154  | 
lemma [code]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
155  | 
"n div m = fst (div_mod_code_int n m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
156  | 
unfolding div_mod_code_int_def by simp  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
157  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
158  | 
lemma [code]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
159  | 
"n mod m = snd (div_mod_code_int n m)"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
160  | 
unfolding div_mod_code_int_def by simp  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
161  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
162  | 
lemma int_of_code [code]:  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
163  | 
"int_of k = (if k = 0 then 0  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
164  | 
else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
165  | 
proof -  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
166  | 
have 1: "(int_of k div 2) * 2 + int_of k mod 2 = int_of k"  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
167  | 
by (rule mod_div_equality)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
168  | 
have "int_of k mod 2 = 0 \<or> int_of k mod 2 = 1" by auto  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
169  | 
from this show ?thesis  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
170  | 
apply auto  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
171  | 
apply (insert 1) by (auto simp add: mult_ac)  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
172  | 
qed  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
173  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
174  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
175  | 
code_instance code_numeral :: equal  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
176  | 
(Haskell_Quickcheck -)  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
177  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
178  | 
setup {* fold (Numeral.add_code @{const_name number_code_int_inst.number_of_code_int}
 | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
179  | 
false Code_Printer.literal_numeral) ["Haskell_Quickcheck"] *}  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
180  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
181  | 
code_const "0 \<Colon> code_int"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
182  | 
(Haskell_Quickcheck "0")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
183  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
184  | 
code_const "1 \<Colon> code_int"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
185  | 
(Haskell_Quickcheck "1")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
186  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
187  | 
code_const "minus \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> code_int"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
188  | 
(Haskell_Quickcheck "(_/ -/ _)")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
189  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
190  | 
code_const div_mod_code_int  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
191  | 
(Haskell_Quickcheck "divMod")  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
192  | 
|
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
193  | 
code_const "HOL.equal \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
194  | 
(Haskell_Quickcheck infix 4 "==")  | 
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
195  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
196  | 
code_const "op \<le> \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
197  | 
(Haskell_Quickcheck infix 4 "<=")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
198  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
199  | 
code_const "op < \<Colon> code_int \<Rightarrow> code_int \<Rightarrow> bool"  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
200  | 
(Haskell_Quickcheck infix 4 "<")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
201  | 
|
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
202  | 
code_type code_int  | 
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
203  | 
(Haskell_Quickcheck "Int")  | 
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
204  | 
|
| 
42021
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
205  | 
code_abort of_int  | 
| 
 
52551c0a3374
extending code_int type more; adding narrowing instance for type int; added test case for int instance
 
bulwahn 
parents: 
41965 
diff
changeset
 | 
206  | 
|
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
207  | 
subsubsection {* Narrowing's deep representation of types and terms *}
 | 
| 41905 | 208  | 
|
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
209  | 
datatype narrowing_type = SumOfProd "narrowing_type list list"  | 
| 
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
210  | 
datatype narrowing_term = Var "code_int list" narrowing_type | Ctr code_int "narrowing_term list"  | 
| 
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
211  | 
datatype 'a cons = C narrowing_type "(narrowing_term list => 'a) list"  | 
| 41905 | 212  | 
|
| 43356 | 213  | 
primrec map_cons :: "('a => 'b) => 'a cons => 'b cons"
 | 
214  | 
where  | 
|
215  | 
"map_cons f (C ty cs) = C ty (map (%c. f o c) cs)"  | 
|
216  | 
||
| 43341 | 217  | 
subsubsection {* From narrowing's deep representation of terms to @{theory Code_Evaluation}'s terms *}
 | 
| 
42980
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
218  | 
|
| 
 
859fe9cc0838
improving code_int setup in Quickcheck_Narrowing; adding partial_term_of class in Quickcheck_Narrowing
 
bulwahn 
parents: 
42024 
diff
changeset
 | 
219  | 
class partial_term_of = typerep +  | 
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
220  | 
fixes partial_term_of :: "'a itself => narrowing_term => Code_Evaluation.term"  | 
| 
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
221  | 
|
| 
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
222  | 
lemma partial_term_of_anything: "partial_term_of x nt \<equiv> t"  | 
| 
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
223  | 
by (rule eq_reflection) (cases "partial_term_of x nt", cases t, simp)  | 
| 43356 | 224  | 
|
| 
41964
 
13904699c859
tuned subsubsection names in Quickcheck_Narrowing
 
bulwahn 
parents: 
41962 
diff
changeset
 | 
225  | 
subsubsection {* Auxilary functions for Narrowing *}
 | 
| 41905 | 226  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
227  | 
consts nth :: "'a list => code_int => 'a"  | 
| 41905 | 228  | 
|
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
229  | 
code_const nth (Haskell_Quickcheck infixl 9 "!!")  | 
| 41905 | 230  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
231  | 
consts error :: "char list => 'a"  | 
| 41905 | 232  | 
|
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
233  | 
code_const error (Haskell_Quickcheck "error")  | 
| 41905 | 234  | 
|
| 
41908
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
235  | 
consts toEnum :: "code_int => char"  | 
| 
 
3bd9a21366d2
changing invocation of ghc from interactive mode to compilation increases the performance of lazysmallcheck by a factor of twenty; changing Integer type to Int reduces by another 50 percent
 
bulwahn 
parents: 
41905 
diff
changeset
 | 
236  | 
|
| 
43308
 
fd6cc1378fec
compilation of Haskell in its own target for Quickcheck; passing options by arguments in Narrowing_Generators
 
bulwahn 
parents: 
43237 
diff
changeset
 | 
237  | 
code_const toEnum (Haskell_Quickcheck "toEnum")  | 
| 41905 | 238  | 
|
| 43316 | 239  | 
consts marker :: "char"  | 
| 41905 | 240  | 
|
| 43316 | 241  | 
code_const marker (Haskell_Quickcheck "''\\0'")  | 
242  | 
||
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
243  | 
subsubsection {* Narrowing's basic operations *}
 | 
| 41905 | 244  | 
|
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
245  | 
type_synonym 'a narrowing = "code_int => 'a cons"  | 
| 41905 | 246  | 
|
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
247  | 
definition empty :: "'a narrowing"  | 
| 41905 | 248  | 
where  | 
249  | 
"empty d = C (SumOfProd []) []"  | 
|
250  | 
||
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
251  | 
definition cons :: "'a => 'a narrowing"  | 
| 41905 | 252  | 
where  | 
253  | 
"cons a d = (C (SumOfProd [[]]) [(%_. a)])"  | 
|
254  | 
||
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
255  | 
fun conv :: "(narrowing_term list => 'a) list => narrowing_term => 'a"  | 
| 41905 | 256  | 
where  | 
| 43316 | 257  | 
"conv cs (Var p _) = error (marker # map toEnum p)"  | 
| 41905 | 258  | 
| "conv cs (Ctr i xs) = (nth cs i) xs"  | 
259  | 
||
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
260  | 
fun nonEmpty :: "narrowing_type => bool"  | 
| 41905 | 261  | 
where  | 
262  | 
"nonEmpty (SumOfProd ps) = (\<not> (List.null ps))"  | 
|
263  | 
||
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
264  | 
definition "apply" :: "('a => 'b) narrowing => 'a narrowing => 'b narrowing"
 | 
| 41905 | 265  | 
where  | 
266  | 
"apply f a d =  | 
|
267  | 
(case f d of C (SumOfProd ps) cfs =>  | 
|
268  | 
case a (d - 1) of C ta cas =>  | 
|
269  | 
let  | 
|
270  | 
shallow = (d > 0 \<and> nonEmpty ta);  | 
|
271  | 
cs = [(%xs'. (case xs' of [] => undefined | x # xs => cf xs (conv cas x))). shallow, cf <- cfs]  | 
|
272  | 
in C (SumOfProd [ta # p. shallow, p <- ps]) cs)"  | 
|
273  | 
||
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
274  | 
definition sum :: "'a narrowing => 'a narrowing => 'a narrowing"  | 
| 41905 | 275  | 
where  | 
276  | 
"sum a b d =  | 
|
277  | 
(case a d of C (SumOfProd ssa) ca =>  | 
|
278  | 
case b d of C (SumOfProd ssb) cb =>  | 
|
279  | 
C (SumOfProd (ssa @ ssb)) (ca @ cb))"  | 
|
280  | 
||
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
281  | 
lemma [fundef_cong]:  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
282  | 
assumes "a d = a' d" "b d = b' d" "d = d'"  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
283  | 
shows "sum a b d = sum a' b' d'"  | 
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
284  | 
using assms unfolding sum_def by (auto split: cons.split narrowing_type.split)  | 
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
285  | 
|
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
286  | 
lemma [fundef_cong]:  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
287  | 
assumes "f d = f' d" "(\<And>d'. 0 <= d' & d' < d ==> a d' = a' d')"  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
288  | 
assumes "d = d'"  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
289  | 
shows "apply f a d = apply f' a' d'"  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
290  | 
proof -  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
291  | 
note assms moreover  | 
| 
41930
 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 
bulwahn 
parents: 
41929 
diff
changeset
 | 
292  | 
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)))"  | 
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
293  | 
by (simp add: of_int_inverse)  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
294  | 
moreover  | 
| 
41930
 
1e008cc4883a
renaming lazysmallcheck ML file to Quickcheck_Narrowing
 
bulwahn 
parents: 
41929 
diff
changeset
 | 
295  | 
have "int_of (of_int (int_of d' - int_of (of_int 1))) < int_of d'"  | 
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
296  | 
by (simp add: of_int_inverse)  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
297  | 
ultimately show ?thesis  | 
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
298  | 
unfolding apply_def by (auto split: cons.split narrowing_type.split simp add: Let_def)  | 
| 
41912
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
299  | 
qed  | 
| 
 
1848775589e5
adding termination proofs to series functions in LSC; commenting out momentarily unused term refinement functions in LSC
 
bulwahn 
parents: 
41910 
diff
changeset
 | 
300  | 
|
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
301  | 
subsubsection {* Narrowing generator type class *}
 | 
| 41905 | 302  | 
|
| 
41961
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
303  | 
class narrowing =  | 
| 
 
fdd37cfcd4a3
renaming series and serial to narrowing in Quickcheck_Narrowing
 
bulwahn 
parents: 
41943 
diff
changeset
 | 
304  | 
fixes narrowing :: "code_int => 'a cons"  | 
| 41905 | 305  | 
|
| 
43237
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
306  | 
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  | 
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
307  | 
|
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
308  | 
(* FIXME: hard-wired maximal depth of 100 here *)  | 
| 
43315
 
893de45ac28d
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
 
bulwahn 
parents: 
43314 
diff
changeset
 | 
309  | 
definition exists :: "('a :: {narrowing, partial_term_of} => property) => property"
 | 
| 
43237
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
310  | 
where  | 
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
311  | 
  "exists f = (case narrowing (100 :: code_int) of C ty cs => Existential ty (\<lambda> t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 | 
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
312  | 
|
| 
43315
 
893de45ac28d
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
 
bulwahn 
parents: 
43314 
diff
changeset
 | 
313  | 
definition "all" :: "('a :: {narrowing, partial_term_of} => property) => property"
 | 
| 
43237
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
314  | 
where  | 
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
315  | 
  "all f = (case narrowing (100 :: code_int) of C ty cs => Universal ty (\<lambda>t. f (conv cs t)) (partial_term_of (TYPE('a))))"
 | 
| 
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
316  | 
|
| 41943 | 317  | 
subsubsection {* class @{text is_testable} *}
 | 
| 41905 | 318  | 
|
| 41943 | 319  | 
text {* The class @{text is_testable} ensures that all necessary type instances are generated. *}
 | 
| 41905 | 320  | 
|
321  | 
class is_testable  | 
|
322  | 
||
323  | 
instance bool :: is_testable ..  | 
|
324  | 
||
| 
43047
 
26774ccb1c74
automatic derivation of partial_term_of functions; renaming type and term to longer names narrowing_type and narrowing_term; hiding constant C; adding overlord option
 
bulwahn 
parents: 
42980 
diff
changeset
 | 
325  | 
instance "fun" :: ("{term_of, narrowing, partial_term_of}", is_testable) is_testable ..
 | 
| 41905 | 326  | 
|
327  | 
definition ensure_testable :: "'a :: is_testable => 'a :: is_testable"  | 
|
328  | 
where  | 
|
329  | 
"ensure_testable f = f"  | 
|
330  | 
||
| 
41910
 
709c04e7b703
adding serial instance of finite_4 in lazysmallcheck; changing code equations for implies
 
bulwahn 
parents: 
41909 
diff
changeset
 | 
331  | 
|
| 
42022
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
332  | 
subsubsection {* Defining a simple datatype to represent functions in an incomplete and redundant way *}
 | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
333  | 
|
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
334  | 
datatype ('a, 'b) ffun = Constant 'b | Update 'a 'b "('a, 'b) ffun"
 | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
335  | 
|
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
336  | 
primrec eval_ffun :: "('a, 'b) ffun => 'a => 'b"
 | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
337  | 
where  | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
338  | 
"eval_ffun (Constant c) x = c"  | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
339  | 
| "eval_ffun (Update x' y f) x = (if x = x' then y else eval_ffun f x)"  | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
340  | 
|
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
341  | 
hide_type (open) ffun  | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
342  | 
hide_const (open) Constant Update eval_ffun  | 
| 
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
343  | 
|
| 
42024
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
344  | 
datatype 'b cfun = Constant 'b  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
345  | 
|
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
346  | 
primrec eval_cfun :: "'b cfun => 'a => 'b"  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
347  | 
where  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
348  | 
"eval_cfun (Constant c) y = c"  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
349  | 
|
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
350  | 
hide_type (open) cfun  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
351  | 
hide_const (open) Constant eval_cfun  | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
352  | 
|
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
353  | 
subsubsection {* Setting up the counterexample generator *}
 | 
| 
43237
 
8f5c3c6c2909
adding compilation that allows existentials in Quickcheck_Narrowing
 
bulwahn 
parents: 
43047 
diff
changeset
 | 
354  | 
|
| 
43702
 
24fb44c1086a
more abstract Thy_Load.load_file/use_file for external theory resources;
 
wenzelm 
parents: 
43562 
diff
changeset
 | 
355  | 
use "Tools/Quickcheck/narrowing_generators.ML"  | 
| 
42024
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
356  | 
|
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
357  | 
setup {* Narrowing_Generators.setup *}
 | 
| 
 
51df23535105
handling a quite restricted set of functions in Quickcheck_Narrowing by an easy transformation
 
bulwahn 
parents: 
42022 
diff
changeset
 | 
358  | 
|
| 43356 | 359  | 
subsection {* Narrowing for integers *}
 | 
360  | 
||
361  | 
||
362  | 
definition drawn_from :: "'a list => 'a cons"  | 
|
363  | 
where "drawn_from xs = C (SumOfProd (map (%_. []) xs)) (map (%x y. x) xs)"  | 
|
364  | 
||
365  | 
function around_zero :: "int => int list"  | 
|
366  | 
where  | 
|
367  | 
"around_zero i = (if i < 0 then [] else (if i = 0 then [0] else around_zero (i - 1) @ [i, -i]))"  | 
|
368  | 
by pat_completeness auto  | 
|
369  | 
termination by (relation "measure nat") auto  | 
|
370  | 
||
371  | 
declare around_zero.simps[simp del]  | 
|
372  | 
||
373  | 
lemma length_around_zero:  | 
|
374  | 
assumes "i >= 0"  | 
|
375  | 
shows "length (around_zero i) = 2 * nat i + 1"  | 
|
376  | 
proof (induct rule: int_ge_induct[OF assms])  | 
|
377  | 
case 1  | 
|
378  | 
from 1 show ?case by (simp add: around_zero.simps)  | 
|
379  | 
next  | 
|
380  | 
case (2 i)  | 
|
381  | 
from 2 show ?case  | 
|
382  | 
by (simp add: around_zero.simps[of "i + 1"])  | 
|
383  | 
qed  | 
|
384  | 
||
385  | 
instantiation int :: narrowing  | 
|
386  | 
begin  | 
|
387  | 
||
388  | 
definition  | 
|
389  | 
"narrowing_int d = (let (u :: _ => _ => unit) = conv; i = Quickcheck_Narrowing.int_of d in drawn_from (around_zero i))"  | 
|
390  | 
||
391  | 
instance ..  | 
|
392  | 
||
393  | 
end  | 
|
394  | 
||
395  | 
lemma [code, code del]: "partial_term_of (ty :: int itself) t == undefined"  | 
|
396  | 
by (rule partial_term_of_anything)+  | 
|
397  | 
||
398  | 
lemma [code]:  | 
|
399  | 
"partial_term_of (ty :: int itself) (Var p t) == Code_Evaluation.Free (STR ''_'') (Typerep.Typerep (STR ''Int.int'') [])"  | 
|
400  | 
"partial_term_of (ty :: int itself) (Ctr i []) == (if i mod 2 = 0 then  | 
|
401  | 
Code_Evaluation.term_of (- (int_of i) div 2) else Code_Evaluation.term_of ((int_of i + 1) div 2))"  | 
|
402  | 
by (rule partial_term_of_anything)+  | 
|
403  | 
||
404  | 
text {* Defining integers by positive and negative copy of naturals *}
 | 
|
405  | 
(*  | 
|
406  | 
datatype simple_int = Positive nat | Negative nat  | 
|
407  | 
||
408  | 
primrec int_of_simple_int :: "simple_int => int"  | 
|
409  | 
where  | 
|
410  | 
"int_of_simple_int (Positive n) = int n"  | 
|
411  | 
| "int_of_simple_int (Negative n) = (-1 - int n)"  | 
|
412  | 
||
413  | 
instantiation int :: narrowing  | 
|
414  | 
begin  | 
|
415  | 
||
416  | 
definition narrowing_int :: "code_int => int cons"  | 
|
417  | 
where  | 
|
418  | 
"narrowing_int d = map_cons int_of_simple_int ((narrowing :: simple_int narrowing) d)"  | 
|
419  | 
||
420  | 
instance ..  | 
|
421  | 
||
422  | 
end  | 
|
423  | 
||
424  | 
text {* printing the partial terms *}
 | 
|
425  | 
||
426  | 
lemma [code]:  | 
|
427  | 
"partial_term_of (ty :: int itself) t == Code_Evaluation.App (Code_Evaluation.Const (STR ''Quickcheck_Narrowing.int_of_simple_int'')  | 
|
428  | 
(Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''Quickcheck_Narrowing.simple_int'') [], Typerep.Typerep (STR ''Int.int'') []])) (partial_term_of (TYPE(simple_int)) t)"  | 
|
429  | 
by (rule partial_term_of_anything)  | 
|
430  | 
||
431  | 
*)  | 
|
432  | 
||
| 
43315
 
893de45ac28d
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
 
bulwahn 
parents: 
43314 
diff
changeset
 | 
433  | 
hide_type code_int narrowing_type narrowing_term cons property  | 
| 43316 | 434  | 
hide_const int_of of_int nth error toEnum marker empty  | 
| 
43315
 
893de45ac28d
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
 
bulwahn 
parents: 
43314 
diff
changeset
 | 
435  | 
C cons conv nonEmpty "apply" sum ensure_testable all exists  | 
| 
43562
 
2c55eac2e5a9
hide rather short auxiliary names, which can easily occur in user theories;
 
wenzelm 
parents: 
43542 
diff
changeset
 | 
436  | 
hide_const (open) Var Ctr  | 
| 
43315
 
893de45ac28d
removing unneccessary manual instantiations and dead definitions; hiding more constants and facts
 
bulwahn 
parents: 
43314 
diff
changeset
 | 
437  | 
hide_fact empty_def cons_def conv.simps nonEmpty.simps apply_def sum_def ensure_testable_def all_def exists_def  | 
| 
42022
 
101ce92333f4
adding a simple datatype for representing functions in Quickcheck_Narrowing
 
bulwahn 
parents: 
42021 
diff
changeset
 | 
438  | 
|
| 43356 | 439  | 
|
| 41905 | 440  | 
end  |