| author | wenzelm | 
| Tue, 08 Feb 2011 19:57:11 +0100 | |
| changeset 41731 | 2fb760843e17 | 
| parent 41115 | 2c362ff5daf4 | 
| child 45117 | 3911cf09899a | 
| permissions | -rw-r--r-- | 
| 31596 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 26348 | 2 | |
| 3 | header {* Finite types as explicit enumerations *}
 | |
| 4 | ||
| 5 | theory Enum | |
| 40650 
d40b347d5b0b
adding Enum to HOL-Main image and removing it from HOL-Library
 bulwahn parents: 
40649diff
changeset | 6 | imports Map String | 
| 26348 | 7 | begin | 
| 8 | ||
| 9 | subsection {* Class @{text enum} *}
 | |
| 10 | ||
| 29797 | 11 | class enum = | 
| 26348 | 12 | fixes enum :: "'a list" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 13 |   fixes enum_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 14 |   fixes enum_ex  :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 33635 | 15 | assumes UNIV_enum: "UNIV = set enum" | 
| 26444 | 16 | and enum_distinct: "distinct enum" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 17 | assumes enum_all : "enum_all P = (\<forall> x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 18 | assumes enum_ex : "enum_ex P = (\<exists> x. P x)" | 
| 26348 | 19 | begin | 
| 20 | ||
| 29797 | 21 | subclass finite proof | 
| 22 | qed (simp add: UNIV_enum) | |
| 26444 | 23 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 24 | lemma enum_UNIV: "set enum = UNIV" unfolding UNIV_enum .. | 
| 26444 | 25 | |
| 40683 | 26 | lemma in_enum: "x \<in> set enum" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 27 | unfolding enum_UNIV by auto | 
| 26348 | 28 | |
| 29 | lemma enum_eq_I: | |
| 30 | assumes "\<And>x. x \<in> set xs" | |
| 31 | shows "set enum = set xs" | |
| 32 | proof - | |
| 33 | from assms UNIV_eq_I have "UNIV = set xs" by auto | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 34 | with enum_UNIV show ?thesis by simp | 
| 26348 | 35 | qed | 
| 36 | ||
| 37 | end | |
| 38 | ||
| 39 | ||
| 40 | subsection {* Equality and order on functions *}
 | |
| 41 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 42 | instantiation "fun" :: (enum, equal) equal | 
| 26513 | 43 | begin | 
| 26348 | 44 | |
| 26513 | 45 | definition | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 46 | "HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)" | 
| 26513 | 47 | |
| 31464 | 48 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 49 | qed (simp_all add: equal_fun_def enum_UNIV fun_eq_iff) | 
| 26513 | 50 | |
| 51 | end | |
| 26348 | 52 | |
| 40898 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 bulwahn parents: 
40683diff
changeset | 53 | lemma [code]: | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 54 | "HOL.equal f g \<longleftrightarrow> enum_all (%x. f x = g x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 55 | by (auto simp add: equal enum_all fun_eq_iff) | 
| 40898 
882e860a1e83
changed order of lemmas to overwrite the general code equation with the nbe-specific one
 bulwahn parents: 
40683diff
changeset | 56 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 57 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 58 | "HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 59 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 60 | |
| 28562 | 61 | lemma order_fun [code]: | 
| 26348 | 62 | fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 63 | shows "f \<le> g \<longleftrightarrow> enum_all (\<lambda>x. f x \<le> g x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 64 | and "f < g \<longleftrightarrow> f \<le> g \<and> enum_ex (\<lambda>x. f x \<noteq> g x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 65 | by (simp_all add: enum_all enum_ex fun_eq_iff le_fun_def order_less_le) | 
| 26968 | 66 | |
| 67 | ||
| 68 | subsection {* Quantifiers *}
 | |
| 69 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 70 | lemma all_code [code]: "(\<forall>x. P x) \<longleftrightarrow> enum_all P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 71 | by (simp add: enum_all) | 
| 26968 | 72 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 73 | lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> enum_ex P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 74 | by (simp add: enum_ex) | 
| 26348 | 75 | |
| 40652 | 76 | lemma exists1_code[code]: "(\<exists>!x. P x) \<longleftrightarrow> list_ex1 P enum" | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 77 | unfolding list_ex1_iff enum_UNIV by auto | 
| 40652 | 78 | |
| 26348 | 79 | |
| 80 | subsection {* Default instances *}
 | |
| 81 | ||
| 26444 | 82 | primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where | 
| 83 | "n_lists 0 xs = [[]]" | |
| 84 | | "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))" | |
| 85 | ||
| 86 | lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])" | |
| 87 | by (induct n) simp_all | |
| 88 | ||
| 89 | lemma length_n_lists: "length (n_lists n xs) = length xs ^ n" | |
| 33639 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 hoelzl parents: 
33635diff
changeset | 90 | by (induct n) (auto simp add: length_concat o_def listsum_triv) | 
| 26444 | 91 | |
| 92 | lemma length_n_lists_elem: "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n" | |
| 93 | by (induct n arbitrary: ys) auto | |
| 94 | ||
| 95 | lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 96 | proof (rule set_eqI) | 
| 26444 | 97 | fix ys :: "'a list" | 
| 98 |   show "ys \<in> set (n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | |
| 99 | proof - | |
| 100 | have "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n" | |
| 101 | by (induct n arbitrary: ys) auto | |
| 102 | moreover have "\<And>x. ys \<in> set (n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs" | |
| 103 | by (induct n arbitrary: ys) auto | |
| 104 | moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (n_lists (length ys) xs)" | |
| 105 | by (induct ys) auto | |
| 106 | ultimately show ?thesis by auto | |
| 107 | qed | |
| 108 | qed | |
| 109 | ||
| 110 | lemma distinct_n_lists: | |
| 111 | assumes "distinct xs" | |
| 112 | shows "distinct (n_lists n xs)" | |
| 113 | proof (rule card_distinct) | |
| 114 | from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) | |
| 115 | have "card (set (n_lists n xs)) = card (set xs) ^ n" | |
| 116 | proof (induct n) | |
| 117 | case 0 then show ?case by simp | |
| 118 | next | |
| 119 | case (Suc n) | |
| 120 | moreover have "card (\<Union>ys\<in>set (n_lists n xs). (\<lambda>y. y # ys) ` set xs) | |
| 121 | = (\<Sum>ys\<in>set (n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))" | |
| 122 | by (rule card_UN_disjoint) auto | |
| 123 | moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)" | |
| 124 | by (rule card_image) (simp add: inj_on_def) | |
| 125 | ultimately show ?case by auto | |
| 126 | qed | |
| 127 | also have "\<dots> = length xs ^ n" by (simp add: card_length) | |
| 128 | finally show "card (set (n_lists n xs)) = length (n_lists n xs)" | |
| 129 | by (simp add: length_n_lists) | |
| 130 | qed | |
| 131 | ||
| 132 | lemma map_of_zip_enum_is_Some: | |
| 133 | assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 134 | shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y" | |
| 135 | proof - | |
| 136 | from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow> | |
| 137 | (\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)" | |
| 138 | by (auto intro!: map_of_zip_is_Some) | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 139 | then show ?thesis using enum_UNIV by auto | 
| 26444 | 140 | qed | 
| 141 | ||
| 142 | lemma map_of_zip_enum_inject: | |
| 143 | fixes xs ys :: "'b\<Colon>enum list" | |
| 144 | assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 145 | "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 146 | and map_of: "the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys)" | |
| 147 | shows "xs = ys" | |
| 148 | proof - | |
| 149 | have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)" | |
| 150 | proof | |
| 151 | fix x :: 'a | |
| 152 | from length map_of_zip_enum_is_Some obtain y1 y2 | |
| 153 | where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1" | |
| 154 | and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast | |
| 155 | moreover from map_of have "the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x) = the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x)" | |
| 156 | by (auto dest: fun_cong) | |
| 157 | ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x" | |
| 158 | by simp | |
| 159 | qed | |
| 160 | with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) | |
| 161 | qed | |
| 162 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 163 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 164 |   all_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 165 | where | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 166 | "all_n_lists P n = (\<forall>xs \<in> set (n_lists n enum). P xs)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 167 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 168 | lemma [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 169 | "all_n_lists P n = (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 170 | unfolding all_n_lists_def enum_all | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 171 | by (cases n) (auto simp add: enum_UNIV) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 172 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 173 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 174 |   ex_n_lists :: "(('a :: enum) list \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> bool"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 175 | where | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 176 | "ex_n_lists P n = (\<exists>xs \<in> set (n_lists n enum). P xs)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 177 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 178 | lemma [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 179 | "ex_n_lists P n = (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 180 | unfolding ex_n_lists_def enum_ex | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 181 | by (cases n) (auto simp add: enum_UNIV) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 182 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 183 | |
| 26444 | 184 | instantiation "fun" :: (enum, enum) enum | 
| 185 | begin | |
| 186 | ||
| 187 | definition | |
| 37765 | 188 | "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)" | 
| 26444 | 189 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 190 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 191 | "enum_all P = all_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 192 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 193 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 194 | "enum_ex P = ex_n_lists (\<lambda>bs. P (the o map_of (zip enum bs))) (length (enum :: 'a list))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 195 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 196 | |
| 26444 | 197 | instance proof | 
| 198 |   show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 199 | proof (rule UNIV_eq_I) | |
| 200 | fix f :: "'a \<Rightarrow> 'b" | |
| 201 | have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | |
| 40683 | 202 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 26444 | 203 | then show "f \<in> set enum" | 
| 40683 | 204 | by (auto simp add: enum_fun_def set_n_lists intro: in_enum) | 
| 26444 | 205 | qed | 
| 206 | next | |
| 207 | from map_of_zip_enum_inject | |
| 208 |   show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 209 | by (auto intro!: inj_onI simp add: enum_fun_def | |
| 210 | distinct_map distinct_n_lists enum_distinct set_n_lists enum_all) | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 211 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 212 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 213 |   show "enum_all (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = (\<forall>x. P x)"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 214 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 215 | assume "enum_all P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 216 | show "\<forall>x. P x" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 217 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 218 | fix f :: "'a \<Rightarrow> 'b" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 219 | have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 220 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 221 | from `enum_all P` have "P (the \<circ> map_of (zip enum (map f enum)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 222 | unfolding enum_all_fun_def all_n_lists_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 223 | apply (simp add: set_n_lists) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 224 | apply (erule_tac x="map f enum" in allE) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 225 | apply (auto intro!: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 226 | done | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 227 | from this f show "P f" by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 228 | qed | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 229 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 230 | assume "\<forall>x. P x" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 231 | from this show "enum_all P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 232 | unfolding enum_all_fun_def all_n_lists_def by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 233 | qed | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 234 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 235 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 236 |   show "enum_ex (P :: ('a \<Rightarrow> 'b) \<Rightarrow> bool) = (\<exists>x. P x)"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 237 | proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 238 | assume "enum_ex P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 239 | from this show "\<exists>x. P x" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 240 | unfolding enum_ex_fun_def ex_n_lists_def by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 241 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 242 | assume "\<exists>x. P x" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 243 | from this obtain f where "P f" .. | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 244 | have f: "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 245 | by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 246 | from `P f` this have "P (the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 247 | by auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 248 | from this show "enum_ex P" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 249 | unfolding enum_ex_fun_def ex_n_lists_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 250 | apply (auto simp add: set_n_lists) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 251 | apply (rule_tac x="map f enum" in exI) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 252 | apply (auto intro!: in_enum) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 253 | done | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 254 | qed | 
| 26444 | 255 | qed | 
| 256 | ||
| 257 | end | |
| 258 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 259 | lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
 | 
| 28245 
9767dd8e1e54
celver code lemma avoid type ambiguity problem with Haskell
 haftmann parents: 
27487diff
changeset | 260 | in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))" | 
| 
9767dd8e1e54
celver code lemma avoid type ambiguity problem with Haskell
 haftmann parents: 
27487diff
changeset | 261 | by (simp add: enum_fun_def Let_def) | 
| 26444 | 262 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 263 | lemma enum_all_fun_code [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 264 |   "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 265 | in all_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 266 | by (simp add: enum_all_fun_def Let_def) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 267 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 268 | lemma enum_ex_fun_code [code]: | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 269 |   "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list)
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 270 | in ex_n_lists (\<lambda>bs. P (the o map_of (zip enum_a bs))) (length enum_a))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 271 | by (simp add: enum_ex_fun_def Let_def) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 272 | |
| 26348 | 273 | instantiation unit :: enum | 
| 274 | begin | |
| 275 | ||
| 276 | definition | |
| 277 | "enum = [()]" | |
| 278 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 279 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 280 | "enum_all P = P ()" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 281 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 282 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 283 | "enum_ex P = P ()" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 284 | |
| 31464 | 285 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 286 | qed (auto simp add: enum_unit_def UNIV_unit enum_all_unit_def enum_ex_unit_def intro: unit.exhaust) | 
| 26348 | 287 | |
| 288 | end | |
| 289 | ||
| 290 | instantiation bool :: enum | |
| 291 | begin | |
| 292 | ||
| 293 | definition | |
| 294 | "enum = [False, True]" | |
| 295 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 296 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 297 | "enum_all P = (P False \<and> P True)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 298 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 299 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 300 | "enum_ex P = (P False \<or> P True)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 301 | |
| 31464 | 302 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 303 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 304 | show "enum_all (P :: bool \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 305 | unfolding enum_all_bool_def by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 306 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 307 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 308 | show "enum_ex (P :: bool \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 309 | unfolding enum_ex_bool_def by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 310 | qed (auto simp add: enum_bool_def UNIV_bool) | 
| 26348 | 311 | |
| 312 | end | |
| 313 | ||
| 314 | primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
 | |
| 315 | "product [] _ = []" | |
| 316 | | "product (x#xs) ys = map (Pair x) ys @ product xs ys" | |
| 317 | ||
| 318 | lemma product_list_set: | |
| 319 | "set (product xs ys) = set xs \<times> set ys" | |
| 320 | by (induct xs) auto | |
| 321 | ||
| 26444 | 322 | lemma distinct_product: | 
| 323 | assumes "distinct xs" and "distinct ys" | |
| 324 | shows "distinct (product xs ys)" | |
| 325 | using assms by (induct xs) | |
| 326 | (auto intro: inj_onI simp add: product_list_set distinct_map) | |
| 327 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37601diff
changeset | 328 | instantiation prod :: (enum, enum) enum | 
| 26348 | 329 | begin | 
| 330 | ||
| 331 | definition | |
| 332 | "enum = product enum enum" | |
| 333 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 334 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 335 | "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 336 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 337 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 338 | "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 339 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 340 | |
| 26348 | 341 | instance by default | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 342 | (simp_all add: enum_prod_def product_list_set distinct_product | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 343 | enum_UNIV enum_distinct enum_all_prod_def enum_all enum_ex_prod_def enum_ex) | 
| 26348 | 344 | |
| 345 | end | |
| 346 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37601diff
changeset | 347 | instantiation sum :: (enum, enum) enum | 
| 26348 | 348 | begin | 
| 349 | ||
| 350 | definition | |
| 351 | "enum = map Inl enum @ map Inr enum" | |
| 352 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 353 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 354 | "enum_all P = (enum_all (%x. P (Inl x)) \<and> enum_all (%x. P (Inr x)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 355 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 356 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 357 | "enum_ex P = (enum_ex (%x. P (Inl x)) \<or> enum_ex (%x. P (Inr x)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 358 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 359 | instance proof | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 360 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 361 |   show "enum_all (P :: ('a + 'b) \<Rightarrow> bool) = (\<forall>x. P x)"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 362 | unfolding enum_all_sum_def enum_all | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 363 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 364 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 365 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 366 |   show "enum_ex (P :: ('a + 'b) \<Rightarrow> bool) = (\<exists>x. P x)"
 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 367 | unfolding enum_ex_sum_def enum_ex | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 368 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 369 | qed (auto simp add: enum_UNIV enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) | 
| 26348 | 370 | |
| 371 | end | |
| 372 | ||
| 373 | primrec sublists :: "'a list \<Rightarrow> 'a list list" where | |
| 374 | "sublists [] = [[]]" | |
| 375 | | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)" | |
| 376 | ||
| 26444 | 377 | lemma length_sublists: | 
| 378 | "length (sublists xs) = Suc (Suc (0\<Colon>nat)) ^ length xs" | |
| 379 | by (induct xs) (simp_all add: Let_def) | |
| 380 | ||
| 26348 | 381 | lemma sublists_powset: | 
| 26444 | 382 | "set ` set (sublists xs) = Pow (set xs)" | 
| 26348 | 383 | proof - | 
| 384 | have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A" | |
| 385 | by (auto simp add: image_def) | |
| 26444 | 386 | have "set (map set (sublists xs)) = Pow (set xs)" | 
| 26348 | 387 | by (induct xs) | 
| 33639 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 hoelzl parents: 
33635diff
changeset | 388 | (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map) | 
| 26444 | 389 | then show ?thesis by simp | 
| 390 | qed | |
| 391 | ||
| 392 | lemma distinct_set_sublists: | |
| 393 | assumes "distinct xs" | |
| 394 | shows "distinct (map set (sublists xs))" | |
| 395 | proof (rule card_distinct) | |
| 396 | have "finite (set xs)" by rule | |
| 397 | then have "card (Pow (set xs)) = Suc (Suc 0) ^ card (set xs)" by (rule card_Pow) | |
| 398 | with assms distinct_card [of xs] | |
| 399 | have "card (Pow (set xs)) = Suc (Suc 0) ^ length xs" by simp | |
| 400 | then show "card (set (map set (sublists xs))) = length (map set (sublists xs))" | |
| 401 | by (simp add: sublists_powset length_sublists) | |
| 26348 | 402 | qed | 
| 403 | ||
| 404 | instantiation nibble :: enum | |
| 405 | begin | |
| 406 | ||
| 407 | definition | |
| 408 | "enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, | |
| 409 | Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" | |
| 410 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 411 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 412 | "enum_all P = (P Nibble0 \<and> P Nibble1 \<and> P Nibble2 \<and> P Nibble3 \<and> P Nibble4 \<and> P Nibble5 \<and> P Nibble6 \<and> P Nibble7 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 413 | \<and> P Nibble8 \<and> P Nibble9 \<and> P NibbleA \<and> P NibbleB \<and> P NibbleC \<and> P NibbleD \<and> P NibbleE \<and> P NibbleF)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 414 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 415 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 416 | "enum_ex P = (P Nibble0 \<or> P Nibble1 \<or> P Nibble2 \<or> P Nibble3 \<or> P Nibble4 \<or> P Nibble5 \<or> P Nibble6 \<or> P Nibble7 | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 417 | \<or> P Nibble8 \<or> P Nibble9 \<or> P NibbleA \<or> P NibbleB \<or> P NibbleC \<or> P NibbleD \<or> P NibbleE \<or> P NibbleF)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 418 | |
| 31464 | 419 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 420 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 421 | show "enum_all (P :: nibble \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 422 | unfolding enum_all_nibble_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 423 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 424 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 425 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 426 | show "enum_ex (P :: nibble \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 427 | unfolding enum_ex_nibble_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 428 | by (auto, case_tac x) auto | 
| 31464 | 429 | qed (simp_all add: enum_nibble_def UNIV_nibble) | 
| 26348 | 430 | |
| 431 | end | |
| 432 | ||
| 433 | instantiation char :: enum | |
| 434 | begin | |
| 435 | ||
| 436 | definition | |
| 37765 | 437 | "enum = map (split Char) (product enum enum)" | 
| 26444 | 438 | |
| 31482 | 439 | lemma enum_chars [code]: | 
| 440 | "enum = chars" | |
| 441 | unfolding enum_char_def chars_def enum_nibble_def by simp | |
| 26348 | 442 | |
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 443 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 444 | "enum_all P = list_all P chars" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 445 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 446 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 447 | "enum_ex P = list_ex P chars" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 448 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 449 | lemma set_enum_char: "set (enum :: char list) = UNIV" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 450 | by (auto intro: char.exhaust simp add: enum_char_def product_list_set enum_UNIV full_SetCompr_eq [symmetric]) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 451 | |
| 31464 | 452 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 453 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 454 | show "enum_all (P :: char \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 455 | unfolding enum_all_char_def enum_chars[symmetric] | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 456 | by (auto simp add: list_all_iff set_enum_char) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 457 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 458 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 459 | show "enum_ex (P :: char \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 460 | unfolding enum_ex_char_def enum_chars[symmetric] | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 461 | by (auto simp add: list_ex_iff set_enum_char) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 462 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 463 | show "distinct (enum :: char list)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 464 | by (auto intro: inj_onI simp add: enum_char_def product_list_set distinct_map distinct_product enum_distinct) | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 465 | qed (auto simp add: set_enum_char) | 
| 26348 | 466 | |
| 467 | end | |
| 468 | ||
| 29024 | 469 | instantiation option :: (enum) enum | 
| 470 | begin | |
| 471 | ||
| 472 | definition | |
| 473 | "enum = None # map Some enum" | |
| 474 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 475 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 476 | "enum_all P = (P None \<and> enum_all (%x. P (Some x)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 477 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 478 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 479 | "enum_ex P = (P None \<or> enum_ex (%x. P (Some x)))" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 480 | |
| 31464 | 481 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 482 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 483 | show "enum_all (P :: 'a option \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 484 | unfolding enum_all_option_def enum_all | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 485 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 486 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 487 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 488 | show "enum_ex (P :: 'a option \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 489 | unfolding enum_ex_option_def enum_ex | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 490 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 491 | qed (auto simp add: enum_UNIV enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct) | 
| 29024 | 492 | |
| 493 | end | |
| 494 | ||
| 40647 | 495 | subsection {* Small finite types *}
 | 
| 496 | ||
| 497 | text {* We define small finite types for the use in Quickcheck *}
 | |
| 498 | ||
| 499 | datatype finite_1 = a\<^isub>1 | |
| 500 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 501 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 502 | |
| 40647 | 503 | instantiation finite_1 :: enum | 
| 504 | begin | |
| 505 | ||
| 506 | definition | |
| 507 | "enum = [a\<^isub>1]" | |
| 508 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 509 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 510 | "enum_all P = P a\<^isub>1" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 511 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 512 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 513 | "enum_ex P = P a\<^isub>1" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 514 | |
| 40647 | 515 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 516 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 517 | show "enum_all (P :: finite_1 \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 518 | unfolding enum_all_finite_1_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 519 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 520 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 521 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 522 | show "enum_ex (P :: finite_1 \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 523 | unfolding enum_ex_finite_1_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 524 | by (auto, case_tac x) auto | 
| 40647 | 525 | qed (auto simp add: enum_finite_1_def intro: finite_1.exhaust) | 
| 526 | ||
| 29024 | 527 | end | 
| 40647 | 528 | |
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 529 | instantiation finite_1 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 530 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 531 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 532 | definition less_eq_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 533 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 534 | "less_eq_finite_1 x y = True" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 535 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 536 | definition less_finite_1 :: "finite_1 \<Rightarrow> finite_1 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 537 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 538 | "less_finite_1 x y = False" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 539 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 540 | instance | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 541 | apply (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 542 | apply (auto simp add: less_finite_1_def less_eq_finite_1_def) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 543 | apply (metis finite_1.exhaust) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 544 | done | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 545 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 546 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 547 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 548 | hide_const (open) a\<^isub>1 | 
| 40657 | 549 | |
| 40647 | 550 | datatype finite_2 = a\<^isub>1 | a\<^isub>2 | 
| 551 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 552 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 553 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 554 | |
| 40647 | 555 | instantiation finite_2 :: enum | 
| 556 | begin | |
| 557 | ||
| 558 | definition | |
| 559 | "enum = [a\<^isub>1, a\<^isub>2]" | |
| 560 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 561 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 562 | "enum_all P = (P a\<^isub>1 \<and> P a\<^isub>2)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 563 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 564 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 565 | "enum_ex P = (P a\<^isub>1 \<or> P a\<^isub>2)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 566 | |
| 40647 | 567 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 568 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 569 | show "enum_all (P :: finite_2 \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 570 | unfolding enum_all_finite_2_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 571 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 572 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 573 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 574 | show "enum_ex (P :: finite_2 \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 575 | unfolding enum_ex_finite_2_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 576 | by (auto, case_tac x) auto | 
| 40647 | 577 | qed (auto simp add: enum_finite_2_def intro: finite_2.exhaust) | 
| 578 | ||
| 579 | end | |
| 580 | ||
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 581 | instantiation finite_2 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 582 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 583 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 584 | definition less_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 585 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 586 | "less_finite_2 x y = ((x = a\<^isub>1) & (y = a\<^isub>2))" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 587 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 588 | definition less_eq_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 589 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 590 | "less_eq_finite_2 x y = ((x = y) \<or> (x < y))" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 591 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 592 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 593 | instance | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 594 | apply (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 595 | apply (auto simp add: less_finite_2_def less_eq_finite_2_def) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 596 | apply (metis finite_2.distinct finite_2.nchotomy)+ | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 597 | done | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 598 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 599 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 600 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 601 | hide_const (open) a\<^isub>1 a\<^isub>2 | 
| 40657 | 602 | |
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 603 | |
| 40647 | 604 | datatype finite_3 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | 
| 605 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 606 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 607 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 608 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 609 | |
| 40647 | 610 | instantiation finite_3 :: enum | 
| 611 | begin | |
| 612 | ||
| 613 | definition | |
| 614 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]" | |
| 615 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 616 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 617 | "enum_all P = (P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 618 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 619 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 620 | "enum_ex P = (P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 621 | |
| 40647 | 622 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 623 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 624 | show "enum_all (P :: finite_3 \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 625 | unfolding enum_all_finite_3_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 626 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 627 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 628 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 629 | show "enum_ex (P :: finite_3 \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 630 | unfolding enum_ex_finite_3_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 631 | by (auto, case_tac x) auto | 
| 40647 | 632 | qed (auto simp add: enum_finite_3_def intro: finite_3.exhaust) | 
| 633 | ||
| 634 | end | |
| 635 | ||
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 636 | instantiation finite_3 :: linorder | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 637 | begin | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 638 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 639 | definition less_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 640 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 641 | "less_finite_3 x y = (case x of a\<^isub>1 => (y \<noteq> a\<^isub>1) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 642 | | a\<^isub>2 => (y = a\<^isub>3)| a\<^isub>3 => False)" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 643 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 644 | definition less_eq_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 645 | where | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 646 | "less_eq_finite_3 x y = ((x = y) \<or> (x < y))" | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 647 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 648 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 649 | instance proof (intro_classes) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 650 | qed (auto simp add: less_finite_3_def less_eq_finite_3_def split: finite_3.split_asm) | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 651 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 652 | end | 
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 653 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 654 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 | 
| 40657 | 655 | |
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 656 | |
| 40647 | 657 | datatype finite_4 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | 
| 658 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 659 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 660 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 661 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 662 | notation (output) a\<^isub>4  ("a\<^isub>4")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 663 | |
| 40647 | 664 | instantiation finite_4 :: enum | 
| 665 | begin | |
| 666 | ||
| 667 | definition | |
| 668 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]" | |
| 669 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 670 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 671 | "enum_all P = (P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 672 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 673 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 674 | "enum_ex P = (P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 675 | |
| 40647 | 676 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 677 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 678 | show "enum_all (P :: finite_4 \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 679 | unfolding enum_all_finite_4_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 680 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 681 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 682 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 683 | show "enum_ex (P :: finite_4 \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 684 | unfolding enum_ex_finite_4_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 685 | by (auto, case_tac x) auto | 
| 40647 | 686 | qed (auto simp add: enum_finite_4_def intro: finite_4.exhaust) | 
| 687 | ||
| 688 | end | |
| 689 | ||
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 690 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 | 
| 40651 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 691 | |
| 
9752ba7348b5
adding code equation for function equality; adding some instantiations for the finite types
 bulwahn parents: 
40650diff
changeset | 692 | |
| 40647 | 693 | datatype finite_5 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | a\<^isub>5 | 
| 694 | ||
| 40900 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 695 | notation (output) a\<^isub>1  ("a\<^isub>1")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 696 | notation (output) a\<^isub>2  ("a\<^isub>2")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 697 | notation (output) a\<^isub>3  ("a\<^isub>3")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 698 | notation (output) a\<^isub>4  ("a\<^isub>4")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 699 | notation (output) a\<^isub>5  ("a\<^isub>5")
 | 
| 
1d5f76d79856
adding shorter output syntax for the finite types of quickcheck
 bulwahn parents: 
40898diff
changeset | 700 | |
| 40647 | 701 | instantiation finite_5 :: enum | 
| 702 | begin | |
| 703 | ||
| 704 | definition | |
| 705 | "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]" | |
| 706 | ||
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 707 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 708 | "enum_all P = (P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4 \<and> P a\<^isub>5)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 709 | |
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 710 | definition | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 711 | "enum_ex P = (P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4 \<or> P a\<^isub>5)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 712 | |
| 40647 | 713 | instance proof | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 714 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 715 | show "enum_all (P :: finite_5 \<Rightarrow> bool) = (\<forall>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 716 | unfolding enum_all_finite_5_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 717 | by (auto, case_tac x) auto | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 718 | next | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 719 | fix P | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 720 | show "enum_ex (P :: finite_5 \<Rightarrow> bool) = (\<exists>x. P x)" | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 721 | unfolding enum_ex_finite_5_def | 
| 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 722 | by (auto, case_tac x) auto | 
| 40647 | 723 | qed (auto simp add: enum_finite_5_def intro: finite_5.exhaust) | 
| 724 | ||
| 725 | end | |
| 726 | ||
| 41115 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 727 | subsection {* An executable THE operator on finite types *}
 | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 728 | |
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 729 | definition | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 730 | [code del]: "enum_the P = The P" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 731 | |
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 732 | lemma [code]: | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 733 | "The P = (case filter P enum of [x] => x | _ => enum_the P)" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 734 | proof - | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 735 |   {
 | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 736 | fix a | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 737 | assume filter_enum: "filter P enum = [a]" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 738 | have "The P = a" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 739 | proof (rule the_equality) | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 740 | fix x | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 741 | assume "P x" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 742 | show "x = a" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 743 | proof (rule ccontr) | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 744 | assume "x \<noteq> a" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 745 | from filter_enum obtain us vs | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 746 | where enum_eq: "enum = us @ [a] @ vs" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 747 | and "\<forall> x \<in> set us. \<not> P x" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 748 | and "\<forall> x \<in> set vs. \<not> P x" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 749 | and "P a" | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 750 | by (auto simp add: filter_eq_Cons_iff) (simp only: filter_empty_conv[symmetric]) | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 751 | with `P x` in_enum[of x, unfolded enum_eq] `x \<noteq> a` show "False" by auto | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 752 | qed | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 753 | next | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 754 | from filter_enum show "P a" by (auto simp add: filter_eq_Cons_iff) | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 755 | qed | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 756 | } | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 757 | from this show ?thesis | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 758 | unfolding enum_the_def by (auto split: list.split) | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 759 | qed | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 760 | |
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 761 | code_abort enum_the | 
| 
2c362ff5daf4
adding an executable THE operator on finite types
 bulwahn parents: 
41085diff
changeset | 762 | |
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 763 | hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5 | 
| 40657 | 764 | |
| 765 | ||
| 41085 
a549ff1d4070
adding a smarter enumeration scheme for finite functions
 bulwahn parents: 
41078diff
changeset | 766 | hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5 | 
| 41078 
051251fde456
adding more efficient implementations for quantifiers in Enum
 bulwahn parents: 
40900diff
changeset | 767 | hide_const (open) enum enum_all enum_ex n_lists all_n_lists ex_n_lists product | 
| 40647 | 768 | |
| 769 | end |