| author | paulson <lp15@cam.ac.uk> | 
| Fri, 02 Oct 2015 15:07:41 +0100 | |
| changeset 61306 | 9dd394c866fc | 
| parent 57921 | 9225b2761126 | 
| child 63167 | 0909deb8059b | 
| permissions | -rw-r--r-- | 
| 46879 | 1  | 
(* Title: HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy  | 
| 
37916
 
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  | 
| 
57921
 
9225b2761126
Quickcheck_Types is no longer needed due to 51aa30c9ee4e
 
Andreas Lochbihler 
parents: 
57573 
diff
changeset
 | 
7  | 
imports Main  | 
| 
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  | 
|
| 
57921
 
9225b2761126
Quickcheck_Types is no longer needed due to 51aa30c9ee4e
 
Andreas Lochbihler 
parents: 
57573 
diff
changeset
 | 
10  | 
declare [[quickcheck_finite_type_size=5]]  | 
| 
 
9225b2761126
Quickcheck_Types is no longer needed due to 51aa30c9ee4e
 
Andreas Lochbihler 
parents: 
57573 
diff
changeset
 | 
11  | 
|
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
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
 | 
13  | 
  the standard default type @{typ int} is insufficient. *}
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
notation  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
less_eq (infix "\<sqsubseteq>" 50) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
17  | 
less (infix "\<sqsubset>" 50) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
  top ("\<top>") and
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
  bot ("\<bottom>") and
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
inf (infixl "\<sqinter>" 70) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
sup (infixl "\<squnion>" 65)  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
|
| 
44994
 
a915907a67d5
increasing quickcheck timeout to reduce spurious test failures due to massive parallel invocations and bad scheduling
 
bulwahn 
parents: 
44845 
diff
changeset
 | 
23  | 
declare [[quickcheck_narrowing_active = false, quickcheck_timeout = 3600]]  | 
| 
43888
 
ee4be704c2a4
adapting an experimental setup to changes in quickcheck's infrastructure
 
bulwahn 
parents: 
41413 
diff
changeset
 | 
24  | 
|
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
subsection {* Distributive lattices *}
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
lemma sup_inf_distrib2:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
"((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: 
37916 
diff
changeset
 | 
29  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
by(simp add: inf_sup_aci sup_inf_distrib1)  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
lemma sup_inf_distrib2_1:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
"((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: 
37916 
diff
changeset
 | 
34  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
lemma sup_inf_distrib2_2:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
"((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: 
37916 
diff
changeset
 | 
39  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
lemma inf_sup_distrib1_1:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
43  | 
"(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: 
37916 
diff
changeset
 | 
44  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
47  | 
lemma inf_sup_distrib2_1:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
"((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: 
37916 
diff
changeset
 | 
49  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
oops  | 
| 
 
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  | 
subsection {* Bounded lattices *}
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
54  | 
lemma inf_bot_left [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
55  | 
"\<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: 
37916 
diff
changeset
 | 
56  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
by (rule inf_absorb1) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
lemma inf_bot_left_1:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
60  | 
"\<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: 
37916 
diff
changeset
 | 
61  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
62  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
64  | 
lemma inf_bot_left_2:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
65  | 
"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: 
37916 
diff
changeset
 | 
66  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
69  | 
lemma inf_bot_left_3:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
70  | 
"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: 
37916 
diff
changeset
 | 
71  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
74  | 
lemma inf_bot_right [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
75  | 
"(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: 
37916 
diff
changeset
 | 
76  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
by (rule inf_absorb2) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
lemma inf_bot_right_1:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
80  | 
"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: 
37916 
diff
changeset
 | 
81  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
82  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
84  | 
lemma inf_bot_right_2:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
"(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: 
37916 
diff
changeset
 | 
86  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
|
| 57573 | 89  | 
lemma sup_bot_right [simp]:  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
"(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: 
37916 
diff
changeset
 | 
91  | 
quickcheck[expect = counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
oops  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
lemma sup_bot_left [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
"\<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: 
37916 
diff
changeset
 | 
96  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
by (rule sup_absorb2) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
|
| 57573 | 99  | 
lemma sup_bot_right_2 [simp]:  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
100  | 
"(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: 
37916 
diff
changeset
 | 
101  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
102  | 
by (rule sup_absorb1) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
lemma sup_eq_bot_iff [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
"(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: 
37916 
diff
changeset
 | 
106  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
107  | 
by (simp add: eq_iff)  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
lemma sup_top_left [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
110  | 
"\<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: 
37916 
diff
changeset
 | 
111  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
112  | 
by (rule sup_absorb1) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
lemma sup_top_right [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
115  | 
"(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: 
37916 
diff
changeset
 | 
116  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
117  | 
by (rule sup_absorb2) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
119  | 
lemma inf_top_left [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
120  | 
"\<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: 
37916 
diff
changeset
 | 
121  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
122  | 
by (rule inf_absorb2) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
124  | 
lemma inf_top_right [simp]:  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
125  | 
"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: 
37916 
diff
changeset
 | 
126  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
127  | 
by (rule inf_absorb1) simp  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
129  | 
lemma inf_eq_top_iff [simp]:  | 
| 44845 | 130  | 
"(x :: 'a :: bounded_lattice_top) \<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: 
37916 
diff
changeset
 | 
131  | 
quickcheck[expect = no_counterexample]  | 
| 
37916
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
132  | 
by (simp add: eq_iff)  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
135  | 
no_notation  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
136  | 
less_eq (infix "\<sqsubseteq>" 50) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
137  | 
less (infix "\<sqsubset>" 50) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
inf (infixl "\<sqinter>" 70) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
139  | 
sup (infixl "\<squnion>" 65) and  | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
  top ("\<top>") and
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
141  | 
  bot ("\<bottom>")
 | 
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
142  | 
|
| 
 
cb55fd65faa6
adding Quickcheck examples for other quickcheck default types
 
bulwahn 
parents:  
diff
changeset
 | 
143  | 
end  |