src/HOL/Induct/Mutil.ML
author paulson
Mon, 27 Mar 2000 16:25:53 +0200
changeset 8589 a24f7e5ee7ef
parent 8558 6c4860b1828d
child 8703 816d8f6513be
permissions -rw-r--r--
simplified constant "colored"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3423
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
     1
(*  Title:      HOL/Induct/Mutil
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
The Mutilated Chess Board Problem, formalized inductively
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
     9
Addsimps (tiling.intrs @ domino.intrs);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    10
AddIs    tiling.intrs;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    11
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    12
(** The union of two disjoint tilings is a tiling **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    13
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    14
Goal "t: tiling A ==> u: tiling A --> t <= -u --> t Un u : tiling A";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    15
by (etac tiling.induct 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    16
by (simp_tac (simpset() addsimps [Un_assoc]) 2);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    17
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    18
qed_spec_mp "tiling_UnI";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    19
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    20
AddIs [tiling_UnI];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    21
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    22
(*** Chess boards ***)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    23
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    24
Goalw [below_def] "(i: below k) = (i<k)";
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    25
by Auto_tac;
3423
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    26
qed "below_less_iff";
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    27
AddIffs [below_less_iff];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    28
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    29
Goalw [below_def] "below 0 = {}";
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    30
by Auto_tac;
3423
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    31
qed "below_0";
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    32
Addsimps [below_0];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    33
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    34
Goalw [below_def]
3423
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    35
    "below(Suc n) Times B = ({n} Times B) Un ((below n) Times B)";
7401
paulson
parents: 5628
diff changeset
    36
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    37
qed "Sigma_Suc1";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    38
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    39
Goalw [below_def]
3423
3684a4420a67 Much polishing of proofs
paulson
parents: 3414
diff changeset
    40
    "A Times below(Suc n) = (A Times {n}) Un (A Times (below n))";
7401
paulson
parents: 5628
diff changeset
    41
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    42
qed "Sigma_Suc2";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    43
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    44
Addsimps [Sigma_Suc1, Sigma_Suc2];
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    45
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    46
Goal "({i} Times {n}) Un ({i} Times {m}) = {(i,m), (i,n)}";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    47
by Auto_tac;
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    48
qed "sing_Times_lemma";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    49
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    50
Goal "{i} Times below(n+n) : tiling domino";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5143
diff changeset
    51
by (induct_tac "n" 1);
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    52
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_assoc RS sym])));
8558
6c4860b1828d now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents: 8553
diff changeset
    53
by (rtac tiling.Un 1);
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    54
by (ALLGOALS (asm_simp_tac (simpset() addsimps [sing_Times_lemma])));
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    55
qed "dominoes_tile_row";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    56
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    57
AddSIs [dominoes_tile_row]; 
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    58
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    59
Goal "(below m) Times below(n+n) : tiling domino";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5143
diff changeset
    60
by (induct_tac "m" 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    61
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    62
qed "dominoes_tile_matrix";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    63
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    64
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    65
(*** "colored" and Dominoes ***)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    66
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    67
Goalw [colored_def]
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    68
    "colored b Int (insert (i,j) C) = \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    69
\      (if (i+j) mod 2 = b then insert (i,j) (colored b Int C) \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    70
\                          else colored b Int C)";
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    71
by Auto_tac;
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    72
qed "colored_insert";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    73
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    74
Addsimps [colored_insert];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    75
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    76
Goal "d:domino ==> finite d";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    77
by (etac domino.elim 1);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    78
by Auto_tac;
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    79
qed "domino_finite";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    80
Addsimps [domino_finite];
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    81
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    82
Goal "d:domino ==> (EX i j. colored 0 Int d = {(i,j)}) & \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    83
\                  (EX k l. colored 1 Int d = {(k,l)})";
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    84
by (etac domino.elim 1);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    85
by (auto_tac (claset(), simpset() addsimps [mod_Suc]));
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    86
qed "domino_singletons";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    87
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    88
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    89
(*** Tilings of dominoes ***)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    90
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
    91
Goal "t:tiling domino ==> finite t";
8558
6c4860b1828d now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents: 8553
diff changeset
    92
by (etac tiling.induct 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    93
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    94
qed "tiling_domino_finite";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    95
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    96
Addsimps [tiling_domino_finite, Int_Un_distrib, Diff_Int_distrib];
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
    97
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
    98
Goal "t: tiling domino ==> card(colored 0 Int t) = card(colored 1 Int t)";
8558
6c4860b1828d now exclusively uses rtac/dtac/etac rather than the long forms
paulson
parents: 8553
diff changeset
    99
by (etac tiling.induct 1);
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   100
by (dtac domino_singletons 2);
7401
paulson
parents: 5628
diff changeset
   101
by Auto_tac;
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   102
(*this lemma tells us that both "inserts" are non-trivial*)
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   103
by (subgoal_tac "ALL p C. C Int a = {p} --> p ~: t" 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   104
by (Asm_simp_tac 1);
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   105
by (blast_tac (claset() addEs [equalityE]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   106
qed "tiling_domino_0_1";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   107
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   108
(*Final argument is surprisingly complex*)
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   109
Goal "[| t : tiling domino;       \
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   110
\        (i+j) mod 2 = 0;  (k+l) mod 2 = 0;  \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   111
\        {(i,j),(k,l)} <= t;  l ~= j |]    \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   112
\     ==> (t - {(i,j)} - {(k,l)}) ~: tiling domino";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   113
by (rtac notI 1);
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   114
by (subgoal_tac "card (colored 0 Int (t - {(i,j)} - {(k,l)})) < \
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8558
diff changeset
   115
\                card (colored 1 Int (t - {(i,j)} - {(k,l)}))" 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   116
by (force_tac (claset(), HOL_ss addsimps [tiling_domino_0_1]) 1);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   117
by (asm_simp_tac (simpset() addsimps [tiling_domino_0_1 RS sym]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   118
by (rtac less_trans 1);
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   119
by (ALLGOALS (force_tac (claset() addSIs [card_Diff1_less], 
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   120
			 simpset() addsimps [colored_def])));
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   121
qed "gen_mutil_not_tiling";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   122
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   123
(*Apply the general theorem to the well-known case*)
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   124
Goal "[| t = below(Suc m + Suc m) Times below(Suc n + Suc n) |] \
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   125
\     ==> t - {(0,0)} - {(Suc(m+m), Suc(n+n))} ~: tiling domino";
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   126
by (rtac gen_mutil_not_tiling 1);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   127
by (blast_tac (claset() addSIs [dominoes_tile_matrix]) 1);
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   128
by Auto_tac;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   129
qed "mutil_not_tiling";
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8339
diff changeset
   130