| author | wenzelm | 
| Thu, 30 Oct 1997 17:00:34 +0100 | |
| changeset 4047 | 67b5552b1067 | 
| parent 3919 | c036caebfc75 | 
| child 4089 | 96fba19bcbe2 | 
| permissions | -rw-r--r-- | 
| 1269 | 1 | (* Title: HOL/Lambda/Lambda.ML | 
| 1120 | 2 | ID: $Id$ | 
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1995 TU Muenchen | |
| 5 | ||
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 6 | Substitution-lemmas. | 
| 1120 | 7 | *) | 
| 8 | ||
| 9 | (*** Lambda ***) | |
| 10 | ||
| 11 | open Lambda; | |
| 12 | ||
| 1974 
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
 nipkow parents: 
1759diff
changeset | 13 | Delsimps [subst_Var]; | 
| 1269 | 14 | Addsimps ([if_not_P, not_less_eq] @ beta.intrs); | 
| 1120 | 15 | |
| 1302 | 16 | (* don't add r_into_rtrancl! *) | 
| 1974 
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
 nipkow parents: 
1759diff
changeset | 17 | AddSIs beta.intrs; | 
| 1120 | 18 | |
| 2159 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 19 | val dB_case_distinction = | 
| 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
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: 
1723diff
changeset | 21 | |
| 1120 | 22 | (*** Congruence rules for ->> ***) | 
| 23 | ||
| 2159 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 24 | goal Lambda.thy "!!s. s ->> s' ==> Abs s ->> Abs s'"; | 
| 1465 | 25 | by (etac rtrancl_induct 1); | 
| 2891 | 26 | by (ALLGOALS (blast_tac (!claset addIs [rtrancl_into_rtrancl]))); | 
| 2159 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 27 | qed "rtrancl_beta_Abs"; | 
| 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 28 | AddSIs [rtrancl_beta_Abs]; | 
| 1120 | 29 | |
| 30 | goal Lambda.thy "!!s. s ->> s' ==> s @ t ->> s' @ t"; | |
| 1465 | 31 | by (etac rtrancl_induct 1); | 
| 2891 | 32 | by (ALLGOALS (blast_tac (!claset addIs [rtrancl_into_rtrancl]))); | 
| 1120 | 33 | qed "rtrancl_beta_AppL"; | 
| 34 | ||
| 35 | goal Lambda.thy "!!s. t ->> t' ==> s @ t ->> s @ t'"; | |
| 1465 | 36 | by (etac rtrancl_induct 1); | 
| 2891 | 37 | by (ALLGOALS (blast_tac (!claset addIs [rtrancl_into_rtrancl]))); | 
| 1120 | 38 | qed "rtrancl_beta_AppR"; | 
| 39 | ||
| 40 | goal Lambda.thy "!!s. [| s ->> s'; t ->> t' |] ==> s @ t ->> s' @ t'"; | |
| 2922 | 41 | by (blast_tac (!claset addSIs [rtrancl_beta_AppL, rtrancl_beta_AppR] | 
| 42 | addIs [rtrancl_trans]) 1); | |
| 1120 | 43 | qed "rtrancl_beta_App"; | 
| 1974 
b50f96543dec
Removed refs to clasets like rel_cs etc. Used implicit claset.
 nipkow parents: 
1759diff
changeset | 44 | AddIs [rtrancl_beta_App]; | 
| 1120 | 45 | |
| 46 | (*** subst and lift ***) | |
| 47 | ||
| 3919 | 48 | fun addsplit ss = ss addsimps [subst_Var] | 
| 49 | setloop (split_inside_tac [expand_if]); | |
| 1120 | 50 | |
| 1153 | 51 | goal Lambda.thy "(Var k)[u/k] = u"; | 
| 1269 | 52 | by (asm_full_simp_tac(addsplit(!simpset)) 1); | 
| 1120 | 53 | qed "subst_eq"; | 
| 54 | ||
| 1153 | 55 | goal Lambda.thy "!!s. i<j ==> (Var j)[u/i] = Var(pred j)"; | 
| 1269 | 56 | by (asm_full_simp_tac(addsplit(!simpset)) 1); | 
| 1120 | 57 | qed "subst_gt"; | 
| 58 | ||
| 1153 | 59 | goal Lambda.thy "!!s. j<i ==> (Var j)[u/i] = Var(j)"; | 
| 1269 | 60 | by (asm_full_simp_tac (addsplit(!simpset) addsimps | 
| 1120 | 61 | [less_not_refl2 RS not_sym,less_SucI]) 1); | 
| 62 | qed "subst_lt"; | |
| 63 | ||
| 1266 | 64 | Addsimps [subst_eq,subst_gt,subst_lt]; | 
| 1120 | 65 | |
| 66 | goal Lambda.thy | |
| 1172 | 67 | "!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: 
2116diff
changeset | 68 | by (dB.induct_tac "t" 1); | 
| 3919 | 69 | by (ALLGOALS(asm_simp_tac (!simpset addsplits [expand_if] | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2159diff
changeset | 70 | addSolver cut_trans_tac))); | 
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 71 | by (safe_tac HOL_cs); | 
| 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 72 | by (ALLGOALS trans_tac); | 
| 1486 | 73 | qed_spec_mp "lift_lift"; | 
| 1120 | 74 | |
| 1153 | 75 | goal Lambda.thy "!i j s. j < Suc i --> \ | 
| 76 | \ 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: 
2116diff
changeset | 77 | by (dB.induct_tac "t" 1); | 
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 78 | by (ALLGOALS(asm_simp_tac (!simpset addsimps [pred_def,subst_Var,lift_lift] | 
| 3919 | 79 | addsplits [expand_if,expand_nat_case] | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2159diff
changeset | 80 | addSolver cut_trans_tac))); | 
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 81 | by (safe_tac HOL_cs); | 
| 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 82 | by (ALLGOALS trans_tac); | 
| 1120 | 83 | qed "lift_subst"; | 
| 1266 | 84 | Addsimps [lift_subst]; | 
| 1120 | 85 | |
| 86 | goal Lambda.thy | |
| 1172 | 87 | "!i j s. i < Suc j -->\ | 
| 88 | \ 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: 
2116diff
changeset | 89 | by (dB.induct_tac "t" 1); | 
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 90 | by (ALLGOALS(asm_simp_tac (!simpset addsimps [subst_Var,lift_lift] | 
| 3919 | 91 | addsplits [expand_if] | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2159diff
changeset | 92 | addSolver cut_trans_tac))); | 
| 2922 | 93 | by (safe_tac (HOL_cs addSEs [nat_neqE])); | 
| 94 | by (ALLGOALS trans_tac); | |
| 1120 | 95 | qed "lift_subst_lt"; | 
| 96 | ||
| 1172 | 97 | goal Lambda.thy "!k s. (lift t k)[s/k] = t"; | 
| 2159 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 98 | by (dB.induct_tac "t" 1); | 
| 3919 | 99 | by (ALLGOALS (asm_full_simp_tac (!simpset addsplits [expand_if]))); | 
| 1120 | 100 | qed "subst_lift"; | 
| 1266 | 101 | Addsimps [subst_lift]; | 
| 1120 | 102 | |
| 103 | ||
| 1172 | 104 | goal Lambda.thy "!i j u v. i < Suc j --> \ | 
| 105 | \ 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: 
2116diff
changeset | 106 | by (dB.induct_tac "t" 1); | 
| 2116 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 107 | by (ALLGOALS(asm_simp_tac | 
| 
73bbf2cc7651
Used trans_tac (see Provers/nat_transitive.ML) to automate arithmetic.
 nipkow parents: 
2031diff
changeset | 108 | (!simpset addsimps [pred_def,subst_Var,lift_lift RS sym,lift_subst_lt] | 
| 3919 | 109 | addsplits [expand_if,expand_nat_case] | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2159diff
changeset | 110 | addSolver cut_trans_tac))); | 
| 2922 | 111 | by (safe_tac (HOL_cs addSEs [nat_neqE])); | 
| 112 | by (ALLGOALS trans_tac); | |
| 1486 | 113 | qed_spec_mp "subst_subst"; | 
| 1153 | 114 | |
| 115 | ||
| 116 | (*** Equivalence proof for optimized substitution ***) | |
| 117 | ||
| 118 | goal Lambda.thy "!k. liftn 0 t k = t"; | |
| 2159 
e650a3f6f600
Used nat_trans_tac. New Eta. various smaller changes.
 nipkow parents: 
2116diff
changeset | 119 | by (dB.induct_tac "t" 1); | 
| 2031 | 120 | by (ALLGOALS(asm_simp_tac(addsplit(!simpset)))); | 
| 1153 | 121 | qed "liftn_0"; | 
| 1266 | 122 | Addsimps [liftn_0]; | 
| 1153 | 123 | |
| 124 | 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: 
2116diff
changeset | 125 | by (dB.induct_tac "t" 1); | 
| 2031 | 126 | by (ALLGOALS(asm_simp_tac(addsplit(!simpset)))); | 
| 2891 | 127 | by (blast_tac (!claset addDs [add_lessD1]) 1); | 
| 1153 | 128 | qed "liftn_lift"; | 
| 1266 | 129 | Addsimps [liftn_lift]; | 
| 1153 | 130 | |
| 131 | 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: 
2116diff
changeset | 132 | by (dB.induct_tac "t" 1); | 
| 2031 | 133 | by (ALLGOALS(asm_simp_tac(addsplit(!simpset)))); | 
| 1153 | 134 | qed "substn_subst_n"; | 
| 1266 | 135 | Addsimps [substn_subst_n]; | 
| 1153 | 136 | |
| 137 | goal Lambda.thy "substn t s 0 = t[s/0]"; | |
| 2031 | 138 | by (Simp_tac 1); | 
| 1153 | 139 | qed "substn_subst_0"; |