src/HOL/Induct/Mutil.thy
author huffman
Sat, 16 Sep 2006 02:35:58 +0200
changeset 20551 ba543692bfa1
parent 19736 d8d0f8f51d69
child 21404 eb85850d3eb7
permissions -rw-r--r--
add theorem norm_diff_triangle_ineq
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14276
diff changeset
     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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    18
consts tiling :: "'a set set => 'a set set"
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    19
inductive "tiling A"
11637
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    20
  intros
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    21
    empty [simp, intro]: "{} \<in> tiling A"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    22
    Un [simp, intro]:    "[| a \<in> A; t \<in> tiling A; a \<inter> t = {} |] 
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    23
                          ==> a \<union> t \<in> tiling A"
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    24
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    25
consts domino :: "(nat \<times> nat) set set"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    26
inductive domino
11637
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    27
  intros
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    28
    horiz [simp]: "{(i, j), (i, Suc j)} \<in> domino"
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    29
    vertl [simp]: "{(i, j), (Suc i, j)} \<in> domino"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    30
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    31
text {* \medskip Sets of squares of the given colour*}
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    32
19736
wenzelm
parents: 18242
diff changeset
    33
definition
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    34
  coloured :: "nat => (nat \<times> nat) set"
19736
wenzelm
parents: 18242
diff changeset
    35
  "coloured b = {(i, j). (i + j) mod 2 = b}"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    36
19736
wenzelm
parents: 18242
diff changeset
    37
abbreviation
wenzelm
parents: 18242
diff changeset
    38
  whites  :: "(nat \<times> nat) set"
wenzelm
parents: 18242
diff changeset
    39
  "whites == coloured 0"
wenzelm
parents: 18242
diff changeset
    40
  blacks  :: "(nat \<times> nat) set"
wenzelm
parents: 18242
diff changeset
    41
  "blacks == coloured (Suc 0)"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    42
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    43
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    44
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
    45
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    46
lemma tiling_UnI [intro]:
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    47
     "[|t \<in> tiling A; u \<in> tiling A; t \<inter> u = {} |] ==>  t \<union> u \<in> tiling A"
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    48
  apply (induct set: tiling)
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    49
  apply (auto simp add: Un_assoc)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    50
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    51
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    52
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    53
text {* \medskip Chess boards *}
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
lemma Sigma_Suc1 [simp]:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    56
     "lessThan (Suc n) \<times> B = ({n} \<times> B) \<union> ((lessThan n) \<times> B)"
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    57
  by (auto simp add: lessThan_def)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    58
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    59
lemma Sigma_Suc2 [simp]:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    60
     "A \<times> lessThan (Suc n) = (A \<times> {n}) \<union> (A \<times> (lessThan n))"
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    61
  by (auto simp add: lessThan_def)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    62
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    63
lemma sing_Times_lemma: "({i} \<times> {n}) \<union> ({i} \<times> {m}) = {(i, m), (i, n)}"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    64
  by auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    65
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    66
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
    67
  apply (induct n)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    68
   apply (simp_all add: Un_assoc [symmetric])
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    69
  apply (rule tiling.Un)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    70
    apply (auto simp add: sing_Times_lemma)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    71
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    72
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    73
lemma dominoes_tile_matrix: "(lessThan m) \<times> lessThan (2 * n) \<in> tiling domino"
18242
2215049cd29c tuned induct proofs;
wenzelm
parents: 16417
diff changeset
    74
  by (induct m) auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    75
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    76
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    77
text {* \medskip @{term coloured} and Dominoes *}
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
lemma coloured_insert [simp]:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    80
     "coloured b \<inter> (insert (i, j) t) =
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    81
      (if (i + j) mod 2 = b then insert (i, j) (coloured b \<inter> t)
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    82
       else coloured b \<inter> t)"
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    83
  by (auto simp add: coloured_def)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    84
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    85
lemma domino_singletons:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    86
     "d \<in> domino ==>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    87
       (\<exists>i j. whites \<inter> d = {(i, j)}) \<and>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    88
       (\<exists>m n. blacks \<inter> d = {(m, n)})";
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    89
  apply (erule domino.cases)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    90
   apply (auto simp add: mod_Suc)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    91
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    92
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    93
lemma domino_finite [simp]: "d \<in> domino ==> finite d"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    94
  by (erule domino.cases, auto)
11046
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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    97
text {* \medskip Tilings of dominoes *}
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
lemma tiling_domino_finite [simp]: "t \<in> tiling domino ==> finite t"
18242
2215049cd29c tuned induct proofs;
wenzelm
parents: 16417
diff changeset
   100
  by (induct set: tiling) auto
11046
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
declare
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   103
  Int_Un_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   104
  Diff_Int_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   105
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   106
lemma tiling_domino_0_1:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   107
     "t \<in> tiling domino ==> card(whites \<inter> t) = card(blacks \<inter> t)"
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11704
diff changeset
   108
  apply (induct set: tiling)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   109
   apply (drule_tac [2] domino_singletons)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   110
   apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   111
  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
   112
    -- {* this lemma tells us that both ``inserts'' are non-trivial *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   113
   apply (simp (no_asm_simp))
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   114
  apply blast
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   115
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   116
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   117
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   118
text {* \medskip Final argument is surprisingly complex *}
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
theorem gen_mutil_not_tiling:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   121
       "t \<in> tiling domino ==>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   122
	 (i + j) mod 2 = 0 ==> (m + n) mod 2 = 0 ==>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   123
	 {(i, j), (m, n)} \<subseteq> t
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   124
       ==> (t - {(i, j)} - {(m, n)}) \<notin> tiling domino"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   125
  apply (rule notI)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   126
  apply (subgoal_tac
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   127
          "card (whites \<inter> (t - {(i, j)} - {(m, n)})) <
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   128
           card (blacks \<inter> (t - {(i, j)} - {(m, n)}))")
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   129
   apply (force simp only: tiling_domino_0_1)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   130
  apply (simp add: tiling_domino_0_1 [symmetric])
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   131
  apply (simp add: coloured_def card_Diff2_less)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   132
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   133
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   134
text {* Apply the general theorem to the well-known case *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   135
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   136
theorem mutil_not_tiling:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   137
       "t = lessThan (2 * Suc m) \<times> lessThan (2 * Suc n)
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   138
	 ==> 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
   139
  apply (rule gen_mutil_not_tiling)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   140
     apply (blast intro!: dominoes_tile_matrix)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   141
    apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   142
  done
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   143
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   144
end