src/HOL/Lambda/Eta.thy
author nipkow
Thu, 30 May 2002 10:12:52 +0200
changeset 13187 e5434b822a96
parent 12011 1a3a7b3cd9bb
child 15522 ec0fd05b2f2c
permissions -rw-r--r--
Modifications due to enhanced linear arithmetic.
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$
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
     4
    Copyright   1995 TU Muenchen
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
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
     9
theory Eta = ParRed:
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
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    24
syntax 
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    25
  "_eta" :: "[dB, dB] => bool"   (infixl "-e>" 50)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    26
  "_eta_rtrancl" :: "[dB, dB] => bool"   (infixl "-e>>" 50)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    27
  "_eta_reflcl" :: "[dB, dB] => bool"   (infixl "-e>=" 50)
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
    28
translations
9827
wenzelm
parents: 9811
diff changeset
    29
  "s -e> t" == "(s, t) \<in> eta"
wenzelm
parents: 9811
diff changeset
    30
  "s -e>> t" == "(s, t) \<in> eta^*"
wenzelm
parents: 9811
diff changeset
    31
  "s -e>= t" == "(s, t) \<in> eta^="
1269
ee011b365770 New version with eta reduction.
nipkow
parents:
diff changeset
    32
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1376
diff changeset
    33
inductive eta
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
    34
  intros
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    35
    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
    36
    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
    37
    appR [simp, intro]: "s -e> t ==> u \<degree> s -e> u \<degree> t"
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
    38
    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
    39
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    40
inductive_cases eta_cases [elim!]:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    41
  "Abs s -e> z"
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
    42
  "s \<degree> t -e> u"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    43
  "Var i -e> t"
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
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    46
subsection "Properties of eta, subst and free"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    47
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    48
lemma subst_not_free [rule_format, simp]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    49
    "\<forall>i t u. \<not> free s i --> s[t/i] = s[u/i]"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    50
  apply (induct_tac s)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    51
    apply (simp_all add: subst_Var)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    52
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    53
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    54
lemma free_lift [simp]:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    55
    "\<forall>i k. free (lift t k) i =
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    56
      (i < k \<and> free t i \<or> k < i \<and> free t (i - 1))"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    57
  apply (induct_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    58
    apply (auto cong: conj_cong)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    59
  apply arith
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    60
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    61
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    62
lemma free_subst [simp]:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    63
    "\<forall>i k t. free (s[t/k]) i =
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    64
      (free s k \<and> free t i \<or> free s (if i < k then i else i + 1))"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    65
  apply (induct_tac s)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    66
    prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    67
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    68
    apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    69
   prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    70
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    71
  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
    72
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    73
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    74
lemma free_eta [rule_format]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    75
    "s -e> t ==> \<forall>i. free t i = free s i"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    76
  apply (erule eta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    77
     apply (simp_all cong: conj_cong)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    78
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    79
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    80
lemma not_free_eta:
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    81
    "[| s -e> t; \<not> free s i |] ==> \<not> free t i"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    82
  apply (simp add: free_eta)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    83
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    84
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    85
lemma eta_subst [rule_format, simp]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    86
    "s -e> t ==> \<forall>u i. s[u/i] -e> t[u/i]"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    87
  apply (erule eta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    88
  apply (simp_all add: subst_subst [symmetric])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    89
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    90
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    91
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    92
subsection "Confluence of eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    93
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    94
lemma square_eta: "square eta eta (eta^=) (eta^=)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    95
  apply (unfold square_def id_def)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    96
  apply (rule impI [THEN allI [THEN allI]])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    97
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    98
  apply (erule eta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
    99
     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
   100
    apply safe
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   101
       prefer 5
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   102
       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
   103
      apply blast+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   104
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   105
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   106
theorem eta_confluent: "confluent eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   107
  apply (rule square_eta [THEN square_reflcl_confluent])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   108
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   109
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   110
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   111
subsection "Congruence rules for eta*"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   112
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   113
lemma rtrancl_eta_Abs: "s -e>> s' ==> Abs s -e>> Abs s'"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   114
  apply (erule rtrancl_induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   115
   apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   116
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   117
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   118
lemma rtrancl_eta_AppL: "s -e>> s' ==> s \<degree> t -e>> s' \<degree> t"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   119
  apply (erule rtrancl_induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   120
   apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   121
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   122
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   123
lemma rtrancl_eta_AppR: "t -e>> t' ==> s \<degree> t -e>> s \<degree> t'"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   124
  apply (erule rtrancl_induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   125
   apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   126
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   127
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   128
lemma rtrancl_eta_App:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   129
    "[| s -e>> s'; t -e>> t' |] ==> s \<degree> t -e>> s' \<degree> t'"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   130
  apply (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtrancl_trans)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   131
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   132
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   133
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   134
subsection "Commutation of beta and eta"
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
   135
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   136
lemma free_beta [rule_format]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   137
    "s -> t ==> \<forall>i. free t i --> free s i"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   138
  apply (erule beta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   139
     apply simp_all
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   140
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   141
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   142
lemma beta_subst [rule_format, intro]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   143
    "s -> t ==> \<forall>u i. s[u/i] -> t[u/i]"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   144
  apply (erule beta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   145
     apply (simp_all add: subst_subst [symmetric])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   146
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   147
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   148
lemma subst_Var_Suc [simp]: "\<forall>i. t[Var i/i] = t[Var(i)/i + 1]"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   149
  apply (induct_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   150
  apply (auto elim!: linorder_neqE simp: subst_Var)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   151
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   152
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   153
lemma eta_lift [rule_format, simp]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   154
    "s -e> t ==> \<forall>i. lift s i -e> lift t i"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   155
  apply (erule eta.induct)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   156
     apply simp_all
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   157
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   158
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   159
lemma rtrancl_eta_subst [rule_format]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   160
    "\<forall>s t i. s -e> t --> u[s/i] -e>> u[t/i]"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   161
  apply (induct_tac u)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   162
    apply (simp_all add: subst_Var)
11183
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   163
    apply (blast)
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   164
   apply (blast intro: rtrancl_eta_App)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   165
  apply (blast intro!: rtrancl_eta_Abs eta_lift)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   166
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   167
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   168
lemma square_beta_eta: "square beta eta (eta^*) (beta^=)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   169
  apply (unfold square_def)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   170
  apply (rule impI [THEN allI [THEN allI]])
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   171
  apply (erule beta.induct)
11183
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   172
     apply (slowsimp intro: rtrancl_eta_subst eta_subst)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   173
    apply (blast intro: rtrancl_eta_AppL)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   174
   apply (blast intro: rtrancl_eta_AppR)
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   175
  apply simp;
0476c6e07bca minor tweaks to speed the proofs
paulson
parents: 9941
diff changeset
   176
  apply (slowsimp intro: rtrancl_eta_Abs free_beta
9858
c3ac6128b649 use 'iff' modifier;
wenzelm
parents: 9827
diff changeset
   177
    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
   178
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   179
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   180
lemma confluent_beta_eta: "confluent (beta \<union> eta)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   181
  apply (assumption |
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   182
    rule square_rtrancl_reflcl_commute confluent_Un
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   183
      beta_confluent eta_confluent square_beta_eta)+
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   184
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   185
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   186
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   187
subsection "Implicit definition of eta"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   188
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   189
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
   190
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   191
lemma not_free_iff_lifted [rule_format]:
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   192
    "\<forall>i. (\<not> free s i) = (\<exists>t. s = lift t i)"
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   193
  apply (induct_tac s)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   194
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   195
    apply clarify
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   196
    apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   197
     apply (erule linorder_neqE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   198
      apply (rule_tac x = "Var nat" in exI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   199
      apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   200
     apply (rule_tac x = "Var (nat - 1)" in exI)
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 clarify
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   203
    apply (rule notE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   204
     prefer 2
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   205
     apply assumption
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   206
    apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   207
    apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   208
      apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   209
     apply simp
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 (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   213
   apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   214
   apply (rule allI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   215
   apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   216
    apply (elim conjE exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   217
    apply (rename_tac u1 u2)
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   218
    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
   219
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   220
   apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   221
   apply (erule rev_mp)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   222
   apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   223
     apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   224
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   225
    apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   226
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   227
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   228
  apply (erule thin_rl)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   229
  apply (rule allI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   230
  apply (rule iffI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   231
   apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   232
   apply (rule_tac x = "Abs t" in exI)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   233
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   234
  apply (erule exE)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   235
  apply (erule rev_mp)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   236
  apply (case_tac t)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   237
    apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   238
   apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   239
  apply simp
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   240
  apply blast
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   241
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   242
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   243
theorem explicit_is_implicit:
12011
1a3a7b3cd9bb tuned notation (degree instead of dollar);
wenzelm
parents: 11638
diff changeset
   244
  "(\<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
   245
    (\<forall>s. R (Abs (lift s 0 \<degree> Var 0)) s)"
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   246
  apply (auto simp add: not_free_iff_lifted)
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   247
  done
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9102
diff changeset
   248
11638
2c3dee321b4b inductive: no collective atts;
wenzelm
parents: 11183
diff changeset
   249
end