src/ZF/ex/Mutil.ML
author nipkow
Tue, 10 Mar 1998 19:02:53 +0100
changeset 4723 9e2609b1bfb1
parent 4152 451104c223e2
child 5068 fb28eaa07e01
permissions -rw-r--r--
Adapted proofs because of new simplification tactics.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/Mutil
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     2
    ID:         $Id$
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     5
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     6
The Mutilated Checkerboard Problem, formalized inductively
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     7
*)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     8
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
     9
open Mutil;
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    10
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    11
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    12
(** Basic properties of evnodd **)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    13
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
    14
goalw thy [evnodd_def] "<i,j>: evnodd(A,b) <-> <i,j>: A & (i#+j) mod 2 = b";
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2493
diff changeset
    15
by (Blast_tac 1);
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
    16
qed "evnodd_iff";
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    17
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    18
goalw thy [evnodd_def] "evnodd(A, b) <= A";
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2493
diff changeset
    19
by (Blast_tac 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    20
qed "evnodd_subset";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    21
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    22
(* Finite(X) ==> Finite(evnodd(X,b)) *)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    23
bind_thm("Finite_evnodd", evnodd_subset RS subset_imp_lepoll RS lepoll_Finite);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    24
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    25
goalw thy [evnodd_def] "evnodd(A Un B, b) = evnodd(A,b) Un evnodd(B,b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    26
by (simp_tac (simpset() addsimps [Collect_Un]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    27
qed "evnodd_Un";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    28
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    29
goalw thy [evnodd_def] "evnodd(A - B, b) = evnodd(A,b) - evnodd(B,b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    30
by (simp_tac (simpset() addsimps [Collect_Diff]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    31
qed "evnodd_Diff";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    32
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    33
goalw thy [evnodd_def]
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    34
    "evnodd(cons(<i,j>,C), b) = \
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    35
\    if((i#+j) mod 2 = b, cons(<i,j>, evnodd(C,b)), evnodd(C,b))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    36
by (asm_simp_tac (simpset() addsimps [evnodd_def, Collect_cons] 
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    37
                        setloop split_tac [expand_if]) 1);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    38
qed "evnodd_cons";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    39
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    40
goalw thy [evnodd_def] "evnodd(0, b) = 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    41
by (simp_tac (simpset() addsimps [evnodd_def]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    42
qed "evnodd_0";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    43
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
    44
Addsimps [evnodd_cons, evnodd_0];
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    45
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    46
(*** Dominoes ***)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    47
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    48
goal thy "!!d. d:domino ==> Finite(d)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    49
by (blast_tac (claset() addSIs [Finite_cons, Finite_0] addEs [domino.elim]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    50
qed "domino_Finite";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    51
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    52
goal thy "!!d. [| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    53
by (eresolve_tac [domino.elim] 1);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    54
by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 2);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    55
by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 1);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    56
by (REPEAT_FIRST (ares_tac [add_type]));
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    57
(*Four similar cases: case (i#+j) mod 2 = b, 2#-b, ...*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    58
by (REPEAT (asm_simp_tac (simpset() addsimps [mod_succ, succ_neq_self] 
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    59
                                   setloop split_tac [expand_if]) 1
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    60
           THEN blast_tac (claset() addDs [ltD]) 1));
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    61
qed "domino_singleton";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    62
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    63
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    64
(*** Tilings ***)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    65
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    66
(** The union of two disjoint tilings is a tiling **)
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    67
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    68
goal thy "!!t. t: tiling(A) ==> \
1630
8c84606672fe Simplified proof of tiling_UnI
paulson
parents: 1624
diff changeset
    69
\              u: tiling(A) --> t Int u = 0 --> t Un u : tiling(A)";
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    70
by (etac tiling.induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    71
by (simp_tac (simpset() addsimps tiling.intrs) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    72
by (asm_full_simp_tac (simpset() addsimps [Un_assoc,
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2493
diff changeset
    73
					  subset_empty_iff RS iff_sym]) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    74
by (blast_tac (claset() addIs tiling.intrs) 1);
3732
c6abd2c3373f Now using qed_spec_mp
paulson
parents: 2896
diff changeset
    75
qed_spec_mp "tiling_UnI";
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    76
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    77
goal thy "!!t. t:tiling(domino) ==> Finite(t)";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    78
by (eresolve_tac [tiling.induct] 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    79
by (rtac Finite_0 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    80
by (blast_tac (claset() addSIs [Finite_Un] addIs [domino_Finite]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    81
qed "tiling_domino_Finite";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    82
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    83
goal thy "!!t. t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    84
by (eresolve_tac [tiling.induct] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    85
by (simp_tac (simpset() addsimps [evnodd_def]) 1);
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
    86
by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
    87
by (Simp_tac 2 THEN assume_tac 1);
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
    88
by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
    89
by (Simp_tac 2 THEN assume_tac 1);
2896
86cc7ef9b30c Another blast_tac call
paulson
parents: 2875
diff changeset
    90
by (Step_tac 1);
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
    91
by (subgoal_tac "ALL p b. p:evnodd(a,b) --> p~:evnodd(ta,b)" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    92
by (asm_simp_tac (simpset() addsimps [evnodd_Un, Un_cons, tiling_domino_Finite,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
    93
                                  evnodd_subset RS subset_Finite,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
    94
                                  Finite_imp_cardinal_cons]) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
    95
by (blast_tac (claset() addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    96
qed "tiling_domino_0_1";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    97
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    98
goal thy "!!i n. [| i: nat;  n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
    99
by (nat_ind_tac "n" [] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   100
by (simp_tac (simpset() addsimps tiling.intrs) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   101
by (asm_simp_tac (simpset() addsimps [Un_assoc RS sym, Sigma_succ2]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   102
by (resolve_tac tiling.intrs 1);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   103
by (assume_tac 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
   104
by (subgoal_tac    (*seems the easiest way of turning one to the other*)
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   105
    "{i}*{succ(n1#+n1)} Un {i}*{n1#+n1} = {<i,n1#+n1>, <i,succ(n1#+n1)>}" 1);
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2493
diff changeset
   106
by (Blast_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   107
by (asm_simp_tac (simpset() addsimps [domino.horiz]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   108
by (blast_tac (claset() addEs [mem_irrefl, mem_asym]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   109
qed "dominoes_tile_row";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   110
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   111
goal thy "!!m n. [| m: nat;  n: nat |] ==> m * (n #+ n) : tiling(domino)";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   112
by (nat_ind_tac "m" [] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   113
by (simp_tac (simpset() addsimps tiling.intrs) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   114
by (asm_simp_tac (simpset() addsimps [Sigma_succ1]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   115
by (blast_tac (claset() addIs [tiling_UnI, dominoes_tile_row] 
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   116
                    addEs [mem_irrefl]) 1);
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   117
qed "dominoes_tile_matrix";
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   118
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   119
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   120
goal thy "!!m n. [| m: nat;  n: nat;  \
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   121
\                   t = (succ(m)#+succ(m))*(succ(n)#+succ(n));  \
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   122
\                   t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>} |] ==> \
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   123
\                t' ~: tiling(domino)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   124
by (rtac notI 1);
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   125
by (dtac tiling_domino_0_1 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   126
by (subgoal_tac "|evnodd(t',0)| < |evnodd(t',1)|" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   127
by (asm_full_simp_tac (simpset() addsimps [lt_not_refl]) 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   128
by (subgoal_tac "t : tiling(domino)" 1);
1624
e2a6102b9369 Alternative proof removes dependence upon AC
paulson
parents: 1606
diff changeset
   129
(*Requires a small simpset that won't move the succ applications*)
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   130
by (asm_simp_tac (ZF_ss addsimps [nat_succI, add_type, 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
   131
                                  dominoes_tile_matrix]) 2);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   132
by (subgoal_tac "(m#+m)#+(n#+n) = (m#+n)#+(m#+n)" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   133
by (asm_simp_tac (simpset() addsimps add_ac) 2);
4723
9e2609b1bfb1 Adapted proofs because of new simplification tactics.
nipkow
parents: 4152
diff changeset
   134
by (asm_lr_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   135
    (simpset() addsimps [evnodd_Diff, mod2_add_self,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
   136
                        mod2_succ_succ, tiling_domino_0_1 RS sym]) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   137
by (rtac lt_trans 1);
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   138
by (REPEAT
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   139
    (rtac Finite_imp_cardinal_Diff 1 
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   140
     THEN
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   141
     asm_simp_tac (simpset() addsimps [tiling_domino_Finite, Finite_evnodd, 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
   142
                                      Finite_Diff]) 1 
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   143
     THEN
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3732
diff changeset
   144
     asm_simp_tac (simpset() addsimps [evnodd_iff, nat_0_le RS ltD, 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2067
diff changeset
   145
                                      mod2_add_self]) 1));
1606
dd66bed09592 New example: mutilated checkerboard
paulson
parents:
diff changeset
   146
qed "mutil_not_tiling";