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/Lazy/fix.thy
ID: $Id$
Author: Martin Coen
Copyright 1993 University of Cambridge
Tentative attempt at including fixed point induction.
Justified by Smith.
*)
Fix = Type +
consts
idgen :: "[i]=>i"
INCL :: "[i=>o]=>o"
rules
idgen_def
"idgen(f) == lam t.case(t,true,false,%x y.<f`x, f`y>,%u.lam x.f ` u(x))"
INCL_def "INCL(%x.P(x)) == (ALL f.(ALL n:Nat.P(f^n`bot)) --> P(fix(f)))"
po_INCL "INCL(%x.a(x) [= b(x))"
INCL_subst "INCL(P) ==> INCL(%x.P((g::i=>i)(x)))"
end