src/HOL/Lambda/Lambda.ML
author paulson
Thu, 18 Jan 1996 10:38:29 +0100
changeset 1444 23ceb1dc9755
parent 1302 ddfdcc9ce667
child 1465 5d7a7e439cec
permissions -rw-r--r--
trivial updates
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
     1
(*  Title:      HOL/Lambda/Lambda.ML
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     4
    Copyright   1995 TU Muenchen
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     5
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     6
Substitution-lemmas.  Most of the proofs, esp. those about natural numbers,
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     7
are ported from Ole Rasmussen's ZF development.  In ZF, m<=n is syntactic
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     8
sugar for m<Suc(n). In HOL <= is a separate operator. Hence we have to prove
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     9
some compatibility lemmas.
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    10
*)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    11
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    12
(*** Nat ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    13
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    14
goal Nat.thy "!!i. [| i < Suc j; j < k |] ==> i < k";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    15
br le_less_trans 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    16
ba 2;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    17
by(asm_full_simp_tac (!simpset addsimps [le_def]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    18
by(fast_tac (HOL_cs addEs [less_asym,less_anti_refl]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    19
qed "lt_trans1";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    20
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    21
goal Nat.thy "!!i. [| i < j; j < Suc(k) |] ==> i < k";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
be less_le_trans 1;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    23
by(asm_full_simp_tac (!simpset addsimps [le_def]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    24
by(fast_tac (HOL_cs addEs [less_asym,less_anti_refl]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    25
qed "lt_trans2";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    26
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    27
val major::prems = goal Nat.thy
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    28
  "[| i < Suc j; i < j ==> P; i = j ==> P |] ==> P";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    29
br (major RS lessE) 1;
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    30
by(ALLGOALS Asm_full_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    31
by(resolve_tac prems 1 THEN etac sym 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    32
by(fast_tac (HOL_cs addIs prems) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    33
qed "leqE";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    34
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    35
goal Arith.thy "!!i. Suc i < j ==> i < pred j ";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    36
by (resolve_tac [Suc_less_SucD] 1);
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    37
by (Asm_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    38
qed "lt_pred";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    39
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    40
goal Arith.thy "!!i. [| i < Suc j; k < i |] ==> pred i < j ";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    41
by (resolve_tac [Suc_less_SucD] 1);
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    42
by (Asm_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    43
qed "gt_pred";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    44
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    45
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    46
(*** Lambda ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    47
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    48
open Lambda;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    49
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    50
Delsimps [less_Suc_eq, subst_Var];
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    51
Addsimps ([if_not_P, not_less_eq] @ beta.intrs);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    52
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    53
(* don't add r_into_rtrancl! *)
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    54
val lambda_cs = trancl_cs addSIs beta.intrs;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    55
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    56
(*** Congruence rules for ->> ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    57
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    58
goal Lambda.thy "!!s. s ->> s' ==> Fun s ->> Fun s'";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    59
be rtrancl_induct 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    60
by (ALLGOALS(fast_tac (lambda_cs addIs [rtrancl_refl,rtrancl_into_rtrancl])));
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    61
qed "rtrancl_beta_Fun";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    62
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    63
goal Lambda.thy "!!s. s ->> s' ==> s @ t ->> s' @ t";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    64
be rtrancl_induct 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    65
by (ALLGOALS(fast_tac (lambda_cs addIs [rtrancl_refl,rtrancl_into_rtrancl])));
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    66
qed "rtrancl_beta_AppL";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    67
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    68
goal Lambda.thy "!!s. t ->> t' ==> s @ t ->> s @ t'";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    69
be rtrancl_induct 1;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    70
by (ALLGOALS(fast_tac (lambda_cs addIs [rtrancl_refl,rtrancl_into_rtrancl])));
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    71
qed "rtrancl_beta_AppR";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    72
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    73
goal Lambda.thy "!!s. [| s ->> s'; t ->> t' |] ==> s @ t ->> s' @ t'";
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    74
by (fast_tac (lambda_cs addSIs [rtrancl_beta_AppL,rtrancl_beta_AppR]
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    75
                        addIs  [rtrancl_trans]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    76
qed "rtrancl_beta_App";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    77
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    78
(*** subst and lift ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    79
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    80
fun addsplit ss = ss addsimps [subst_Var] setloop (split_tac [expand_if]);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    81
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    82
goal Lambda.thy "(Var k)[u/k] = u";
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    83
by (asm_full_simp_tac(addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    84
qed "subst_eq";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    85
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    86
goal Lambda.thy "!!s. i<j ==> (Var j)[u/i] = Var(pred j)";
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    87
by (asm_full_simp_tac(addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    88
qed "subst_gt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    89
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    90
goal Lambda.thy "!!s. j<i ==> (Var j)[u/i] = Var(j)";
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    91
by (asm_full_simp_tac (addsplit(!simpset) addsimps
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    92
                          [less_not_refl2 RS not_sym,less_SucI]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    93
qed "subst_lt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    94
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    95
Addsimps [subst_eq,subst_gt,subst_lt];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    96
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    97
goal Lambda.thy
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    98
  "!i k. i < Suc k --> lift (lift t i) (Suc k) = lift (lift t k) i";
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    99
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   100
by(ALLGOALS Asm_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   101
by(strip_tac 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   102
by (excluded_middle_tac "nat < i" 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   103
by ((forward_tac [lt_trans2] 2) THEN (assume_tac 2));
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   104
by (ALLGOALS(asm_full_simp_tac (addsplit(!simpset) addsimps [less_SucI])));
1288
6eb89a693e05 Improved a few proofs.
nipkow
parents: 1269
diff changeset
   105
bind_thm("lift_lift", result() RS spec RS spec RS mp);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   106
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   107
goal Lambda.thy "!i j s. j < Suc i --> \
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   108
\         lift (t[s/j]) i = (lift t (Suc i)) [lift s i / j]";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   109
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   110
by(ALLGOALS(asm_simp_tac (!simpset addsimps [lift_lift])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   111
by(strip_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   112
by (excluded_middle_tac "nat < j" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   113
by (Asm_full_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   114
by (eres_inst_tac [("j","nat")] leqE 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   115
by (asm_full_simp_tac (addsplit(!simpset)
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
   116
                       addsimps [less_SucI,gt_pred]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   117
by (hyp_subst_tac 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   118
by (Asm_simp_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   119
by (forw_inst_tac [("j","j")] lt_trans2 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   120
by (assume_tac 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   121
by (asm_full_simp_tac (addsplit(!simpset) addsimps [less_SucI]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   122
qed "lift_subst";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   123
Addsimps [lift_subst];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   124
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   125
goal Lambda.thy
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   126
  "!i j s. i < Suc j -->\
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   127
\         lift (t[s/j]) i = (lift t i) [lift s i / Suc j]";
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   128
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   129
by(ALLGOALS(asm_simp_tac (!simpset addsimps [lift_lift])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   130
by(strip_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   131
by (excluded_middle_tac "nat < j" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   132
by (Asm_full_simp_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   133
by (eres_inst_tac [("i","j")] leqE 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   134
by (forward_tac  [lt_trans1] 1 THEN assume_tac 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   135
by (ALLGOALS(asm_full_simp_tac
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
   136
               (!simpset addsimps [less_SucI,gt_pred])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   137
by (hyp_subst_tac 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   138
by (asm_full_simp_tac (!simpset addsimps [less_SucI]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   139
by(split_tac [expand_if] 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   140
by (asm_full_simp_tac (!simpset addsimps [less_SucI]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   141
qed "lift_subst_lt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   142
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   143
goal Lambda.thy "!k s. (lift t k)[s/k] = t";
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   144
by(db.induct_tac "t" 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   145
by(ALLGOALS Asm_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   146
by(split_tac [expand_if] 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   147
by(ALLGOALS Asm_full_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   148
qed "subst_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   149
Addsimps [subst_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   150
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   151
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   152
goal Lambda.thy "!i j u v. i < Suc j --> \
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   153
\ t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]";
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   154
by(db.induct_tac "t" 1);
1288
6eb89a693e05 Improved a few proofs.
nipkow
parents: 1269
diff changeset
   155
by(ALLGOALS(asm_simp_tac(!simpset addsimps [lift_lift RS sym,lift_subst_lt])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   156
by(strip_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   157
by (excluded_middle_tac "nat < Suc(Suc j)" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   158
by(Asm_full_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   159
by (forward_tac [lessI RS less_trans] 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   160
by (eresolve_tac [leqE] 1);
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
   161
by (asm_simp_tac (!simpset addsimps [lt_pred]) 2);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   162
by (forward_tac [Suc_mono RS less_trans] 1 THEN assume_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   163
by (forw_inst_tac [("i","i")] (lessI RS less_trans) 1);
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
   164
by (asm_simp_tac (!simpset addsimps [lt_pred]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   165
by (eres_inst_tac [("i","nat")] leqE 1);
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
   166
by (asm_full_simp_tac (!simpset addsimps [less_SucI]) 2);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   167
by (excluded_middle_tac "nat < i" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   168
by (Asm_full_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   169
by (eres_inst_tac [("j","nat")] leqE 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   170
by (asm_simp_tac (!simpset addsimps [gt_pred]) 1);
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   171
by (Asm_simp_tac 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   172
by (forward_tac [lt_trans2] 1 THEN assume_tac 1);
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   173
by (asm_simp_tac (!simpset addsimps [gt_pred]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   174
bind_thm("subst_subst", result() RS spec RS spec RS spec RS spec RS mp);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   175
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   176
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   177
(*** Equivalence proof for optimized substitution ***)
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   178
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   179
goal Lambda.thy "!k. liftn 0 t k = t";
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   180
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   181
by(ALLGOALS(asm_simp_tac(addsplit(!simpset))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   182
qed "liftn_0";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   183
Addsimps [liftn_0];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   184
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   185
goal Lambda.thy "!k. liftn (Suc n) t k = lift (liftn n t k) k";
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   186
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   187
by(ALLGOALS(asm_simp_tac(addsplit(!simpset))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   188
by(fast_tac (HOL_cs addDs [add_lessD1]) 1);
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   189
qed "liftn_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   190
Addsimps [liftn_lift];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   191
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   192
goal Lambda.thy "!n. substn t s n = t[liftn n s 0 / n]";
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   193
by(db.induct_tac "t" 1);
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
   194
by(ALLGOALS(asm_simp_tac(addsplit(!simpset))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   195
qed "substn_subst_n";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   196
Addsimps [substn_subst_n];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   197
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   198
goal Lambda.thy "substn t s 0 = t[s/0]";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   199
by(Simp_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   200
qed "substn_subst_0";