src/HOL/Lambda/Lambda.thy
author urbanc
Thu, 23 Nov 2006 17:52:48 +0100
changeset 21488 e1b260d204a0
parent 21404 eb85850d3eb7
child 22271 51a80e238b29
permissions -rw-r--r--
fixed some typos
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1269
ee011b365770 New version with eta reduction.
nipkow
parents: 1153
diff changeset
     1
(*  Title:      HOL/Lambda/Lambda.thy
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
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
     7
header {* Basic definitions of Lambda-calculus *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14065
diff changeset
     9
theory Lambda imports Main begin
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    10
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    11
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    12
subsection {* Lambda-terms in de Bruijn notation and substitution *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    13
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    14
datatype dB =
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    15
    Var nat
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11943
diff changeset
    16
  | App dB dB (infixl "\<degree>" 200)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    17
  | Abs dB
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    18
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    19
consts
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    20
  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    21
  lift :: "[dB, nat] => dB"
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    22
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    23
primrec
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    24
  "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11943
diff changeset
    25
  "lift (s \<degree> t) k = lift s k \<degree> lift t k"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    26
  "lift (Abs s) k = Abs (lift s (k + 1))"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    27
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    28
primrec  (* FIXME base names *)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    29
  subst_Var: "(Var i)[s/k] =
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    30
    (if k < i then Var (i - 1) else if i = k then s else Var i)"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11943
diff changeset
    31
  subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    32
  subst_Abs: "(Abs t)[s/k] = Abs (t[lift s 0 / k+1])"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    33
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    34
declare subst_Var [simp del]
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    35
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    36
text {* Optimized versions of @{term subst} and @{term lift}. *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    37
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    38
consts
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    39
  substn :: "[dB, dB, nat] => dB"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    40
  liftn :: "[nat, dB, nat] => dB"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    41
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    42
primrec
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    43
  "liftn n (Var i) k = (if i < k then Var i else Var (i + n))"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11943
diff changeset
    44
  "liftn n (s \<degree> t) k = liftn n s k \<degree> liftn n t k"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    45
  "liftn n (Abs s) k = Abs (liftn n s (k + 1))"
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    46
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5146
diff changeset
    47
primrec
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    48
  "substn (Var i) s k =
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    49
    (if k < i then Var (i - 1) else if i = k then liftn k s 0 else Var i)"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11943
diff changeset
    50
  "substn (t \<degree> u) s k = substn t s k \<degree> substn u s k"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    51
  "substn (Abs t) s k = Abs (substn t s (k + 1))"
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    52
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    53
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    54
subsection {* Beta-reduction *}
1153
5c5daf97cf2d Simplified the confluence proofs.
nipkow
parents: 1124
diff changeset
    55
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    56
consts
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    57
  beta :: "(dB \<times> dB) set"
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    58
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    59
abbreviation
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    60
  beta_red :: "[dB, dB] => bool"  (infixl "->" 50) where
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    61
  "s -> t == (s, t) \<in> beta"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    62
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    63
abbreviation
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    64
  beta_reds :: "[dB, dB] => bool"  (infixl "->>" 50) where
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    65
  "s ->> t == (s, t) \<in> beta^*"
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18257
diff changeset
    66
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 20503
diff changeset
    67
notation (latex)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    68
  beta_red  (infixl "\<rightarrow>\<^sub>\<beta>" 50) and
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
    69
  beta_reds  (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50)
1120
ff7dd80513e6 Lambda calculus in de Bruijn notation.
nipkow
parents:
diff changeset
    70
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1376
diff changeset
    71
inductive beta
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 10851
diff changeset
    72
  intros
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    73
    beta [simp, intro!]: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    74
    appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    75
    appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    76
    abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs s \<rightarrow>\<^sub>\<beta> Abs t"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    77
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    78
inductive_cases beta_cases [elim!]:
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    79
  "Var i \<rightarrow>\<^sub>\<beta> t"
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    80
  "Abs r \<rightarrow>\<^sub>\<beta> s"
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    81
  "s \<degree> t \<rightarrow>\<^sub>\<beta> u"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    82
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    83
declare if_not_P [simp] not_less_eq [simp]
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    84
  -- {* don't add @{text "r_into_rtrancl[intro!]"} *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    85
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    86
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    87
subsection {* Congruence rules *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    88
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    89
lemma rtrancl_beta_Abs [intro!]:
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    90
    "s \<rightarrow>\<^sub>\<beta>\<^sup>* s' ==> Abs s \<rightarrow>\<^sub>\<beta>\<^sup>* Abs s'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
    91
  by (induct set: rtrancl) (blast intro: rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    92
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    93
lemma rtrancl_beta_AppL:
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    94
    "s \<rightarrow>\<^sub>\<beta>\<^sup>* s' ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
    95
  by (induct set: rtrancl) (blast intro: rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    96
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
    97
lemma rtrancl_beta_AppR:
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
    98
    "t \<rightarrow>\<^sub>\<beta>\<^sup>* t' ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s \<degree> t'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
    99
  by (induct set: rtrancl) (blast intro: rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   100
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   101
lemma rtrancl_beta_App [intro]:
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   102
    "[| s \<rightarrow>\<^sub>\<beta>\<^sup>* s'; t \<rightarrow>\<^sub>\<beta>\<^sup>* t' |] ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'"
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
   103
  by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtrancl_trans)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   104
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   105
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   106
subsection {* Substitution-lemmas *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   107
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   108
lemma subst_eq [simp]: "(Var k)[u/k] = u"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   109
  by (simp add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   110
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   111
lemma subst_gt [simp]: "i < j ==> (Var j)[u/i] = Var (j - 1)"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   112
  by (simp add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   113
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   114
lemma subst_lt [simp]: "j < i ==> (Var j)[u/i] = Var j"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   115
  by (simp add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   116
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   117
lemma lift_lift:
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   118
    "i < k + 1 \<Longrightarrow> lift (lift t i) (Suc k) = lift (lift t k) i"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   119
  by (induct t arbitrary: i k) auto
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   120
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   121
lemma lift_subst [simp]:
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   122
    "j < i + 1 \<Longrightarrow> lift (t[s/j]) i = (lift t (i + 1)) [lift s i / j]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   123
  by (induct t arbitrary: i j s)
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   124
    (simp_all add: diff_Suc subst_Var lift_lift split: nat.split)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   125
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   126
lemma lift_subst_lt:
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   127
    "i < j + 1 \<Longrightarrow> lift (t[s/j]) i = (lift t i) [lift s i / j + 1]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   128
  by (induct t arbitrary: i j s) (simp_all add: subst_Var lift_lift)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   129
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   130
lemma subst_lift [simp]:
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   131
    "(lift t k)[s/k] = t"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   132
  by (induct t arbitrary: k s) simp_all
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   133
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   134
lemma subst_subst:
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   135
    "i < j + 1 \<Longrightarrow> t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   136
  by (induct t arbitrary: i j u v)
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   137
    (simp_all add: diff_Suc subst_Var lift_lift [symmetric] lift_subst_lt
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   138
      split: nat.split)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   139
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   140
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   141
subsection {* Equivalence proof for optimized substitution *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   142
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   143
lemma liftn_0 [simp]: "liftn 0 t k = t"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   144
  by (induct t arbitrary: k) (simp_all add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   145
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   146
lemma liftn_lift [simp]: "liftn (Suc n) t k = lift (liftn n t k) k"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   147
  by (induct t arbitrary: k) (simp_all add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   148
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   149
lemma substn_subst_n [simp]: "substn t s n = t[liftn n s 0 / n]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   150
  by (induct t arbitrary: n) (simp_all add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   151
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   152
theorem substn_subst_0: "substn t s 0 = t[s/0]"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   153
  by simp
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   154
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   155
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   156
subsection {* Preservation theorems *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   157
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   158
text {* Not used in Church-Rosser proof, but in Strong
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   159
  Normalization. \medskip *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   160
13915
28ccb51bd2f3 Eliminated most occurrences of rule_format attribute.
berghofe
parents: 13187
diff changeset
   161
theorem subst_preserves_beta [simp]:
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   162
    "r \<rightarrow>\<^sub>\<beta> s ==> r[t/i] \<rightarrow>\<^sub>\<beta> s[t/i]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   163
  by (induct arbitrary: t i set: beta) (simp_all add: subst_subst [symmetric])
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   164
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   165
theorem subst_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> r[t/i] \<rightarrow>\<^sub>\<beta>\<^sup>* s[t/i]"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   166
  apply (induct set: rtrancl)
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   167
   apply (rule rtrancl_refl)
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   168
  apply (erule rtrancl_into_rtrancl)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   169
  apply (erule subst_preserves_beta)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   170
  done
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   171
13915
28ccb51bd2f3 Eliminated most occurrences of rule_format attribute.
berghofe
parents: 13187
diff changeset
   172
theorem lift_preserves_beta [simp]:
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   173
    "r \<rightarrow>\<^sub>\<beta> s ==> lift r i \<rightarrow>\<^sub>\<beta> lift s i"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   174
  by (induct arbitrary: i set: beta) auto
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   175
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   176
theorem lift_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> lift r i \<rightarrow>\<^sub>\<beta>\<^sup>* lift s i"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   177
  apply (induct set: rtrancl)
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   178
   apply (rule rtrancl_refl)
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   179
  apply (erule rtrancl_into_rtrancl)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   180
  apply (erule lift_preserves_beta)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   181
  done
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   182
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   183
theorem subst_preserves_beta2 [simp]: "r \<rightarrow>\<^sub>\<beta> s ==> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 19656
diff changeset
   184
  apply (induct t arbitrary: r s i)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   185
    apply (simp add: subst_Var r_into_rtrancl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   186
   apply (simp add: rtrancl_beta_App)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   187
  apply (simp add: rtrancl_beta_Abs)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   188
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 6453
diff changeset
   189
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   190
theorem subst_preserves_beta2': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   191
  apply (induct set: rtrancl)
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
   192
   apply (rule rtrancl_refl)
14065
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   193
  apply (erule rtrancl_trans)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   194
  apply (erule subst_preserves_beta2)
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   195
  done
8abaf978c9c2 Nicer syntax for beta reduction.
berghofe
parents: 13915
diff changeset
   196
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 10851
diff changeset
   197
end