author | paulson |
Tue, 10 Jan 2023 11:06:20 +0000 | |
changeset 76942 | c732fa27b60f |
parent 76215 | a642599ffdea |
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 |
||
60770 | 6 |
section \<open>The Mutilated Chess Board Problem, formalized inductively\<close> |
12173 | 7 |
|
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
60770
diff
changeset
|
8 |
theory Mutil imports ZF begin |
12173 | 9 |
|
60770 | 10 |
text \<open> |
12173 | 11 |
Originator is Max Black, according to J A Robinson. Popularized as |
12 |
the Mutilated Checkerboard Problem by J McCarthy. |
|
60770 | 13 |
\<close> |
12173 | 14 |
|
12088 | 15 |
consts |
12173 | 16 |
domino :: i |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
17 |
tiling :: "i \<Rightarrow> i" |
12173 | 18 |
|
19 |
inductive |
|
20 |
domains "domino" \<subseteq> "Pow(nat \<times> nat)" |
|
21 |
intros |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
22 |
horiz: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> {\<langle>i,j\<rangle>, <i,succ(j)>} \<in> domino" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
23 |
vertl: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> {\<langle>i,j\<rangle>, <succ(i),j>} \<in> domino" |
12173 | 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:
35762
diff
changeset
|
27 |
domains "tiling(A)" \<subseteq> "Pow(\<Union>(A))" |
12173 | 28 |
intros |
29 |
empty: "0 \<in> tiling(A)" |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
30 |
Un: "\<lbrakk>a \<in> A; t \<in> tiling(A); a \<inter> t = 0\<rbrakk> \<Longrightarrow> 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 |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
35 |
evnodd :: "[i, i] \<Rightarrow> i" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
36 |
"evnodd(A,b) \<equiv> {z \<in> A. \<exists>i j. z = \<langle>i,j\<rangle> \<and> (i #+ j) mod 2 = b}" |
12173 | 37 |
|
38 |
||
60770 | 39 |
subsection \<open>Basic properties of evnodd\<close> |
12173 | 40 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
41 |
lemma evnodd_iff: "\<langle>i,j\<rangle>: evnodd(A,b) \<longleftrightarrow> \<langle>i,j\<rangle>: A \<and> (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 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
47 |
lemma Finite_evnodd: "Finite(X) \<Longrightarrow> Finite(evnodd(X,b))" |
12173 | 48 |
by (rule lepoll_Finite, rule subset_imp_lepoll, rule evnodd_subset) |
49 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
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]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
57 |
"evnodd(cons(\<langle>i,j\<rangle>,C), b) = |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
58 |
(if (i#+j) mod 2 = b then cons(\<langle>i,j\<rangle>, evnodd(C,b)) else evnodd(C,b))" |
12173 | 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 |
||
60770 | 65 |
subsection \<open>Dominoes\<close> |
12173 | 66 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
67 |
lemma domino_Finite: "d \<in> domino \<Longrightarrow> Finite(d)" |
12173 | 68 |
by (blast intro!: Finite_cons Finite_0 elim: domino.cases) |
69 |
||
12185 | 70 |
lemma domino_singleton: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
71 |
"\<lbrakk>d \<in> domino; b<2\<rbrakk> \<Longrightarrow> \<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 |
||
60770 | 81 |
subsection \<open>Tilings\<close> |
12173 | 82 |
|
60770 | 83 |
text \<open>The union of two disjoint tilings is a tiling\<close> |
12173 | 84 |
|
85 |
lemma tiling_UnI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
86 |
"t \<in> tiling(A) \<Longrightarrow> u \<in> tiling(A) \<Longrightarrow> t \<inter> u = 0 \<Longrightarrow> 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 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
93 |
lemma tiling_domino_Finite: "t \<in> tiling(domino) \<Longrightarrow> 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 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
99 |
lemma tiling_domino_0_1: "t \<in> tiling(domino) \<Longrightarrow> |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:
35762
diff
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: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
118 |
"\<lbrakk>i \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> {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:
35762
diff
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: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
134 |
"\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> 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 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
141 |
lemma eq_lt_E: "\<lbrakk>x=y; x<y\<rbrakk> \<Longrightarrow> P" |
12173 | 142 |
by auto |
143 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
144 |
theorem mutil_not_tiling: "\<lbrakk>m \<in> nat; n \<in> nat; |
12173 | 145 |
t = (succ(m)#+succ(m))*(succ(n)#+succ(n)); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
146 |
t' = t - {\<langle>0,0\<rangle>} - {<succ(m#+m), succ(n#+n)>}\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
147 |
\<Longrightarrow> t' \<notin> tiling(domino)" |
12173 | 148 |
apply (rule notI) |
149 |
apply (drule tiling_domino_0_1) |
|
59788 | 150 |
apply (erule_tac x = "|A|" for A in eq_lt_E) |
12173 | 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 |