nipkow@2525: (* Title: HOL/MiniML/Generalize.ML nipkow@2525: ID: $Id$ nipkow@2525: Author: Wolfgang Naraschewski and Tobias Nipkow nipkow@2525: Copyright 1996 TU Muenchen nipkow@2525: *) nipkow@2525: nipkow@2525: AddSEs [equalityE]; nipkow@2525: nipkow@2525: goal thy "!!A B. free_tv A = free_tv B ==> gen A t = gen B t"; nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (ALLGOALS Asm_simp_tac); nipkow@2525: qed "gen_eq_on_free_tv"; nipkow@2525: nipkow@2525: goal thy "!!t.(free_tv t) <= (free_tv sch) --> gen sch t = (mk_scheme t)"; nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: qed_spec_mp "gen_without_effect"; nipkow@2525: nipkow@2525: Addsimps [gen_without_effect]; nipkow@2525: nipkow@2525: goal thy "!!A. free_tv (gen ($ S A) t) = free_tv t Int free_tv ($ S A)"; nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (Simp_tac 1); nipkow@2525: by (case_tac "nat : free_tv ($ S A)" 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: by (Asm_full_simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: qed "free_tv_gen"; nipkow@2525: nipkow@2525: Addsimps [free_tv_gen]; nipkow@2525: nipkow@2525: goal thy "!!A. free_tv (gen ($ S A) t # $ S A) = free_tv ($ S A)"; nipkow@2525: by (Simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: qed "free_tv_gen_cons"; nipkow@2525: nipkow@2525: Addsimps [free_tv_gen_cons]; nipkow@2525: nipkow@2525: goal thy "!!A. bound_tv (gen A t1) = (free_tv t1) - (free_tv A)"; nipkow@2525: by (typ.induct_tac "t1" 1); nipkow@2525: by (Simp_tac 1); nipkow@2525: by (case_tac "nat : free_tv A" 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: qed "bound_tv_gen"; nipkow@2525: nipkow@2525: Addsimps [bound_tv_gen]; nipkow@2525: nipkow@2525: goal thy "!!A. new_tv n t --> new_tv n (gen A t)"; nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (Simp_tac 1); nipkow@2525: by (case_tac "nat : free_tv A" 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Asm_full_simp_tac 1); nipkow@2525: qed_spec_mp "new_tv_compatible_gen"; nipkow@2525: nipkow@2525: goalw thy [gen_ML_def] "!!A. gen A t = gen_ML A t"; nipkow@2525: by (typ.induct_tac "t" 1); wenzelm@4089: by (simp_tac (simpset() addsimps [free_tv_ML_scheme_list]) 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [free_tv_ML_scheme_list]) 1); nipkow@2525: qed "gen_eq_gen_ML"; nipkow@2525: nipkow@2525: goal thy "!!S. (free_tv S) Int ((free_tv t) - (free_tv A)) = {} \ nipkow@2525: \ --> gen ($ S A) ($ S t) = $ S (gen A t)"; nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (strip_tac 1); nipkow@2525: by (case_tac "nat:(free_tv A)" 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Asm_full_simp_tac 1); nipkow@2525: by (subgoal_tac "nat ~: free_tv S" 1); nipkow@2525: by (Fast_tac 2); wenzelm@4089: by (asm_full_simp_tac (simpset() addsimps [free_tv_subst,dom_def,de_Morgan_disj]) 1); nipkow@2525: by (split_tac [expand_if] 1); nipkow@2525: by (cut_facts_tac [free_tv_app_subst_scheme_list] 1); nipkow@2525: by (Fast_tac 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: by (Best_tac 1); nipkow@2525: qed_spec_mp "gen_subst_commutes"; nipkow@2525: nipkow@2525: goal Generalize.thy "free_tv(t::typ) <= free_tv(A) --> bound_typ_inst S (gen A t) = t"; paulson@3018: by (typ.induct_tac "t" 1); paulson@3018: by (ALLGOALS Asm_simp_tac); paulson@3018: by (Fast_tac 1); nipkow@2525: qed_spec_mp "bound_typ_inst_gen"; nipkow@2525: Addsimps [bound_typ_inst_gen]; nipkow@2525: nipkow@2525: goalw Generalize.thy [le_type_scheme_def,is_bound_typ_instance] nipkow@2525: "gen ($ S A) ($ S t) <= $ S (gen A t)"; paulson@4153: by Safe_tac; paulson@3018: by (rename_tac "R" 1); paulson@3018: by (res_inst_tac [("x", "(%a. bound_typ_inst R (gen ($S A) (S a)))")] exI 1); paulson@3018: by (typ.induct_tac "t" 1); wenzelm@4089: by (simp_tac (simpset() addsplits [expand_if]) 1); paulson@3018: by (Asm_simp_tac 1); nipkow@2525: qed "gen_bound_typ_instance"; nipkow@2525: nipkow@2525: goalw Generalize.thy [le_type_scheme_def,is_bound_typ_instance] nipkow@2525: "!!A B. free_tv B <= free_tv A ==> gen A t <= gen B t"; paulson@4153: by Safe_tac; paulson@3018: by (rename_tac "S" 1); paulson@3018: by (res_inst_tac [("x","%b. if b:free_tv A then TVar b else S b")] exI 1); nipkow@2525: by (typ.induct_tac "t" 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); paulson@3018: by (Fast_tac 1); paulson@3018: by (Asm_simp_tac 1); nipkow@2525: qed "free_tv_subset_gen_le"; nipkow@2525: nipkow@2525: goalw thy [le_type_scheme_def,is_bound_typ_instance] nipkow@2525: "!!A. new_tv n A --> gen A t <= gen A ($ (%x. TVar (if x : free_tv A then x else n + x)) t)"; nipkow@2525: by (strip_tac 1); nipkow@2525: by (etac exE 1); nipkow@2525: by (hyp_subst_tac 1); nipkow@2525: by (res_inst_tac [("x","(%x. S (if n <= x then x - n else x))")] exI 1); nipkow@2525: by (typ.induct_tac "t" 1); nipkow@2525: by (Simp_tac 1); nipkow@2525: by (case_tac "nat : free_tv A" 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); nipkow@2525: by (subgoal_tac "n <= n + nat" 1); nipkow@2525: by (forw_inst_tac [("t","A")] new_tv_le 1); paulson@3018: by (assume_tac 1); nipkow@2525: by (dtac new_tv_not_free_tv 1); nipkow@2525: by (dtac new_tv_not_free_tv 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [diff_add_inverse]) 1); wenzelm@4089: by (simp_tac (simpset() addsimps [le_add1]) 1); nipkow@2525: by (Asm_simp_tac 1); nipkow@2525: qed_spec_mp "gen_t_le_gen_alpha_t"; nipkow@2525: nipkow@2525: Addsimps [gen_t_le_gen_alpha_t];