| author | krauss | 
| Sun, 12 Dec 2010 21:40:59 +0100 | |
| changeset 41113 | b223fa19af3c | 
| parent 37929 | 22e0797857e6 | 
| child 41413 | 64cd30d6b0b8 | 
| permissions | -rw-r--r-- | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 1 | (* Title: HOL/ex/Quickcheck_Lattice_Examples.thy | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 2 | Author: Lukas Bulwahn | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 3 | Copyright 2010 TU Muenchen | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 4 | *) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 5 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 6 | theory Quickcheck_Lattice_Examples | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 7 | imports Quickcheck_Types | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 8 | begin | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 9 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 10 | text {* We show how other default types help to find counterexamples to propositions if
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 11 |   the standard default type @{typ int} is insufficient. *}
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 12 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 13 | notation | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 14 | less_eq (infix "\<sqsubseteq>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 15 | less (infix "\<sqsubset>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 16 |   top ("\<top>") and
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 17 |   bot ("\<bottom>") and
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 18 | inf (infixl "\<sqinter>" 70) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 19 | sup (infixl "\<squnion>" 65) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 20 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 21 | subsection {* Distributive lattices *}
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 22 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 23 | lemma sup_inf_distrib2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 24 | "((y :: 'a :: distrib_lattice) \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 25 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 26 | by(simp add: inf_sup_aci sup_inf_distrib1) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 27 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 28 | lemma sup_inf_distrib2_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 29 | "((y :: 'a :: lattice) \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 30 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 31 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 32 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 33 | lemma sup_inf_distrib2_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 34 | "((y :: 'a :: distrib_lattice) \<sqinter> z') \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 35 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 36 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 37 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 38 | lemma inf_sup_distrib1_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 39 | "(x :: 'a :: distrib_lattice) \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x' \<sqinter> z)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 40 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 41 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 42 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 43 | lemma inf_sup_distrib2_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 44 | "((y :: 'a :: distrib_lattice) \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (y \<sqinter> x)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 45 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 46 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 47 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 48 | subsection {* Bounded lattices *}
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 49 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 50 | lemma inf_bot_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 51 | "\<bottom> \<sqinter> (x :: 'a :: bounded_lattice_bot) = \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 52 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 53 | by (rule inf_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 54 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 55 | lemma inf_bot_left_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 56 | "\<bottom> \<sqinter> (x :: 'a :: bounded_lattice_bot) = x" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 57 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 58 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 59 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 60 | lemma inf_bot_left_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 61 | "y \<sqinter> (x :: 'a :: bounded_lattice_bot) = \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 62 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 63 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 64 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 65 | lemma inf_bot_left_3: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 66 | "x \<noteq> \<bottom> ==> y \<sqinter> (x :: 'a :: bounded_lattice_bot) \<noteq> \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 67 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 68 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 69 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 70 | lemma inf_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 71 | "(x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> = \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 72 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 73 | by (rule inf_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 74 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 75 | lemma inf_bot_right_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 76 | "x \<noteq> \<bottom> ==> (x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> = y" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 77 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 78 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 79 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 80 | lemma inf_bot_right_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 81 | "(x :: 'a :: bounded_lattice_bot) \<sqinter> \<bottom> ~= \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 82 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 83 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 84 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 85 | lemma inf_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 86 | "(x :: 'a :: bounded_lattice_bot) \<squnion> \<bottom> = \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 87 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 88 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 89 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 90 | lemma sup_bot_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 91 | "\<bottom> \<squnion> (x :: 'a :: bounded_lattice_bot) = x" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 92 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 93 | by (rule sup_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 94 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 95 | lemma sup_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 96 | "(x :: 'a :: bounded_lattice_bot) \<squnion> \<bottom> = x" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 97 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 98 | by (rule sup_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 99 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 100 | lemma sup_eq_bot_iff [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 101 | "(x :: 'a :: bounded_lattice_bot) \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 102 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 103 | by (simp add: eq_iff) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 104 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 105 | lemma sup_top_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 106 | "\<top> \<squnion> (x :: 'a :: bounded_lattice_top) = \<top>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 107 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 108 | by (rule sup_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 109 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 110 | lemma sup_top_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 111 | "(x :: 'a :: bounded_lattice_top) \<squnion> \<top> = \<top>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 112 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 113 | by (rule sup_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 114 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 115 | lemma inf_top_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 116 | "\<top> \<sqinter> x = (x :: 'a :: bounded_lattice_top)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 117 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 118 | by (rule inf_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 119 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 120 | lemma inf_top_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 121 | "x \<sqinter> \<top> = (x :: 'a :: bounded_lattice_top)" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 122 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 123 | by (rule inf_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 124 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 125 | lemma inf_eq_top_iff [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 126 | "x \<sqinter> y = \<top> \<longleftrightarrow> x = \<top> \<and> y = \<top>" | 
| 37929 
22e0797857e6
adding checking of expected result for the tool quickcheck; annotated a few quickcheck examples
 bulwahn parents: 
37916diff
changeset | 127 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 128 | by (simp add: eq_iff) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 129 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 130 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 131 | no_notation | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 132 | less_eq (infix "\<sqsubseteq>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 133 | less (infix "\<sqsubset>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 134 | inf (infixl "\<sqinter>" 70) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 135 | sup (infixl "\<squnion>" 65) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 136 |   top ("\<top>") and
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 137 |   bot ("\<bottom>")
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 138 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 139 | end |