| author | wenzelm | 
| Wed, 15 May 2013 20:28:43 +0200 | |
| changeset 52008 | bdb82afdcb92 | 
| parent 46822 | 95f1e700b712 | 
| child 58871 | c399ae4b836f | 
| permissions | -rw-r--r-- | 
| 12173 | 1 | (* Title: ZF/Induct/Mutil.thy | 
| 12088 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1996 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 12173 | 6 | header {* The Mutilated Chess Board Problem, formalized inductively *}
 | 
| 7 | ||
| 16417 | 8 | theory Mutil imports Main begin | 
| 12173 | 9 | |
| 10 | text {*
 | |
| 11 | Originator is Max Black, according to J A Robinson. Popularized as | |
| 12 | the Mutilated Checkerboard Problem by J McCarthy. | |
| 13 | *} | |
| 14 | ||
| 12088 | 15 | consts | 
| 12173 | 16 | domino :: i | 
| 17 | tiling :: "i => i" | |
| 18 | ||
| 19 | inductive | |
| 20 | domains "domino" \<subseteq> "Pow(nat \<times> nat)" | |
| 21 | intros | |
| 22 |     horiz: "[| i \<in> nat;  j \<in> nat |] ==> {<i,j>, <i,succ(j)>} \<in> domino"
 | |
| 23 |     vertl: "[| i \<in> nat;  j \<in> nat |] ==> {<i,j>, <succ(i),j>} \<in> domino"
 | |
| 24 | type_intros empty_subsetI cons_subsetI PowI SigmaI nat_succI | |
| 12088 | 25 | |
| 26 | inductive | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 27 | domains "tiling(A)" \<subseteq> "Pow(\<Union>(A))" | 
| 12173 | 28 | intros | 
| 29 | empty: "0 \<in> tiling(A)" | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 30 | Un: "[| a \<in> A; t \<in> tiling(A); a \<inter> t = 0 |] ==> a \<union> t \<in> tiling(A)" | 
| 12173 | 31 | type_intros empty_subsetI Union_upper Un_least PowI | 
| 32 | type_elims PowD [elim_format] | |
| 33 | ||
| 24893 | 34 | definition | 
| 35 | evnodd :: "[i, i] => i" where | |
| 12173 | 36 |   "evnodd(A,b) == {z \<in> A. \<exists>i j. z = <i,j> \<and> (i #+ j) mod 2 = b}"
 | 
| 37 | ||
| 38 | ||
| 12185 | 39 | subsection {* Basic properties of evnodd *}
 | 
| 12173 | 40 | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 41 | lemma evnodd_iff: "<i,j>: evnodd(A,b) \<longleftrightarrow> <i,j>: A & (i#+j) mod 2 = b" | 
| 12173 | 42 | by (unfold evnodd_def) blast | 
| 43 | ||
| 44 | lemma evnodd_subset: "evnodd(A, b) \<subseteq> A" | |
| 45 | by (unfold evnodd_def) blast | |
| 46 | ||
| 47 | lemma Finite_evnodd: "Finite(X) ==> Finite(evnodd(X,b))" | |
| 48 | by (rule lepoll_Finite, rule subset_imp_lepoll, rule evnodd_subset) | |
| 49 | ||
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 50 | lemma evnodd_Un: "evnodd(A \<union> B, b) = evnodd(A,b) \<union> evnodd(B,b)" | 
| 12173 | 51 | by (simp add: evnodd_def Collect_Un) | 
| 52 | ||
| 53 | lemma evnodd_Diff: "evnodd(A - B, b) = evnodd(A,b) - evnodd(B,b)" | |
| 54 | by (simp add: evnodd_def Collect_Diff) | |
| 55 | ||
| 56 | lemma evnodd_cons [simp]: | |
| 57 | "evnodd(cons(<i,j>,C), b) = | |
| 58 | (if (i#+j) mod 2 = b then cons(<i,j>, evnodd(C,b)) else evnodd(C,b))" | |
| 59 | by (simp add: evnodd_def Collect_cons) | |
| 60 | ||
| 61 | lemma evnodd_0 [simp]: "evnodd(0, b) = 0" | |
| 62 | by (simp add: evnodd_def) | |
| 63 | ||
| 64 | ||
| 12185 | 65 | subsection {* Dominoes *}
 | 
| 12173 | 66 | |
| 67 | lemma domino_Finite: "d \<in> domino ==> Finite(d)" | |
| 68 | by (blast intro!: Finite_cons Finite_0 elim: domino.cases) | |
| 69 | ||
| 12185 | 70 | lemma domino_singleton: | 
| 71 |     "[| d \<in> domino; b<2 |] ==> \<exists>i' j'. evnodd(d,b) = {<i',j'>}"
 | |
| 12173 | 72 | apply (erule domino.cases) | 
| 73 | apply (rule_tac [2] k1 = "i#+j" in mod2_cases [THEN disjE]) | |
| 74 | apply (rule_tac k1 = "i#+j" in mod2_cases [THEN disjE]) | |
| 75 | apply (rule add_type | assumption)+ | |
| 76 | (*Four similar cases: case (i#+j) mod 2 = b, 2#-b, ...*) | |
| 77 | apply (auto simp add: mod_succ succ_neq_self dest: ltD) | |
| 78 | done | |
| 12088 | 79 | |
| 80 | ||
| 12185 | 81 | subsection {* Tilings *}
 | 
| 12173 | 82 | |
| 83 | text {* The union of two disjoint tilings is a tiling *}
 | |
| 84 | ||
| 85 | lemma tiling_UnI: | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 86 | "t \<in> tiling(A) ==> u \<in> tiling(A) ==> t \<inter> u = 0 ==> t \<union> u \<in> tiling(A)" | 
| 12173 | 87 | apply (induct set: tiling) | 
| 88 | apply (simp add: tiling.intros) | |
| 89 | apply (simp add: Un_assoc subset_empty_iff [THEN iff_sym]) | |
| 90 | apply (blast intro: tiling.intros) | |
| 91 | done | |
| 92 | ||
| 93 | lemma tiling_domino_Finite: "t \<in> tiling(domino) ==> Finite(t)" | |
| 18415 | 94 | apply (induct set: tiling) | 
| 12173 | 95 | apply (rule Finite_0) | 
| 96 | apply (blast intro!: Finite_Un intro: domino_Finite) | |
| 97 | done | |
| 98 | ||
| 99 | lemma tiling_domino_0_1: "t \<in> tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|" | |
| 18415 | 100 | apply (induct set: tiling) | 
| 12173 | 101 | apply (simp add: evnodd_def) | 
| 102 | apply (rule_tac b1 = 0 in domino_singleton [THEN exE]) | |
| 103 | prefer 2 | |
| 104 | apply simp | |
| 105 | apply assumption | |
| 106 | apply (rule_tac b1 = 1 in domino_singleton [THEN exE]) | |
| 107 | prefer 2 | |
| 108 | apply simp | |
| 109 | apply assumption | |
| 110 | apply safe | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 111 | apply (subgoal_tac "\<forall>p b. p \<in> evnodd (a,b) \<longrightarrow> p\<notin>evnodd (t,b)") | 
| 12173 | 112 | apply (simp add: evnodd_Un Un_cons tiling_domino_Finite | 
| 113 | evnodd_subset [THEN subset_Finite] Finite_imp_cardinal_cons) | |
| 114 | apply (blast dest!: evnodd_subset [THEN subsetD] elim: equalityE) | |
| 115 | done | |
| 12088 | 116 | |
| 12185 | 117 | lemma dominoes_tile_row: | 
| 118 |     "[| i \<in> nat;  n \<in> nat |] ==> {i} * (n #+ n) \<in> tiling(domino)"
 | |
| 12173 | 119 | apply (induct_tac n) | 
| 120 | apply (simp add: tiling.intros) | |
| 121 | apply (simp add: Un_assoc [symmetric] Sigma_succ2) | |
| 122 | apply (rule tiling.intros) | |
| 123 | prefer 2 apply assumption | |
| 124 | apply (rename_tac n') | |
| 125 | apply (subgoal_tac (*seems the easiest way of turning one to the other*) | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
35762diff
changeset | 126 |      "{i}*{succ (n'#+n') } \<union> {i}*{n'#+n'} =
 | 
| 12185 | 127 |        {<i,n'#+n'>, <i,succ (n'#+n') >}")
 | 
| 12173 | 128 | prefer 2 apply blast | 
| 129 | apply (simp add: domino.horiz) | |
| 130 | apply (blast elim: mem_irrefl mem_asym) | |
| 131 | done | |
| 132 | ||
| 12185 | 133 | lemma dominoes_tile_matrix: | 
| 134 | "[| m \<in> nat; n \<in> nat |] ==> m * (n #+ n) \<in> tiling(domino)" | |
| 12173 | 135 | apply (induct_tac m) | 
| 136 | apply (simp add: tiling.intros) | |
| 137 | apply (simp add: Sigma_succ1) | |
| 138 | apply (blast intro: tiling_UnI dominoes_tile_row elim: mem_irrefl) | |
| 139 | done | |
| 140 | ||
| 141 | lemma eq_lt_E: "[| x=y; x<y |] ==> P" | |
| 142 | by auto | |
| 143 | ||
| 144 | theorem mutil_not_tiling: "[| m \<in> nat; n \<in> nat; | |
| 145 | t = (succ(m)#+succ(m))*(succ(n)#+succ(n)); | |
| 146 |          t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>} |]
 | |
| 147 | ==> t' \<notin> tiling(domino)" | |
| 148 | apply (rule notI) | |
| 149 | apply (drule tiling_domino_0_1) | |
| 150 | apply (erule_tac x = "|?A|" in eq_lt_E) | |
| 151 | apply (subgoal_tac "t \<in> tiling (domino)") | |
| 152 | prefer 2 (*Requires a small simpset that won't move the succ applications*) | |
| 153 | apply (simp only: nat_succI add_type dominoes_tile_matrix) | |
| 12185 | 154 | apply (simp add: evnodd_Diff mod2_add_self mod2_succ_succ | 
| 155 | tiling_domino_0_1 [symmetric]) | |
| 12173 | 156 | apply (rule lt_trans) | 
| 157 | apply (rule Finite_imp_cardinal_Diff, | |
| 158 | simp add: tiling_domino_Finite Finite_evnodd Finite_Diff, | |
| 159 | simp add: evnodd_iff nat_0_le [THEN ltD] mod2_add_self)+ | |
| 160 | done | |
| 12088 | 161 | |
| 162 | end |