author | wenzelm |
Sun, 31 Oct 1999 20:11:23 +0100 | |
changeset 7990 | 0a604b2fc2b1 |
parent 7499 | 23e090051cb8 |
child 8339 | bcadeb9c7095 |
permissions | -rw-r--r-- |
3423 | 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 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
9 |
Addsimps tiling.intrs; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
10 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
11 |
(** The union of two disjoint tilings is a tiling **) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
12 |
|
5490 | 13 |
Goal "t: tiling A ==> u: tiling A --> t Int u = {} --> t Un u : tiling A"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
14 |
by (etac tiling.induct 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
15 |
by (Simp_tac 1); |
4089 | 16 |
by (simp_tac (simpset() addsimps [Un_assoc]) 1); |
17 |
by (blast_tac (claset() addIs tiling.intrs) 1); |
|
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 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
20 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
21 |
(*** Chess boards ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
22 |
|
5069 | 23 |
Goalw [below_def] "(i: below k) = (i<k)"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
24 |
by (Blast_tac 1); |
3423 | 25 |
qed "below_less_iff"; |
26 |
AddIffs [below_less_iff]; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
27 |
|
5069 | 28 |
Goalw [below_def] "below 0 = {}"; |
3423 | 29 |
by (Simp_tac 1); |
30 |
qed "below_0"; |
|
31 |
Addsimps [below_0]; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
32 |
|
5069 | 33 |
Goalw [below_def] |
3423 | 34 |
"below(Suc n) Times B = ({n} Times B) Un ((below n) Times B)"; |
7401 | 35 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
36 |
qed "Sigma_Suc1"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
37 |
|
5069 | 38 |
Goalw [below_def] |
3423 | 39 |
"A Times below(Suc n) = (A Times {n}) Un (A Times (below n))"; |
7401 | 40 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
41 |
qed "Sigma_Suc2"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
42 |
|
5069 | 43 |
Goal "{i} Times below(n+n) : tiling domino"; |
5184 | 44 |
by (induct_tac "n" 1); |
7401 | 45 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [Un_assoc RS sym, Sigma_Suc2]))); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
46 |
by (resolve_tac tiling.intrs 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
47 |
by (assume_tac 2); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
48 |
by (subgoal_tac (*seems the easiest way of turning one to the other*) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
49 |
"({i} Times {Suc(n+n)}) Un ({i} Times {n+n}) = \ |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
50 |
\ {(i, n+n), (i, Suc(n+n))}" 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
51 |
by (Blast_tac 2); |
4089 | 52 |
by (asm_simp_tac (simpset() addsimps [domino.horiz]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4387
diff
changeset
|
53 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
54 |
qed "dominoes_tile_row"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
55 |
|
5069 | 56 |
Goal "(below m) Times below(n+n) : tiling domino"; |
5184 | 57 |
by (induct_tac "m" 1); |
4089 | 58 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Sigma_Suc1]))); |
59 |
by (blast_tac (claset() addSIs [tiling_UnI, dominoes_tile_row] |
|
5490 | 60 |
addSEs [below_less_iff RS iffD1 RS less_irrefl]) 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
61 |
qed "dominoes_tile_matrix"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
62 |
|
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 |
(*** Basic properties of evnodd ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
65 |
|
5069 | 66 |
Goalw [evnodd_def] "(i,j): evnodd A b = ((i,j): A & (i+j) mod 2 = b)"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
67 |
by (Simp_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
68 |
qed "evnodd_iff"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
69 |
|
5069 | 70 |
Goalw [evnodd_def] "evnodd A b <= A"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
71 |
by (rtac Int_lower1 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
72 |
qed "evnodd_subset"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
73 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
74 |
(* finite X ==> finite(evnodd X b) *) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
75 |
bind_thm("finite_evnodd", evnodd_subset RS finite_subset); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
76 |
|
5069 | 77 |
Goalw [evnodd_def] "evnodd (A Un B) b = evnodd A b Un evnodd B b"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
78 |
by (Blast_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
79 |
qed "evnodd_Un"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
80 |
|
5069 | 81 |
Goalw [evnodd_def] "evnodd (A - B) b = evnodd A b - evnodd B b"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
82 |
by (Blast_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
83 |
qed "evnodd_Diff"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
84 |
|
5069 | 85 |
Goalw [evnodd_def] "evnodd {} b = {}"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
86 |
by (Simp_tac 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
87 |
qed "evnodd_empty"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
88 |
|
5069 | 89 |
Goalw [evnodd_def] |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
90 |
"evnodd (insert (i,j) C) b = \ |
3423 | 91 |
\ (if (i+j) mod 2 = b then insert (i,j) (evnodd C b) else evnodd C b)"; |
7401 | 92 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
93 |
qed "evnodd_insert"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
94 |
|
3423 | 95 |
Addsimps [finite_evnodd, evnodd_Un, evnodd_Diff, evnodd_empty, evnodd_insert]; |
96 |
||
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
97 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
98 |
(*** Dominoes ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
99 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
100 |
Goal "[| d:domino; b<2 |] ==> EX i j. evnodd d b = {(i,j)}"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
101 |
by (eresolve_tac [domino.elim] 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
102 |
by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 2); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
103 |
by (res_inst_tac [("k1", "i+j")] (mod2_cases RS disjE) 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
104 |
by (REPEAT_FIRST assume_tac); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
105 |
(*Four similar cases: case (i+j) mod 2 = b, 2#-b, ...*) |
7499 | 106 |
by (auto_tac (claset(), simpset() addsimps [less_Suc_eq, mod_Suc])); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
107 |
qed "domino_singleton"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
108 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
109 |
Goal "d:domino ==> finite d"; |
4089 | 110 |
by (blast_tac (claset() addSEs [domino.elim]) 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
111 |
qed "domino_finite"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
112 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
113 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
114 |
(*** Tilings of dominoes ***) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
115 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
116 |
Goal "t:tiling domino ==> finite t"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
117 |
by (eresolve_tac [tiling.induct] 1); |
3414
804c8a178a7f
Modified a few defs and proofs because of the changes to theory Finite.thy.
nipkow
parents:
3357
diff
changeset
|
118 |
by (rtac Finites.emptyI 1); |
4089 | 119 |
by (blast_tac (claset() addSIs [finite_UnI] addIs [domino_finite]) 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
120 |
qed "tiling_domino_finite"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
121 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
122 |
Goal "t: tiling domino ==> card(evnodd t 0) = card(evnodd t 1)"; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
123 |
by (eresolve_tac [tiling.induct] 1); |
4089 | 124 |
by (simp_tac (simpset() addsimps [evnodd_def]) 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
125 |
by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1); |
7401 | 126 |
by Auto_tac; |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
127 |
by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1); |
7401 | 128 |
by Auto_tac; |
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
129 |
by (subgoal_tac "ALL p b. p : evnodd a b --> p ~: evnodd t b" 1); |
4089 | 130 |
by (asm_simp_tac (simpset() addsimps [tiling_domino_finite]) 1); |
4387 | 131 |
by (blast_tac (claset() addSDs [evnodd_subset RS subsetD] |
132 |
addEs [equalityE]) 1); |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
133 |
qed "tiling_domino_0_1"; |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
134 |
|
4387 | 135 |
(*Final argument is surprisingly complex. Note the use of small simpsets |
136 |
to avoid moving Sucs, etc.*) |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
137 |
Goal "[| t = below(Suc m + Suc m) Times below(Suc n + Suc n); \ |
5419 | 138 |
\ t' = t - {(0,0)} - {(Suc(m+m), Suc(n+n))} \ |
139 |
\ |] ==> t' ~: tiling domino"; |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
140 |
by (rtac notI 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
141 |
by (subgoal_tac "card(evnodd t' 0) < card(evnodd t' 1)" 1); |
4387 | 142 |
by (asm_full_simp_tac (HOL_ss addsimps [less_not_refl, tiling_domino_0_1]) 1); |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
143 |
by (subgoal_tac "t : tiling domino" 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
144 |
by (asm_simp_tac (HOL_ss addsimps [dominoes_tile_matrix]) 2); |
4387 | 145 |
by (asm_full_simp_tac (simpset() |
146 |
addsimps [mod_less, tiling_domino_0_1 RS sym]) 1); |
|
147 |
(*Cardinality is smaller because of the two elements fewer*) |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
148 |
by (rtac less_trans 1); |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
149 |
by (REPEAT |
5628 | 150 |
(rtac card_Diff1_less 1 |
4089 | 151 |
THEN asm_simp_tac (simpset() addsimps [tiling_domino_finite]) 1 |
152 |
THEN asm_simp_tac (simpset() addsimps [mod_less, evnodd_iff]) 1)); |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
153 |
qed "mutil_not_tiling"; |