author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 9930 | c02d48a47ed1 |
child 10375 | d943898cc3a9 |
permissions | -rw-r--r-- |
3423 | 1 |
(* Title: HOL/Induct/Mutil |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
4 |
Copyright 1996 University of Cambridge |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
5 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
6 |
The Mutilated Chess Board Problem, formalized inductively |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
7 |
*) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
8 |
|
8589 | 9 |
Addsimps (tiling.intrs @ domino.intrs); |
8399 | 10 |
AddIs tiling.intrs; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
11 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
12 |
(** The union of two disjoint tilings is a tiling **) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
13 |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
14 |
Goal "t: tiling A ==> u: tiling A --> t Int u = {} --> t Un u : tiling A"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
15 |
by (etac tiling.induct 1); |
8399 | 16 |
by (simp_tac (simpset() addsimps [Un_assoc]) 2); |
17 |
by Auto_tac; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
18 |
qed_spec_mp "tiling_UnI"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
19 |
|
8399 | 20 |
AddIs [tiling_UnI]; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
21 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
22 |
(*** Chess boards ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
23 |
|
8950 | 24 |
Goalw [lessThan_def] |
25 |
"lessThan(Suc n) <*> B = ({n} <*> B) Un ((lessThan n) <*> B)"; |
|
7401 | 26 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
27 |
qed "Sigma_Suc1"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
28 |
|
8950 | 29 |
Goalw [lessThan_def] |
30 |
"A <*> lessThan(Suc n) = (A <*> {n}) Un (A <*> (lessThan n))"; |
|
7401 | 31 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
32 |
qed "Sigma_Suc2"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
33 |
|
8399 | 34 |
Addsimps [Sigma_Suc1, Sigma_Suc2]; |
35 |
||
8876
f797816932d7
reverted to old proof of dominoes_tile_row, given new treatment of #2+...
paulson
parents:
8781
diff
changeset
|
36 |
Goal "({i} <*> {n}) Un ({i} <*> {m}) = {(i,m), (i,n)}"; |
8399 | 37 |
by Auto_tac; |
38 |
qed "sing_Times_lemma"; |
|
39 |
||
8950 | 40 |
Goal "{i} <*> lessThan(#2*n) : tiling domino"; |
5184 | 41 |
by (induct_tac "n" 1); |
8876
f797816932d7
reverted to old proof of dominoes_tile_row, given new treatment of #2+...
paulson
parents:
8781
diff
changeset
|
42 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_assoc RS sym]))); |
8558
6c4860b1828d
now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents:
8553
diff
changeset
|
43 |
by (rtac tiling.Un 1); |
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
44 |
by (auto_tac (claset(), simpset() addsimps [sing_Times_lemma])); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
45 |
qed "dominoes_tile_row"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
46 |
|
8399 | 47 |
AddSIs [dominoes_tile_row]; |
48 |
||
8950 | 49 |
Goal "(lessThan m) <*> lessThan(#2*n) : tiling domino"; |
5184 | 50 |
by (induct_tac "m" 1); |
8399 | 51 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
52 |
qed "dominoes_tile_matrix"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
53 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
54 |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
55 |
(*** "coloured" and Dominoes ***) |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
56 |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
57 |
Goalw [coloured_def] |
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
58 |
"coloured b Int (insert (i,j) t) = \ |
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
59 |
\ (if (i+j) mod #2 = b then insert (i,j) (coloured b Int t) \ |
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
60 |
\ else coloured b Int t)"; |
8399 | 61 |
by Auto_tac; |
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
62 |
qed "coloured_insert"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
63 |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
64 |
Addsimps [coloured_insert]; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
65 |
|
8399 | 66 |
Goal "d:domino ==> finite d"; |
67 |
by (etac domino.elim 1); |
|
68 |
by Auto_tac; |
|
69 |
qed "domino_finite"; |
|
70 |
Addsimps [domino_finite]; |
|
71 |
||
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
72 |
Goal "d:domino ==> (EX i j. coloured 0 Int d = {(i,j)}) & \ |
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
73 |
\ (EX k l. coloured 1 Int d = {(k,l)})"; |
8399 | 74 |
by (etac domino.elim 1); |
75 |
by (auto_tac (claset(), simpset() addsimps [mod_Suc])); |
|
8589 | 76 |
qed "domino_singletons"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
77 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
78 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
79 |
(*** Tilings of dominoes ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
80 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
81 |
Goal "t:tiling domino ==> finite t"; |
8558
6c4860b1828d
now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents:
8553
diff
changeset
|
82 |
by (etac tiling.induct 1); |
8399 | 83 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
84 |
qed "tiling_domino_finite"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
85 |
|
8589 | 86 |
Addsimps [tiling_domino_finite, Int_Un_distrib, Diff_Int_distrib]; |
8399 | 87 |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
88 |
Goal "t: tiling domino ==> card(coloured 0 Int t) = card(coloured 1 Int t)"; |
8558
6c4860b1828d
now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents:
8553
diff
changeset
|
89 |
by (etac tiling.induct 1); |
8589 | 90 |
by (dtac domino_singletons 2); |
7401 | 91 |
by Auto_tac; |
8399 | 92 |
(*this lemma tells us that both "inserts" are non-trivial*) |
8589 | 93 |
by (subgoal_tac "ALL p C. C Int a = {p} --> p ~: t" 1); |
8399 | 94 |
by (Asm_simp_tac 1); |
9188 | 95 |
by (Blast_tac 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
96 |
qed "tiling_domino_0_1"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
97 |
|
8399 | 98 |
(*Final argument is surprisingly complex*) |
99 |
Goal "[| t : tiling domino; \ |
|
8781 | 100 |
\ (i+j) mod #2 = 0; (k+l) mod #2 = 0; \ |
8589 | 101 |
\ {(i,j),(k,l)} <= t; l ~= j |] \ |
102 |
\ ==> (t - {(i,j)} - {(k,l)}) ~: tiling domino"; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
103 |
by (rtac notI 1); |
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
104 |
by (subgoal_tac "card (coloured 0 Int (t - {(i,j)} - {(k,l)})) < \ |
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
105 |
\ card (coloured 1 Int (t - {(i,j)} - {(k,l)}))" 1); |
8399 | 106 |
by (force_tac (claset(), HOL_ss addsimps [tiling_domino_0_1]) 1); |
107 |
by (asm_simp_tac (simpset() addsimps [tiling_domino_0_1 RS sym]) 1); |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
108 |
by (rtac less_trans 1); |
8399 | 109 |
by (ALLGOALS (force_tac (claset() addSIs [card_Diff1_less], |
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
110 |
simpset() addsimps [coloured_def]))); |
8399 | 111 |
qed "gen_mutil_not_tiling"; |
112 |
||
113 |
(*Apply the general theorem to the well-known case*) |
|
9930
c02d48a47ed1
tidying and updating for revised Mutilated Chess Board article
paulson
parents:
9188
diff
changeset
|
114 |
Goal "t = lessThan(#2 * Suc m) <*> lessThan(#2 * Suc n) \ |
8781 | 115 |
\ ==> t - {(0,0)} - {(Suc(#2*m), Suc(#2*n))} ~: tiling domino"; |
8399 | 116 |
by (rtac gen_mutil_not_tiling 1); |
117 |
by (blast_tac (claset() addSIs [dominoes_tile_matrix]) 1); |
|
118 |
by Auto_tac; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
119 |
qed "mutil_not_tiling"; |
8399 | 120 |