src/HOL/MiniML/W.ML
author paulson
Thu, 20 Jun 1996 10:27:29 +0200
changeset 1818 ffc20ff80190
parent 1669 e56cdf711729
child 1910 6d572f96fb76
permissions -rw-r--r--
Restored Addsimps [Suc_le_lessD], as it cannot be added to base HOL
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     1
(* Title:     HOL/MiniML/W.ML
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     2
   ID:        $Id$
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     3
   Author:    Dieter Nazareth and Tobias Nipkow
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     4
   Copyright  1995 TU Muenchen
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     5
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     6
Correctness and completeness of type inference algorithm W
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     7
*)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     8
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     9
open W;
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    10
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    11
Addsimps [Suc_le_lessD];
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    12
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    13
(* correctness of W with respect to has_type *)
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
    14
goal W.thy
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
    15
        "!a s t m n . Ok (s,t,m) = W e a n --> $s a |- e :: t";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    16
by (expr.induct_tac "e" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    17
(* case Var n *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    18
by (asm_simp_tac (!simpset setloop (split_tac [expand_if])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    19
(* case Abs e *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    20
by (asm_full_simp_tac (!simpset addsimps [app_subst_list]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    21
                        setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    22
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    23
by (eres_inst_tac [("x","TVar(n) # a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    24
by( fast_tac (HOL_cs addss (!simpset addsimps [eq_sym_conv])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    25
(* case App e1 e2 *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    26
by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    27
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    28
by( rename_tac "sa ta na sb tb nb sc" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    29
by (res_inst_tac [("t2.0","$ sc tb")] has_type.AppI 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    30
by (res_inst_tac [("s1","sc")] (app_subst_TVar RS subst) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    31
by (rtac (app_subst_Fun RS subst) 1);
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
    32
by (res_inst_tac [("t","$sc(tb -> (TVar nb))"),("s","$sc($sb ta)")] subst 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    33
by (Asm_full_simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    34
by (simp_tac (HOL_ss addsimps [subst_comp_tel RS sym]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    35
by ( (rtac has_type_cl_sub 1) THEN (rtac has_type_cl_sub 1));
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    36
by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1);
1669
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
    37
by (asm_full_simp_tac (!simpset addsimps [subst_comp_tel RS sym,o_def,has_type_cl_sub,eq_sym_conv]) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    38
qed_spec_mp "W_correct";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    39
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    40
val has_type_casesE = map(has_type.mk_cases expr.simps)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
    41
        [" s |- Var n :: t"," s |- Abs e :: t","s |- App e1 e2 ::t"];
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    42
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    43
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    44
(* the resulting type variable is always greater or equal than the given one *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    45
goal thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
    46
        "!a n s t m. W e a n  = Ok (s,t,m) --> n<=m";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    47
by (expr.induct_tac "e" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    48
(* case Var(n) *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    49
by (fast_tac (HOL_cs addss (!simpset setloop (split_tac [expand_if]))) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    50
(* case Abs e *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    51
by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    52
by (fast_tac (HOL_cs addDs [Suc_leD]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    53
(* case App e1 e2 *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    54
by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    55
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    56
by (rename_tac "s t na sa ta nb sb sc tb m" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    57
by (eres_inst_tac [("x","a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    58
by (eres_inst_tac [("x","n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    59
by (eres_inst_tac [("x","$ s a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    60
by (eres_inst_tac [("x","s")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    61
by (eres_inst_tac [("x","t")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    62
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    63
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    64
by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    65
by (etac conjE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    66
by (eres_inst_tac [("x","sa")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    67
by (eres_inst_tac [("x","ta")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    68
by (eres_inst_tac [("x","nb")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    69
by (etac conjE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    70
by (res_inst_tac [("j","na")] le_trans 1); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    71
by (Asm_simp_tac 1);
1669
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
    72
by (Asm_simp_tac 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    73
qed_spec_mp "W_var_ge";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    74
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    75
Addsimps [W_var_ge];
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    76
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    77
goal thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
    78
        "!! s. Ok (s,t,m) = W e a n ==> n<=m";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    79
by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    80
qed "W_var_geD";
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    81
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    82
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    83
(* auxiliary lemma *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    84
goal Maybe.thy "(y = Ok x) = (Ok x = y)";
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    85
by( simp_tac (!simpset addsimps [eq_sym_conv]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    86
qed "rotate_Ok";
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    87
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    88
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    89
(* resulting type variable is new *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    90
goal thy
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    91
     "!n a s t m. new_tv n a --> W e a n = Ok (s,t,m) -->    \
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
    92
\                 new_tv m s & new_tv m t";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    93
by (expr.induct_tac "e" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    94
(* case Var n *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    95
by (fast_tac (HOL_cs addss (!simpset 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
    96
        addsimps [id_subst_def,list_all_nth,new_tv_list,new_tv_subst] 
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    97
        setloop (split_tac [expand_if]))) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    98
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    99
(* case Abs e *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   100
by (simp_tac (!simpset addsimps [new_tv_subst,new_tv_Suc_list] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   101
    setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   102
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   103
by (eres_inst_tac [("x","Suc n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   104
by (eres_inst_tac [("x","(TVar n)#a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   105
by (fast_tac (HOL_cs addss (!simpset
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   106
              addsimps [new_tv_subst,new_tv_Suc_list])) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   107
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   108
(* case App e1 e2 *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   109
by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   110
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   111
by (rename_tac "s t na sa ta nb sb sc tb m" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   112
by (eres_inst_tac [("x","n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   113
by (eres_inst_tac [("x","a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   114
by (eres_inst_tac [("x","s")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   115
by (eres_inst_tac [("x","t")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   116
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   117
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   118
by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   119
by (eres_inst_tac [("x","$ s a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   120
by (eres_inst_tac [("x","sa")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   121
by (eres_inst_tac [("x","ta")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   122
by (eres_inst_tac [("x","nb")] allE 1);
1669
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
   123
by( asm_full_simp_tac (!simpset addsimps [o_def,rotate_Ok]) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   124
by (rtac conjI 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   125
by (rtac new_tv_subst_comp_2 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   126
by (rtac new_tv_subst_comp_2 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   127
by (rtac (lessI RS less_imp_le RS new_tv_subst_le) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   128
by (res_inst_tac [("n","na")] new_tv_subst_le 1); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   129
by (asm_full_simp_tac (!simpset addsimps [rotate_Ok]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   130
by (Asm_simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   131
by (fast_tac (HOL_cs addDs [W_var_geD] addIs
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   132
     [new_tv_list_le,new_tv_subst_tel,lessI RS less_imp_le RS new_tv_subst_le])
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   133
    1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   134
by (etac (sym RS mgu_new) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   135
by (fast_tac (HOL_cs addDs [W_var_geD] addIs [new_tv_subst_te,new_tv_list_le,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   136
   new_tv_subst_tel,lessI RS less_imp_le RS new_tv_le,lessI RS less_imp_le RS 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   137
   new_tv_subst_le,new_tv_le]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   138
by (fast_tac (HOL_cs addDs [W_var_geD] addIs
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   139
     [new_tv_list_le,new_tv_subst_tel,new_tv_le] 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   140
        addss (!simpset)) 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   141
by (rtac (lessI RS new_tv_subst_var) 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   142
by (etac (sym RS mgu_new) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   143
by (fast_tac (HOL_cs addSIs [lessI RS less_imp_le RS new_tv_le,new_tv_subst_te]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   144
   addDs [W_var_geD] addIs
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   145
   [new_tv_list_le,new_tv_subst_tel,lessI RS less_imp_le RS
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   146
   new_tv_subst_le,new_tv_le] addss !simpset) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   147
by (fast_tac (HOL_cs addDs [W_var_geD] addIs
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   148
     [new_tv_list_le,new_tv_subst_tel,new_tv_le]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   149
     addss (!simpset)) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   150
qed_spec_mp "new_tv_W";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   151
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   152
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   153
goal thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   154
     "!n a s t m v. W e a n = Ok (s,t,m) -->            \
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   155
\         (v:free_tv s | v:free_tv t) --> v<n --> v:free_tv a";
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   156
by (expr.induct_tac "e" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   157
(* case Var n *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   158
by (fast_tac (HOL_cs addIs [nth_mem,subsetD,ftv_mem_sub_ftv_list] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   159
    addss (!simpset setloop (split_tac [expand_if]))) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   160
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   161
(* case Abs e *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   162
by (asm_full_simp_tac (!simpset addsimps
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   163
    [free_tv_subst] setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   164
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   165
by (rename_tac "s t na sa ta m v" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   166
by (eres_inst_tac [("x","Suc n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   167
by (eres_inst_tac [("x","TVar n # a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   168
by (eres_inst_tac [("x","s")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   169
by (eres_inst_tac [("x","t")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   170
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   171
by (eres_inst_tac [("x","v")] allE 1);
1669
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
   172
by (fast_tac (HOL_cs addIs [cod_app_subst]
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
   173
                     addss (!simpset addsimps [less_Suc_eq])) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   174
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   175
(* case App e1 e2 *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   176
by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   177
by (strip_tac 1); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   178
by (rename_tac "s t na sa ta nb sb sc tb m v" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   179
by (eres_inst_tac [("x","n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   180
by (eres_inst_tac [("x","a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   181
by (eres_inst_tac [("x","s")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   182
by (eres_inst_tac [("x","t")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   183
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   184
by (eres_inst_tac [("x","na")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   185
by (eres_inst_tac [("x","v")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   186
(* second case *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   187
by (eres_inst_tac [("x","$ s a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   188
by (eres_inst_tac [("x","sa")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   189
by (eres_inst_tac [("x","ta")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   190
by (eres_inst_tac [("x","nb")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   191
by (eres_inst_tac [("x","v")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   192
by (safe_tac (empty_cs addSIs [conjI,impI] addSEs [conjE]) ); 
1669
e56cdf711729 inserted Suc_less_eq explictly in a few proofs.
nipkow
parents: 1525
diff changeset
   193
by (asm_full_simp_tac (!simpset addsimps [rotate_Ok,o_def]) 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   194
by (dtac W_var_geD 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   195
by (dtac W_var_geD 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   196
by ( (forward_tac [less_le_trans] 1) THEN (assume_tac 1) );
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   197
by (fast_tac (HOL_cs addDs [free_tv_comp_subst RS subsetD,sym RS mgu_free, 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   198
    codD,free_tv_app_subst_te RS subsetD,free_tv_app_subst_tel RS subsetD,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   199
    less_le_trans,less_not_refl2,subsetD]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   200
  addEs [UnE] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   201
  addss !simpset) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   202
by (Asm_full_simp_tac 1); 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   203
by (dtac (sym RS W_var_geD) 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   204
by (dtac (sym RS W_var_geD) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   205
by ( (forward_tac [less_le_trans] 1) THEN (assume_tac 1) );
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   206
by (fast_tac (HOL_cs addDs [mgu_free, codD,free_tv_subst_var RS subsetD,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   207
    free_tv_app_subst_te RS subsetD,free_tv_app_subst_tel RS subsetD,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   208
    less_le_trans,subsetD]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   209
  addEs [UnE]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   210
  addss !simpset) 1); 
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   211
qed_spec_mp "free_tv_W"; 
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   212
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   213
(* Completeness of W w.r.t. has_type *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   214
goal thy
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   215
 "!s' a t' n. $s' a |- e :: t' --> new_tv n a -->     \
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   216
\             (? s t. (? m. W e a n = Ok (s,t,m)) &  \
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   217
\                     (? r. $s' a = $r ($s a) & t' = $r t))";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   218
by (expr.induct_tac "e" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   219
(* case Var n *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   220
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   221
by (simp_tac (!simpset addcongs [conj_cong] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   222
    setloop (split_tac [expand_if])) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   223
by (eresolve_tac has_type_casesE 1); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   224
by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv,app_subst_list]) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   225
by (res_inst_tac [("x","id_subst")] exI 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   226
by (res_inst_tac [("x","nth nat a")] exI 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   227
by (Simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   228
by (res_inst_tac [("x","s'")] exI 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   229
by (Asm_simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   230
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   231
(* case Abs e *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   232
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   233
by (eresolve_tac has_type_casesE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   234
by (eres_inst_tac [("x","%x.if x=n then t1 else (s' x)")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   235
by (eres_inst_tac [("x","(TVar n)#a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   236
by (eres_inst_tac [("x","t2")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   237
by (eres_inst_tac [("x","Suc n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   238
by (fast_tac (HOL_cs addss (!simpset addcongs [conj_cong]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   239
    setloop (split_tac [expand_bind]))) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   240
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   241
(* case App e1 e2 *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   242
by (strip_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   243
by (eresolve_tac has_type_casesE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   244
by (eres_inst_tac [("x","s'")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   245
by (eres_inst_tac [("x","a")] allE 1);
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   246
by (eres_inst_tac [("x","t2 -> t'")] allE 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   247
by (eres_inst_tac [("x","n")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   248
by (safe_tac HOL_cs);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   249
by (eres_inst_tac [("x","r")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   250
by (eres_inst_tac [("x","$ s a")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   251
by (eres_inst_tac [("x","t2")] allE 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   252
by (eres_inst_tac [("x","m")] allE 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   253
by (dtac asm_rl 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   254
by (dtac asm_rl 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   255
by (dtac asm_rl 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   256
by (Asm_full_simp_tac 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   257
by (safe_tac HOL_cs);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   258
by (fast_tac HOL_cs 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   259
by (fast_tac (HOL_cs addIs [sym RS W_var_geD,new_tv_W RS
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   260
        conjunct1,new_tv_list_le,new_tv_subst_tel]) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   261
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   262
by (subgoal_tac
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   263
  "$ (%x.if x=ma then t' else (if x:(free_tv t - free_tv sa) then r x \
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   264
\        else ra x)) ($ sa t) = \
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   265
\  $ (%x.if x=ma then t' else (if x:(free_tv t - free_tv sa) then r x \
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   266
\        else ra x)) (ta -> (TVar ma))" 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   267
by (res_inst_tac [("t","$ (%x. if x = ma then t' else \
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   268
\   (if x:(free_tv t - free_tv sa) then r x else ra x)) ($ sa t)"),
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   269
    ("s","($ ra ta) -> t'")] ssubst 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   270
by (asm_simp_tac (!simpset addsimps [subst_comp_te]) 2);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   271
by (rtac eq_free_eq_subst_te 2);  
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   272
by (strip_tac 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   273
by (subgoal_tac "na ~=ma" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   274
by (fast_tac (HOL_cs addDs [new_tv_W,sym RS W_var_geD,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   275
    new_tv_not_free_tv,new_tv_le]) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   276
by (case_tac "na:free_tv sa" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   277
(* na ~: free_tv sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   278
by (asm_simp_tac (!simpset addsimps [not_free_impl_id]
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   279
    setloop (split_tac [expand_if])) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   280
(* na : free_tv sa *)
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   281
by (dres_inst_tac [("ts1","$ s a")] (subst_comp_tel RSN (2,trans)) 2);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   282
by (dtac eq_subst_tel_eq_free 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   283
by (fast_tac (HOL_cs addIs [free_tv_W,free_tv_le_new_tv] addDs [new_tv_W]) 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   284
by (Asm_simp_tac 2); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   285
by (case_tac "na:dom sa" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   286
(* na ~: dom sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   287
by (asm_full_simp_tac (!simpset addsimps [dom_def] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   288
    setloop (split_tac [expand_if])) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   289
(* na : dom sa *)
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   290
by (rtac eq_free_eq_subst_te 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   291
by (strip_tac 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   292
by (subgoal_tac "nb ~= ma" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   293
by ((forward_tac [new_tv_W] 3) THEN (atac 3));
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   294
by (etac conjE 3);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   295
by (dtac new_tv_subst_tel 3);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   296
by (fast_tac (HOL_cs addIs [new_tv_list_le] addDs [sym RS W_var_geD]) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   297
by (fast_tac (set_cs addDs [new_tv_W,new_tv_not_free_tv] addss 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   298
    (!simpset addsimps [cod_def,free_tv_subst])) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   299
by (fast_tac (set_cs addss (!simpset addsimps [cod_def,free_tv_subst] 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   300
    setloop (split_tac [expand_if]))) 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   301
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   302
by (Simp_tac 2);  
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   303
by (rtac eq_free_eq_subst_te 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   304
by (strip_tac 2 );
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   305
by (subgoal_tac "na ~= ma" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   306
by ((forward_tac [new_tv_W] 3) THEN (atac 3));
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   307
by (etac conjE 3);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   308
by (dtac (sym RS W_var_geD) 3);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   309
by (fast_tac (HOL_cs addDs [new_tv_list_le,new_tv_subst_tel,new_tv_W,new_tv_not_free_tv]) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   310
by (case_tac "na: free_tv t - free_tv sa" 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   311
(* case na ~: free_tv t - free_tv sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   312
by( asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 3);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   313
(* case na : free_tv t - free_tv sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   314
by( asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 2);
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   315
by (dres_inst_tac [("ts1","$ s a")] (subst_comp_tel RSN (2,trans)) 2);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   316
by (dtac eq_subst_tel_eq_free 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   317
by (fast_tac (HOL_cs addIs [free_tv_W,free_tv_le_new_tv] addDs [new_tv_W]) 2);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   318
by (asm_full_simp_tac (!simpset addsimps [free_tv_subst,dom_def,de_Morgan_disj]) 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   319
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   320
by (asm_simp_tac (!simpset setloop (split_tac [expand_bind])) 1); 
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   321
by (safe_tac HOL_cs );
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   322
by (dtac mgu_Ok 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   323
by( fast_tac (HOL_cs addss !simpset) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   324
by (REPEAT (resolve_tac [exI,conjI] 1));
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   325
by (fast_tac HOL_cs 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   326
by (fast_tac HOL_cs 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   327
by ((dtac mgu_mg 1) THEN (atac 1));
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   328
by (etac exE 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   329
by (res_inst_tac [("x","rb")] exI 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   330
by (rtac conjI 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   331
by (dres_inst_tac [("x","ma")] fun_cong 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   332
by( asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   333
by (simp_tac (!simpset addsimps [subst_comp_tel RS sym]) 1);
1400
5d909faf0e04 Introduced Monad syntax Pat := Val; Cont
nipkow
parents: 1300
diff changeset
   334
by (res_inst_tac [("ts2","($ sa ($ s a))")] ((subst_comp_tel RS sym) RSN 
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   335
    (2,trans)) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   336
by( asm_full_simp_tac (!simpset addsimps [o_def,eq_sym_conv]) 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   337
by (rtac eq_free_eq_subst_tel 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   338
by( safe_tac HOL_cs );
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   339
by (subgoal_tac "ma ~= na" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   340
by ((forward_tac [new_tv_W] 2) THEN (atac 2));
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   341
by (etac conjE 2);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   342
by (dtac new_tv_subst_tel 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   343
by (fast_tac (HOL_cs addIs [new_tv_list_le] addDs [sym RS W_var_geD]) 2);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   344
by (( forw_inst_tac [("n","m")] (sym RSN (2,new_tv_W)) 2) THEN (atac 2));
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   345
by (etac conjE 2);
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   346
by (dtac (free_tv_app_subst_tel RS subsetD) 2);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   347
by (fast_tac (set_cs addDs [W_var_geD,new_tv_list_le,codD,
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   348
    new_tv_not_free_tv]) 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   349
by (case_tac "na: free_tv t - free_tv sa" 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   350
(* case na ~: free_tv t - free_tv sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   351
by (asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 2);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   352
(* case na : free_tv t - free_tv sa *)
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   353
by (asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 1);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1400
diff changeset
   354
by (dtac (free_tv_app_subst_tel RS subsetD) 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   355
by (fast_tac (set_cs addDs [codD,subst_comp_tel RSN (2,trans),
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
   356
    eq_subst_tel_eq_free] addss ((!simpset addsimps 
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   357
    [de_Morgan_disj,free_tv_subst,dom_def]))) 1);
1525
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   358
qed_spec_mp "W_complete_lemma";
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   359
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   360
goal W.thy
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   361
 "!!e. [] |- e :: t' ==>  (? s t. (? m. W e [] n = Ok(s,t,m)) &  \
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   362
\                                 (? r. t' = $r t))";
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   363
by(cut_inst_tac [("a","[]"),("s'","id_subst"),("e","e"),("t'","t'")]
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   364
                W_complete_lemma 1);
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   365
by(ALLGOALS Asm_full_simp_tac);
d127436567d0 modified priorities in syntax
nipkow
parents: 1486
diff changeset
   366
qed "W_complete";