src/ZF/Coind/Types.ML
changeset 6046 2c8a8be36c94
parent 5529 4a54acae6a15
child 6141 a6922171b396
equal deleted inserted replaced
6045:6a9dc67d48f5 6046:2c8a8be36c94
     7 open Types;
     7 open Types;
     8 
     8 
     9 val te_owrE = 
     9 val te_owrE = 
    10   TyEnv.mk_cases TyEnv.con_defs "te_owr(te,f,t):TyEnv";
    10   TyEnv.mk_cases TyEnv.con_defs "te_owr(te,f,t):TyEnv";
    11 
    11 
    12 Goalw TyEnv.con_defs "rank(te) < rank(te_owr(te,x,t))";
    12 Goal "te_app(te_owr(te,x,t),x) = t";
    13 by (simp_tac rank_ss 1);
    13 by (Simp_tac 1);
    14 qed "rank_te_owr1";
       
    15 
       
    16 Goal "te_rec(te_emp,c_te_emp,f_te_owr) = c_te_emp";
       
    17 by (rtac (te_rec_def RS def_Vrec RS trans) 1);
       
    18 by (simp_tac (simpset() addsimps rank_te_owr1::TyEnv.case_eqns) 1);
       
    19 qed "te_rec_emp";
       
    20 
       
    21 Goal " te_rec(te_owr(te,x,t),c_te_emp,f_te_owr) = \
       
    22 \   f_te_owr(te,x,t,te_rec(te,c_te_emp,f_te_owr))";
       
    23 by (rtac (te_rec_def RS def_Vrec RS trans) 1);
       
    24 by (simp_tac (rank_ss addsimps rank_te_owr1::TyEnv.case_eqns) 1);
       
    25 qed "te_rec_owr";
       
    26 
       
    27 Goalw [te_dom_def] "te_dom(te_emp) = 0";
       
    28 by (simp_tac (simpset() addsimps [te_rec_emp]) 1);
       
    29 qed "te_dom_emp";
       
    30 
       
    31 Goalw [te_dom_def] "te_dom(te_owr(te,x,v)) = te_dom(te) Un {x}";
       
    32 by (simp_tac (simpset() addsimps [te_rec_owr]) 1);
       
    33 qed "te_dom_owr";
       
    34 
       
    35 Goalw [te_app_def] "te_app(te_owr(te,x,t),x) = t";
       
    36 by (simp_tac (simpset() addsimps [te_rec_owr]) 1);
       
    37 qed "te_app_owr1";
    14 qed "te_app_owr1";
    38 
    15 
    39 Goalw [te_app_def]
    16 Goal "x ~= y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)";
    40   "x ~= y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)";
    17 by Auto_tac;
    41 by (asm_simp_tac (simpset() addsimps [te_rec_owr,(not_sym RS if_not_P)]) 1);
       
    42 qed "te_app_owr2";
    18 qed "te_app_owr2";
    43 
    19 
    44 Goal "[| te:TyEnv; x:ExVar; x:te_dom(te) |] ==> te_app(te,x):Ty";
    20 Goal "[| te:TyEnv; x:ExVar; x:te_dom(te) |] ==> te_app(te,x):Ty";
    45 by (res_inst_tac [("P","x:te_dom(te)")] impE 1);
    21 by (res_inst_tac [("P","x:te_dom(te)")] impE 1);
    46 by (assume_tac 2);
    22 by (assume_tac 2);
    47 by (assume_tac 2);
    23 by (assume_tac 2);
    48 by (etac TyEnv.induct 1);
    24 by (etac TyEnv.induct 1);
    49 by (simp_tac (simpset() addsimps [te_dom_emp]) 1);
    25 by (Simp_tac 1);
    50 by (rtac impI 1);
    26 by (case_tac "xa = x" 1);
    51 by (rtac (excluded_middle RS disjE) 1);
    27 by Auto_tac;
    52 by (stac te_app_owr2 1);
       
    53 by (assume_tac 1);
       
    54 by (asm_full_simp_tac (simpset() addsimps [te_dom_owr]) 1);
       
    55 by (Fast_tac 1);
       
    56 by (asm_simp_tac (simpset() addsimps [te_app_owr1]) 1);
       
    57 qed "te_appI";
    28 qed "te_appI";
    58 
    29 
    59 
    30 
    60 
    31 
    61 
    32