| author | oheimb | 
| Fri, 31 Jan 1997 15:54:00 +0100 | |
| changeset 2567 | 7a28e02e10b7 | 
| parent 2513 | d708d8cdc8e8 | 
| child 2834 | 9b47fc57ab7a | 
| permissions | -rw-r--r-- | 
| 1621 | 1 | (* Title: HOL/ex/Mutil | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 6 | The Mutilated Chess Board Problem, formalized inductively | 
| 1621 | 7 | *) | 
| 8 | ||
| 9 | open Mutil; | |
| 10 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 11 | Addsimps tiling.intrs; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 12 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 13 | (** The union of two disjoint tilings is a tiling **) | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 14 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 15 | goal thy "!!t. t: tiling A ==> \ | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2031diff
changeset | 16 | \ u: tiling A --> t <= Compl u --> t Un u : tiling A"; | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 17 | by (etac tiling.induct 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 18 | by (Simp_tac 1); | 
| 1820 | 19 | by (fast_tac (!claset addIs tiling.intrs | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2031diff
changeset | 20 | addss (!simpset addsimps [Un_assoc])) 1); | 
| 1911 | 21 | qed_spec_mp "tiling_UnI"; | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 22 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 23 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 24 | (*** Chess boards ***) | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 25 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 26 | val [below_0, below_Suc] = nat_recs below_def; | 
| 1911 | 27 | Addsimps [below_0, below_Suc]; | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 28 | |
| 1911 | 29 | goal thy "ALL i. (i: below k) = (i<k)"; | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 30 | by (nat_ind_tac "k" 1); | 
| 1911 | 31 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [less_Suc_eq]))); | 
| 1820 | 32 | by (Fast_tac 1); | 
| 1911 | 33 | qed_spec_mp "below_less_iff"; | 
| 34 | ||
| 35 | Addsimps [below_less_iff]; | |
| 1621 | 36 | |
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 37 | goal thy "below(Suc n) Times B = ({n} Times B) Un ((below n) Times B)";
 | 
| 1911 | 38 | by (Simp_tac 1); | 
| 39 | by (Fast_tac 1); | |
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 40 | qed "Sigma_Suc1"; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 41 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 42 | goal thy "A Times below(Suc n) = (A Times {n}) Un (A Times (below n))";
 | 
| 1911 | 43 | by (Simp_tac 1); | 
| 44 | by (Fast_tac 1); | |
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 45 | qed "Sigma_Suc2"; | 
| 1621 | 46 | |
| 1911 | 47 | (*Deletion is essential to allow use of Sigma_Suc1,2*) | 
| 48 | Delsimps [below_Suc]; | |
| 49 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 50 | goal thy "{i} Times below(n + n) : tiling domino";
 | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 51 | by (nat_ind_tac "n" 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 52 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Un_assoc RS sym, Sigma_Suc2]))); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 53 | by (resolve_tac tiling.intrs 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 54 | by (assume_tac 2); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 55 | by (subgoal_tac (*seems the easiest way of turning one to the other*) | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 56 |     "({i} Times {Suc(n1+n1)}) Un ({i} Times {n1+n1}) = \
 | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 57 | \    {(i, n1+n1), (i, Suc(n1+n1))}" 1);
 | 
| 1911 | 58 | by (Fast_tac 2); | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 59 | by (asm_simp_tac (!simpset addsimps [domino.horiz]) 1); | 
| 1820 | 60 | by (fast_tac (!claset addIs [equalityI, lessI] addEs [less_irrefl, less_asym] | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 61 | addDs [below_less_iff RS iffD1]) 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 62 | qed "dominoes_tile_row"; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 63 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 64 | goal thy "(below m) Times below(n + n) : tiling domino"; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 65 | by (nat_ind_tac "m" 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 66 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Sigma_Suc1]))); | 
| 1820 | 67 | by (fast_tac (!claset addIs [equalityI, tiling_UnI, dominoes_tile_row] | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 68 | addEs [below_less_iff RS iffD1 RS less_irrefl]) 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 69 | qed "dominoes_tile_matrix"; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 70 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 71 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 72 | (*** Basic properties of evnodd ***) | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 73 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 74 | goalw thy [evnodd_def] "(i,j): evnodd A b = ((i,j): A & (i+j) mod 2 = b)"; | 
| 1621 | 75 | by (Simp_tac 1); | 
| 76 | qed "evnodd_iff"; | |
| 77 | ||
| 78 | goalw thy [evnodd_def] "evnodd A b <= A"; | |
| 79 | by (rtac Int_lower1 1); | |
| 80 | qed "evnodd_subset"; | |
| 81 | ||
| 82 | (* finite X ==> finite(evnodd X b) *) | |
| 83 | bind_thm("finite_evnodd", evnodd_subset RS finite_subset);
 | |
| 84 | ||
| 85 | goalw thy [evnodd_def] "evnodd (A Un B) b = evnodd A b Un evnodd B b"; | |
| 1820 | 86 | by (Fast_tac 1); | 
| 1621 | 87 | qed "evnodd_Un"; | 
| 88 | ||
| 89 | goalw thy [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; | |
| 1820 | 90 | by (Fast_tac 1); | 
| 1621 | 91 | qed "evnodd_Diff"; | 
| 92 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 93 | goalw thy [evnodd_def] "evnodd {} b = {}";
 | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 94 | by (Simp_tac 1); | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 95 | qed "evnodd_empty"; | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 96 | |
| 1621 | 97 | goalw thy [evnodd_def] | 
| 98 | "evnodd (insert (i,j) C) b = \ | |
| 99 | \ (if (i+j) mod 2 = b then insert (i,j) (evnodd C b) else evnodd C b)"; | |
| 100 | by (asm_full_simp_tac (!simpset addsimps [evnodd_def] | |
| 1895 | 101 | setloop (split_tac [expand_if] THEN' Step_tac)) 1); | 
| 1621 | 102 | qed "evnodd_insert"; | 
| 103 | ||
| 104 | ||
| 105 | (*** Dominoes ***) | |
| 106 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 107 | goal thy "!!d. [| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}";
 | 
| 1621 | 108 | by (eresolve_tac [domino.elim] 1); | 
| 109 | by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 2);
 | |
| 110 | by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 1);
 | |
| 111 | by (REPEAT_FIRST assume_tac); | |
| 112 | (*Four similar cases: case (i+j) mod 2 = b, 2#-b, ...*) | |
| 1911 | 113 | by (REPEAT (asm_full_simp_tac (!simpset addsimps | 
| 114 | [less_Suc_eq, evnodd_insert, evnodd_empty, mod_Suc] | |
| 1621 | 115 | setloop split_tac [expand_if]) 1 | 
| 1820 | 116 | THEN Fast_tac 1)); | 
| 1621 | 117 | qed "domino_singleton"; | 
| 118 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 119 | goal thy "!!d. d:domino ==> finite d"; | 
| 1820 | 120 | by (fast_tac (!claset addSIs [finite_insertI, finite_emptyI] | 
| 1911 | 121 | addEs [domino.elim]) 1); | 
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 122 | qed "domino_finite"; | 
| 1621 | 123 | |
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 124 | |
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 125 | (*** Tilings of dominoes ***) | 
| 1621 | 126 | |
| 127 | goal thy "!!t. t:tiling domino ==> finite t"; | |
| 128 | by (eresolve_tac [tiling.induct] 1); | |
| 129 | by (rtac finite_emptyI 1); | |
| 1820 | 130 | by (fast_tac (!claset addIs [domino_finite, finite_UnI]) 1); | 
| 1621 | 131 | qed "tiling_domino_finite"; | 
| 132 | ||
| 133 | goal thy "!!t. t: tiling domino ==> card(evnodd t 0) = card(evnodd t 1)"; | |
| 134 | by (eresolve_tac [tiling.induct] 1); | |
| 135 | by (simp_tac (!simpset addsimps [evnodd_def]) 1); | |
| 136 | by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1);
 | |
| 137 | by (Simp_tac 2 THEN assume_tac 1); | |
| 138 | by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1);
 | |
| 139 | by (Simp_tac 2 THEN assume_tac 1); | |
| 1895 | 140 | by (Step_tac 1); | 
| 1621 | 141 | by (subgoal_tac "ALL p b. p : evnodd a b --> p ~: evnodd ta b" 1); | 
| 142 | by (asm_simp_tac (!simpset addsimps [evnodd_Un, Un_insert_left, | |
| 143 | tiling_domino_finite, | |
| 144 | evnodd_subset RS finite_subset, | |
| 145 | card_insert_disjoint]) 1); | |
| 1820 | 146 | by (fast_tac (!claset addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1); | 
| 1621 | 147 | qed "tiling_domino_0_1"; | 
| 148 | ||
| 1684 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 149 | goal thy "!!m n. [| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ | 
| 
3eaf3ab53082
Rearrangement and polishing to look for for publication
 paulson parents: 
1673diff
changeset | 150 | \                   t' = t - {(0,0)} - {(Suc(m+m), Suc(n+n))}              \
 | 
| 1621 | 151 | \ |] ==> t' ~: tiling domino"; | 
| 152 | by (rtac notI 1); | |
| 153 | by (dtac tiling_domino_0_1 1); | |
| 154 | by (subgoal_tac "card(evnodd t' 0) < card(evnodd t' 1)" 1); | |
| 155 | by (Asm_full_simp_tac 1); | |
| 156 | by (subgoal_tac "t : tiling domino" 1); | |
| 157 | (*Requires a small simpset that won't move the Suc applications*) | |
| 158 | by (asm_simp_tac (HOL_ss addsimps [dominoes_tile_matrix]) 2); | |
| 159 | by (subgoal_tac "(m+m)+(n+n) = (m+n)+(m+n)" 1); | |
| 160 | by (asm_simp_tac (!simpset addsimps add_ac) 2); | |
| 161 | by (asm_full_simp_tac | |
| 162 | (!simpset addsimps [evnodd_Diff, evnodd_insert, evnodd_empty, | |
| 163 | mod_less, tiling_domino_0_1 RS sym]) 1); | |
| 164 | by (rtac less_trans 1); | |
| 165 | by (REPEAT | |
| 166 | (rtac card_Diff 1 | |
| 167 | THEN | |
| 168 | asm_simp_tac (!simpset addsimps [tiling_domino_finite, finite_evnodd]) 1 | |
| 169 | THEN | |
| 1911 | 170 | asm_simp_tac (!simpset addsimps [mod_less, evnodd_iff]) 1)); | 
| 1621 | 171 | qed "mutil_not_tiling"; | 
| 172 |