| author | wenzelm | 
| Wed, 20 Feb 2002 00:54:54 +0100 | |
| changeset 12903 | 58da1dc2720c | 
| parent 11232 | 558a4feebb04 | 
| permissions | -rw-r--r-- | 
| 2520 | 1 | (* Title: HOL/W0/I.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Thomas Stauner | |
| 4 | Copyright 1995 TU Muenchen | |
| 5 | ||
| 6 | Equivalence of W and I. | |
| 7 | *) | |
| 8 | ||
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 9 | open I; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 10 | |
| 5349 | 11 | Goal "! a m s s' t n. \ | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 12 | \ (new_tv m a & new_tv m s) --> I e a m s = Ok(s',t,n) --> \ | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 13 | \ ( ? r. W e ($ s a) m = Ok(r, $ s' t, n) & s' = ($ r o s) )"; | 
| 5184 | 14 | by (induct_tac "e" 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 15 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 16 | (* case Var n *) | 
| 4089 | 17 | by (simp_tac (simpset() addsimps [app_subst_list] | 
| 4831 | 18 | setloop (split_inside_tac [split_if])) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 19 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 20 | (* case Abs e *) | 
| 4831 | 21 | by (asm_full_simp_tac (simpset() setloop (split_inside_tac [split_bind])) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 22 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 23 | by (rtac conjI 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 24 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 25 | by (REPEAT (etac allE 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 26 | by (etac impE 1); | 
| 4089 | 27 | by (fast_tac (HOL_cs addss (simpset() addsimps [new_tv_subst])) 2); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 28 | by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 29 | less_imp_le,lessI]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 30 | (** LEVEL 10 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 31 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 32 | by (REPEAT (etac allE 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 33 | by (etac impE 1); | 
| 4089 | 34 | by (fast_tac (HOL_cs addss (simpset() addsimps [new_tv_subst])) 2); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 35 | by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 36 | less_imp_le,lessI]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 37 | (** LEVEL 15 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 38 | (* case App e1 e2 *) | 
| 4831 | 39 | by (simp_tac (simpset() setloop (split_inside_tac [split_bind])) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 40 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 41 | by (rename_tac "s1' t1 n1 s2' t2 n2 sa" 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 42 | by (rtac conjI 1); | 
| 4089 | 43 | by (fast_tac (HOL_cs addss simpset()) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 44 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 45 | by (rename_tac "s1 t1' n1'" 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 46 | by (eres_inst_tac [("x","a")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 47 | by (eres_inst_tac [("x","m")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 48 | by (eres_inst_tac [("x","s")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 49 | by (eres_inst_tac [("x","s1'")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 50 | by (eres_inst_tac [("x","t1")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 51 | by (eres_inst_tac [("x","n1")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 52 | by (eres_inst_tac [("x","a")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 53 | by (eres_inst_tac [("x","n1")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 54 | by (eres_inst_tac [("x","s1'")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 55 | by (eres_inst_tac [("x","s2'")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 56 | by (eres_inst_tac [("x","t2")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 57 | by (eres_inst_tac [("x","n2")] allE 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 58 | (** LEVEL 34 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 59 | by (rtac conjI 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 60 | by (strip_tac 1); | 
| 11232 | 61 | by (rtac notI 1); | 
| 62 | by (Asm_full_simp_tac 1); | |
| 63 | (* by (mp_tac 1); | |
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 64 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 65 | by (etac exE 1); | 
| 11232 | 66 | by (etac conjE 1);*) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 67 | by (etac impE 1); | 
| 7499 | 68 | by ((ftac new_tv_subst_tel 1) THEN (atac 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 69 |   by ((dres_inst_tac [("a","$ s a")] new_tv_W 1) THEN (atac 1));
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 70 | by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] | 
| 4089 | 71 | addss simpset()) 1); | 
| 72 | by (fast_tac (HOL_cs addss (simpset() addsimps [subst_comp_tel])) 1); | |
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 73 | (** LEVEL 45 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 74 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 75 | by (rename_tac "s2 t2' n2'" 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 76 | by (rtac conjI 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 77 | by (strip_tac 1); | 
| 11232 | 78 | by (rtac notI 1); | 
| 79 | by (Asm_full_simp_tac 1); | |
| 80 | (* | |
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 81 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 82 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 83 | by (etac exE 1); | 
| 11232 | 84 | by (etac conjE 1);*) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 85 | by (etac impE 1); | 
| 7499 | 86 | by ((ftac new_tv_subst_tel 1) THEN (atac 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 87 |   by ((dres_inst_tac [("a","$ s a")] new_tv_W 1) THEN (atac 1));
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 88 | by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] | 
| 4089 | 89 | addss simpset()) 1); | 
| 90 | by (fast_tac (HOL_cs addss (simpset() addsimps [subst_comp_tel,subst_comp_te])) 1); | |
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 91 | by (strip_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 92 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 93 | (** LEVEL 60 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 94 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 95 | by (etac exE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 96 | by (etac conjE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 97 | by (etac impE 1); | 
| 7499 | 98 | by ((ftac new_tv_subst_tel 1) THEN (atac 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 99 |  by ((dres_inst_tac [("a","$ s a")] new_tv_W 1) THEN (atac 1));
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 100 | by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] | 
| 4089 | 101 | addss simpset()) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 102 | by (mp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 103 | by (REPEAT (eresolve_tac [exE,conjE] 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 104 | by (REPEAT(EVERY1 | 
| 4089 | 105 | [asm_full_simp_tac (simpset() addsimps [subst_comp_tel,subst_comp_te,o_def]), | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 106 | REPEAT o etac conjE, hyp_subst_tac])); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 107 | (** LEVEL 70 **) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 108 | by (subgoal_tac "new_tv n2 s & new_tv n2 r & new_tv n2 ra" 1); | 
| 4089 | 109 | by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); | 
| 7499 | 110 | by ((ftac new_tv_subst_tel 1) THEN (atac 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 111 | by ((dres_inst_tac [("a","$ s a")] new_tv_W 1) THEN (atac 1));
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 112 | by (safe_tac HOL_cs); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 113 | by (best_tac (HOL_cs addDs[sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] | 
| 4089 | 114 | addss simpset()) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 115 | by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] | 
| 4089 | 116 | addss simpset()) 1); | 
| 2793 
b30c41754c86
Modified proofs because simplifier does not eta-contract any longer.
 nipkow parents: 
2637diff
changeset | 117 | (** LEVEL 77 **) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 118 | by (dres_inst_tac [("e","expr1")] (sym RS W_var_geD) 1);
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 119 | by ((dtac new_tv_subst_tel 1) THEN (atac 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 120 | by ((dres_inst_tac [("ts","$ s a")] new_tv_list_le 1) THEN (atac 1));
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 121 | by ((dtac new_tv_subst_tel 1) THEN (atac 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 122 | by (best_tac (HOL_cs addDs [new_tv_W] | 
| 4089 | 123 | addss (simpset() addsimps [subst_comp_tel])) 1); | 
| 2793 
b30c41754c86
Modified proofs because simplifier does not eta-contract any longer.
 nipkow parents: 
2637diff
changeset | 124 | (** LEVEL 82 **) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 125 | qed_spec_mp "I_correct_wrt_W"; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 126 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 127 | (*** | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 128 | We actually want the corollary | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 129 | |
| 5349 | 130 | Goal "I e [] m id_subst = Ok(s,t,n) --> W e [] m = Ok(s, $s t, n)"; | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 131 | by (cut_facts_tac [(read_instantiate[("x","id_subst")]
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 132 |  (read_instantiate[("x","[]")](thm RS spec)
 | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 133 | RS spec RS spec))] 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 134 | by (Full_simp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 135 | by (fast_tac HOL_cs 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 136 | qed; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 137 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 138 | assuming that thm is the undischarged version of I_correct_wrt_W. | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 139 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 140 | Wait until simplification of thms is possible. | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 141 | ***) | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 142 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 143 | val lemma = I_correct_wrt_W COMP swap_prems_rl; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 144 | |
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 145 | Addsimps [split_paired_Ex]; | 
| 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 146 | |
| 5069 | 147 | Goal "!a m s. \ | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 148 | \ new_tv m a & new_tv m s --> I e a m s = Fail --> W e ($s a) m = Fail"; | 
| 5184 | 149 | by (induct_tac "e" 1); | 
| 4686 | 150 | by (simp_tac (simpset() addsimps [app_subst_list]) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 151 | by (Simp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 152 | by (strip_tac 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 153 | by (subgoal_tac "TVar m # $ s a = $s(TVar m # a)" 1); | 
| 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 154 | by (asm_simp_tac (HOL_ss addsimps | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 155 | [new_tv_Suc_list, lessI RS less_imp_le RS new_tv_subst_le]) 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 156 | by (etac conjE 1); | 
| 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 157 | by (dtac (new_tv_not_free_tv RS not_free_impl_id) 1); | 
| 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 158 | by (Asm_simp_tac 1); | 
| 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 159 | (** LEVEL 9 **) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 160 | by (Asm_simp_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 161 | by (strip_tac 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 162 | by (REPEAT(etac exE 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 163 | by (REPEAT(etac conjE 1)); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 164 | by (dtac lemma 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 165 | by (fast_tac HOL_cs 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 166 | (** LEVEL 15 **) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 167 | by (etac exE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 168 | by (etac conjE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 169 | by (hyp_subst_tac 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 170 | by (Asm_simp_tac 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 171 | by (REPEAT(resolve_tac [exI,conjI,refl] 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 172 | by (etac disjE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 173 | by (rtac disjI1 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 174 | (** LEVEL 22 **) | 
| 4089 | 175 | by (full_simp_tac (simpset() addsimps [o_def,subst_comp_tel]) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 176 | by (EVERY[etac allE 1, etac allE 1, etac allE 1, | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 177 | etac impE 1, etac impE 2, atac 2, atac 2]); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 178 | by (rtac conjI 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 179 | by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 180 | by (rtac new_tv_subst_comp_2 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 181 | by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 182 | by (fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 183 | by (rtac disjI2 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 184 | by (REPEAT(etac exE 1)); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 185 | by (etac conjE 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 186 | (** LEVEL 32 **) | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 187 | by (dtac lemma 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 188 | by (rtac conjI 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 189 | by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 190 | by (rtac new_tv_subst_comp_1 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 191 | by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 192 | by (fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 193 | by (etac exE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 194 | by (etac conjE 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 195 | by (hyp_subst_tac 1); | 
| 3569 
4467015d5080
Simplified a few proofs because of improved simplification.
 nipkow parents: 
3207diff
changeset | 196 | (** LEVEL 41 **) | 
| 4089 | 197 | by (asm_full_simp_tac (simpset() addsimps | 
| 2518 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 198 | [o_def,subst_comp_te RS sym,subst_comp_tel RS sym]) 1); | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 199 | qed_spec_mp "I_complete_wrt_W"; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 200 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 201 | (*** | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 202 | We actually want the corollary | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 203 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 204 | "I e [] m id_subst = Fail ==> W e [] m = Fail"; | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 205 | |
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 206 | Wait until simplification of thms is possible. | 
| 
bee082efaa46
This is the old version og MiniML for the monomorphic case.
 nipkow parents: diff
changeset | 207 | ***) |