src/HOL/Lambda/Eta.thy
author berghofe
Fri, 28 Apr 2006 17:56:20 +0200
changeset 19499 1a082c1257d7
parent 19380 b808efaa5828
child 19656 09be06943252
permissions -rw-r--r--
Added Class, Fsub, and Lambda_mu examples for nominal datatypes.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Lambda/Eta.thy
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
     3
    Author:     Tobias Nipkow and Stefan Berghofer
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
     4
    Copyright   1995, 2005 TU Muenchen
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     5
*)
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     6
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
     7
header {* Eta-reduction *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15522
diff changeset
     9
theory Eta imports ParRed begin
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    10
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    11
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    12
subsection {* Definition of eta-reduction and relatives *}
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
    13
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    14
consts
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    15
  free :: "dB => nat => bool"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    16
primrec
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    17
  "free (Var j) i = (j = i)"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    18
  "free (s \<degree> t) i = (free s i \<or> free t i)"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    19
  "free (Abs s) i = free s (i + 1)"
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
    20
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    21
consts
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    22
  eta :: "(dB \<times> dB) set"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    23
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    24
abbreviation
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18557
diff changeset
    25
  eta_red :: "[dB, dB] => bool"   (infixl "-e>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    26
  "s -e> t == (s, t) \<in> eta"
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18557
diff changeset
    27
  eta_reds :: "[dB, dB] => bool"   (infixl "-e>>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    28
  "s -e>> t == (s, t) \<in> eta^*"
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18557
diff changeset
    29
  eta_red0 :: "[dB, dB] => bool"   (infixl "-e>=" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
    30
  "s -e>= t == (s, t) \<in> eta^="
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
    31
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1376
diff changeset
    32
inductive eta
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
    33
  intros
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    34
    eta [simp, intro]: "\<not> free s 0 ==> Abs (s \<degree> Var 0) -e> s[dummy/0]"
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    35
    appL [simp, intro]: "s -e> t ==> s \<degree> u -e> t \<degree> u"
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    36
    appR [simp, intro]: "s -e> t ==> u \<degree> s -e> u \<degree> t"
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
    37
    abs [simp, intro]: "s -e> t ==> Abs s -e> Abs t"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    38
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    39
inductive_cases eta_cases [elim!]:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    40
  "Abs s -e> z"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    41
  "s \<degree> t -e> u"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    42
  "Var i -e> t"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    43
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    44
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    45
subsection "Properties of eta, subst and free"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    46
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    47
lemma subst_not_free [simp]: "\<not> free s i \<Longrightarrow> s[t/i] = s[u/i]"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    48
  by (induct s fixing: i t u) (simp_all add: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    49
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    50
lemma free_lift [simp]:
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    51
    "free (lift t k) i = (i < k \<and> free t i \<or> k < i \<and> free t (i - 1))"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    52
  apply (induct t fixing: i k)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    53
    apply (auto cong: conj_cong)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    54
  apply arith
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    55
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    56
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    57
lemma free_subst [simp]:
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    58
    "free (s[t/k]) i =
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    59
      (free s k \<and> free t i \<or> free s (if i < k then i else i + 1))"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    60
  apply (induct s fixing: i k t)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    61
    prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    62
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    63
    apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    64
   prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    65
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    66
  apply (simp add: diff_Suc subst_Var split: nat.split)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    67
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    68
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    69
lemma free_eta: "s -e> t ==> free t i = free s i"
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    70
  by (induct fixing: i set: eta) (simp_all cong: conj_cong)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    71
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    72
lemma not_free_eta:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    73
    "[| s -e> t; \<not> free s i |] ==> \<not> free t i"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    74
  by (simp add: free_eta)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    75
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    76
lemma eta_subst [simp]:
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    77
    "s -e> t ==> s[u/i] -e> t[u/i]"
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    78
  by (induct fixing: u i set: eta) (simp_all add: subst_subst [symmetric])
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    79
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    80
theorem lift_subst_dummy: "\<not> free s i \<Longrightarrow> lift (s[dummy/i]) i = s"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
    81
  by (induct s fixing: i dummy) simp_all
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
    82
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    83
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    84
subsection "Confluence of eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    85
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    86
lemma square_eta: "square eta eta (eta^=) (eta^=)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    87
  apply (unfold square_def id_def)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    88
  apply (rule impI [THEN allI [THEN allI]])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    89
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    90
  apply (erule eta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    91
     apply (slowsimp intro: subst_not_free eta_subst free_eta [THEN iffD1])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    92
    apply safe
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    93
       prefer 5
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    94
       apply (blast intro!: eta_subst intro: free_eta [THEN iffD1])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    95
      apply blast+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    96
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    97
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    98
theorem eta_confluent: "confluent eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    99
  apply (rule square_eta [THEN square_reflcl_confluent])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   100
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   101
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   102
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   103
subsection "Congruence rules for eta*"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   104
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   105
lemma rtrancl_eta_Abs: "s -e>> s' ==> Abs s -e>> Abs s'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   106
  by (induct set: rtrancl)
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   107
    (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   108
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   109
lemma rtrancl_eta_AppL: "s -e>> s' ==> s \<degree> t -e>> s' \<degree> t"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   110
  by (induct set: rtrancl)
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   111
    (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   112
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   113
lemma rtrancl_eta_AppR: "t -e>> t' ==> s \<degree> t -e>> s \<degree> t'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   114
  by (induct set: rtrancl) (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   115
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   116
lemma rtrancl_eta_App:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   117
    "[| s -e>> s'; t -e>> t' |] ==> s \<degree> t -e>> s' \<degree> t'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   118
  by (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtrancl_trans)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   119
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   120
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   121
subsection "Commutation of beta and eta"
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
   122
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   123
lemma free_beta:
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   124
    "s -> t ==> free t i \<Longrightarrow> free s i"
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   125
  by (induct fixing: i set: beta) auto
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   126
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   127
lemma beta_subst [intro]: "s -> t ==> s[u/i] -> t[u/i]"
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   128
  by (induct fixing: u i set: beta) (simp_all add: subst_subst [symmetric])
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   129
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   130
lemma subst_Var_Suc [simp]: "t[Var i/i] = t[Var(i)/i + 1]"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   131
  by (induct t fixing: i) (auto elim!: linorder_neqE simp: subst_Var)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   132
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   133
lemma eta_lift [simp]: "s -e> t ==> lift s i -e> lift t i"
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
   134
  by (induct fixing: i set: eta) simp_all
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   135
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   136
lemma rtrancl_eta_subst: "s -e> t \<Longrightarrow> u[s/i] -e>> u[t/i]"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   137
  apply (induct u fixing: s t i)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   138
    apply (simp_all add: subst_Var)
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   139
    apply blast
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   140
   apply (blast intro: rtrancl_eta_App)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   141
  apply (blast intro!: rtrancl_eta_Abs eta_lift)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   142
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   143
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   144
lemma square_beta_eta: "square beta eta (eta^*) (beta^=)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   145
  apply (unfold square_def)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   146
  apply (rule impI [THEN allI [THEN allI]])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   147
  apply (erule beta.induct)
11183
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   148
     apply (slowsimp intro: rtrancl_eta_subst eta_subst)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   149
    apply (blast intro: rtrancl_eta_AppL)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   150
   apply (blast intro: rtrancl_eta_AppR)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   151
  apply simp;
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   152
  apply (slowsimp intro: rtrancl_eta_Abs free_beta
9858
c3ac6128b649 use 'iff' modifier;
wenzelm
parents: 9827
diff changeset
   153
    iff del: dB.distinct simp: dB.distinct)    (*23 seconds?*)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   154
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   155
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   156
lemma confluent_beta_eta: "confluent (beta \<union> eta)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   157
  apply (assumption |
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   158
    rule square_rtrancl_reflcl_commute confluent_Un
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   159
      beta_confluent eta_confluent square_beta_eta)+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   160
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   161
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   162
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   163
subsection "Implicit definition of eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   164
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   165
text {* @{term "Abs (lift s 0 \<degree> Var 0) -e> s"} *}
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   166
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   167
lemma not_free_iff_lifted:
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   168
    "(\<not> free s i) = (\<exists>t. s = lift t i)"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   169
  apply (induct s fixing: i)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   170
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   171
    apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   172
     apply (erule linorder_neqE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   173
      apply (rule_tac x = "Var nat" in exI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   174
      apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   175
     apply (rule_tac x = "Var (nat - 1)" in exI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   176
     apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   177
    apply clarify
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   178
    apply (rule notE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   179
     prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   180
     apply assumption
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   181
    apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   182
    apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   183
      apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   184
     apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   185
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   186
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   187
   apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   188
   apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   189
   apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   190
    apply (elim conjE exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   191
    apply (rename_tac u1 u2)
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   192
    apply (rule_tac x = "u1 \<degree> u2" in exI)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   193
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   194
   apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   195
   apply (erule rev_mp)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   196
   apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   197
     apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   198
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   199
    apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   200
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   201
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   202
  apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   203
  apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   204
   apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   205
   apply (rule_tac x = "Abs t" in exI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   206
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   207
  apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   208
  apply (erule rev_mp)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   209
  apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   210
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   211
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   212
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   213
  apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   214
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   215
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   216
theorem explicit_is_implicit:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   217
  "(\<forall>s u. (\<not> free s 0) --> R (Abs (s \<degree> Var 0)) (s[u/0])) =
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   218
    (\<forall>s. R (Abs (lift s 0 \<degree> Var 0)) s)"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   219
  by (auto simp add: not_free_iff_lifted)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   220
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   221
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   222
subsection {* Parallel eta-reduction *}
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   223
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   224
consts
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   225
  par_eta :: "(dB \<times> dB) set"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   226
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
   227
abbreviation
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18557
diff changeset
   228
  par_eta_red :: "[dB, dB] => bool"   (infixl "=e>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
   229
  "s =e> t == (s, t) \<in> par_eta"
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   230
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
   231
abbreviation (xsymbols)
19380
b808efaa5828 tuned syntax/abbreviations;
wenzelm
parents: 19363
diff changeset
   232
  par_eta_red1 :: "[dB, dB] => bool"   (infixl "\<Rightarrow>\<^sub>\<eta>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19086
diff changeset
   233
  "op \<Rightarrow>\<^sub>\<eta> == op =e>"
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   234
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   235
inductive par_eta
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   236
intros
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   237
  var [simp, intro]: "Var x \<Rightarrow>\<^sub>\<eta> Var x"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   238
  eta [simp, intro]: "\<not> free s 0 \<Longrightarrow> s \<Rightarrow>\<^sub>\<eta> s'\<Longrightarrow> Abs (s \<degree> Var 0) \<Rightarrow>\<^sub>\<eta> s'[dummy/0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   239
  app [simp, intro]: "s \<Rightarrow>\<^sub>\<eta> s' \<Longrightarrow> t \<Rightarrow>\<^sub>\<eta> t' \<Longrightarrow> s \<degree> t \<Rightarrow>\<^sub>\<eta> s' \<degree> t'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   240
  abs [simp, intro]: "s \<Rightarrow>\<^sub>\<eta> t \<Longrightarrow> Abs s \<Rightarrow>\<^sub>\<eta> Abs t"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   241
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   242
lemma free_par_eta [simp]:
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   243
  assumes eta: "s \<Rightarrow>\<^sub>\<eta> t"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   244
  shows "free t i = free s i" using eta
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   245
  by (induct fixing: i) (simp_all cong: conj_cong)
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   246
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   247
lemma par_eta_refl [simp]: "t \<Rightarrow>\<^sub>\<eta> t"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   248
  by (induct t) simp_all
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   249
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   250
lemma par_eta_lift [simp]:
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   251
  assumes eta: "s \<Rightarrow>\<^sub>\<eta> t"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   252
  shows "lift s i \<Rightarrow>\<^sub>\<eta> lift t i" using eta
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   253
  by (induct fixing: i) simp_all
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   254
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   255
lemma par_eta_subst [simp]:
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   256
  assumes eta: "s \<Rightarrow>\<^sub>\<eta> t"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   257
  shows "u \<Rightarrow>\<^sub>\<eta> u' \<Longrightarrow> s[u/i] \<Rightarrow>\<^sub>\<eta> t[u'/i]" using eta
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   258
  by (induct fixing: u u' i) (simp_all add: subst_subst [symmetric] subst_Var)
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   259
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   260
theorem eta_subset_par_eta: "eta \<subseteq> par_eta"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   261
  apply (rule subsetI)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   262
  apply clarify
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   263
  apply (erule eta.induct)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   264
  apply (blast intro!: par_eta_refl)+
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   265
  done
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   266
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   267
theorem par_eta_subset_eta: "par_eta \<subseteq> eta\<^sup>*"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   268
  apply (rule subsetI)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   269
  apply clarify
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   270
  apply (erule par_eta.induct)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   271
  apply blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   272
  apply (rule rtrancl_into_rtrancl)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   273
  apply (rule rtrancl_eta_Abs)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   274
  apply (rule rtrancl_eta_AppL)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   275
  apply assumption
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   276
  apply (rule eta.eta)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   277
  apply simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   278
  apply (rule rtrancl_eta_App)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   279
  apply assumption+
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   280
  apply (rule rtrancl_eta_Abs)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   281
  apply assumption
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   282
  done
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   283
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   284
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   285
subsection {* n-ary eta-expansion *}
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   286
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   287
consts eta_expand :: "nat \<Rightarrow> dB \<Rightarrow> dB"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   288
primrec
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   289
  eta_expand_0: "eta_expand 0 t = t"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   290
  eta_expand_Suc: "eta_expand (Suc n) t = Abs (lift (eta_expand n t) 0 \<degree> Var 0)"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   291
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   292
lemma eta_expand_Suc':
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   293
  "eta_expand (Suc n) t = eta_expand n (Abs (lift t 0 \<degree> Var 0))"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   294
  by (induct n fixing: t) simp_all
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   295
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   296
theorem lift_eta_expand: "lift (eta_expand k t) i = eta_expand k (lift t i)"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   297
  by (induct k) (simp_all add: lift_lift)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   298
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   299
theorem eta_expand_beta:
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   300
  assumes u: "u => u'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   301
  shows "t => t' \<Longrightarrow> eta_expand k (Abs t) \<degree> u => t'[u'/0]"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   302
proof (induct k fixing: t t')
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   303
  case 0 with u show ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   304
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   305
  case (Suc k)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   306
  hence "Abs (lift t (Suc 0)) \<degree> Var 0 => lift t' (Suc 0)[Var 0/0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   307
    by (blast intro: par_beta_lift)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   308
  with Suc show ?case by (simp del: eta_expand_Suc add: eta_expand_Suc')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   309
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   310
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   311
theorem eta_expand_red:
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   312
  assumes t: "t => t'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   313
  shows "eta_expand k t => eta_expand k t'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   314
  by (induct k) (simp_all add: t)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   315
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   316
theorem eta_expand_eta: "t \<Rightarrow>\<^sub>\<eta> t' \<Longrightarrow> eta_expand k t \<Rightarrow>\<^sub>\<eta> t'"
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   317
proof (induct k fixing: t t')
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   318
  case 0
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   319
  show ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   320
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   321
  case (Suc k)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   322
  have "Abs (lift (eta_expand k t) 0 \<degree> Var 0) \<Rightarrow>\<^sub>\<eta> lift t' 0[arbitrary/0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   323
    by (fastsimp intro: par_eta_lift Suc)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   324
  thus ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   325
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   326
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   327
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   328
subsection {* Elimination rules for parallel eta reduction *}
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   329
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   330
theorem par_eta_elim_app: assumes eta: "t \<Rightarrow>\<^sub>\<eta> u"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   331
  shows "u = u1' \<degree> u2' \<Longrightarrow>
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   332
    \<exists>u1 u2 k. t = eta_expand k (u1 \<degree> u2) \<and> u1 \<Rightarrow>\<^sub>\<eta> u1' \<and> u2 \<Rightarrow>\<^sub>\<eta> u2'" using eta
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   333
proof (induct fixing: u1' u2')
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   334
  case (app s s' t t')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   335
  have "s \<degree> t = eta_expand 0 (s \<degree> t)" by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   336
  with app show ?case by blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   337
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   338
  case (eta dummy s s')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   339
  then obtain u1'' u2'' where s': "s' = u1'' \<degree> u2''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   340
    by (cases s') (auto simp add: subst_Var free_par_eta [symmetric] split: split_if_asm)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   341
  then have "\<exists>u1 u2 k. s = eta_expand k (u1 \<degree> u2) \<and> u1 \<Rightarrow>\<^sub>\<eta> u1'' \<and> u2 \<Rightarrow>\<^sub>\<eta> u2''" by (rule eta)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   342
  then obtain u1 u2 k where s: "s = eta_expand k (u1 \<degree> u2)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 16417
diff changeset
   343
    and u1: "u1 \<Rightarrow>\<^sub>\<eta> u1''" and u2: "u2 \<Rightarrow>\<^sub>\<eta> u2''" by iprover
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   344
  from u1 u2 eta s' have "\<not> free u1 0" and "\<not> free u2 0"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   345
    by (simp_all del: free_par_eta add: free_par_eta [symmetric])
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   346
  with s have "Abs (s \<degree> Var 0) = eta_expand (Suc k) (u1[dummy/0] \<degree> u2[dummy/0])"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   347
    by (simp del: lift_subst add: lift_subst_dummy lift_eta_expand)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   348
  moreover from u1 par_eta_refl have "u1[dummy/0] \<Rightarrow>\<^sub>\<eta> u1''[dummy/0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   349
    by (rule par_eta_subst)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   350
  moreover from u2 par_eta_refl have "u2[dummy/0] \<Rightarrow>\<^sub>\<eta> u2''[dummy/0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   351
    by (rule par_eta_subst)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   352
  ultimately show ?case using eta s'
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   353
    by (simp only: subst.simps dB.simps) blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   354
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   355
  case var thus ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   356
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   357
  case abs thus ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   358
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   359
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   360
theorem par_eta_elim_abs: assumes eta: "t \<Rightarrow>\<^sub>\<eta> t'"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   361
  shows "t' = Abs u' \<Longrightarrow>
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   362
    \<exists>u k. t = eta_expand k (Abs u) \<and> u \<Rightarrow>\<^sub>\<eta> u'" using eta
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   363
proof (induct fixing: u')
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   364
  case (abs s t)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   365
  have "Abs s = eta_expand 0 (Abs s)" by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   366
  with abs show ?case by blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   367
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   368
  case (eta dummy s s')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   369
  then obtain u'' where s': "s' = Abs u''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   370
    by (cases s') (auto simp add: subst_Var free_par_eta [symmetric] split: split_if_asm)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   371
  then have "\<exists>u k. s = eta_expand k (Abs u) \<and> u \<Rightarrow>\<^sub>\<eta> u''" by (rule eta)
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 16417
diff changeset
   372
  then obtain u k where s: "s = eta_expand k (Abs u)" and u: "u \<Rightarrow>\<^sub>\<eta> u''" by iprover
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   373
  from eta u s' have "\<not> free u (Suc 0)"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   374
    by (simp del: free_par_eta add: free_par_eta [symmetric])
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   375
  with s have "Abs (s \<degree> Var 0) = eta_expand (Suc k) (Abs (u[lift dummy 0/Suc 0]))"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   376
    by (simp del: lift_subst add: lift_eta_expand lift_subst_dummy)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   377
  moreover from u par_eta_refl have "u[lift dummy 0/Suc 0] \<Rightarrow>\<^sub>\<eta> u''[lift dummy 0/Suc 0]"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   378
    by (rule par_eta_subst)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   379
  ultimately show ?case using eta s' by fastsimp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   380
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   381
  case var thus ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   382
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   383
  case app thus ?case by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   384
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   385
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   386
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   387
subsection {* Eta-postponement theorem *}
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   388
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   389
text {*
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   390
Based on a proof by Masako Takahashi \cite{Takahashi-IandC}.
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   391
*}
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   392
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   393
theorem par_eta_beta: "s \<Rightarrow>\<^sub>\<eta> t \<Longrightarrow> t => u \<Longrightarrow> \<exists>t'. s => t' \<and> t' \<Rightarrow>\<^sub>\<eta> u"
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   394
proof (induct t fixing: s u taking: "size :: dB \<Rightarrow> nat" rule: measure_induct_rule)
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   395
  case (less t)
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   396
  from `t => u`
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   397
  show ?case
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   398
  proof cases
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   399
    case (var n)
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   400
    with less show ?thesis
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   401
      by (auto intro: par_beta_refl)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   402
  next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   403
    case (abs r' r'')
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   404
    with less have "s \<Rightarrow>\<^sub>\<eta> Abs r'" by simp
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   405
    then obtain r k where s: "s = eta_expand k (Abs r)" and rr: "r \<Rightarrow>\<^sub>\<eta> r'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   406
      by (blast dest: par_eta_elim_abs)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   407
    from abs have "size r' < size t" by simp
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   408
    with abs(2) rr obtain t' where rt: "r => t'" and t': "t' \<Rightarrow>\<^sub>\<eta> r''"
18557
60a0f9caa0a2 Provers/classical: stricter checks to ensure that supplied intro, dest and
paulson
parents: 18460
diff changeset
   409
      by (blast dest: less(1))
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   410
    with s abs show ?thesis
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   411
      by (auto intro: eta_expand_red eta_expand_eta)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   412
  next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   413
    case (app q' q'' r' r'')
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   414
    with less have "s \<Rightarrow>\<^sub>\<eta> q' \<degree> r'" by simp
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   415
    then obtain q r k where s: "s = eta_expand k (q \<degree> r)"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   416
      and qq: "q \<Rightarrow>\<^sub>\<eta> q'" and rr: "r \<Rightarrow>\<^sub>\<eta> r'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   417
      by (blast dest: par_eta_elim_app [OF _ refl])
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   418
    from app have "size q' < size t" and "size r' < size t" by simp_all
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   419
    with app(2,3) qq rr obtain t' t'' where "q => t'" and
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   420
      "t' \<Rightarrow>\<^sub>\<eta> q''" and "r => t''" and "t'' \<Rightarrow>\<^sub>\<eta> r''"
18557
60a0f9caa0a2 Provers/classical: stricter checks to ensure that supplied intro, dest and
paulson
parents: 18460
diff changeset
   421
      by (blast dest: less(1))
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   422
    with s app show ?thesis
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   423
      by (fastsimp intro: eta_expand_red eta_expand_eta)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   424
  next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   425
    case (beta q' q'' r' r'')
18460
9a1458cb2956 tuned induct proofs;
wenzelm
parents: 18257
diff changeset
   426
    with less have "s \<Rightarrow>\<^sub>\<eta> Abs q' \<degree> r'" by simp
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   427
    then obtain q r k k' where s: "s = eta_expand k (eta_expand k' (Abs q) \<degree> r)"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   428
      and qq: "q \<Rightarrow>\<^sub>\<eta> q'" and rr: "r \<Rightarrow>\<^sub>\<eta> r'"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   429
      by (blast dest: par_eta_elim_app par_eta_elim_abs)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   430
    from beta have "size q' < size t" and "size r' < size t" by simp_all
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   431
    with beta(2,3) qq rr obtain t' t'' where "q => t'" and
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   432
      "t' \<Rightarrow>\<^sub>\<eta> q''" and "r => t''" and "t'' \<Rightarrow>\<^sub>\<eta> r''"
18557
60a0f9caa0a2 Provers/classical: stricter checks to ensure that supplied intro, dest and
paulson
parents: 18460
diff changeset
   433
      by (blast dest: less(1))
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   434
    with s beta show ?thesis
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   435
      by (auto intro: eta_expand_red eta_expand_beta eta_expand_eta par_eta_subst)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   436
  qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   437
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   438
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   439
theorem eta_postponement': assumes eta: "s -e>> t"
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   440
  shows "t => u \<Longrightarrow> \<exists>t'. s => t' \<and> t' -e>> u"
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   441
  using eta [simplified rtrancl_subset
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   442
    [OF eta_subset_par_eta par_eta_subset_eta, symmetric]]
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 17589
diff changeset
   443
proof (induct fixing: u)
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   444
  case 1
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   445
  thus ?case by blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   446
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   447
  case (2 s' s'' s''')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   448
  from 2 obtain t' where s': "s' => t'" and t': "t' \<Rightarrow>\<^sub>\<eta> s'''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   449
    by (auto dest: par_eta_beta)
18557
60a0f9caa0a2 Provers/classical: stricter checks to ensure that supplied intro, dest and
paulson
parents: 18460
diff changeset
   450
  from s' obtain t'' where s: "s => t''" and t'': "t'' -e>> t'" using 2
60a0f9caa0a2 Provers/classical: stricter checks to ensure that supplied intro, dest and
paulson
parents: 18460
diff changeset
   451
    by blast
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   452
  from par_eta_subset_eta t' have "t' -e>> s'''" ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   453
  with t'' have "t'' -e>> s'''" by (rule rtrancl_trans)
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 16417
diff changeset
   454
  with s show ?case by iprover
15522
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   455
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   456
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   457
theorem eta_postponement:
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   458
  assumes st: "(s, t) \<in> (beta \<union> eta)\<^sup>*"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   459
  shows "(s, t) \<in> eta\<^sup>* O beta\<^sup>*" using st
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   460
proof induct
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   461
  case 1
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   462
  show ?case by blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   463
next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   464
  case (2 s' s'')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   465
  from 2(3) obtain t' where s: "s \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and t': "t' -e>> s'" by blast
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   466
  from 2(2) show ?case
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   467
  proof
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   468
    assume "s' -> s''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   469
    with beta_subset_par_beta have "s' => s''" ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   470
    with t' obtain t'' where st: "t' => t''" and tu: "t'' -e>> s''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   471
      by (auto dest: eta_postponement')
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   472
    from par_beta_subset_beta st have "t' \<rightarrow>\<^sub>\<beta>\<^sup>* t''" ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   473
    with s have "s \<rightarrow>\<^sub>\<beta>\<^sup>* t''" by (rule rtrancl_trans)
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   474
    thus ?thesis using tu ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   475
  next
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   476
    assume "s' -e> s''"
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   477
    with t' have "t' -e>> s''" ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   478
    with s show ?thesis ..
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   479
  qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   480
qed
ec0fd05b2f2c Added proof of eta-postponement theorem (using parallel eta-reduction).
berghofe
parents: 13187
diff changeset
   481
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
   482
end