src/HOL/Induct/Mutil.thy
author wenzelm
Sat, 03 Feb 2001 17:40:16 +0100
changeset 11046 b5f5942781a0
parent 9930 c02d48a47ed1
child 11464 ddea204de5bc
permissions -rw-r--r--
Induct: converted some theories to new-style format;
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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
     9
theory Mutil = Main:
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"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    20
  intros [simp, intro]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    21
    empty: "{} \<in> tiling A"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    22
    Un: "a \<in> A ==> t \<in> tiling A ==> a \<inter> t = {} ==> a \<union> t \<in> tiling A"
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    23
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    24
consts domino :: "(nat \<times> nat) set set"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    25
inductive domino
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    26
  intros [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    27
    horiz: "{(i, j), (i, Suc j)} \<in> domino"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    28
    vertl: "{(i, j), (Suc i, j)} \<in> domino"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    29
8340
paulson
parents: 5931
diff changeset
    30
constdefs
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    31
  coloured :: "nat => (nat \<times> nat) set"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    32
   "coloured b == {(i, j). (i + j) mod #2 = b}"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    33
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    34
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    35
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
    36
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    37
lemma tiling_UnI [rule_format, intro]:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    38
  "t \<in> tiling A ==> u \<in> tiling A --> t \<inter> u = {} --> t \<union> u \<in> tiling A"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    39
  apply (erule tiling.induct)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    40
   prefer 2
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    41
   apply (simp add: Un_assoc)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    42
  apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    43
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    44
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    45
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    46
text {* \medskip Chess boards *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    47
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    48
lemma Sigma_Suc1 [simp]:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    49
  "lessThan (Suc n) \<times> B = ({n} \<times> B) \<union> ((lessThan n) \<times> B)"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    50
  apply (unfold lessThan_def)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    51
  apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    52
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    53
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    54
lemma Sigma_Suc2 [simp]:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    55
  "A \<times> lessThan (Suc n) = (A \<times> {n}) \<union> (A \<times> (lessThan n))"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    56
  apply (unfold lessThan_def)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    57
  apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    58
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    59
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    60
lemma sing_Times_lemma: "({i} \<times> {n}) \<union> ({i} \<times> {m}) = {(i, m), (i, n)}"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    61
  apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    62
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    63
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    64
lemma dominoes_tile_row [intro!]: "{i} \<times> lessThan (#2 * n) \<in> tiling domino"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    65
  apply (induct n)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    66
   apply (simp_all add: Un_assoc [symmetric])
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    67
  apply (rule tiling.Un)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    68
    apply (auto simp add: sing_Times_lemma)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    69
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    70
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    71
lemma dominoes_tile_matrix: "(lessThan m) \<times> lessThan (#2 * n) \<in> tiling domino"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    72
  apply (induct m)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    73
   apply auto
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
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]:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    80
  "coloured b \<inter> (insert (i, j) t) =
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    81
   (if (i + j) mod #2 = b then insert (i, j) (coloured b \<inter> t)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    82
    else coloured b \<inter> t)"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    83
  apply (unfold coloured_def)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    84
  apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    85
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    86
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    87
lemma domino_singletons:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    88
  "d \<in> domino ==>
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    89
    (\<exists>i j. coloured 0 \<inter> d = {(i, j)}) \<and>
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    90
    (\<exists>m n. coloured 1 \<inter> d = {(m, n)})"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    91
  apply (erule domino.cases)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    92
   apply (auto simp add: mod_Suc)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    93
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    94
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    95
lemma domino_finite [simp]: "d \<in> domino ==> finite d"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    96
  apply (erule domino.cases)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    97
   apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    98
  done
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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   101
text {* \medskip Tilings of dominoes *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   102
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   103
lemma tiling_domino_finite [simp]: "t \<in> tiling domino ==> finite t"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   104
  apply (erule tiling.induct)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   105
   apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   106
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   107
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   108
declare
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   109
  Int_Un_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   110
  Diff_Int_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   111
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   112
lemma tiling_domino_0_1:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   113
  "t \<in> tiling domino ==> card (coloured 0 \<inter> t) = card (coloured 1 \<inter> t)"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   114
  apply (erule tiling.induct)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   115
   apply (drule_tac [2] domino_singletons)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   116
   apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   117
  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
   118
    -- {* this lemma tells us that both ``inserts'' are non-trivial *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   119
   apply (simp (no_asm_simp))
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   120
  apply blast
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   121
  done
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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   124
text {* \medskip Final argument is surprisingly complex *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   125
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   126
theorem gen_mutil_not_tiling:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   127
  "t \<in> tiling domino ==>
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   128
    (i + j) mod #2 = 0 ==> (m + n) mod #2 = 0 ==>
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   129
    {(i, j), (m, n)} \<subseteq> t
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   130
  ==> (t - {(i, j)} - {(m, n)}) \<notin> tiling domino"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   131
  apply (rule notI)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   132
  apply (subgoal_tac
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   133
    "card (coloured 0 \<inter> (t - {(i, j)} - {(m, n)})) <
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   134
      card (coloured 1 \<inter> (t - {(i, j)} - {(m, n)}))")
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   135
   apply (force simp only: tiling_domino_0_1)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   136
  apply (simp add: tiling_domino_0_1 [symmetric])
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   137
  apply (simp add: coloured_def card_Diff2_less)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   138
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   139
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   140
text {* Apply the general theorem to the well-known case *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   141
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   142
theorem mutil_not_tiling:
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   143
  "t = lessThan (#2 * Suc m) \<times> lessThan (#2 * Suc n)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   144
    ==> t - {(0, 0)} - {(Suc (#2 * m), Suc (#2 * n))} \<notin> tiling domino"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   145
  apply (rule gen_mutil_not_tiling)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   146
     apply (blast intro!: dominoes_tile_matrix)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   147
    apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   148
  done
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   149
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   150
end