1 (* Title: HOL/ex/Mutil |
|
2 ID: $Id$ |
|
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 Copyright 1996 University of Cambridge |
|
5 |
|
6 The Mutilated Chess Board Problem, formalized inductively |
|
7 *) |
|
8 |
|
9 open Mutil; |
|
10 |
|
11 Addsimps tiling.intrs; |
|
12 |
|
13 (** The union of two disjoint tilings is a tiling **) |
|
14 |
|
15 goal thy "!!t. t: tiling A ==> \ |
|
16 \ u: tiling A --> t <= Compl u --> t Un u : tiling A"; |
|
17 by (etac tiling.induct 1); |
|
18 by (Simp_tac 1); |
|
19 by (simp_tac (!simpset addsimps [Un_assoc]) 1); |
|
20 by (blast_tac (!claset addIs tiling.intrs) 1); |
|
21 qed_spec_mp "tiling_UnI"; |
|
22 |
|
23 |
|
24 (*** Chess boards ***) |
|
25 |
|
26 val [below_0, below_Suc] = nat_recs below_def; |
|
27 Addsimps [below_0, below_Suc]; |
|
28 |
|
29 goal thy "ALL i. (i: below k) = (i<k)"; |
|
30 by (nat_ind_tac "k" 1); |
|
31 by (ALLGOALS (asm_simp_tac (!simpset addsimps [less_Suc_eq]))); |
|
32 by (Blast_tac 1); |
|
33 qed_spec_mp "below_less_iff"; |
|
34 |
|
35 Addsimps [below_less_iff]; |
|
36 |
|
37 goal thy "below(Suc n) Times B = ({n} Times B) Un ((below n) Times B)"; |
|
38 by (Simp_tac 1); |
|
39 by (Blast_tac 1); |
|
40 qed "Sigma_Suc1"; |
|
41 |
|
42 goal thy "A Times below(Suc n) = (A Times {n}) Un (A Times (below n))"; |
|
43 by (Simp_tac 1); |
|
44 by (Blast_tac 1); |
|
45 qed "Sigma_Suc2"; |
|
46 |
|
47 (*Deletion is essential to allow use of Sigma_Suc1,2*) |
|
48 Delsimps [below_Suc]; |
|
49 |
|
50 goal thy "{i} Times below(n + n) : tiling domino"; |
|
51 by (nat_ind_tac "n" 1); |
|
52 by (ALLGOALS (asm_simp_tac (!simpset addsimps [Un_assoc RS sym, Sigma_Suc2]))); |
|
53 by (resolve_tac tiling.intrs 1); |
|
54 by (assume_tac 2); |
|
55 by (subgoal_tac (*seems the easiest way of turning one to the other*) |
|
56 "({i} Times {Suc(n+n)}) Un ({i} Times {n+n}) = \ |
|
57 \ {(i, n+n), (i, Suc(n+n))}" 1); |
|
58 by (Blast_tac 2); |
|
59 by (asm_simp_tac (!simpset addsimps [domino.horiz]) 1); |
|
60 by (blast_tac (!claset addEs [less_irrefl, less_asym] |
|
61 addSDs [below_less_iff RS iffD1]) 1); |
|
62 qed "dominoes_tile_row"; |
|
63 |
|
64 goal thy "(below m) Times below(n + n) : tiling domino"; |
|
65 by (nat_ind_tac "m" 1); |
|
66 by (ALLGOALS (asm_simp_tac (!simpset addsimps [Sigma_Suc1]))); |
|
67 by (blast_tac (!claset addSIs [tiling_UnI, dominoes_tile_row] |
|
68 addSEs [below_less_iff RS iffD1 RS less_irrefl]) 1); |
|
69 qed "dominoes_tile_matrix"; |
|
70 |
|
71 |
|
72 (*** Basic properties of evnodd ***) |
|
73 |
|
74 goalw thy [evnodd_def] "(i,j): evnodd A b = ((i,j): A & (i+j) mod 2 = b)"; |
|
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"; |
|
86 by (Blast_tac 1); |
|
87 qed "evnodd_Un"; |
|
88 |
|
89 goalw thy [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; |
|
90 by (Blast_tac 1); |
|
91 qed "evnodd_Diff"; |
|
92 |
|
93 goalw thy [evnodd_def] "evnodd {} b = {}"; |
|
94 by (Simp_tac 1); |
|
95 qed "evnodd_empty"; |
|
96 |
|
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] |
|
101 setloop (split_tac [expand_if] THEN' Step_tac)) 1); |
|
102 qed "evnodd_insert"; |
|
103 |
|
104 |
|
105 (*** Dominoes ***) |
|
106 |
|
107 goal thy "!!d. [| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}"; |
|
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, ...*) |
|
113 by (REPEAT (asm_full_simp_tac (!simpset addsimps |
|
114 [less_Suc_eq, evnodd_insert, evnodd_empty, mod_Suc] |
|
115 setloop split_tac [expand_if]) 1 |
|
116 THEN Blast_tac 1)); |
|
117 qed "domino_singleton"; |
|
118 |
|
119 goal thy "!!d. d:domino ==> finite d"; |
|
120 by (blast_tac (!claset addSIs [finite_insertI, finite_emptyI] |
|
121 addSEs [domino.elim]) 1); |
|
122 qed "domino_finite"; |
|
123 |
|
124 |
|
125 (*** Tilings of dominoes ***) |
|
126 |
|
127 goal thy "!!t. t:tiling domino ==> finite t"; |
|
128 by (eresolve_tac [tiling.induct] 1); |
|
129 by (rtac finite_emptyI 1); |
|
130 by (blast_tac (!claset addSIs [finite_UnI] addIs [domino_finite]) 1); |
|
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); |
|
140 by (Step_tac 1); |
|
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); |
|
146 by (blast_tac (!claset addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1); |
|
147 qed "tiling_domino_0_1"; |
|
148 |
|
149 goal thy "!!m n. [| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ |
|
150 \ t' = t - {(0,0)} - {(Suc(m+m), Suc(n+n))} \ |
|
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 |
|
170 asm_simp_tac (!simpset addsimps [mod_less, evnodd_iff]) 1)); |
|
171 qed "mutil_not_tiling"; |
|
172 |
|