| author | wenzelm | 
| Tue, 03 Apr 2012 20:37:52 +0200 | |
| changeset 47319 | 8aa23a259ab2 | 
| parent 47230 | 6584098d5378 | 
| child 51143 | 0a2371e7ced3 | 
| permissions | -rw-r--r-- | 
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 1 | (* Title: HOL/Quickcheck_Examples/Completeness.thy | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 2 | Author: Lukas Bulwahn | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 3 | Copyright 2012 TU Muenchen | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 4 | *) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 5 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 6 | header {* Proving completeness of exhaustive generators *}
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 7 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 8 | theory Completeness | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 9 | imports Main | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 10 | begin | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 11 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 12 | subsection {* Preliminaries *}
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 13 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 14 | primrec is_some | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 15 | where | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 16 | "is_some (Some t) = True" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 17 | | "is_some None = False" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 18 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 19 | setup {* Exhaustive_Generators.setup_exhaustive_datatype_interpretation *} 
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 20 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 21 | subsection {* Defining the size of values *}
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 22 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 23 | hide_const size | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 24 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 25 | class size = | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 26 | fixes size :: "'a => nat" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 27 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 28 | instantiation int :: size | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 29 | begin | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 30 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 31 | definition size_int :: "int => nat" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 32 | where | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 33 | "size_int n = nat (abs n)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 34 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 35 | instance .. | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 36 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 37 | end | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 38 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 39 | instantiation code_numeral :: size | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 40 | begin | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 41 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 42 | definition size_code_numeral :: "code_numeral => nat" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 43 | where | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 44 | "size_code_numeral = Code_Numeral.nat_of" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 45 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 46 | instance .. | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 47 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 48 | end | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 49 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 50 | instantiation nat :: size | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 51 | begin | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 52 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 53 | definition size_nat :: "nat => nat" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 54 | where | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 55 | "size_nat n = n" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 56 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 57 | instance .. | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 58 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 59 | end | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 60 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 61 | instantiation list :: (size) size | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 62 | begin | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 63 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 64 | primrec size_list :: "'a list => nat" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 65 | where | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 66 | "size [] = 1" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 67 | | "size (x # xs) = max (size x) (size xs) + 1" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 68 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 69 | instance .. | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 70 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 71 | end | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 72 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 73 | subsection {* Completeness *}
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 74 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 75 | class complete = exhaustive + size + | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 76 | assumes | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 77 | complete: "(\<exists> v. size v \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 78 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 79 | lemma complete_imp1: | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 80 | "size (v :: 'a :: complete) \<le> n \<Longrightarrow> is_some (f v) \<Longrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 81 | by (metis complete) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 82 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 83 | lemma complete_imp2: | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 84 | assumes "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 85 | obtains v where "size (v :: 'a :: complete) \<le> n" "is_some (f v)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 86 | using assms by (metis complete) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 87 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 88 | subsubsection {* Instance Proofs *}
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 89 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 90 | declare exhaustive'.simps [simp del] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 91 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 92 | lemma complete_exhaustive': | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 93 | "(\<exists> i. j <= i & i <= k & is_some (f i)) \<longleftrightarrow> is_some (Quickcheck_Exhaustive.exhaustive' f k j)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 94 | proof (induct rule: Quickcheck_Exhaustive.exhaustive'.induct[of _ f k j]) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 95 | case (1 f d i) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 96 | show ?case | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 97 | proof (cases "f i") | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 98 | case None | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 99 | from this 1 show ?thesis | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 100 | unfolding exhaustive'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 101 | apply auto | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 102 | apply (metis is_some.simps(2) order_le_neq_trans zless_imp_add1_zle) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 103 | apply (metis add1_zle_eq less_int_def) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 104 | done | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 105 | next | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 106 | case Some | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 107 | from this show ?thesis | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 108 | unfolding exhaustive'.simps[of _ _ i] Quickcheck_Exhaustive.orelse_def by auto | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 109 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 110 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 111 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 112 | lemma int_of_nat: | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 113 | "Code_Numeral.int_of (Code_Numeral.of_nat n) = int n" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 114 | unfolding int_of_def by simp | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 115 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 116 | instance int :: complete | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 117 | proof | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 118 | fix n f | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 119 | show "(\<exists>v. size (v :: int) \<le> n \<and> is_some (f v)) = is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 120 | unfolding exhaustive_int_def complete_exhaustive'[symmetric] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 121 | apply auto apply (rule_tac x="v" in exI) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 122 | unfolding size_int_def by (metis int_of_nat abs_le_iff minus_le_iff nat_le_iff)+ | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 123 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 124 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 125 | declare exhaustive_code_numeral'.simps[simp del] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 126 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 127 | lemma complete_code_numeral': | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 128 | "(\<exists>n. j \<le> n \<and> n \<le> k \<and> is_some (f n)) = | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 129 | is_some (Quickcheck_Exhaustive.exhaustive_code_numeral' f k j)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 130 | proof (induct rule: exhaustive_code_numeral'.induct[of _ f k j]) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 131 | case (1 f k j) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 132 | show "(\<exists>n\<ge>j. n \<le> k \<and> is_some (f n)) = is_some (Quickcheck_Exhaustive.exhaustive_code_numeral' f k j)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 133 | unfolding exhaustive_code_numeral'.simps[of f k j] Quickcheck_Exhaustive.orelse_def | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 134 | apply auto | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 135 | apply (auto split: option.split) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 136 | apply (insert 1[symmetric]) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 137 | apply simp | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 138 | apply (metis is_some.simps(2) less_eq_code_numeral_def not_less_eq_eq order_antisym) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 139 | apply simp | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 140 | apply (split option.split_asm) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 141 | defer apply fastforce | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 142 | apply simp by (metis Suc_leD) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 143 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 144 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 145 | instance code_numeral :: complete | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 146 | proof | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 147 | fix n f | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 148 | show "(\<exists>v. size (v :: code_numeral) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 149 | unfolding exhaustive_code_numeral_def complete_code_numeral'[symmetric] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 150 | by (auto simp add: size_code_numeral_def) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 151 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 152 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 153 | instance nat :: complete | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 154 | proof | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 155 | fix n f | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 156 | show "(\<exists>v. size (v :: nat) \<le> n \<and> is_some (f v)) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 157 | unfolding exhaustive_nat_def complete[of n "%x. f (Code_Numeral.nat_of x)", symmetric] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 158 | apply auto | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 159 | apply (rule_tac x="Code_Numeral.of_nat v" in exI) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 160 | apply (auto simp add: size_code_numeral_def size_nat_def) done | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 161 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 162 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 163 | instance list :: (complete) complete | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 164 | proof | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 165 | fix n f | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 166 | show "(\<exists> v. size (v :: 'a list) \<le> n \<and> is_some (f (v :: 'a list))) \<longleftrightarrow> is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 167 | proof (induct n arbitrary: f) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 168 | case 0 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 169 |     { fix v have "size (v :: 'a list) > 0" by (induct v) auto }
 | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 170 | from this show ?case by (simp add: list.exhaustive_list.simps) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 171 | next | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 172 | case (Suc n) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 173 | show ?case | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 174 | proof | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 175 | assume "\<exists>v. Completeness.size_class.size v \<le> Suc n \<and> is_some (f v)" | 
| 47230 | 176 | then obtain v where v: "size v \<le> Suc n" "is_some (f v)" by blast | 
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 177 | show "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat (Suc n)))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 178 | proof (cases "v") | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 179 | case Nil | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 180 | from this v show ?thesis | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 181 | unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 182 | by (auto split: option.split) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 183 | next | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 184 | case (Cons v' vs') | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 185 | from Cons v have size_v': "Completeness.size_class.size v' \<le> n" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 186 | and "Completeness.size_class.size vs' \<le> n" by auto | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 187 | from Suc v Cons this have "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 188 | by metis | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 189 | from complete_imp1[OF size_v', of "%x. (exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (Code_Numeral.of_nat n))", OF this] | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 190 | show ?thesis | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 191 | unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 192 | by (auto split: option.split) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 193 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 194 | next | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 195 | assume is_some: "is_some (exhaustive_class.exhaustive f (Code_Numeral.of_nat (Suc n)))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 196 | show "\<exists>v. size v \<le> Suc n \<and> is_some (f v)" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 197 | proof (cases "f []") | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 198 | case Some | 
| 47230 | 199 | then show ?thesis | 
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 200 | by (metis Suc_eq_plus1 is_some.simps(1) le_add2 size_list.simps(1)) | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 201 | next | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 202 | case None | 
| 47230 | 203 | with is_some have | 
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 204 | "is_some (exhaustive_class.exhaustive (\<lambda>x. exhaustive_class.exhaustive (\<lambda>xs. f (x # xs)) (Code_Numeral.of_nat n)) (Code_Numeral.of_nat n))" | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 205 | unfolding list.exhaustive_list.simps[of _ "Code_Numeral.of_nat (Suc n)"] Quickcheck_Exhaustive.orelse_def | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 206 | by simp | 
| 47230 | 207 | then obtain v' where | 
| 208 | v: "size v' \<le> n" | |
| 209 | "is_some (exhaustive_class.exhaustive (\<lambda>xs. f (v' # xs)) (Code_Numeral.of_nat n))" | |
| 210 | by (rule complete_imp2) | |
| 211 | with Suc[of "%xs. f (v' # xs)"] | |
| 212 | obtain vs' where vs': "size vs' \<le> n" "is_some (f (v' # vs'))" | |
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 213 | by metis | 
| 47230 | 214 | with v have "size (v' # vs') \<le> Suc n" by auto | 
| 215 | with vs' v show ?thesis by metis | |
| 47205 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 216 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 217 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 218 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 219 | qed | 
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 220 | |
| 
34e8b7347dda
adding theory to prove completeness of the exhaustive generators
 bulwahn parents: diff
changeset | 221 | end |