src/CCL/wfd.thy
author lcp
Mon, 20 Sep 1993 17:02:11 +0200
changeset 8 c3d2c6dcf3f0
parent 0 a5a9c433f639
child 227 5415c6ad0028
permissions -rw-r--r--
Installation of new simplfier. Previously appeared to set up the old simplifier to rewrite with the partial ordering [=, something not possible with the new simplifier. But such rewriting appears not to have actually been used, and there were few complications. In terms.ML setloop was used to avoid infinite rewriting with the letrec rule. Congruence rules were deleted, and an occasional SIMP_TAC had to become asm_simp_tac.

(*  Title: 	CCL/wf.thy
    ID:         $Id$
    Author: 	Martin Coen, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

Well-founded relations in CCL.
*)

Wfd = Trancl + Type +

consts
      (*** Predicates ***)
  Wfd        ::       "[i set] => o"
      (*** Relations ***)
  wf         ::       "[i set] => i set"
  wmap       ::       "[i=>i,i set] => i set"
  "**"       ::       "[i set,i set] => i set"      (infixl 70)
  NatPR      ::       "i set"
  ListPR     ::       "i set => i set"

rules

  Wfd_def
  "Wfd(R) == ALL P.(ALL x.(ALL y.<y,x> : R --> y:P) --> x:P) --> (ALL a.a:P)"

  wf_def         "wf(R) == {x.x:R & Wfd(R)}"

  wmap_def       "wmap(f,R) == {p. EX x y. p=<x,y>  &  <f(x),f(y)> : R}"
  lex_def
  "ra**rb == {p. EX a a' b b'.p = <<a,b>,<a',b'>> & (<a,a'> : ra | (a=a' & <b,b'> : rb))}"

  NatPR_def      "NatPR == {p.EX x:Nat. p=<x,succ(x)>}"
  ListPR_def     "ListPR(A) == {p.EX h:A.EX t:List(A). p=<t,h.t>}"
end