author | wenzelm |
Sat, 18 Aug 2007 13:32:28 +0200 | |
changeset 24325 | 5c29e8822f50 |
parent 23746 | a455e69c31cc |
child 28718 | ef16499edaab |
permissions | -rw-r--r-- |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
1 |
(* Title: HOL/Induct/Mutil.thy |
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 |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
7 |
header {* The Mutilated Chess Board Problem *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
8 |
|
16417 | 9 |
theory Mutil imports Main begin |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
10 |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
11 |
text {* |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
12 |
The Mutilated Chess Board Problem, formalized inductively. |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
13 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
14 |
Originator is Max Black, according to J A Robinson. Popularized as |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
15 |
the Mutilated Checkerboard Problem by J McCarthy. |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
16 |
*} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
17 |
|
23746 | 18 |
inductive_set |
19 |
tiling :: "'a set set => 'a set set" |
|
20 |
for A :: "'a set set" |
|
21 |
where |
|
11637 | 22 |
empty [simp, intro]: "{} \<in> tiling A" |
23746 | 23 |
| Un [simp, intro]: "[| a \<in> A; t \<in> tiling A; a \<inter> t = {} |] |
14276 | 24 |
==> a \<union> t \<in> tiling A" |
8399 | 25 |
|
23746 | 26 |
inductive_set |
27 |
domino :: "(nat \<times> nat) set set" |
|
28 |
where |
|
11637 | 29 |
horiz [simp]: "{(i, j), (i, Suc j)} \<in> domino" |
23746 | 30 |
| vertl [simp]: "{(i, j), (Suc i, j)} \<in> domino" |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
31 |
|
14276 | 32 |
text {* \medskip Sets of squares of the given colour*} |
33 |
||
19736 | 34 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
35 |
coloured :: "nat => (nat \<times> nat) set" where |
19736 | 36 |
"coloured b = {(i, j). (i + j) mod 2 = b}" |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
37 |
|
19736 | 38 |
abbreviation |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
39 |
whites :: "(nat \<times> nat) set" where |
19736 | 40 |
"whites == coloured 0" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
41 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
42 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
43 |
blacks :: "(nat \<times> nat) set" where |
19736 | 44 |
"blacks == coloured (Suc 0)" |
14276 | 45 |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
46 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
47 |
text {* \medskip The union of two disjoint tilings is a tiling *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
48 |
|
14276 | 49 |
lemma tiling_UnI [intro]: |
50 |
"[|t \<in> tiling A; u \<in> tiling A; t \<inter> u = {} |] ==> t \<union> u \<in> tiling A" |
|
51 |
apply (induct set: tiling) |
|
52 |
apply (auto simp add: Un_assoc) |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
53 |
done |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
54 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
55 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
56 |
text {* \medskip Chess boards *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
57 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
58 |
lemma Sigma_Suc1 [simp]: |
14276 | 59 |
"lessThan (Suc n) \<times> B = ({n} \<times> B) \<union> ((lessThan n) \<times> B)" |
23506 | 60 |
by auto |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
61 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
62 |
lemma Sigma_Suc2 [simp]: |
14276 | 63 |
"A \<times> lessThan (Suc n) = (A \<times> {n}) \<union> (A \<times> (lessThan n))" |
23506 | 64 |
by auto |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
65 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
66 |
lemma sing_Times_lemma: "({i} \<times> {n}) \<union> ({i} \<times> {m}) = {(i, m), (i, n)}" |
14276 | 67 |
by auto |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
68 |
|
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
69 |
lemma dominoes_tile_row [intro!]: "{i} \<times> lessThan (2 * n) \<in> tiling domino" |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
70 |
apply (induct n) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
71 |
apply (simp_all add: Un_assoc [symmetric]) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
72 |
apply (rule tiling.Un) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
73 |
apply (auto simp add: sing_Times_lemma) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
74 |
done |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
75 |
|
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
76 |
lemma dominoes_tile_matrix: "(lessThan m) \<times> lessThan (2 * n) \<in> tiling domino" |
18242 | 77 |
by (induct m) auto |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
78 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
79 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
80 |
text {* \medskip @{term coloured} and Dominoes *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
81 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
82 |
lemma coloured_insert [simp]: |
14276 | 83 |
"coloured b \<inter> (insert (i, j) t) = |
84 |
(if (i + j) mod 2 = b then insert (i, j) (coloured b \<inter> t) |
|
85 |
else coloured b \<inter> t)" |
|
86 |
by (auto simp add: coloured_def) |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
87 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
88 |
lemma domino_singletons: |
14276 | 89 |
"d \<in> domino ==> |
90 |
(\<exists>i j. whites \<inter> d = {(i, j)}) \<and> |
|
91 |
(\<exists>m n. blacks \<inter> d = {(m, n)})"; |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
92 |
apply (erule domino.cases) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
93 |
apply (auto simp add: mod_Suc) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
94 |
done |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
95 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
96 |
lemma domino_finite [simp]: "d \<in> domino ==> finite d" |
14276 | 97 |
by (erule domino.cases, auto) |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
98 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
99 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
100 |
text {* \medskip Tilings of dominoes *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
101 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
102 |
lemma tiling_domino_finite [simp]: "t \<in> tiling domino ==> finite t" |
18242 | 103 |
by (induct set: tiling) auto |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
104 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
105 |
declare |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
106 |
Int_Un_distrib [simp] |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
107 |
Diff_Int_distrib [simp] |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
108 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
109 |
lemma tiling_domino_0_1: |
14276 | 110 |
"t \<in> tiling domino ==> card(whites \<inter> t) = card(blacks \<inter> t)" |
12171 | 111 |
apply (induct set: tiling) |
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
112 |
apply (drule_tac [2] domino_singletons) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
113 |
apply auto |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
114 |
apply (subgoal_tac "\<forall>p C. C \<inter> a = {p} --> p \<notin> t") |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
115 |
-- {* this lemma tells us that both ``inserts'' are non-trivial *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
116 |
apply (simp (no_asm_simp)) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
117 |
apply blast |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
118 |
done |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
119 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
120 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
121 |
text {* \medskip Final argument is surprisingly complex *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
122 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
123 |
theorem gen_mutil_not_tiling: |
14276 | 124 |
"t \<in> tiling domino ==> |
125 |
(i + j) mod 2 = 0 ==> (m + n) mod 2 = 0 ==> |
|
126 |
{(i, j), (m, n)} \<subseteq> t |
|
127 |
==> (t - {(i, j)} - {(m, n)}) \<notin> tiling domino" |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
128 |
apply (rule notI) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
129 |
apply (subgoal_tac |
14276 | 130 |
"card (whites \<inter> (t - {(i, j)} - {(m, n)})) < |
131 |
card (blacks \<inter> (t - {(i, j)} - {(m, n)}))") |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
132 |
apply (force simp only: tiling_domino_0_1) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
133 |
apply (simp add: tiling_domino_0_1 [symmetric]) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
134 |
apply (simp add: coloured_def card_Diff2_less) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
135 |
done |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
136 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
137 |
text {* Apply the general theorem to the well-known case *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
138 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
139 |
theorem mutil_not_tiling: |
14276 | 140 |
"t = lessThan (2 * Suc m) \<times> lessThan (2 * Suc n) |
141 |
==> t - {(0, 0)} - {(Suc (2 * m), Suc (2 * n))} \<notin> tiling domino" |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
142 |
apply (rule gen_mutil_not_tiling) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
143 |
apply (blast intro!: dominoes_tile_matrix) |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
144 |
apply auto |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
9930
diff
changeset
|
145 |
done |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
146 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
147 |
end |