src/HOL/Induct/Mutil.thy
author nipkow
Thu, 06 Nov 2008 11:52:50 +0100
changeset 28720 a08c37b478b2
parent 28718 ef16499edaab
permissions -rw-r--r--
tuned
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
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
     4
                Tobias Nipkow, TUM (part 2)
3120
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
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    11
subsection{* The Mutilated Chess Board Cannot be Tiled by Dominoes *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    12
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    13
text {*
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    14
  The Mutilated Chess Board Problem, formalized inductively.
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    15
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    16
  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
    17
  the Mutilated Checkerboard Problem by J McCarthy.
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    18
*}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    19
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    20
inductive_set
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    21
  tiling :: "'a set set => 'a set set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    22
  for A :: "'a set set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    23
  where
11637
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    24
    empty [simp, intro]: "{} \<in> tiling A"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    25
  | Un [simp, intro]:    "[| a \<in> A; t \<in> tiling A; a \<inter> t = {} |] 
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    26
                          ==> a \<union> t \<in> tiling A"
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    27
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    28
inductive_set
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    29
  domino :: "(nat \<times> nat) set set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    30
  where
11637
647e6c84323c inductive: no collective atts;
wenzelm
parents: 11464
diff changeset
    31
    horiz [simp]: "{(i, j), (i, Suc j)} \<in> domino"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 23506
diff changeset
    32
  | vertl [simp]: "{(i, j), (Suc i, j)} \<in> domino"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    33
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    34
text {* \medskip Sets of squares of the given colour*}
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    35
19736
wenzelm
parents: 18242
diff changeset
    36
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    37
  coloured :: "nat => (nat \<times> nat) set" where
19736
wenzelm
parents: 18242
diff changeset
    38
  "coloured b = {(i, j). (i + j) mod 2 = b}"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    39
19736
wenzelm
parents: 18242
diff changeset
    40
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    41
  whites  :: "(nat \<times> nat) set" where
19736
wenzelm
parents: 18242
diff changeset
    42
  "whites == coloured 0"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    43
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    44
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    45
  blacks  :: "(nat \<times> nat) set" where
19736
wenzelm
parents: 18242
diff changeset
    46
  "blacks == coloured (Suc 0)"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    47
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    48
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    49
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
    50
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    51
lemma tiling_UnI [intro]:
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    52
     "[|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
    53
  apply (induct set: tiling)
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    54
  apply (auto simp add: Un_assoc)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    55
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    56
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    57
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    58
lemma tiling_Diff1E [intro]:
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    59
assumes "t-a \<in> tiling A"  "a \<in> A"  "a \<subseteq> t"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    60
shows "t \<in> tiling A"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    61
proof -
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    62
  from assms(2-3) have  "EX r. t = r Un a & r Int a = {}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    63
    by (metis Diff_disjoint Int_commute Un_Diff_cancel Un_absorb1 Un_commute)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    64
  thus ?thesis using assms(1,2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    65
    by (auto simp:Un_Diff)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    66
       (metis Compl_Diff_eq Diff_Compl Diff_empty Int_commute Un_Diff_cancel
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    67
              Un_commute double_complement tiling.Un)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    68
qed
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    69
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    70
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    71
text {* \medskip Chess boards *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    72
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    73
lemma Sigma_Suc1 [simp]:
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    74
     "{0..< Suc n} \<times> B = ({n} \<times> B) \<union> ({0..<n} \<times> B)"
23506
332a9f5c7c29 simplified
paulson
parents: 21404
diff changeset
    75
  by auto
11046
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
lemma Sigma_Suc2 [simp]:
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    78
     "A \<times> {0..< Suc n} = (A \<times> {n}) \<union> (A \<times> {0..<n})"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    79
  by auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    80
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    81
lemma dominoes_tile_row [intro!]: "{i} \<times> {0..< 2*n} \<in> tiling domino"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    82
apply (induct n)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    83
apply (simp_all del:Un_insert_left add: Un_assoc [symmetric])
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    84
done
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    85
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
    86
lemma dominoes_tile_matrix: "{0..<m} \<times> {0..< 2*n} \<in> tiling domino"
18242
2215049cd29c tuned induct proofs;
wenzelm
parents: 16417
diff changeset
    87
  by (induct m) auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    88
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    89
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    90
text {* \medskip @{term coloured} and Dominoes *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    91
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    92
lemma coloured_insert [simp]:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    93
     "coloured b \<inter> (insert (i, j) t) =
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    94
      (if (i + j) mod 2 = b then insert (i, j) (coloured b \<inter> t)
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    95
       else coloured b \<inter> t)"
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    96
  by (auto simp add: coloured_def)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    97
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
    98
lemma domino_singletons:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
    99
     "d \<in> domino ==>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   100
       (\<exists>i j. whites \<inter> d = {(i, j)}) \<and>
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   101
       (\<exists>m n. blacks \<inter> d = {(m, n)})";
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   102
  apply (erule domino.cases)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   103
   apply (auto simp add: mod_Suc)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   104
  done
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 domino_finite [simp]: "d \<in> domino ==> finite d"
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   107
  by (erule domino.cases, auto)
11046
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
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   110
text {* \medskip Tilings of dominoes *}
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_finite [simp]: "t \<in> tiling domino ==> finite t"
18242
2215049cd29c tuned induct proofs;
wenzelm
parents: 16417
diff changeset
   113
  by (induct set: tiling) auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   114
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   115
declare
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   116
  Int_Un_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   117
  Diff_Int_distrib [simp]
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   118
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   119
lemma tiling_domino_0_1:
14276
950c12139016 stylistic changes
paulson
parents: 12171
diff changeset
   120
     "t \<in> tiling domino ==> card(whites \<inter> t) = card(blacks \<inter> t)"
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11704
diff changeset
   121
  apply (induct set: tiling)
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   122
   apply (drule_tac [2] domino_singletons)
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   123
   apply auto
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   124
  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
   125
    -- {* this lemma tells us that both ``inserts'' are non-trivial *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   126
   apply (simp (no_asm_simp))
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   127
  apply blast
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   128
  done
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   129
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   130
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   131
text {* \medskip Final argument is surprisingly complex *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   132
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   133
theorem gen_mutil_not_tiling:
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   134
  "t \<in> tiling domino ==>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   135
  (i + j) mod 2 = 0 ==> (m + n) mod 2 = 0 ==>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   136
  {(i, j), (m, n)} \<subseteq> t
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   137
  ==> (t - {(i,j)} - {(m,n)}) \<notin> tiling domino"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   138
apply (rule notI)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   139
apply (subgoal_tac
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   140
  "card (whites \<inter> (t - {(i,j)} - {(m,n)})) <
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   141
   card (blacks \<inter> (t - {(i,j)} - {(m,n)}))")
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   142
 apply (force simp only: tiling_domino_0_1)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   143
apply (simp add: tiling_domino_0_1 [symmetric])
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   144
apply (simp add: coloured_def card_Diff2_less)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   145
done
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   146
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   147
text {* Apply the general theorem to the well-known case *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   148
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 9930
diff changeset
   149
theorem mutil_not_tiling:
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   150
  "t = {0..< 2 * Suc m} \<times> {0..< 2 * Suc n}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   151
   ==> t - {(0,0)} - {(Suc(2 * m), Suc(2 * n))} \<notin> tiling domino"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   152
apply (rule gen_mutil_not_tiling)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   153
 apply (blast intro!: dominoes_tile_matrix)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   154
apply auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   155
done
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   156
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   157
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   158
subsection{* The Mutilated Chess Board Can be Tiled by Ls *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   159
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   160
text{* We remove any square from a chess board of size $2^n \times 2^n$.
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   161
The result can be tiled by L-shaped tiles.
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   162
Found in Velleman's \emph{How to Prove it}.
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   163
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   164
The four possible L-shaped tiles are obtained by dropping
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   165
one of the four squares from $\{(x,y),(x+1,y),(x,y+1),(x+1,y+1)\}$: *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   166
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   167
definition "L2 (x::nat) (y::nat) = {(x,y), (x+1,y), (x, y+1)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   168
definition "L3 (x::nat) (y::nat) = {(x,y), (x+1,y), (x+1, y+1)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   169
definition "L0 (x::nat) (y::nat) = {(x+1,y), (x,y+1), (x+1, y+1)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   170
definition "L1 (x::nat) (y::nat) = {(x,y), (x,y+1), (x+1, y+1)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   171
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   172
text{* All tiles: *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   173
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   174
definition Ls :: "(nat * nat) set set" where
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   175
"Ls \<equiv> { L0 x y | x y. True} \<union> { L1 x y | x y. True} \<union>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   176
      { L2 x y | x y. True} \<union> { L3 x y | x y. True}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   177
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   178
lemma LinLs: "L0 i j : Ls & L1 i j : Ls & L2 i j : Ls & L3 i j : Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   179
by(fastsimp simp:Ls_def)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   180
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   181
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   182
text{* Square $2^n \times 2^n$ grid, shifted by $i$ and $j$: *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   183
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   184
definition "square2 (n::nat) (i::nat) (j::nat) = {i..< 2^n+i} \<times> {j..< 2^n+j}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   185
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   186
lemma in_square2[simp]:
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   187
  "(a,b) : square2 n i j \<longleftrightarrow> i\<le>a \<and> a<2^n+i \<and> j\<le>b \<and> b<2^n+j"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   188
by(simp add:square2_def)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   189
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   190
lemma square2_Suc: "square2 (Suc n) i j =
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   191
  square2 n i j \<union> square2 n (2^n + i) j \<union> square2 n i (2^n + j) \<union>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   192
  square2 n (2^n + i) (2^n + j)"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   193
by(auto simp:square2_def)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   194
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   195
lemma square2_disj: "square2 n i j \<inter> square2 n x y = {} \<longleftrightarrow>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   196
  (2^n+i \<le> x \<or> 2^n+x \<le> i) \<or> (2^n+j \<le> y \<or> 2^n+y \<le> j)" (is "?A = ?B")
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   197
proof-
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   198
  { assume ?B hence ?A by(auto simp:square2_def) }
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   199
  moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   200
  { assume "\<not> ?B"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   201
    hence "(max i x, max j y) : square2 n i j \<inter> square2 n x y" by simp
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   202
    hence "\<not> ?A" by blast }
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   203
  ultimately show ?thesis by blast
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   204
qed
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   205
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   206
text{* Some specific lemmas: *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   207
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   208
lemma pos_pow2: "(0::nat) < 2^(n::nat)"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   209
by simp
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   210
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   211
declare nat_zero_less_power_iff[simp del] zero_less_power[simp del]
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   212
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   213
lemma Diff_insert_if: shows
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   214
  "B \<noteq> {} \<Longrightarrow> a:A \<Longrightarrow> A - insert a B = (A-B - {a})" and
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   215
  "B \<noteq> {} \<Longrightarrow> a ~: A \<Longrightarrow> A - insert a B = A-B"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   216
by auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   217
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   218
lemma DisjI1: "A Int B = {} \<Longrightarrow> (A-X) Int B = {}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   219
by blast
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   220
lemma DisjI2: "A Int B = {} \<Longrightarrow> A Int (B-X) = {}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   221
by blast
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   222
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   223
text{* The main theorem: *}
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   224
28720
nipkow
parents: 28718
diff changeset
   225
declare [[max_clauses = 200]]
nipkow
parents: 28718
diff changeset
   226
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   227
theorem Ls_can_tile: "i \<le> a \<Longrightarrow> a < 2^n + i \<Longrightarrow> j \<le> b \<Longrightarrow> b < 2^n + j
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   228
  \<Longrightarrow> square2 n i j - {(a,b)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   229
proof(induct n arbitrary: a b i j)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   230
  case 0 thus ?case by (simp add:square2_def)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   231
next
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   232
  case (Suc n) note IH = Suc(1) and a = Suc(2-3) and b = Suc(4-5)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   233
  hence "a<2^n+i \<and> b<2^n+j \<or>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   234
         2^n+i\<le>a \<and> a<2^(n+1)+i \<and> b<2^n+j \<or>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   235
         a<2^n+i \<and> 2^n+j\<le>b \<and> b<2^(n+1)+j \<or>
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   236
         2^n+i\<le>a \<and> a<2^(n+1)+i \<and> 2^n+j\<le>b \<and> b<2^(n+1)+j" (is "?A|?B|?C|?D")
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   237
    by simp arith
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   238
  moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   239
  { assume "?A"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   240
    hence "square2 n i j - {(a,b)} : tiling Ls" using IH a b by auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   241
    moreover have "square2 n (2^n+i) j - {(2^n+i,2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   242
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   243
    moreover have "square2 n i (2^n+j) - {(2^n+i - 1, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   244
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   245
    moreover have "square2 n (2^n+i) (2^n+j) - {(2^n+i, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   246
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   247
    ultimately
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   248
    have "square2 (n+1) i j - {(a,b)} - L0 (2^n+i - 1) (2^n+j - 1) \<in> tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   249
      using  a b `?A`
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   250
      by (clarsimp simp: square2_Suc L0_def Un_Diff Diff_insert_if)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   251
         (fastsimp intro!: tiling_UnI DisjI1 DisjI2 square2_disj[THEN iffD2]
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   252
                   simp:Int_Un_distrib2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   253
  } moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   254
  { assume "?B"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   255
    hence "square2 n (2^n+i) j - {(a,b)} : tiling Ls" using IH a b by auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   256
    moreover have "square2 n i j - {(2^n+i - 1,2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   257
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   258
    moreover have "square2 n i (2^n+j) - {(2^n+i - 1, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   259
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   260
    moreover have "square2 n (2^n+i) (2^n+j) - {(2^n+i, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   261
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   262
    ultimately
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   263
    have "square2 (n+1) i j - {(a,b)} - L1 (2^n+i - 1) (2^n+j - 1) \<in> tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   264
      using  a b `?B`
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   265
      by (simp add: square2_Suc L1_def Un_Diff Diff_insert_if le_diff_conv2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   266
         (fastsimp intro!: tiling_UnI DisjI1 DisjI2 square2_disj[THEN iffD2]
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   267
                   simp:Int_Un_distrib2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   268
  } moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   269
  { assume "?C"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   270
    hence "square2 n i (2^n+j) - {(a,b)} : tiling Ls" using IH a b by auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   271
    moreover have "square2 n i j - {(2^n+i - 1,2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   272
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   273
    moreover have "square2 n (2^n+i) j - {(2^n+i, 2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   274
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   275
    moreover have "square2 n (2^n+i) (2^n+j) - {(2^n+i, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   276
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   277
    ultimately
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   278
    have "square2 (n+1) i j - {(a,b)} - L3 (2^n+i - 1) (2^n+j - 1) \<in> tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   279
      using  a b `?C`
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   280
      by (simp add: square2_Suc L3_def Un_Diff Diff_insert_if le_diff_conv2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   281
         (fastsimp intro!: tiling_UnI DisjI1 DisjI2 square2_disj[THEN iffD2]
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   282
                   simp:Int_Un_distrib2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   283
  } moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   284
  { assume "?D"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   285
    hence "square2 n (2^n+i) (2^n+j) -{(a,b)} : tiling Ls" using IH a b by auto
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   286
    moreover have "square2 n i j - {(2^n+i - 1,2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   287
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   288
    moreover have "square2 n (2^n+i) j - {(2^n+i, 2^n+j - 1)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   289
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   290
    moreover have "square2 n i (2^n+j) - {(2^n+i - 1, 2^n+j)} : tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   291
      by(rule IH)(insert pos_pow2[of n], auto)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   292
    ultimately
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   293
    have "square2 (n+1) i j - {(a,b)} - L2 (2^n+i - 1) (2^n+j - 1) \<in> tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   294
      using  a b `?D`
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   295
      by (simp add: square2_Suc L2_def Un_Diff Diff_insert_if le_diff_conv2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   296
         (fastsimp intro!: tiling_UnI DisjI1 DisjI2 square2_disj[THEN iffD2]
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   297
                   simp:Int_Un_distrib2)
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   298
  } moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   299
  have "?A \<Longrightarrow> L0 (2^n + i - 1) (2^n + j - 1) \<subseteq> square2 (n+1) i j - {(a, b)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   300
    using a b by(simp add:L0_def) arith moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   301
  have "?B \<Longrightarrow> L1 (2^n + i - 1) (2^n + j - 1) \<subseteq> square2 (n+1) i j - {(a, b)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   302
    using a b by(simp add:L1_def) arith moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   303
  have "?C \<Longrightarrow> L3 (2^n + i - 1) (2^n + j - 1) \<subseteq> square2 (n+1) i j - {(a, b)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   304
    using a b by(simp add:L3_def) arith moreover
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   305
  have "?D \<Longrightarrow> L2 (2^n + i - 1) (2^n + j - 1) \<subseteq> square2 (n+1) i j - {(a, b)}"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   306
    using a b by(simp add:L2_def) arith
28720
nipkow
parents: 28718
diff changeset
   307
  ultimately show ?case by simp (metis LinLs tiling_Diff1E)
28718
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   308
qed
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   309
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   310
corollary Ls_can_tile00:
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   311
  "a < 2^n \<Longrightarrow> b < 2^n \<Longrightarrow> square2 n 0 0 - {(a, b)} \<in> tiling Ls"
ef16499edaab Added second tiling example.
nipkow
parents: 23746
diff changeset
   312
by(rule Ls_can_tile) auto
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   313
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   314
end