author | paulson |
Tue, 16 Jul 1996 15:49:46 +0200 | |
changeset 1868 | 836950047d85 |
parent 1824 | 44254696843a |
child 2513 | d708d8cdc8e8 |
permissions | -rw-r--r-- |
1621 | 1 |
(* Title: HOL/ex/Mutil |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
1684
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
6 |
The Mutilated Chess Board Problem, formalized inductively |
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
7 |
Originator is Max Black, according to J A Robinson. |
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
8 |
Popularized as the Mutilated Checkerboard Problem by J McCarthy |
1621 | 9 |
*) |
10 |
||
11 |
Mutil = Finite + |
|
12 |
consts |
|
13 |
domino :: "(nat*nat)set set" |
|
14 |
tiling :: 'a set set => 'a set set |
|
1684
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
15 |
below :: nat => nat set |
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
16 |
evnodd :: "[(nat*nat)set, nat] => (nat*nat)set" |
1621 | 17 |
|
1789 | 18 |
inductive domino |
1621 | 19 |
intrs |
20 |
horiz "{(i, j), (i, Suc j)} : domino" |
|
21 |
vertl "{(i, j), (Suc i, j)} : domino" |
|
22 |
||
23 |
inductive "tiling A" |
|
24 |
intrs |
|
25 |
empty "{} : tiling A" |
|
26 |
Un "[| a: A; t: tiling A; a Int t = {} |] ==> a Un t : tiling A" |
|
27 |
||
1684
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
28 |
defs |
1824 | 29 |
below_def "below n == nat_rec {} insert n" |
1684
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
30 |
evnodd_def "evnodd A b == A Int {(i,j). (i+j) mod 2 = b}" |
3eaf3ab53082
Rearrangement and polishing to look for for publication
paulson
parents:
1621
diff
changeset
|
31 |
|
1621 | 32 |
end |