src/HOL/ex/Mutil.thy
author paulson
Thu, 06 Jun 1996 14:39:44 +0200
changeset 1789 aade046ec6d5
parent 1684 3eaf3ab53082
child 1824 44254696843a
permissions -rw-r--r--
Quotes now optional around inductive set

(*  Title:      HOL/ex/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 = Finite +
consts
  domino  :: "(nat*nat)set set"
  tiling  :: 'a set set => 'a set set
  below   :: nat => nat set
  evnodd  :: "[(nat*nat)set, nat] => (nat*nat)set"

inductive domino
  intrs
    horiz  "{(i, j), (i, Suc j)} : domino"
    vertl  "{(i, j), (Suc i, j)} : domino"

inductive "tiling A"
  intrs
    empty  "{} : tiling A"
    Un     "[| a: A;  t: tiling A;  a Int t = {} |] ==> a Un t : tiling A"

defs
  below_def  "below n    == nat_rec n {} insert"
  evnodd_def "evnodd A b == A Int {(i,j). (i+j) mod 2 = b}"

end