Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
split_weak_cong, nat_case_weak_cong, nat_rec_weak_cong. Used in llist.ML
to make simplifications faster.
HOL/gfp: re-ordered premises to put mono(f) early (first or right after
A==gfp(f) in the def_ rules). Renamed some variables in rules, A to X and
h to A. Renamed coinduct to weak_coinduct, coinduct2 to coinduct.
Strengthened coinduct as suggested by j. Frost, to have the premise X <= f(X
Un gfp(f)).
HOL/llist: used stronger coinduct rule to strengthen LList_coinduct,
LList_equalityI, llist_equalityI, llist_fun_equalityI and to delete the "2"
form of those rules. Proved List_Fun_LList_I, LListD_Fun_diag_I and
llistD_Fun_range_I to help use the new coinduction rules; most proofs
involving them required small changes. Proved prod_fun_range_eq_diag as
lemma for llist_equalityI.
(* Title: HOL/wf.ML
ID: $Id$
Author: Tobias Nipkow
Copyright 1992 University of Cambridge
Well-founded Recursion
*)
WF = Trancl +
consts
wf :: "('a * 'a)set => bool"
cut :: "['a => 'b, ('a * 'a)set, 'a] => 'a => 'b"
wftrec,wfrec :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'b"
is_recfun :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b, 'a=>'b] => bool"
the_recfun :: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'a=>'b"
rules
wf_def "wf(r) == (!P. (!x. (!y. <y,x>:r --> P(y)) --> P(x)) --> (!x.P(x)))"
cut_def "cut(f,r,x) == (%y. if(<y,x>:r, f(y), @z.True))"
is_recfun_def "is_recfun(r,a,H,f) == (f = cut(%x.H(x, cut(f,r,x)), r, a))"
the_recfun_def "the_recfun(r,a,H) == (@f.is_recfun(r,a,H,f))"
wftrec_def "wftrec(r,a,H) == H(a, the_recfun(r,a,H))"
(*version not requiring transitivity*)
wfrec_def "wfrec(r,a,H) == wftrec(trancl(r), a, %x f. H(x, cut(f,r,x)))"
end