src/HOL/Lambda/Lambda.ML
author nipkow
Mon, 04 Jan 1999 15:07:47 +0100
changeset 6055 fdf4638bf726
parent 5983 79e301a6a51b
child 6141 a6922171b396
permissions -rw-r--r--
Version 1.0 of linear nat arithmetic.
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
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
    11
val beta_cases = map (beta.mk_cases dB.simps)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
    12
    ["Var i -> t", "Abs r -> s", "s $ t -> u"];
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
    13
AddSEs beta_cases;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    14
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    15
Delsimps [subst_Var];
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1266
diff changeset
    16
Addsimps ([if_not_P, not_less_eq] @ beta.intrs);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    17
1302
ddfdcc9ce667 Moved some thms to Arith and to Trancl.
nipkow
parents: 1288
diff changeset
    18
(* don't add r_into_rtrancl! *)
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    19
AddSIs beta.intrs;
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    20
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    21
(*** Congruence rules for ->> ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    22
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    23
Goal "s ->> s' ==> Abs s ->> Abs s'";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    24
by (etac rtrancl_induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    25
by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl])));
2159
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    26
qed "rtrancl_beta_Abs";
e650a3f6f600 Used nat_trans_tac. New Eta. various smaller changes.
nipkow
parents: 2116
diff changeset
    27
AddSIs [rtrancl_beta_Abs];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    28
5146
nipkow
parents: 5117
diff changeset
    29
Goal "s ->> s' ==> s $ t ->> s' $ t";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    30
by (etac rtrancl_induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    31
by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    32
qed "rtrancl_beta_AppL";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    33
5146
nipkow
parents: 5117
diff changeset
    34
Goal "t ->> t' ==> s $ t ->> s $ t'";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1302
diff changeset
    35
by (etac rtrancl_induct 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    36
by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_rtrancl])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    37
qed "rtrancl_beta_AppR";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    38
5146
nipkow
parents: 5117
diff changeset
    39
Goal "[| s ->> s'; t ->> t' |] ==> s $ t ->> s' $ t'";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    40
by (blast_tac (claset() addSIs [rtrancl_beta_AppL, rtrancl_beta_AppR]
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
    41
                       addIs  [rtrancl_trans]) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    42
qed "rtrancl_beta_App";
1974
b50f96543dec Removed refs to clasets like rel_cs etc. Used implicit claset.
nipkow
parents: 1759
diff changeset
    43
AddIs [rtrancl_beta_App];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    44
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    45
(*** subst and lift ***)
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    46
3919
c036caebfc75 setloop split_tac -> addsplits
nipkow
parents: 2922
diff changeset
    47
fun addsplit ss = ss addsimps [subst_Var]
4831
dae4d63a1318 Renamed expand_const -> split_const.
nipkow
parents: 4686
diff changeset
    48
                     delsplits [split_if]
dae4d63a1318 Renamed expand_const -> split_const.
nipkow
parents: 4686
diff changeset
    49
                     setloop (split_inside_tac [split_if]);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    50
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    51
Goal "(Var k)[u/k] = u";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    52
by (asm_full_simp_tac(addsplit(simpset())) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    53
qed "subst_eq";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    54
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    55
Goal "i<j ==> (Var j)[u/i] = Var(j-1)";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
    56
by (asm_full_simp_tac(addsplit(simpset())) 1);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    57
qed "subst_gt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    58
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    59
Goal "j<i ==> (Var j)[u/i] = Var(j)";
5351
paulson
parents: 5326
diff changeset
    60
by (asm_full_simp_tac (addsplit(simpset()) addsimps [less_not_refl3]) 1);
1120
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
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    65
Goal
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";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    67
by (induct_tac "t" 1);
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    68
by (Auto_tac);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    69
qed_spec_mp "lift_lift";
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    70
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    71
Goal "!i j s. j < Suc i --> lift (t[s/j]) i = (lift t (Suc i)) [lift s i / j]";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    72
by (induct_tac "t" 1);
4360
40e5c97e988d pred n -> n-1
nipkow
parents: 4089
diff changeset
    73
by (ALLGOALS(asm_simp_tac (simpset() addsimps [diff_Suc,subst_Var,lift_lift]
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    74
                                addsplits [nat.split])));
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    75
by (Auto_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    76
qed "lift_subst";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    77
Addsimps [lift_subst];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    78
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    79
Goal
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    80
  "!i j s. i < Suc j --> lift (t[s/j]) i = (lift t i) [lift s i / Suc j]";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    81
by (induct_tac "t" 1);
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    82
by (ALLGOALS(asm_simp_tac (simpset() addsimps [subst_Var,lift_lift])));
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    83
qed "lift_subst_lt";
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    84
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    85
Goal "!k s. (lift t k)[s/k] = t";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    86
by (induct_tac "t" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4360
diff changeset
    87
by (ALLGOALS Asm_full_simp_tac);
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    88
qed "subst_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
    89
Addsimps [subst_lift];
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    90
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    91
5117
7b5efef2ca74 Removed leading !! in goals.
nipkow
parents: 5069
diff changeset
    92
Goal "!i j u v. i < Suc j --> t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    93
by (induct_tac "t" 1);
2116
73bbf2cc7651 Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
nipkow
parents: 2031
diff changeset
    94
by (ALLGOALS(asm_simp_tac
4360
40e5c97e988d pred n -> n-1
nipkow
parents: 4089
diff changeset
    95
      (simpset() addsimps [diff_Suc,subst_Var,lift_lift RS sym,lift_subst_lt]
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    96
                 addsplits [nat.split])));
2922
580647a879cf Using Blast_tac
paulson
parents: 2891
diff changeset
    97
by (safe_tac (HOL_cs addSEs [nat_neqE]));
5983
79e301a6a51b At last: linear arithmetic for nat!
nipkow
parents: 5351
diff changeset
    98
by (ALLGOALS Simp_tac);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    99
qed_spec_mp "subst_subst";
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   100
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   101
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   102
(*** Equivalence proof for optimized substitution ***)
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   103
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   104
Goal "!k. liftn 0 t k = t";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
   105
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   106
by (ALLGOALS(asm_simp_tac(addsplit(simpset()))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   107
qed "liftn_0";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   108
Addsimps [liftn_0];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   109
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   110
Goal "!k. liftn (Suc n) t k = lift (liftn n t k) k";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
   111
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   112
by (ALLGOALS(asm_simp_tac(addsplit(simpset()))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   113
qed "liftn_lift";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   114
Addsimps [liftn_lift];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   115
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   116
Goal "!n. substn t s n = t[liftn n s 0 / n]";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
   117
by (induct_tac "t" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3919
diff changeset
   118
by (ALLGOALS(asm_simp_tac(addsplit(simpset()))));
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   119
qed "substn_subst_n";
1266
3ae9fe3c0f68 added local simpsets
clasohm
parents: 1172
diff changeset
   120
Addsimps [substn_subst_n];
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   121
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   122
Goal "substn t s 0 = t[s/0]";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1974
diff changeset
   123
by (Simp_tac 1);
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
   124
qed "substn_subst_0";
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   125
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   126
(*** Preservation thms ***)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   127
(* Not used in CR-proof but in SN-proof *)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   128
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   129
Goal "r -> s ==> !t i. r[t/i] -> s[t/i]";
5326
8f9056ce5dfb expandshort
paulson
parents: 5318
diff changeset
   130
by (etac beta.induct 1);
5318
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   131
by (ALLGOALS (asm_simp_tac (simpset() addsimps [subst_subst RS sym])));
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   132
qed_spec_mp "subst_preserves_beta";
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   133
Addsimps [subst_preserves_beta];
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   134
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   135
Goal "r -> s ==> !i. lift r i -> lift s i";
5326
8f9056ce5dfb expandshort
paulson
parents: 5318
diff changeset
   136
by (etac beta.induct 1);
5318
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   137
by (Auto_tac);
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   138
qed_spec_mp "lift_preserves_beta";
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   139
Addsimps [lift_preserves_beta];
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   140
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   141
Goal "!r s i. r -> s --> t[r/i] ->> t[s/i]";
5318
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   142
by (induct_tac "t" 1);
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   143
by (asm_simp_tac (addsplit(simpset() addsimps [r_into_rtrancl])) 1);
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   144
by (asm_simp_tac (simpset() addsimps [rtrancl_beta_App]) 1);
72bf8039b53f expandshort
paulson
parents: 5261
diff changeset
   145
by (asm_simp_tac (simpset() addsimps [rtrancl_beta_Abs]) 1);
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   146
qed_spec_mp "subst_preserves_beta2";
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents: 5184
diff changeset
   147
Addsimps [subst_preserves_beta2];