src/HOL/Induct/Mutil.thy
author paulson
Tue, 02 May 2000 18:55:11 +0200
changeset 8781 d0c2bd57a9fb
parent 8589 a24f7e5ee7ef
child 8950 3e858b72fac9
permissions -rw-r--r--
modified for new simprocs

(*  Title:      HOL/Induct/Mutil
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1996  University of Cambridge

The Mutilated Chess Board Problem, formalized inductively
  Originator is Max Black, according to J A Robinson.
  Popularized as the Mutilated Checkerboard Problem by J McCarthy
*)

Mutil = Main +

consts     tiling :: "'a set set => 'a set set"
inductive "tiling A"
  intrs
    empty  "{} : tiling A"
    Un     "[| a: A;  t: tiling A;  a <= -t |] ==> a Un t : tiling A"

consts    domino :: "(nat*nat)set set"
inductive domino
  intrs
    horiz  "{(i, j), (i, Suc j)} : domino"
    vertl  "{(i, j), (Suc i, j)} : domino"

constdefs
  below   :: "nat => nat set"
   "below n   == {i. i<n}"
  
  colored  :: "nat => (nat*nat)set"
   "colored b == {(i,j). (i+j) mod #2 = b}"

end