Moved theorems Ord_cases_lemma and Ord_cases here from Univ,
strengthening the succ(j) case to include Ord(j). Proved trans_induct3,
le_implies_UN_le_UN, Ord_1, lt_Ord2, le_Ord2.
(* Title: LCF/simpdata
ID: $Id$
Author: Tobias Nipkow, Cambridge University Computer Laboratory
Copyright 1993 University of Cambridge
Simplification data for LCF
*)
val LCF_ss = FOL_ss addsimps
[minimal,
UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm,
not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU,
not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF,
not_FF_eq_UU,not_FF_eq_TT,
COND_UU,COND_TT,COND_FF,
surj_pairing,FST,SND];