src/HOL/MiniML/Instance.ML
author oheimb
Mon, 21 Sep 1998 23:12:31 +0200
changeset 5521 7970832271cc
parent 5350 165b9c212caf
child 5983 79e301a6a51b
permissions -rw-r--r--
added wrapper for bspec
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     1
(* Title:     HOL/MiniML/Instance.ML
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     2
   ID:        $Id$
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     3
   Author:    Wolfgang Naraschewski and Tobias Nipkow
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     4
   Copyright  1996 TU Muenchen
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     5
*)
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     6
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     7
(* lemmatas for instatiation *)
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     8
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
     9
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    10
(* lemmatas for bound_typ_inst *)
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    11
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    12
Goal "bound_typ_inst S (mk_scheme t) = t";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    13
by (induct_tac "t" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    14
by (ALLGOALS Asm_simp_tac);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    15
qed "bound_typ_inst_mk_scheme";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    16
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    17
Addsimps [bound_typ_inst_mk_scheme];
2625
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    18
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    19
Goal "bound_typ_inst ($S o R) ($S sch) = $S (bound_typ_inst R sch)";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    20
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    21
by (ALLGOALS Asm_full_simp_tac);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    22
qed "bound_typ_inst_composed_subst";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    23
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    24
Addsimps [bound_typ_inst_composed_subst];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    25
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    26
Goal "S = S' ==> sch = sch' ==> bound_typ_inst S sch = bound_typ_inst S' sch'";
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    27
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    28
qed "bound_typ_inst_eq";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    29
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    30
2625
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    31
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    32
(* lemmatas for bound_scheme_inst *)
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    33
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    34
Goal "bound_scheme_inst B (mk_scheme t) = mk_scheme t";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    35
by (induct_tac "t" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    36
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    37
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    38
qed "bound_scheme_inst_mk_scheme";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    39
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    40
Addsimps [bound_scheme_inst_mk_scheme];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    41
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    42
Goal "$S (bound_scheme_inst B sch) = (bound_scheme_inst ($S o B) ($ S sch))";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    43
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    44
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    45
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    46
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    47
qed "substitution_lemma";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    48
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    49
Goal "!t. mk_scheme t = bound_scheme_inst B sch --> \
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    50
\         (? S. !x:bound_tv sch. B x = mk_scheme (S x))";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    51
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    52
by (Simp_tac 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
    53
by Safe_tac;
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    54
by (rtac exI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    55
by (rtac ballI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    56
by (rtac sym 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    57
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    58
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    59
by (dtac mk_scheme_Fun 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    60
by (REPEAT (etac exE 1));
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    61
by (etac conjE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    62
by (dtac sym 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    63
by (dtac sym 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    64
by (REPEAT ((dtac mp 1) THEN (Fast_tac 1)));
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
    65
by Safe_tac;
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    66
by (rename_tac "S1 S2" 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    67
by (res_inst_tac [("x","%x. if x:bound_tv type_scheme1 then (S1 x) else (S2 x)")] exI 1);
5521
7970832271cc added wrapper for bspec
oheimb
parents: 5350
diff changeset
    68
by (Auto_tac);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    69
qed_spec_mp "bound_scheme_inst_type";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    70
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    71
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    72
(* lemmas for subst_to_scheme *)
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    73
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    74
Goal "new_tv n sch --> subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) \
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    75
\                                                 (bound_typ_inst (%k. TVar (k + n)) sch) = sch";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    76
by (induct_tac "sch" 1);
5350
165b9c212caf Must remove leD from simpset
paulson
parents: 5184
diff changeset
    77
by (simp_tac (simpset() addsimps [le_def]) 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4641
diff changeset
    78
by (simp_tac (simpset() addsimps [le_add2,diff_add_inverse2]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    79
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    80
qed_spec_mp "subst_to_scheme_inverse";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    81
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    82
Goal "t = t' ==> \
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    83
\     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t = \
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    84
\     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) t'";
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    85
by (Fast_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    86
val aux = result ();
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    87
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    88
Goal "new_tv n sch --> \
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    89
\     subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k) (bound_typ_inst S sch) = \
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    90
\      bound_scheme_inst ((subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k)) o S) sch";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
    91
by (induct_tac "sch" 1);
5350
165b9c212caf Must remove leD from simpset
paulson
parents: 5184
diff changeset
    92
by Auto_tac;
165b9c212caf Must remove leD from simpset
paulson
parents: 5184
diff changeset
    93
by (ALLGOALS trans_tac);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    94
val aux2 = result () RS mp;
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    95
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    96
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    97
(* lemmata for <= *)
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
    98
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
    99
Goalw [le_type_scheme_def,is_bound_typ_instance]
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   100
  "!!(sch::type_scheme) sch'. \
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   101
\  (sch' <= sch) = (? B. sch' = bound_scheme_inst B sch)";
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   102
by (rtac iffI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   103
by (cut_inst_tac [("sch","sch")] fresh_variable_type_schemes 1); 
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   104
by (cut_inst_tac [("sch","sch'")] fresh_variable_type_schemes 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   105
by (dtac make_one_new_out_of_two 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   106
by (assume_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   107
by (thin_tac "? n. new_tv n sch'" 1); 
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   108
by (etac exE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   109
by (etac allE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   110
by (dtac mp 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   111
by (res_inst_tac [("x","(%k. TVar (k + n))")] exI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   112
by (rtac refl 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   113
by (etac exE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   114
by (REPEAT (etac conjE 1));
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   115
by (dres_inst_tac [("n","n")] aux 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   116
by (asm_full_simp_tac (simpset() addsimps [subst_to_scheme_inverse]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   117
by (res_inst_tac [("x","(subst_to_scheme (%k. if n <= k then BVar (k - n) else FVar k)) o S")] exI 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   118
by (asm_simp_tac (simpset() addsimps [aux2]) 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   119
by Safe_tac;
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   120
by (res_inst_tac [("x","%n. bound_typ_inst S (B n)")] exI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   121
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   122
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   123
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   124
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   125
qed "le_type_scheme_def2";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   126
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   127
Goalw [is_bound_typ_instance] "(mk_scheme t) <= sch = t <| sch";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   128
by (simp_tac (simpset() addsimps [le_type_scheme_def2]) 1); 
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   129
by (rtac iffI 1); 
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   130
by (etac exE 1); 
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   131
by (forward_tac [bound_scheme_inst_type] 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   132
by (etac exE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   133
by (rtac exI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   134
by (rtac mk_scheme_injective 1); 
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   135
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   136
by (rotate_tac 1 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   137
by (rtac mp 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   138
by (assume_tac 2);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   139
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   140
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   141
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   142
by (Fast_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   143
by (strip_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   144
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   145
by (etac exE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   146
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   147
by (rtac exI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   148
by (induct_tac "sch" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   149
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   150
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   151
by (Asm_full_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   152
qed_spec_mp "le_type_eq_is_bound_typ_instance";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   153
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   154
Goalw [le_env_def]
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   155
  "(sch # A <= sch' # B) = (sch <= (sch'::type_scheme) & A <= B)";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   156
by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   157
by (rtac iffI 1);
4153
e534c4c32d54 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4089
diff changeset
   158
 by (SELECT_GOAL Safe_tac 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   159
  by (eres_inst_tac [("x","0")] allE 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   160
  by (Asm_full_simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   161
 by (eres_inst_tac [("x","Suc i")] allE 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   162
 by (Asm_full_simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   163
by (rtac conjI 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   164
 by (Fast_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   165
by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   166
by (induct_tac "i" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   167
by (ALLGOALS Asm_simp_tac);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   168
qed "le_env_Cons";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   169
AddIffs [le_env_Cons];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   170
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   171
Goalw [is_bound_typ_instance]"t <| sch ==> $S t <| $S sch";
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   172
by (etac exE 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   173
by (rename_tac "SA" 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   174
by (hyp_subst_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   175
by (res_inst_tac [("x","$S o SA")] exI 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   176
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   177
qed "is_bound_typ_instance_closed_subst";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   178
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   179
Goal "!!(sch::type_scheme) sch'. sch' <= sch ==> $S sch' <= $ S sch";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   180
by (asm_full_simp_tac (simpset() addsimps [le_type_scheme_def2]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   181
by (etac exE 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   182
by (asm_full_simp_tac (simpset() addsimps [substitution_lemma]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   183
by (Fast_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   184
qed "S_compatible_le_scheme";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   185
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   186
Goalw [le_env_def,app_subst_list]
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   187
 "!!(A::type_scheme list) A'. A' <= A ==> $S A' <= $ S A";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   188
by (simp_tac (simpset() addcongs [conj_cong]) 1);
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   189
by (fast_tac (claset() addSIs [S_compatible_le_scheme]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   190
qed "S_compatible_le_scheme_lists";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   191
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   192
Goalw [le_type_scheme_def] "[| t <| sch; sch <= sch' |] ==> t <| sch'";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   193
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   194
qed "bound_typ_instance_trans";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   195
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   196
Goalw [le_type_scheme_def] "sch <= (sch::type_scheme)";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   197
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   198
qed "le_type_scheme_refl";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   199
AddIffs [le_type_scheme_refl];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   200
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   201
Goalw [le_env_def] "A <= (A::type_scheme list)";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   202
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   203
qed "le_env_refl";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   204
AddIffs [le_env_refl];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   205
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   206
Goalw [le_type_scheme_def,is_bound_typ_instance] "sch <= BVar n";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   207
by (strip_tac 1);
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3018
diff changeset
   208
by (res_inst_tac [("x","%a. t")]exI 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   209
by (Simp_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   210
qed "bound_typ_instance_BVar";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   211
AddIffs [bound_typ_instance_BVar];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   212
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   213
Goalw [le_type_scheme_def,is_bound_typ_instance]
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   214
 "(sch <= FVar n) = (sch = FVar n)";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   215
by (induct_tac "sch" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   216
  by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   217
 by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   218
 by (Fast_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   219
by (Asm_full_simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   220
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   221
qed "le_FVar";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   222
Addsimps [le_FVar];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   223
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   224
Goalw [le_type_scheme_def,is_bound_typ_instance] "~(FVar n <= sch1 =-> sch2)";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   225
by (Simp_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   226
qed "not_FVar_le_Fun";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   227
AddIffs [not_FVar_le_Fun];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   228
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   229
Goalw [le_type_scheme_def,is_bound_typ_instance] "~(BVar n <= sch1 =-> sch2)";
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   230
by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   231
by (res_inst_tac [("x","TVar n")] exI 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   232
by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   233
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   234
qed "not_BVar_le_Fun";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   235
AddIffs [not_BVar_le_Fun];
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   236
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   237
Goalw [le_type_scheme_def,is_bound_typ_instance]
5118
6b995dad8a9d Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
   238
  "(sch1 =-> sch2 <= sch1' =-> sch2') ==> sch1 <= sch1' & sch2 <= sch2'";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   239
by (fast_tac (claset() addss simpset()) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   240
qed "Fun_le_FunD";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   241
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   242
Goal "(sch' <= sch1 =-> sch2) --> (? sch'1 sch'2. sch' = sch'1 =-> sch'2)";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   243
by (induct_tac "sch'" 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   244
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   245
by (Asm_simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   246
by (Fast_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   247
qed_spec_mp "scheme_le_Fun";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   248
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   249
Goal "!sch'::type_scheme. sch <= sch' --> free_tv sch' <= free_tv sch";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   250
by (induct_tac "sch" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   251
  by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   252
  by (induct_tac "sch'" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   253
    by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   254
   by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   255
  by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   256
 by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   257
 by (induct_tac "sch'" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   258
   by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   259
  by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   260
 by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   261
by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   262
by (induct_tac "sch'" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   263
  by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   264
 by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   265
by (Asm_full_simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   266
by (strip_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   267
by (dtac Fun_le_FunD 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   268
by (Fast_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   269
qed_spec_mp "le_type_scheme_free_tv";
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   270
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4686
diff changeset
   271
Goal "!A::type_scheme list. A <= B --> free_tv B <= free_tv A";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   272
by (induct_tac "B" 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   273
 by (Simp_tac 1);
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   274
by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5118
diff changeset
   275
by (induct_tac "A" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   276
 by (simp_tac (simpset() addsimps [le_env_def]) 1);
3018
e65b60b28341 Ran expandshort
paulson
parents: 2625
diff changeset
   277
by (Simp_tac 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   278
by (fast_tac (claset() addDs [le_type_scheme_free_tv]) 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents:
diff changeset
   279
qed_spec_mp "le_env_free_tv";