author | oheimb |
Fri, 20 Feb 1998 16:00:18 +0100 | |
changeset 4637 | bac998af6ea2 |
parent 4089 | 96fba19bcbe2 |
child 4686 | 74a12e86b20b |
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 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
11 |
goal thy |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
12 |
"! a m s s' t n. \ |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
13 |
\ (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
|
14 |
\ ( ? r. W e ($ s a) m = Ok(r, $ s' t, n) & s' = ($ r o s) )"; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
15 |
by (expr.induct_tac "e" 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
16 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
17 |
(* case Var n *) |
4089 | 18 |
by (simp_tac (simpset() addsimps [app_subst_list] |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
19 |
setloop (split_inside_tac [expand_if])) 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
20 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
21 |
(* case Abs e *) |
4089 | 22 |
by (asm_full_simp_tac (simpset() setloop (split_inside_tac [expand_bind])) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
23 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
24 |
by (rtac conjI 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
25 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
26 |
by (REPEAT (etac allE 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
27 |
by (etac impE 1); |
4089 | 28 |
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
|
29 |
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
|
30 |
less_imp_le,lessI]) 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
31 |
(** LEVEL 10 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
32 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
33 |
by (REPEAT (etac allE 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
34 |
by (etac impE 1); |
4089 | 35 |
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
|
36 |
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
|
37 |
less_imp_le,lessI]) 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
38 |
(** LEVEL 15 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
39 |
(* case App e1 e2 *) |
4089 | 40 |
by (simp_tac (simpset() setloop (split_inside_tac [expand_bind])) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
41 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
42 |
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
|
43 |
by (rtac conjI 1); |
4089 | 44 |
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
|
45 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
46 |
by (rename_tac "s1 t1' n1'" 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
47 |
by (eres_inst_tac [("x","a")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
48 |
by (eres_inst_tac [("x","m")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
49 |
by (eres_inst_tac [("x","s")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
50 |
by (eres_inst_tac [("x","s1'")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
51 |
by (eres_inst_tac [("x","t1")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
52 |
by (eres_inst_tac [("x","n1")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
53 |
by (eres_inst_tac [("x","a")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
54 |
by (eres_inst_tac [("x","n1")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
55 |
by (eres_inst_tac [("x","s1'")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
56 |
by (eres_inst_tac [("x","s2'")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
57 |
by (eres_inst_tac [("x","t2")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
58 |
by (eres_inst_tac [("x","n2")] allE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
59 |
(** LEVEL 34 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
60 |
by (rtac conjI 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
61 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
62 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
63 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
64 |
by (etac exE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
65 |
by (etac conjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
66 |
by (etac impE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
67 |
by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
68 |
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
|
69 |
by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] |
4089 | 70 |
addss simpset()) 1); |
71 |
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
|
72 |
(** LEVEL 45 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
73 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
74 |
by (rename_tac "s2 t2' n2'" 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
75 |
by (rtac conjI 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
76 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
77 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
78 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
79 |
by (etac exE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
80 |
by (etac conjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
81 |
by (etac impE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
82 |
by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
83 |
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
|
84 |
by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] |
4089 | 85 |
addss simpset()) 1); |
86 |
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
|
87 |
by (strip_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
88 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
89 |
(** LEVEL 60 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
90 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
91 |
by (etac exE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
92 |
by (etac conjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
93 |
by (etac impE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
94 |
by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
95 |
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
|
96 |
by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] |
4089 | 97 |
addss simpset()) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
98 |
by (mp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
99 |
by (REPEAT (eresolve_tac [exE,conjE] 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
100 |
by (REPEAT(EVERY1 |
4089 | 101 |
[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
|
102 |
REPEAT o etac conjE, hyp_subst_tac])); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
103 |
(** LEVEL 70 **) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
104 |
by (subgoal_tac "new_tv n2 s & new_tv n2 r & new_tv n2 ra" 1); |
4089 | 105 |
by (asm_full_simp_tac (simpset() addsimps [new_tv_subst]) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
106 |
by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
107 |
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
|
108 |
by (safe_tac HOL_cs); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
109 |
by (best_tac (HOL_cs addDs[sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] |
4089 | 110 |
addss simpset()) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
111 |
by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] |
4089 | 112 |
addss simpset()) 1); |
2793
b30c41754c86
Modified proofs because simplifier does not eta-contract any longer.
nipkow
parents:
2637
diff
changeset
|
113 |
(** LEVEL 77 **) |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
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
|
118 |
by (best_tac (HOL_cs addDs [new_tv_W] |
4089 | 119 |
addss (simpset() addsimps [subst_comp_tel])) 1); |
2793
b30c41754c86
Modified proofs because simplifier does not eta-contract any longer.
nipkow
parents:
2637
diff
changeset
|
120 |
(** LEVEL 82 **) |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
121 |
qed_spec_mp "I_correct_wrt_W"; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
122 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
123 |
(*** |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
124 |
We actually want the corollary |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
125 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
126 |
goal I.thy |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
127 |
"I e [] m id_subst = Ok(s,t,n) --> W e [] m = Ok(s, $s t, n)"; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
128 |
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
|
129 |
(read_instantiate[("x","[]")](thm RS spec) |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
130 |
RS spec RS spec))] 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
131 |
by (Full_simp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
132 |
by (fast_tac HOL_cs 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
133 |
qed; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
134 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
135 |
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
|
136 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
137 |
Wait until simplification of thms is possible. |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
138 |
***) |
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 |
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
|
141 |
|
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
142 |
Addsimps [split_paired_Ex]; |
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
143 |
|
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
144 |
goal I.thy "!a m s. \ |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
145 |
\ new_tv m a & new_tv m s --> I e a m s = Fail --> W e ($s a) m = Fail"; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
146 |
by (expr.induct_tac "e" 1); |
4089 | 147 |
by (simp_tac (simpset() addsimps [app_subst_list] addsplits [expand_if]) 1); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
148 |
by (Simp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
149 |
by (strip_tac 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
150 |
by (subgoal_tac "TVar m # $ s a = $s(TVar m # a)" 1); |
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
151 |
by (asm_simp_tac (HOL_ss addsimps |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
152 |
[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:
3207
diff
changeset
|
153 |
by (etac conjE 1); |
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
154 |
by (dtac (new_tv_not_free_tv RS not_free_impl_id) 1); |
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
155 |
by (Asm_simp_tac 1); |
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
156 |
(** LEVEL 9 **) |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
157 |
by (Asm_simp_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
158 |
by (strip_tac 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
159 |
by (REPEAT(etac exE 1)); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
160 |
by (REPEAT(etac conjE 1)); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
161 |
by (dtac lemma 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
162 |
by (fast_tac HOL_cs 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
163 |
(** LEVEL 15 **) |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
164 |
by (etac exE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
165 |
by (etac conjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
166 |
by (hyp_subst_tac 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
167 |
by (Asm_simp_tac 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
168 |
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
|
169 |
by (etac disjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
170 |
by (rtac disjI1 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
171 |
(** LEVEL 22 **) |
4089 | 172 |
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
|
173 |
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
|
174 |
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
|
175 |
by (rtac conjI 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
176 |
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
|
177 |
by (rtac new_tv_subst_comp_2 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
178 |
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
|
179 |
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
|
180 |
by (rtac disjI2 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
181 |
by (REPEAT(etac exE 1)); |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
182 |
by (etac conjE 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
183 |
(** LEVEL 32 **) |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
184 |
by (dtac lemma 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
185 |
by (rtac conjI 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
186 |
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
|
187 |
by (rtac new_tv_subst_comp_1 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
188 |
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
|
189 |
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
|
190 |
by (etac exE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
191 |
by (etac conjE 1); |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
192 |
by (hyp_subst_tac 1); |
3569
4467015d5080
Simplified a few proofs because of improved simplification.
nipkow
parents:
3207
diff
changeset
|
193 |
(** LEVEL 41 **) |
4089 | 194 |
by (asm_full_simp_tac (simpset() addsimps |
2518
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
195 |
[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
|
196 |
qed_spec_mp "I_complete_wrt_W"; |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
197 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
198 |
(*** |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
199 |
We actually want the corollary |
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 |
"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
|
202 |
|
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
203 |
Wait until simplification of thms is possible. |
bee082efaa46
This is the old version og MiniML for the monomorphic case.
nipkow
parents:
diff
changeset
|
204 |
***) |