src/HOL/Induct/Mutil.thy
changeset 3120 c58423c20740
child 3424 bf466159ef84
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Induct/Mutil.thy	Wed May 07 12:50:26 1997 +0200
@@ -0,0 +1,32 @@
+(*  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 <= Compl t |] ==> a Un t : tiling A"
+
+defs
+  below_def  "below n    == nat_rec {} insert n"
+  evnodd_def "evnodd A b == A Int {(i,j). (i+j) mod 2 = b}"
+
+end