# HG changeset patch # User bulwahn # Date 1279728711 -7200 # Node ID e709e764f20cfff50da86dd4fa4ed083e303b44f # Parent 49b908e43d6185a253cce1ed293b56042b742c1c adding Library theory for other quickcheck default types diff -r 49b908e43d61 -r e709e764f20c src/HOL/Library/Quickcheck_Types.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Quickcheck_Types.thy Wed Jul 21 18:11:51 2010 +0200 @@ -0,0 +1,261 @@ +(* Title: HOL/Library/Quickcheck_Types.thy + Author: Lukas Bulwahn + Copyright 2010 TU Muenchen +*) + +theory Quickcheck_Types +imports Main +begin + +text {* +This theory provides some default types for the quickcheck execution. +In most cases, the default type @{typ "int"} meets the sort constraints +of the proposition. +But for the type classes bot and top, the type @{typ "int"} is insufficient. +Hence, we provide other types than @{typ "int"} as further default types. +*} + +subsection {* A non-distributive lattice *} + +datatype non_distrib_lattice = Zero | A | B | C | One + +instantiation non_distrib_lattice :: order +begin + +definition less_eq_non_distrib_lattice +where + "a <= b = (case a of Zero => True | A => (b = A) \ (b = One) | B => (b = B) \ (b = One) | C => (b = C) \ (b = One) | One => (b = One))" + +definition less_non_distrib_lattice +where + "a < b = (case a of Zero => (b \ Zero) | A => (b = One) | B => (b = One) | C => (b = One) | One => False)" + +instance proof +qed (auto simp add: less_eq_non_distrib_lattice_def less_non_distrib_lattice_def split: non_distrib_lattice.split non_distrib_lattice.split_asm) + +end + +instantiation non_distrib_lattice :: lattice +begin + + +definition sup_non_distrib_lattice +where + "sup a b = (if a = b then a else (if a = Zero then b else (if b = Zero then a else One)))" + +definition inf_non_distrib_lattice +where + "inf a b = (if a = b then a else (if a = One then b else (if b = One then a else Zero)))" + +instance proof +qed (auto simp add: inf_non_distrib_lattice_def sup_non_distrib_lattice_def less_eq_non_distrib_lattice_def split: split_if non_distrib_lattice.split non_distrib_lattice.split_asm) + +end + +subsection {* Values extended by a bottom element *} + +datatype 'a bot = Value 'a | Bot + +instantiation bot :: (preorder) preorder +begin + +definition less_eq_bot where + "x \ y \ (case x of Bot \ True | Value x \ (case y of Bot \ False | Value y \ x \ y))" + +definition less_bot where + "x < y \ (case y of Bot \ False | Value y \ (case x of Bot \ True | Value x \ x < y))" + +lemma less_eq_bot_Bot [simp]: "Bot \ x" + by (simp add: less_eq_bot_def) + +lemma less_eq_bot_Bot_code [code]: "Bot \ x \ True" + by simp + +lemma less_eq_bot_Bot_is_Bot: "x \ Bot \ x = Bot" + by (cases x) (simp_all add: less_eq_bot_def) + +lemma less_eq_bot_Value_Bot [simp, code]: "Value x \ Bot \ False" + by (simp add: less_eq_bot_def) + +lemma less_eq_bot_Value [simp, code]: "Value x \ Value y \ x \ y" + by (simp add: less_eq_bot_def) + +lemma less_bot_Bot [simp, code]: "x < Bot \ False" + by (simp add: less_bot_def) + +lemma less_bot_Bot_is_Value: "Bot < x \ \z. x = Value z" + by (cases x) (simp_all add: less_bot_def) + +lemma less_bot_Bot_Value [simp]: "Bot < Value x" + by (simp add: less_bot_def) + +lemma less_bot_Bot_Value_code [code]: "Bot < Value x \ True" + by simp + +lemma less_bot_Value [simp, code]: "Value x < Value y \ x < y" + by (simp add: less_bot_def) + +instance proof +qed (auto simp add: less_eq_bot_def less_bot_def less_le_not_le elim: order_trans split: bot.splits) + +end + +instance bot :: (order) order proof +qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits) + +instance bot :: (linorder) linorder proof +qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits) + +instantiation bot :: (preorder) bot +begin + +definition "bot = Bot" + +instance proof +qed (simp add: bot_bot_def) + +end + +instantiation bot :: (top) top +begin + +definition "top = Value top" + +instance proof +qed (simp add: top_bot_def less_eq_bot_def split: bot.split) + +end + +instantiation bot :: (semilattice_inf) semilattice_inf +begin + +definition inf_bot +where + "inf x y = (case x of Bot => Bot | Value v => (case y of Bot => Bot | Value v' => Value (inf v v')))" + +instance proof +qed (auto simp add: inf_bot_def less_eq_bot_def split: bot.splits) + +end + +instantiation bot :: (semilattice_sup) semilattice_sup +begin + +definition sup_bot +where + "sup x y = (case x of Bot => y | Value v => (case y of Bot => x | Value v' => Value (sup v v')))" + +instance proof +qed (auto simp add: sup_bot_def less_eq_bot_def split: bot.splits) + +end + +instance bot :: (lattice) bounded_lattice_bot .. + +section {* Values extended by a top element *} + +datatype 'a top = Value 'a | Top + +instantiation top :: (preorder) preorder +begin + +definition less_eq_top where + "x \ y \ (case y of Top \ True | Value y \ (case x of Top \ False | Value x \ x \ y))" + +definition less_top where + "x < y \ (case x of Top \ False | Value x \ (case y of Top \ True | Value y \ x < y))" + +lemma less_eq_top_Top [simp]: "x <= Top" + by (simp add: less_eq_top_def) + +lemma less_eq_top_Top_code [code]: "x \ Top \ True" + by simp + +lemma less_eq_top_is_Top: "Top \ x \ x = Top" + by (cases x) (simp_all add: less_eq_top_def) + +lemma less_eq_top_Top_Value [simp, code]: "Top \ Value x \ False" + by (simp add: less_eq_top_def) + +lemma less_eq_top_Value_Value [simp, code]: "Value x \ Value y \ x \ y" + by (simp add: less_eq_top_def) + +lemma less_top_Top [simp, code]: "Top < x \ False" + by (simp add: less_top_def) + +lemma less_top_Top_is_Value: "x < Top \ \z. x = Value z" + by (cases x) (simp_all add: less_top_def) + +lemma less_top_Value_Top [simp]: "Value x < Top" + by (simp add: less_top_def) + +lemma less_top_Value_Top_code [code]: "Value x < Top \ True" + by simp + +lemma less_top_Value [simp, code]: "Value x < Value y \ x < y" + by (simp add: less_top_def) + +instance proof +qed (auto simp add: less_eq_top_def less_top_def less_le_not_le elim: order_trans split: top.splits) + +end + +instance top :: (order) order proof +qed (auto simp add: less_eq_top_def less_top_def split: top.splits) + +instance top :: (linorder) linorder proof +qed (auto simp add: less_eq_top_def less_top_def split: top.splits) + +instantiation top :: (preorder) top +begin + +definition "top = Top" + +instance proof +qed (simp add: top_top_def) + +end + +instantiation top :: (bot) bot +begin + +definition "bot = Value bot" + +instance proof +qed (simp add: bot_top_def less_eq_top_def split: top.split) + +end + +instantiation top :: (semilattice_inf) semilattice_inf +begin + +definition inf_top +where + "inf x y = (case x of Top => y | Value v => (case y of Top => x | Value v' => Value (inf v v')))" + +instance proof +qed (auto simp add: inf_top_def less_eq_top_def split: top.splits) + +end + +instantiation top :: (semilattice_sup) semilattice_sup +begin + +definition sup_top +where + "sup x y = (case x of Top => Top | Value v => (case y of Top => Top | Value v' => Value (sup v v')))" + +instance proof +qed (auto simp add: sup_top_def less_eq_top_def split: top.splits) + +end + +instance top :: (lattice) bounded_lattice_top .. + +section {* Quickcheck configuration *} + +quickcheck_params[default_type = ["int", "non_distrib_lattice", "int bot", "int top"]] + +hide_type non_distrib_lattice bot top + +end \ No newline at end of file