src/CCL/Wfd.thy
author wenzelm
Sat, 17 Sep 2005 17:35:26 +0200
changeset 17456 bcf7544875b2
parent 3837 d7f033c74b38
child 20140 98acc6d0fab6
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     1
(*  Title:      CCL/Wfd.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1474
3f7d67927fe2 expanded tabs
clasohm
parents: 289
diff changeset
     3
    Author:     Martin Coen, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     7
header {* Well-founded relations in CCL *}
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     8
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     9
theory Wfd
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    10
imports Trancl Type Hered
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    11
uses ("wfd.ML") ("genrec.ML") ("typecheck.ML") ("eval.ML")
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    12
begin
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
consts
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
      (*** Predicates ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
  Wfd        ::       "[i set] => o"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
      (*** Relations ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
  wf         ::       "[i set] => i set"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  wmap       ::       "[i=>i,i set] => i set"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
  "**"       ::       "[i set,i set] => i set"      (infixl 70)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
  NatPR      ::       "i set"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
  ListPR     ::       "i set => i set"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    24
axioms
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    26
  Wfd_def:
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1474
diff changeset
    27
  "Wfd(R) == ALL P.(ALL x.(ALL y.<y,x> : R --> y:P) --> x:P) --> (ALL a. a:P)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    29
  wf_def:         "wf(R) == {x. x:R & Wfd(R)}"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    31
  wmap_def:       "wmap(f,R) == {p. EX x y. p=<x,y>  &  <f(x),f(y)> : R}"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    32
  lex_def:
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1474
diff changeset
    33
  "ra**rb == {p. EX a a' b b'. p = <<a,b>,<a',b'>> & (<a,a'> : ra | (a=a' & <b,b'> : rb))}"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    35
  NatPR_def:      "NatPR == {p. EX x:Nat. p=<x,succ(x)>}"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    36
  ListPR_def:     "ListPR(A) == {p. EX h:A. EX t:List(A). p=<t,h$t>}"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    37
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    38
ML {* use_legacy_bindings (the_context ()) *}
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    39
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    40
use "wfd.ML"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    41
use "genrec.ML"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    42
use "typecheck.ML"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    43
use "eval.ML"
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    44
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
end