ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
domrange/image_subset,vimage_subset: deleted needless premise!
misc: This slightly simplifies two proofs in Schroeder-Bernstein Theorem
ind-syntax/rule_concl: recoded to avoid exceptions
intr-elim: now checks conclusions of introduction rules
func/fun_disjoint_Un: now uses ex_ex1I
list-fn/hd,tl,drop: new
simpdata/bquant_simps: new
list/list_case_type: restored!
bool.thy: changed 1 from a "def" to a translation
Removed occurreces of one_def in bool.ML, nat.ML, univ.ML, ex/integ.ML
nat/succ_less_induct: new induction principle
arith/add_mono: new results about monotonicity
simpdata/mem_simps: removed the ones for succ and cons; added succI1,
consI2 to ZF_ss
upair/succ_iff: new, for use with simp_tac (cons_iff already existed)
ordinal/Ord_0_in_succ: renamed from Ord_0_mem_succ
nat/nat_0_in_succ: new
ex/prop-log/hyps_thms_if: split up the fast_tac call for more speed
(* Title: ZF/list.ML
ID: $Id$
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1993 University of Cambridge
Datatype definition of Lists
*)
structure List = Datatype_Fun
(val thy = Univ.thy;
val rec_specs =
[("list", "univ(A)",
[(["Nil"], "i"),
(["Cons"], "[i,i]=>i")])];
val rec_styp = "i=>i";
val ext = None
val sintrs =
["Nil : list(A)",
"[| a: A; l: list(A) |] ==> Cons(a,l) : list(A)"];
val monos = [];
val type_intrs = data_typechecks
val type_elims = []);
val [NilI, ConsI] = List.intrs;
(*An elimination rule, for type-checking*)
val ConsE = List.mk_cases List.con_defs "Cons(a,l) : list(A)";
(*Proving freeness results*)
val Cons_iff = List.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'";
val Nil_Cons_iff = List.mk_free "~ Nil=Cons(a,l)";
(*Perform induction on l, then prove the major premise using prems. *)
fun list_ind_tac a prems i =
EVERY [res_inst_tac [("x",a)] List.induct i,
rename_last_tac a ["1"] (i+2),
ares_tac prems i];
(** Lemmas to justify using "list" in other recursive type definitions **)
goalw List.thy List.defs "!!A B. A<=B ==> list(A) <= list(B)";
by (rtac lfp_mono 1);
by (REPEAT (rtac List.bnd_mono 1));
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
val list_mono = result();
(*There is a similar proof by list induction.*)
goalw List.thy (List.defs@List.con_defs) "list(univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac (A_subset_univ RS univ_mono) 2);
by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
Pair_in_univ]) 1);
val list_univ = result();
val list_subset_univ = standard
(list_mono RS (list_univ RSN (2,subset_trans)));
val major::prems = goal List.thy
"[| l: list(A); \
\ c: C(Nil); \
\ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C(Cons(x,y)) \
\ |] ==> list_case(c,h,l) : C(l)";
by (rtac (major RS List.induct) 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (List.case_eqns @ prems))));
val list_case_type = result();
(** For recursion **)
goalw List.thy List.con_defs "rank(a) : rank(Cons(a,l))";
by (simp_tac rank_ss 1);
val rank_Cons1 = result();
goalw List.thy List.con_defs "rank(l) : rank(Cons(a,l))";
by (simp_tac rank_ss 1);
val rank_Cons2 = result();