src/HOL/Lambda/Lambda.ML
author wenzelm
Wed, 18 Dec 1996 15:56:58 +0100
changeset 2448 61337170db84
parent 2159 e650a3f6f600
child 2637 e9b203f854ae
permissions -rw-r--r--
IsaMakefile for HOL;
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
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
     6
Substitution-lemmas.
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     7
*)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     8
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
     9
(*** Lambda ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    10
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    11
open Lambda;
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    12
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    13
Delsimps [subst_Var];
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    14
Addsimps ([if_not_P, not_less_eq] @ beta.intrs);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    15
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    16
(* don't add r_into_rtrancl! *)
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    17
AddSIs beta.intrs;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    18
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    19
val dB_case_distinction =
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    20
  rule_by_tactic(EVERY[etac thin_rl 2,etac thin_rl 2,etac thin_rl 3])dB.induct;
1759
a42d6c537f4a Added comparison with implicit rule Fun(lift s 0 @ Var 0) -e> s
nipkow
parents: 1723
diff changeset
    21
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
(*** Congruence rules for ->> ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    23
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    24
goal Lambda.thy "!!s. s ->> s' ==> Abs s ->> Abs s'";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    25
by (etac rtrancl_induct 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    26
by (ALLGOALS(fast_tac (!claset addIs [rtrancl_into_rtrancl])));
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    27
qed "rtrancl_beta_Abs";
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    28
AddSIs [rtrancl_beta_Abs];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    29
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    30
goal Lambda.thy "!!s. s ->> s' ==> s @ t ->> s' @ t";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    31
by (etac rtrancl_induct 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    32
by (ALLGOALS(fast_tac (!claset addIs [rtrancl_into_rtrancl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    33
qed "rtrancl_beta_AppL";
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 Lambda.thy "!!s. t ->> t' ==> s @ t ->> s @ t'";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    36
by (etac rtrancl_induct 1);
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    37
by (ALLGOALS(fast_tac (!claset addIs [rtrancl_into_rtrancl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    38
qed "rtrancl_beta_AppR";
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 Lambda.thy "!!s. [| s ->> s'; t ->> t' |] ==> s @ t ->> s' @ t'";
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    41
by (deepen_tac (!claset addSIs [rtrancl_beta_AppL,rtrancl_beta_AppR]
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    42
                        addIs  [rtrancl_trans]) 3 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    43
qed "rtrancl_beta_App";
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    44
AddIs [rtrancl_beta_App];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    45
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    46
(*** subst and lift ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    47
1723
286f9b6370ab Replaced split_tac by split_inside_tac.
berghofe
parents: 1673
diff changeset
    48
fun addsplit ss = ss addsimps [subst_Var] setloop (split_inside_tac [expand_if]);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    49
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    50
goal Lambda.thy "(Var k)[u/k] = u";
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    51
by (asm_full_simp_tac(addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    52
qed "subst_eq";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    53
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    54
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
    55
by (asm_full_simp_tac(addsplit(!simpset)) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    56
qed "subst_gt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    57
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    58
goal Lambda.thy "!!s. j<i ==> (Var j)[u/i] = Var(j)";
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    59
by (asm_full_simp_tac (addsplit(!simpset) addsimps
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    60
                          [less_not_refl2 RS not_sym,less_SucI]) 1);
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    61
qed "subst_lt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    62
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    63
Addsimps [subst_eq,subst_gt,subst_lt];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    64
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    65
goal Lambda.thy
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    66
  "!i k. i < Suc k --> lift (lift t i) (Suc k) = lift (lift t k) i";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    67
by (dB.induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    68
by (ALLGOALS(asm_simp_tac (!simpset setloop (split_tac [expand_if])
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    69
                                    addsolver (cut_trans_tac))));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    70
by (safe_tac HOL_cs);
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    71
by (ALLGOALS trans_tac);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    72
qed_spec_mp "lift_lift";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    73
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    74
goal Lambda.thy "!i j s. j < Suc i --> \
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    75
\         lift (t[s/j]) i = (lift t (Suc i)) [lift s i / j]";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    76
by (dB.induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    77
by (ALLGOALS(asm_simp_tac (!simpset addsimps [pred_def,subst_Var,lift_lift]
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    78
                                setloop (split_tac [expand_if,expand_nat_case])
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    79
                                addsolver (cut_trans_tac))));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    80
by (safe_tac HOL_cs);
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    81
by (ALLGOALS trans_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    82
qed "lift_subst";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    83
Addsimps [lift_subst];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    84
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    85
goal Lambda.thy
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    86
  "!i j s. i < Suc j -->\
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    87
\         lift (t[s/j]) i = (lift t i) [lift s i / Suc j]";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    88
by (dB.induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    89
by (ALLGOALS(asm_simp_tac (!simpset addsimps [subst_Var,lift_lift]
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    90
                                setloop (split_tac [expand_if])
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    91
                                addsolver (cut_trans_tac))));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    92
by(safe_tac (HOL_cs addSEs [nat_neqE]));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    93
by(ALLGOALS trans_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    94
qed "lift_subst_lt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    95
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
    96
goal Lambda.thy "!k s. (lift t k)[s/k] = t";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    97
by (dB.induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    98
by (ALLGOALS (asm_full_simp_tac (!simpset setloop (split_tac[expand_if]))));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    99
qed "subst_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   100
Addsimps [subst_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   101
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
   102
1172
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   103
goal Lambda.thy "!i j u v. i < Suc j --> \
ab725b698cb2 Renamed some vars.
nipkow
parents: 1153
diff changeset
   104
\ t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
   105
by (dB.induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   106
by (ALLGOALS(asm_simp_tac
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   107
      (!simpset addsimps [pred_def,subst_Var,lift_lift RS sym,lift_subst_lt]
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   108
                setloop (split_tac [expand_if,expand_nat_case])
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   109
                addsolver (cut_trans_tac))));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   110
by(safe_tac (HOL_cs addSEs [nat_neqE]));
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
   111
by(ALLGOALS trans_tac);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
   112
qed_spec_mp "subst_subst";
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   113
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   114
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   115
(*** Equivalence proof for optimized substitution ***)
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   116
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   117
goal Lambda.thy "!k. liftn 0 t k = t";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
   118
by (dB.induct_tac "t" 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   119
by (ALLGOALS(asm_simp_tac(addsplit(!simpset))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   120
qed "liftn_0";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   121
Addsimps [liftn_0];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   122
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   123
goal Lambda.thy "!k. liftn (Suc n) t k = lift (liftn n t k) k";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
   124
by (dB.induct_tac "t" 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   125
by (ALLGOALS(asm_simp_tac(addsplit(!simpset))));
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   126
by (fast_tac (HOL_cs addDs [add_lessD1]) 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   127
qed "liftn_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   128
Addsimps [liftn_lift];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   129
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   130
goal Lambda.thy "!n. substn t s n = t[liftn n s 0 / n]";
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
   131
by (dB.induct_tac "t" 1);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   132
by (ALLGOALS(asm_simp_tac(addsplit(!simpset))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   133
qed "substn_subst_n";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   134
Addsimps [substn_subst_n];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   135
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   136
goal Lambda.thy "substn t s 0 = t[s/0]";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   137
by (Simp_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   138
qed "substn_subst_0";