# HG changeset patch # User bulwahn # Date 1279728711 -7200 # Node ID cb55fd65faa6ca886fd67903cd3d0d6babe2d989 # Parent e709e764f20cfff50da86dd4fa4ed083e303b44f adding Quickcheck examples for other quickcheck default types diff -r e709e764f20c -r cb55fd65faa6 src/HOL/ex/Quickcheck_Lattice_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Quickcheck_Lattice_Examples.thy Wed Jul 21 18:11:51 2010 +0200 @@ -0,0 +1,139 @@ +(* Title: HOL/ex/Quickcheck_Lattice_Examples.thy + Author: Lukas Bulwahn + Copyright 2010 TU Muenchen +*) + +theory Quickcheck_Lattice_Examples +imports Quickcheck_Types +begin + +text {* We show how other default types help to find counterexamples to propositions if + the standard default type @{typ int} is insufficient. *} + +notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + top ("\") and + bot ("\") and + inf (infixl "\" 70) and + sup (infixl "\" 65) + +subsection {* Distributive lattices *} + +lemma sup_inf_distrib2: + "((y :: 'a :: distrib_lattice) \ z) \ x = (y \ x) \ (z \ x)" + quickcheck +by(simp add: inf_sup_aci sup_inf_distrib1) + +lemma sup_inf_distrib2_1: + "((y :: 'a :: lattice) \ z) \ x = (y \ x) \ (z \ x)" + quickcheck + oops + +lemma sup_inf_distrib2_2: + "((y :: 'a :: distrib_lattice) \ z') \ x = (y \ x) \ (z \ x)" + quickcheck + oops + +lemma inf_sup_distrib1_1: + "(x :: 'a :: distrib_lattice) \ (y \ z) = (x \ y) \ (x' \ z)" + quickcheck + oops + +lemma inf_sup_distrib2_1: + "((y :: 'a :: distrib_lattice) \ z) \ x = (y \ x) \ (y \ x)" + quickcheck + oops + +subsection {* Bounded lattices *} + +lemma inf_bot_left [simp]: + "\ \ (x :: 'a :: bounded_lattice_bot) = \" + quickcheck + by (rule inf_absorb1) simp + +lemma inf_bot_left_1: + "\ \ (x :: 'a :: bounded_lattice_bot) = x" + quickcheck + oops + +lemma inf_bot_left_2: + "y \ (x :: 'a :: bounded_lattice_bot) = \" + quickcheck + oops + +lemma inf_bot_left_3: + "x \ \ ==> y \ (x :: 'a :: bounded_lattice_bot) \ \" + quickcheck + oops + +lemma inf_bot_right [simp]: + "(x :: 'a :: bounded_lattice_bot) \ \ = \" + quickcheck + by (rule inf_absorb2) simp + +lemma inf_bot_right_1: + "x \ \ ==> (x :: 'a :: bounded_lattice_bot) \ \ = y" + quickcheck + oops + +lemma inf_bot_right_2: + "(x :: 'a :: bounded_lattice_bot) \ \ ~= \" + quickcheck + oops + +lemma inf_bot_right [simp]: + "(x :: 'a :: bounded_lattice_bot) \ \ = \" + quickcheck + oops + +lemma sup_bot_left [simp]: + "\ \ (x :: 'a :: bounded_lattice_bot) = x" + quickcheck + by (rule sup_absorb2) simp + +lemma sup_bot_right [simp]: + "(x :: 'a :: bounded_lattice_bot) \ \ = x" + quickcheck + by (rule sup_absorb1) simp + +lemma sup_eq_bot_iff [simp]: + "(x :: 'a :: bounded_lattice_bot) \ y = \ \ x = \ \ y = \" + quickcheck + by (simp add: eq_iff) + +lemma sup_top_left [simp]: + "\ \ (x :: 'a :: bounded_lattice_top) = \" + quickcheck + by (rule sup_absorb1) simp + +lemma sup_top_right [simp]: + "(x :: 'a :: bounded_lattice_top) \ \ = \" + quickcheck + by (rule sup_absorb2) simp + +lemma inf_top_left [simp]: + "\ \ x = (x :: 'a :: bounded_lattice_top)" + quickcheck + by (rule inf_absorb2) simp + +lemma inf_top_right [simp]: + "x \ \ = (x :: 'a :: bounded_lattice_top)" + quickcheck + by (rule inf_absorb1) simp + +lemma inf_eq_top_iff [simp]: + "x \ y = \ \ x = \ \ y = \" + quickcheck + by (simp add: eq_iff) + + +no_notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + top ("\") and + bot ("\") + +end