diff -r 75e163863e16 -r 8ce8c4d13d4d src/ZF/ListFn.ML --- a/src/ZF/ListFn.ML Fri Sep 17 12:53:53 1993 +0200 +++ b/src/ZF/ListFn.ML Fri Sep 17 16:16:38 1993 +0200 @@ -11,19 +11,14 @@ (** list_rec -- by Vset recursion **) -(*Used to verify list_rec*) -val list_rec_ss = ZF_ss - addcongs (mk_typed_congs ListFn.thy [("h", "[i,i,i]=>i")]) - addrews List.case_eqns; - goal ListFn.thy "list_rec(Nil,c,h) = c"; by (rtac (list_rec_def RS def_Vrec RS trans) 1); -by (SIMP_TAC list_rec_ss 1); +by (simp_tac (ZF_ss addsimps List.case_eqns) 1); val list_rec_Nil = result(); goal ListFn.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; by (rtac (list_rec_def RS def_Vrec RS trans) 1); -by (SIMP_TAC (list_rec_ss addrews [Vset_rankI, rank_Cons2]) 1); +by (simp_tac (ZF_ss addsimps (List.case_eqns @ [Vset_rankI, rank_Cons2])) 1); val list_rec_Cons = result(); (*Type checking -- proved by induction, as usual*) @@ -33,8 +28,8 @@ \ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \ \ |] ==> list_rec(l,c,h) : C(l)"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC - (ZF_ss addrews (prems@[list_rec_Nil,list_rec_Cons])))); +by (ALLGOALS (asm_simp_tac + (ZF_ss addsimps (prems@[list_rec_Nil,list_rec_Cons])))); val list_rec_type = result(); (** Versions for use with definitions **) @@ -122,47 +117,42 @@ map_type, map_type2, app_type, length_type, rev_type, flat_type, list_add_type]; -val list_congs = - List.congs @ - mk_congs ListFn.thy - ["list_rec","map","op @","length","rev","flat","list_add"]; - val list_ss = arith_ss - addcongs list_congs - addrews List.case_eqns - addrews list_typechecks - addrews [list_rec_Nil, list_rec_Cons, + addsimps List.case_eqns + addsimps [list_rec_Nil, list_rec_Cons, map_Nil, map_Cons, app_Nil, app_Cons, length_Nil, length_Cons, rev_Nil, rev_Cons, flat_Nil, flat_Cons, - list_add_Nil, list_add_Cons]; + list_add_Nil, list_add_Cons] + setsolver (type_auto_tac list_typechecks); +(*Could also rewrite using the list_typechecks...*) (*** theorems about map ***) val prems = goal ListFn.thy "l: list(A) ==> map(%u.u, l) = l"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val map_ident = result(); val prems = goal ListFn.thy "l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val map_compose = result(); val prems = goal ListFn.thy "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"; by (list_ind_tac "xs" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val map_app_distrib = result(); val prems = goal ListFn.thy "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"; by (list_ind_tac "ls" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib]))); val map_flat = result(); val prems = goal ListFn.thy @@ -170,9 +160,7 @@ \ list_rec(map(h,l), c, d) = \ \ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))"; by (list_ind_tac "l" prems 1); -by (ALLGOALS - (ASM_SIMP_TAC - (list_ss addcongs (mk_typed_congs ListFn.thy [("d", "[i,i,i]=>i")])))); +by (ALLGOALS (asm_simp_tac list_ss)); val list_rec_map = result(); (** theorems about list(Collect(A,P)) -- used in ex/term.ML **) @@ -183,7 +171,7 @@ val prems = goal ListFn.thy "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val map_list_Collect = result(); (*** theorems about length ***) @@ -191,13 +179,13 @@ val prems = goal ListFn.thy "xs: list(A) ==> length(map(h,xs)) = length(xs)"; by (list_ind_tac "xs" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val length_map = result(); val prems = goal ListFn.thy "xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)"; by (list_ind_tac "xs" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val length_app = result(); (* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m @@ -207,60 +195,59 @@ val prems = goal ListFn.thy "xs: list(A) ==> length(rev(xs)) = length(xs)"; by (list_ind_tac "xs" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app, add_commute_succ]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app, add_commute_succ]))); val length_rev = result(); val prems = goal ListFn.thy "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"; by (list_ind_tac "ls" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [length_app]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app]))); val length_flat = result(); (*** theorems about app ***) val [major] = goal ListFn.thy "xs: list(A) ==> xs@Nil=xs"; by (rtac (major RS List.induct) 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val app_right_Nil = result(); val prems = goal ListFn.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"; by (list_ind_tac "xs" prems 1); -by (ALLGOALS (ASM_SIMP_TAC list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); val app_assoc = result(); val prems = goal ListFn.thy "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"; by (list_ind_tac "ls" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_assoc]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_assoc]))); val flat_app_distrib = result(); (*** theorems about rev ***) val prems = goal ListFn.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [map_app_distrib]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib]))); val rev_map_distrib = result(); (*Simplifier needs the premises as assumptions because rewriting will not instantiate the variable ?A in the rules' typing conditions; note that rev_type does not instantiate ?A. Only the premises do. *) -val prems = goal ListFn.thy - "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; -by (cut_facts_tac prems 1); +goal ListFn.thy + "!!xs. [| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; by (etac List.induct 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [app_right_Nil,app_assoc]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_right_Nil,app_assoc]))); val rev_app_distrib = result(); val prems = goal ListFn.thy "l: list(A) ==> rev(rev(l))=l"; by (list_ind_tac "l" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [rev_app_distrib]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [rev_app_distrib]))); val rev_rev_ident = result(); val prems = goal ListFn.thy "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"; by (list_ind_tac "ls" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews +by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil]))); val rev_flat = result(); @@ -273,22 +260,22 @@ by (cut_facts_tac prems 1); by (list_ind_tac "xs" prems 1); by (ALLGOALS - (ASM_SIMP_TAC (list_ss addrews [add_0_right, add_assoc RS sym]))); -by (resolve_tac arith_congs 1); -by (REPEAT (ares_tac [refl, list_add_type, add_commute] 1)); + (asm_simp_tac (list_ss addsimps [add_0_right, add_assoc RS sym]))); +by (rtac (add_commute RS subst_context) 1); +by (REPEAT (ares_tac [refl, list_add_type] 1)); val list_add_app = result(); val prems = goal ListFn.thy "l: list(nat) ==> list_add(rev(l)) = list_add(l)"; by (list_ind_tac "l" prems 1); by (ALLGOALS - (ASM_SIMP_TAC (list_ss addrews [list_add_app, add_0_right]))); + (asm_simp_tac (list_ss addsimps [list_add_app, add_0_right]))); val list_add_rev = result(); val prems = goal ListFn.thy "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"; by (list_ind_tac "ls" prems 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews [list_add_app]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [list_add_app]))); by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1)); val list_add_flat = result(); @@ -301,6 +288,6 @@ \ |] ==> P(l)"; by (rtac (major RS rev_rev_ident RS subst) 1); by (rtac (major RS rev_type RS List.induct) 1); -by (ALLGOALS (ASM_SIMP_TAC (list_ss addrews prems))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps prems))); val list_append_induct = result();