diff -r 000000000000 -r 7949f97df77a List.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/List.ML Thu Sep 16 12:21:07 1993 +0200 @@ -0,0 +1,362 @@ +(* Title: HOL/list + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +For list.thy. +*) + +open List; + +(** the list functional **) + +goalw List.thy [List_Fun_def] "mono(List_Fun(A))"; +by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono] 1)); +val List_Fun_mono = result(); + +goalw List.thy [List_Fun_def] + "!!A B. A<=B ==> List_Fun(A,Z) <= List_Fun(B,Z)"; +by (REPEAT (ares_tac [subset_refl, usum_mono, uprod_mono] 1)); +val List_Fun_mono2 = result(); + +(*This justifies using List in other recursive type definitions*) +goalw List.thy [List_def] "!!A B. A<=B ==> List(A) <= List(B)"; +by (rtac lfp_mono 1); +by (etac List_Fun_mono2 1); +val List_mono = result(); + +(** Type checking rules -- List creates well-founded sets **) + +val prems = goalw List.thy [List_def,List_Fun_def] "List(Sexp) <= Sexp"; +by (rtac lfp_lowerbound 1); +by (fast_tac (univ_cs addIs [Sexp_NumbI,Sexp_In0I,Sexp_In1I,Sexp_SconsI]) 1); +val List_Sexp = result(); + +(* A <= Sexp ==> List(A) <= Sexp *) +val List_subset_Sexp = standard + (List_mono RS (List_Sexp RSN (2,subset_trans))); + +(** Induction **) + +(*Induction for the set List(A) *) +val major::prems = goalw List.thy [NIL_def,CONS_def] + "[| M: List(A); P(NIL); \ +\ !!M N. [| M: A; N: List(A); P(N) |] ==> P(CONS(M,N)) |] \ +\ ==> P(M)"; +by (rtac (major RS (List_def RS def_induct)) 1); +by (rtac List_Fun_mono 1); +by (rewtac List_Fun_def); +by (fast_tac (set_cs addIs prems addEs [usumE,uprodE]) 1); +val List_induct = result(); + +(*Induction for the type 'a list *) +val prems = goalw List.thy [Nil_def,Cons_def] + "[| P(Nil); \ +\ !!x xs. P(xs) ==> P(Cons(x,xs)) |] ==> P(l)"; +by (rtac (Rep_List_inverse RS subst) 1); (*types force good instantiation*) +by (rtac (Rep_List RS List_induct) 1); +by (REPEAT (ares_tac prems 1 + ORELSE eresolve_tac [rangeE, ssubst, Abs_List_inverse RS subst] 1)); +val list_induct = result(); + +(*Perform induction on xs. *) +fun list_ind_tac a M = + EVERY [res_inst_tac [("l",a)] list_induct M, + rename_last_tac a ["1"] (M+1)]; + +(** Introduction rules for List constructors **) + +val List_unfold = rewrite_rule [List_Fun_def] + (List_Fun_mono RS (List_def RS def_lfp_Tarski)); + +(* c : {Numb(0)} <+> A <*> List(A) ==> c : List(A) *) +val ListI = List_unfold RS equalityD2 RS subsetD; + +(* NIL is a List -- this also justifies the type definition*) +goalw List.thy [NIL_def] "NIL: List(A)"; +by (rtac (singletonI RS usum_In0I RS ListI) 1); +val NIL_I = result(); + +goalw List.thy [CONS_def] + "!!a A M. [| a: A; M: List(A) |] ==> CONS(a,M) : List(A)"; +by (REPEAT (ares_tac [uprodI RS usum_In1I RS ListI] 1)); +val CONS_I = result(); + +(*** Isomorphisms ***) + +goal List.thy "inj(Rep_List)"; +by (rtac inj_inverseI 1); +by (rtac Rep_List_inverse 1); +val inj_Rep_List = result(); + +goal List.thy "inj_onto(Abs_List,List(range(Leaf)))"; +by (rtac inj_onto_inverseI 1); +by (etac Abs_List_inverse 1); +val inj_onto_Abs_List = result(); + +(** Distinctness of constructors **) + +goalw List.thy [NIL_def,CONS_def] "~ CONS(M,N) = NIL"; +by (rtac In1_not_In0 1); +val CONS_not_NIL = result(); +val NIL_not_CONS = standard (CONS_not_NIL RS not_sym); + +val CONS_neq_NIL = standard (CONS_not_NIL RS notE); +val NIL_neq_CONS = sym RS CONS_neq_NIL; + +goalw List.thy [Nil_def,Cons_def] "~ Cons(x,xs) = Nil"; +by (rtac (CONS_not_NIL RS (inj_onto_Abs_List RS inj_onto_contraD)) 1); +by (REPEAT (resolve_tac [rangeI, NIL_I, CONS_I, Rep_List] 1)); +val Cons_not_Nil = result(); + +val Nil_not_Cons = standard (Cons_not_Nil RS not_sym); + +val Cons_neq_Nil = standard (Cons_not_Nil RS notE); +val Nil_neq_Cons = sym RS Cons_neq_Nil; + +(** Injectiveness of CONS and Cons **) + +goalw List.thy [CONS_def] "(CONS(K,M)=CONS(L,N)) = (K=L & M=N)"; +by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); +val CONS_CONS_eq = result(); + +val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE); + +(*For reasoning about abstract list constructors*) +val List_cs = set_cs addIs [Rep_List, NIL_I, CONS_I] + addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] + addSDs [inj_onto_Abs_List RS inj_ontoD, + inj_Rep_List RS injD, Leaf_inject]; + +goalw List.thy [Cons_def] "(Cons(x,xs)=Cons(y,ys)) = (x=y & xs=ys)"; +by (fast_tac List_cs 1); +val Cons_Cons_eq = result(); +val Cons_inject = standard (Cons_Cons_eq RS iffD1 RS conjE); + +val [major] = goal List.thy "CONS(M,N): List(A) ==> M: A & N: List(A)"; +by (rtac (major RS setup_induction) 1); +by (etac List_induct 1); +by (ALLGOALS (fast_tac List_cs)); +val CONS_D = result(); + +val prems = goalw List.thy [CONS_def,In1_def] + "CONS(M,N): Sexp ==> M: Sexp & N: Sexp"; +by (cut_facts_tac prems 1); +by (fast_tac (set_cs addSDs [Scons_D]) 1); +val Sexp_CONS_D = result(); + + +(*Basic ss with constructors and their freeness*) +val list_free_simps = [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq, + CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq, + NIL_I, CONS_I]; +val list_free_ss = HOL_ss addsimps list_free_simps; + +goal List.thy "!!N. N: List(A) ==> !M. ~ N = CONS(M,N)"; +by (etac List_induct 1); +by (ALLGOALS (asm_simp_tac list_free_ss)); +val not_CONS_self = result(); + +goal List.thy "!x. ~ l=Cons(x,l)"; +by (list_ind_tac "l" 1); +by (ALLGOALS (asm_simp_tac list_free_ss)); +val not_Cons_self = result(); + + +(** Conversion rules for List_case: case analysis operator **) + +goalw List.thy [List_case_def,NIL_def] "List_case(NIL,c,h) = c"; +by (rtac Case_In0 1); +val List_case_NIL = result(); + +goalw List.thy [List_case_def,CONS_def] "List_case(CONS(M,N), c, h) = h(M,N)"; +by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); +val List_case_CONS = result(); + +(*** List_rec -- by wf recursion on pred_Sexp ***) + +(* The trancl(pred_sexp) is essential because pred_Sexp_CONS_I1,2 would not + hold if pred_Sexp^+ were changed to pred_Sexp. *) + +val List_rec_unfold = wf_pred_Sexp RS wf_trancl RS (List_rec_def RS def_wfrec); + +(** pred_Sexp lemmas **) + +goalw List.thy [CONS_def,In1_def] + "!!M. [| M: Sexp; N: Sexp |] ==> : pred_Sexp^+"; +by (asm_simp_tac pred_Sexp_ss 1); +val pred_Sexp_CONS_I1 = result(); + +goalw List.thy [CONS_def,In1_def] + "!!M. [| M: Sexp; N: Sexp |] ==> : pred_Sexp^+"; +by (asm_simp_tac pred_Sexp_ss 1); +val pred_Sexp_CONS_I2 = result(); + +val [prem] = goal List.thy + " : pred_Sexp^+ ==> \ +\ : pred_Sexp^+ & : pred_Sexp^+"; +by (rtac (prem RS (pred_Sexp_subset_Sigma RS trancl_subset_Sigma RS + subsetD RS SigmaE2)) 1); +by (etac (Sexp_CONS_D RS conjE) 1); +by (REPEAT (ares_tac [conjI, pred_Sexp_CONS_I1, pred_Sexp_CONS_I2, + prem RSN (2, trans_trancl RS transD)] 1)); +val pred_Sexp_CONS_D = result(); + +(** Conversion rules for List_rec **) + +goal List.thy "List_rec(NIL,c,h) = c"; +by (rtac (List_rec_unfold RS trans) 1); +by (rtac List_case_NIL 1); +val List_rec_NIL = result(); + +goal List.thy "!!M. [| M: Sexp; N: Sexp |] ==> \ +\ List_rec(CONS(M,N), c, h) = h(M, N, List_rec(N,c,h))"; +by (rtac (List_rec_unfold RS trans) 1); +by (rtac (List_case_CONS RS trans) 1); +by(asm_simp_tac(HOL_ss addsimps [CONS_I, pred_Sexp_CONS_I2, cut_apply])1); +val List_rec_CONS = result(); + +(*** list_rec -- by List_rec ***) + +val Rep_List_in_Sexp = + Rep_List RS (range_Leaf_subset_Sexp RS List_subset_Sexp RS subsetD); + +local + val list_rec_simps = list_free_simps @ + [List_rec_NIL, List_rec_CONS, + Abs_List_inverse, Rep_List_inverse, + Rep_List, rangeI, inj_Leaf, Inv_f_f, + Sexp_LeafI, Rep_List_in_Sexp] +in + val list_rec_Nil = prove_goalw List.thy [list_rec_def, Nil_def] + "list_rec(Nil,c,h) = c" + (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); + + val list_rec_Cons = prove_goalw List.thy [list_rec_def, Cons_def] + "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))" + (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); +end; + +val list_simps = [List_rec_NIL, List_rec_CONS, + list_rec_Nil, list_rec_Cons]; +val list_ss = list_free_ss addsimps list_simps; + + +(*Type checking. Useful?*) +val major::A_subset_Sexp::prems = goal List.thy + "[| M: List(A); \ +\ A<=Sexp; \ +\ c: C(NIL); \ +\ !!x y r. [| x: A; y: List(A); r: C(y) |] ==> h(x,y,r): C(CONS(x,y)) \ +\ |] ==> List_rec(M,c,h) : C(M :: 'a node set)"; +val Sexp_ListA_I = A_subset_Sexp RS List_subset_Sexp RS subsetD; +val Sexp_A_I = A_subset_Sexp RS subsetD; +by (rtac (major RS List_induct) 1); +by (ALLGOALS(asm_simp_tac (list_ss addsimps ([Sexp_A_I,Sexp_ListA_I]@prems)))); +val List_rec_type = result(); + +(** Generalized map functionals **) + +goalw List.thy [Rep_map_def] "Rep_map(f,Nil) = NIL"; +by (rtac list_rec_Nil 1); +val Rep_map_Nil = result(); + +goalw List.thy [Rep_map_def] + "Rep_map(f, Cons(x,xs)) = CONS(f(x), Rep_map(f,xs))"; +by (rtac list_rec_Cons 1); +val Rep_map_Cons = result(); + +goalw List.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map(f,xs): List(A)"; +by (rtac list_induct 1); +by(ALLGOALS(asm_simp_tac list_ss)); +val Rep_map_type = result(); + +goalw List.thy [Abs_map_def] "Abs_map(g,NIL) = Nil"; +by (rtac List_rec_NIL 1); +val Abs_map_NIL = result(); + +val prems = goalw List.thy [Abs_map_def] + "[| M: Sexp; N: Sexp |] ==> \ +\ Abs_map(g, CONS(M,N)) = Cons(g(M), Abs_map(g,N))"; +by (REPEAT (resolve_tac (List_rec_CONS::prems) 1)); +val Abs_map_CONS = result(); + +(** null, hd, tl, list_case **) + +goalw List.thy [null_def] "null([]) = True"; +by (rtac list_rec_Nil 1); +val null_Nil = result(); + +goalw List.thy [null_def] "null(Cons(x,xs)) = False"; +by (rtac list_rec_Cons 1); +val null_Cons = result(); + + +goalw List.thy [hd_def] "hd(Cons(x,xs)) = x"; +by (rtac list_rec_Cons 1); +val hd_Cons = result(); + + +goalw List.thy [tl_def] "tl(Cons(x,xs)) = xs"; +by (rtac list_rec_Cons 1); +val tl_Cons = result(); + + +goalw List.thy [list_case_def] "list_case([],a,f) = a"; +by (rtac list_rec_Nil 1); +val list_case_Nil = result(); + +goalw List.thy [list_case_def] "list_case(Cons(x,xs),a,f) = f(x,xs)"; +by (rtac list_rec_Cons 1); +val list_case_Cons = result(); + + +(** The functional "map" **) + +goalw List.thy [map_def] "map(f,Nil) = Nil"; +by (rtac list_rec_Nil 1); +val map_Nil = result(); + +goalw List.thy [map_def] "map(f, Cons(x,xs)) = Cons(f(x), map(f,xs))"; +by (rtac list_rec_Cons 1); +val map_Cons = result(); + +val map_simps = [Abs_map_NIL, Abs_map_CONS, + Rep_map_Nil, Rep_map_Cons, + map_Nil, map_Cons]; +val map_ss = list_free_ss addsimps map_simps; + +val [major,A_subset_Sexp,minor] = goal List.thy + "[| M: List(A); A<=Sexp; !!z. z: A ==> f(g(z)) = z |] \ +\ ==> Rep_map(f, Abs_map(g,M)) = M"; +by (rtac (major RS List_induct) 1); +by (ALLGOALS (asm_simp_tac(map_ss addsimps [Sexp_A_I,Sexp_ListA_I,minor]))); +val Abs_map_inverse = result(); + +(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*) + + +(** The functional "list_all" -- creates predicates over lists **) + +goalw List.thy [list_all_def] "list_all(P,Nil) = True"; +by (rtac list_rec_Nil 1); +val list_all_Nil = result(); + +goalw List.thy [list_all_def] + "list_all(P, Cons(x,xs)) = (P(x) & list_all(P,xs))"; +by (rtac list_rec_Cons 1); +val list_all_Cons = result(); + +(** Additional mapping lemmas **) + +goal List.thy "map(%x.x, xs) = xs"; +by (list_ind_tac "xs" 1); +by (ALLGOALS (asm_simp_tac map_ss)); +val map_ident = result(); + +goal List.thy "!!f. (!!x. f(x): Sexp) ==> \ +\ Abs_map(g, Rep_map(f,xs)) = map(%t. g(f(t)), xs)"; +by (list_ind_tac "xs" 1); +by(ALLGOALS(asm_simp_tac(map_ss addsimps + [Rep_map_type,List_Sexp RS subsetD]))); +val Abs_Rep_map = result();