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
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3424
bf466159ef84 Tidying and simplification of declarations
paulson
parents: 3120
diff changeset
     1
(*  Title:      HOL/Induct/Mutil
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     2
    ID:         $Id$
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     5
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     6
The Mutilated Chess Board Problem, formalized inductively
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
  Originator is Max Black, according to J A Robinson.
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
  Popularized as the Mutilated Checkerboard Problem by J McCarthy
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     9
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    10
8340
paulson
parents: 5931
diff changeset
    11
Mutil = Main +
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    12
8399
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    13
consts     tiling :: "'a set set => 'a set set"
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    14
inductive "tiling A"
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    15
  intrs
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    16
    empty  "{} : tiling A"
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    17
    Un     "[| a: A;  t: tiling A;  a <= -t |] ==> a Un t : tiling A"
86b04d47b853 nicely tarted up Mutil
paulson
parents: 8340
diff changeset
    18
8340
paulson
parents: 5931
diff changeset
    19
consts    domino :: "(nat*nat)set set"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    20
inductive domino
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    21
  intrs
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    22
    horiz  "{(i, j), (i, Suc j)} : domino"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    23
    vertl  "{(i, j), (Suc i, j)} : domino"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    24
8340
paulson
parents: 5931
diff changeset
    25
constdefs
paulson
parents: 5931
diff changeset
    26
  below   :: "nat => nat set"
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8399
diff changeset
    27
   "below n   == {i. i<n}"
8340
paulson
parents: 5931
diff changeset
    28
  
8589
a24f7e5ee7ef simplified constant "colored"
paulson
parents: 8399
diff changeset
    29
  colored  :: "nat => (nat*nat)set"
8781
d0c2bd57a9fb modified for new simprocs
paulson
parents: 8589
diff changeset
    30
   "colored b == {(i,j). (i+j) mod #2 = b}"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    31
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    32
end