src/HOL/ex/Mutil.thy
author paulson
Tue, 16 Jul 1996 15:49:46 +0200
changeset 1868 836950047d85
parent 1824 44254696843a
child 2513 d708d8cdc8e8
permissions -rw-r--r--
Put in minimal simpset to avoid excessive simplification, just as in revision 1.9 of HOL/indrule.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1621
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     1
(*  Title:      HOL/ex/Mutil
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     2
    ID:         $Id$
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     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
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
     9
*)
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    10
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    11
Mutil = Finite +
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    12
consts
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    13
  domino  :: "(nat*nat)set set"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    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
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    17
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1684
diff changeset
    18
inductive domino
1621
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    19
  intrs
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    20
    horiz  "{(i, j), (i, Suc j)} : domino"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    21
    vertl  "{(i, j), (Suc i, j)} : domino"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    22
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    23
inductive "tiling A"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    24
  intrs
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    25
    empty  "{} : tiling A"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    26
    Un     "[| a: A;  t: tiling A;  a Int t = {} |] ==> a Un t : tiling A"
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    27
1684
3eaf3ab53082 Rearrangement and polishing to look for for publication
paulson
parents: 1621
diff changeset
    28
defs
1824
44254696843a Changed argument order of nat_rec.
berghofe
parents: 1789
diff changeset
    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
d92f42acdb26 New mutilated checkerboard example
paulson
parents:
diff changeset
    32
end