| author | huffman | 
| Wed, 10 Aug 2011 17:02:03 -0700 | |
| changeset 44141 | 0697c01ff3ea | 
| parent 43888 | ee4be704c2a4 | 
| child 44845 | 5e51075cbd97 | 
| 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 | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
37929diff
changeset | 7 | imports "~~/src/HOL/Library/Quickcheck_Types" | 
| 37916 
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 | |
| 43888 
ee4be704c2a4
adapting an experimental setup to changes in quickcheck's infrastructure
 bulwahn parents: 
41413diff
changeset | 21 | declare [[quickcheck_narrowing_active = false]] | 
| 
ee4be704c2a4
adapting an experimental setup to changes in quickcheck's infrastructure
 bulwahn parents: 
41413diff
changeset | 22 | |
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 23 | subsection {* Distributive lattices *}
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 24 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 25 | lemma sup_inf_distrib2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 26 | "((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 | 27 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 28 | by(simp add: inf_sup_aci sup_inf_distrib1) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 29 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 30 | lemma sup_inf_distrib2_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 31 | "((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 | 32 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 33 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 34 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 35 | lemma sup_inf_distrib2_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 36 | "((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 | 37 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 38 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 39 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 40 | lemma inf_sup_distrib1_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 41 | "(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 | 42 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 43 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 44 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 45 | lemma inf_sup_distrib2_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 46 | "((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 | 47 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 48 | oops | 
| 
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 | subsection {* Bounded lattices *}
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 51 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 52 | lemma inf_bot_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 53 | "\<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 | 54 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 55 | by (rule inf_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 56 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 57 | lemma inf_bot_left_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 58 | "\<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 | 59 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 60 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 61 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 62 | lemma inf_bot_left_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 63 | "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 | 64 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 65 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 66 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 67 | lemma inf_bot_left_3: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 68 | "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 | 69 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 70 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 71 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 72 | lemma inf_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 73 | "(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 | 74 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 75 | by (rule inf_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 76 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 77 | lemma inf_bot_right_1: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 78 | "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 | 79 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 80 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 81 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 82 | lemma inf_bot_right_2: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 83 | "(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 | 84 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 85 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 86 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 87 | lemma inf_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 88 | "(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 | 89 | quickcheck[expect = counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 90 | oops | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 91 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 92 | lemma sup_bot_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 93 | "\<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 | 94 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 95 | by (rule sup_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 96 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 97 | lemma sup_bot_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 98 | "(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 | 99 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 100 | by (rule sup_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 101 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 102 | lemma sup_eq_bot_iff [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 103 | "(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 | 104 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 105 | by (simp add: eq_iff) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 106 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 107 | lemma sup_top_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 108 | "\<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 | 109 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 110 | by (rule sup_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 111 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 112 | lemma sup_top_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 113 | "(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 | 114 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 115 | by (rule sup_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 116 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 117 | lemma inf_top_left [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 118 | "\<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 | 119 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 120 | by (rule inf_absorb2) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 121 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 122 | lemma inf_top_right [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 123 | "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 | 124 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 125 | by (rule inf_absorb1) simp | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 126 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 127 | lemma inf_eq_top_iff [simp]: | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 128 | "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 | 129 | quickcheck[expect = no_counterexample] | 
| 37916 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 130 | by (simp add: eq_iff) | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 131 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 132 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 133 | no_notation | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 134 | less_eq (infix "\<sqsubseteq>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 135 | less (infix "\<sqsubset>" 50) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 136 | inf (infixl "\<sqinter>" 70) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 137 | sup (infixl "\<squnion>" 65) and | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 138 |   top ("\<top>") and
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 139 |   bot ("\<bottom>")
 | 
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 140 | |
| 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 bulwahn parents: diff
changeset | 141 | end |