src/HOL/Nominal/Examples/Standardization.thy
author paulson <lp15@cam.ac.uk>
Mon, 22 Apr 2024 10:43:57 +0100
changeset 80141 022a9c26b14f
parent 71989 bad75618fb82
child 80914 d97fdabd9e2b
permissions -rw-r--r--
Tidied up another messy theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Nominal/Examples/Standardization.thy
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer and Tobias Nipkow
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     3
    Copyright   2005, 2008 TU Muenchen
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     4
*)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     5
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
     6
section \<open>Standardization\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     7
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
     8
theory Standardization
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63882
diff changeset
     9
imports "HOL-Nominal.Nominal"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    10
begin
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    11
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    12
text \<open>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    13
The proof of the standardization theorem, as well as most of the theorems about
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    14
lambda calculus in the following sections, are taken from \<open>HOL/Lambda\<close>.
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    15
\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    16
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    17
subsection \<open>Lambda terms\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    18
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    19
atom_decl name
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    20
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    21
nominal_datatype lam =
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    22
  Var "name"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    23
| App "lam" "lam" (infixl "\<degree>" 200)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    24
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [0, 10] 10)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    25
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    26
instantiation lam :: size
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    27
begin
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    28
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    29
nominal_primrec size_lam
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    30
where
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    31
  "size (Var n) = 0"
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    32
| "size (t \<degree> u) = size t + size u + 1"
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    33
| "size (Lam [x].t) = size t + 1"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    34
  by (finite_guess | simp add: fresh_nat | fresh_guess)+
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    35
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    36
instance ..
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    37
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    38
end
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    39
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    40
nominal_primrec
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    41
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [300, 0, 0] 300)
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    42
where
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    43
  subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 47397
diff changeset
    44
| subst_App: "(t\<^sub>1 \<degree> t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \<degree> t\<^sub>2[y::=s]"
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 27623
diff changeset
    45
| subst_Lam: "x \<sharp> (y, s) \<Longrightarrow> (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    46
  by (finite_guess | simp add: abs_fresh | fresh_guess)+
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    47
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    48
lemma subst_eqvt [eqvt]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    49
  "(pi::name prm) \<bullet> (t[x::=u]) = (pi \<bullet> t)[(pi \<bullet> x)::=(pi \<bullet> u)]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    50
  by (nominal_induct t avoiding: x u rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    51
     (perm_simp add: fresh_bij)+
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    52
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    53
lemma subst_rename:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    54
  "y \<sharp> t \<Longrightarrow> ([(y, x)] \<bullet> t)[y::=u] = t[x::=u]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    55
  by (nominal_induct t avoiding: x y u rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    56
     (simp_all add: fresh_atm calc_atm abs_fresh)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    57
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    58
lemma fresh_subst: 
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    59
  "(x::name) \<sharp> t \<Longrightarrow> x \<sharp> u \<Longrightarrow> x \<sharp> t[y::=u]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    60
  by (nominal_induct t avoiding: x y u rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    61
     (auto simp add: abs_fresh fresh_atm)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    62
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    63
lemma fresh_subst': 
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    64
  "(x::name) \<sharp> u \<Longrightarrow> x \<sharp> t[x::=u]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    65
  by (nominal_induct t avoiding: x u rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    66
     (auto simp add: abs_fresh fresh_atm)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    67
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    68
lemma subst_forget: "(x::name) \<sharp> t \<Longrightarrow> t[x::=u] = t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    69
  by (nominal_induct t avoiding: x u rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    70
     (auto simp add: abs_fresh fresh_atm)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    71
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    72
lemma subst_subst:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    73
  "x \<noteq> y \<Longrightarrow> x \<sharp> v \<Longrightarrow> t[y::=v][x::=u[y::=v]] = t[x::=u][y::=v]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    74
  by (nominal_induct t avoiding: x y u v rule: lam.strong_induct)
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
    75
     (auto simp add: fresh_subst subst_forget)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    76
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    77
declare subst_Var [simp del]
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    78
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    79
lemma subst_eq [simp]: "(Var x)[x::=u] = u"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    80
  by (simp add: subst_Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    81
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    82
lemma subst_neq [simp]: "x \<noteq> y \<Longrightarrow> (Var x)[y::=u] = Var x"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    83
  by (simp add: subst_Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    84
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    85
inductive beta :: "lam \<Rightarrow> lam \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    86
  where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    87
    beta: "x \<sharp> t \<Longrightarrow> (Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    88
  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    89
  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    90
  | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta> (Lam [x].t)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    91
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    92
equivariance beta
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    93
nominal_inductive beta
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    94
  by (simp_all add: abs_fresh fresh_subst')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    95
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    96
lemma better_beta [simp, intro!]: "(Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    97
proof -
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    98
  obtain y::name where y: "y \<sharp> (x, s, t)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
    99
    by (rule exists_fresh) (rule fin_supp)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   100
  then have "y \<sharp> t" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   101
  then have "(Lam [y]. [(y, x)] \<bullet> s) \<degree> t \<rightarrow>\<^sub>\<beta> ([(y, x)] \<bullet> s)[y::=t]"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   102
    by (rule beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   103
  moreover from y have "(Lam [x].s) = (Lam [y]. [(y, x)] \<bullet> s)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   104
    by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   105
  ultimately show ?thesis using y by (simp add: subst_rename)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   106
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   107
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   108
abbreviation
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   109
  beta_reds :: "lam \<Rightarrow> lam \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50) where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   110
  "s \<rightarrow>\<^sub>\<beta>\<^sup>* t \<equiv> beta\<^sup>*\<^sup>* s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   111
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   112
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   113
subsection \<open>Application of a term to a list of terms\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   114
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   115
abbreviation
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   116
  list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam"  (infixl "\<degree>\<degree>" 150) where
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   117
  "t \<degree>\<degree> ts \<equiv> foldl (\<degree>) t ts"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   118
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   119
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   120
  by (induct ts rule: rev_induct) (auto simp add: lam.inject)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   121
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   122
lemma Var_eq_apps_conv [iff]: "(Var m = s \<degree>\<degree> ss) = (Var m = s \<and> ss = [])"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   123
  by (induct ss arbitrary: s) auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   124
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   125
lemma Var_apps_eq_Var_apps_conv [iff]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   126
    "(Var m \<degree>\<degree> rs = Var n \<degree>\<degree> ss) = (m = n \<and> rs = ss)"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   127
proof (induct rs arbitrary: ss rule: rev_induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   128
  case Nil then show ?case by (auto simp add: lam.inject)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   129
next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   130
  case (snoc x xs) then show ?case 
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   131
  by (induct ss rule: rev_induct) (auto simp add: lam.inject)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   132
qed
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   133
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   134
lemma App_eq_foldl_conv:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   135
  "(r \<degree> s = t \<degree>\<degree> ts) =
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   136
    (if ts = [] then r \<degree> s = t
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   137
    else (\<exists>ss. ts = ss @ [s] \<and> r = t \<degree>\<degree> ss))"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   138
  by (rule rev_exhaust [of ts]) (auto simp: lam.inject)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   139
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   140
lemma Abs_eq_apps_conv [iff]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   141
    "((Lam [x].r) = s \<degree>\<degree> ss) = ((Lam [x].r) = s \<and> ss = [])"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   142
  by (induct ss rule: rev_induct) auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   143
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   144
lemma apps_eq_Abs_conv [iff]: "(s \<degree>\<degree> ss = (Lam [x].r)) = (s = (Lam [x].r) \<and> ss = [])"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   145
  by (induct ss rule: rev_induct) auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   146
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   147
lemma Abs_App_neq_Var_apps [iff]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   148
    "(Lam [x].s) \<degree> t \<noteq> Var n \<degree>\<degree> ss"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   149
  by (induct ss arbitrary: s t rule: rev_induct) (auto simp add: lam.inject)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   150
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   151
lemma Var_apps_neq_Abs_apps [iff]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   152
    "Var n \<degree>\<degree> ts \<noteq> (Lam [x].r) \<degree>\<degree> ss"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   153
proof (induct ss arbitrary: ts rule: rev_induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   154
  case Nil then show ?case by simp
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   155
next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   156
  case (snoc x xs) then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   157
  by (induct ts rule: rev_induct) (auto simp add: lam.inject)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   158
qed
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   159
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   160
lemma ex_head_tail:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   161
  "\<exists>ts h. t = h \<degree>\<degree> ts \<and> ((\<exists>n. h = Var n) \<or> (\<exists>x u. h = (Lam [x].u)))"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   162
proof (induct t rule: lam.induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   163
  case (App lam1 lam2)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   164
  then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   165
    by (metis foldl_Cons foldl_Nil foldl_append)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   166
qed auto
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   167
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   168
lemma size_apps [simp]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   169
  "size (r \<degree>\<degree> rs) = size r + foldl (+) 0 (map size rs) + length rs"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   170
  by (induct rs rule: rev_induct) auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   171
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   172
lemma lem0: "(0::nat) < k \<Longrightarrow> m \<le> n \<Longrightarrow> m < n + k"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   173
  by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   174
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   175
lemma subst_map [simp]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   176
    "(t \<degree>\<degree> ts)[x::=u] = t[x::=u] \<degree>\<degree> map (\<lambda>t. t[x::=u]) ts"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   177
  by (induct ts arbitrary: t) simp_all
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   178
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   179
lemma app_last: "(t \<degree>\<degree> ts) \<degree> u = t \<degree>\<degree> (ts @ [u])"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   180
  by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   181
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   182
lemma perm_apps [eqvt]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   183
  "(pi::name prm) \<bullet> (t \<degree>\<degree> ts) = ((pi \<bullet> t) \<degree>\<degree> (pi \<bullet> ts))"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   184
  by (induct ts rule: rev_induct) (auto simp add: append_eqvt)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   185
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   186
lemma fresh_apps [simp]: "(x::name) \<sharp> (t \<degree>\<degree> ts) = (x \<sharp> t \<and> x \<sharp> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   187
  by (induct ts rule: rev_induct)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   188
    (auto simp add: fresh_list_append fresh_list_nil fresh_list_cons)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   189
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   190
text \<open>A customized induction schema for \<open>\<degree>\<degree>\<close>.\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   191
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   192
lemma Apps_lam_induct_aux:
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   193
  assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   194
    and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   195
  shows "size t = n \<Longrightarrow> P z t"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   196
proof (induct n arbitrary: t z rule: less_induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   197
  case (less n)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   198
  obtain ts h where t: "t = h \<degree>\<degree> ts" and D: "(\<exists>a. h = Var a) \<or> (\<exists>x u. h = (Lam [x].u))"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   199
    using ex_head_tail [of t] by metis
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   200
  show ?case 
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   201
    using D
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   202
  proof (elim exE disjE)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   203
    fix a :: name
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   204
    assume h: "h = Var a"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   205
    have "P z t" if "t \<in> set ts" for z t
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   206
    proof -
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   207
      have "size t < length ts + fold (+) (map size ts) 0"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   208
        using that
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   209
        by (fastforce simp add: sum_list_map_remove1 fold_plus_sum_list_rev)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   210
      then have "size t < size (Var a \<degree>\<degree> ts)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   211
        by simp (simp add: add.commute foldl_conv_fold)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   212
      then show ?thesis
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   213
        using h less.hyps less.prems t by blast
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   214
    qed
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   215
    then show "P z t"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   216
      unfolding t h by (blast intro: assms)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   217
  next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   218
    fix x u
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   219
    assume h: "h = (Lam [x].u)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   220
    obtain y::name where y: "y \<sharp> (x, u, z)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   221
      by (metis exists_fresh' fin_supp) 
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   222
    then have eq: "(Lam [x].u) = (Lam [y].([(y, x)] \<bullet> u))"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   223
      by (metis alpha' fresh_prod lam.inject(3) perm_fresh_fresh)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   224
    show "P z t"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   225
      unfolding t h eq
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   226
    proof (intro assms strip)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   227
      show "y \<sharp> z"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   228
        by (simp add: y)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   229
      have "size ([(y, x)] \<bullet> u) < size ((Lam [x].u) \<degree>\<degree> ts)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   230
        by (simp add: eq)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   231
      then show "P z ([(y, x)] \<bullet> u)" for z
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   232
        using h less.hyps less.prems t by blast
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   233
      show "P z t" if "t\<in>set ts" for z t
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   234
      proof -
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   235
        have 2: "size t < size ((Lam [x].u) \<degree>\<degree> ts)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   236
          using that
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   237
          apply (simp add: eq)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   238
          apply (simp only: foldl_conv_fold add.commute fold_plus_sum_list_rev)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   239
          apply (fastforce simp add: sum_list_map_remove1)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   240
          done
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   241
        then show ?thesis
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   242
          using h less.hyps less.prems t by blast
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   243
      qed
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   244
    qed
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   245
  qed
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   246
qed
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   247
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   248
theorem Apps_lam_induct:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   249
  assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   250
    and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   251
  shows "P z t"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   252
  using Apps_lam_induct_aux [of P] assms by blast
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   253
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   254
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   255
subsection \<open>Congruence rules\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   256
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   257
lemma apps_preserves_beta [simp]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   258
    "r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> r \<degree>\<degree> ss \<rightarrow>\<^sub>\<beta> s \<degree>\<degree> ss"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   259
  by (induct ss rule: rev_induct) auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   260
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   261
lemma rtrancl_beta_Abs [intro!]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   262
    "s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta>\<^sup>* (Lam [x].s')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   263
  by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   264
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   265
lemma rtrancl_beta_AppL:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   266
    "s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   267
  by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   268
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   269
lemma rtrancl_beta_AppR:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   270
    "t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s \<degree> t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   271
  by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   272
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   273
lemma rtrancl_beta_App [intro]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   274
    "s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   275
  by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtranclp_trans)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   276
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   277
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   278
subsection \<open>Lifting an order to lists of elements\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   279
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   280
definition
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   281
  step1 :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   282
  "step1 r \<equiv>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   283
    (\<lambda>ys xs. \<exists>us z z' vs. xs = us @ z # vs \<and> r z' z \<and> ys =
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   284
      us @ z' # vs)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   285
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   286
lemma not_Nil_step1 [iff]: "\<not> step1 r [] xs"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   287
  by (simp add: step1_def)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   288
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   289
lemma not_step1_Nil [iff]: "\<not> step1 r xs []"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   290
  by (simp add: step1_def)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   291
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   292
lemma Cons_step1_Cons [iff]:
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   293
    "step1 r (y # ys) (x # xs) \<longleftrightarrow> r y x \<and> xs = ys \<or> x = y \<and> step1 r ys xs"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   294
  apply (rule )
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   295
   apply (smt (verit, ccfv_SIG) append_eq_Cons_conv list.inject step1_def)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   296
  by (metis append_Cons append_Nil step1_def)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   297
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   298
lemma Cons_step1E [elim!]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   299
  assumes "step1 r ys (x # xs)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   300
    and "\<And>y. ys = y # xs \<Longrightarrow> r y x \<Longrightarrow> R"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   301
    and "\<And>zs. ys = x # zs \<Longrightarrow> step1 r zs xs \<Longrightarrow> R"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   302
  shows R
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   303
  by (metis Cons_step1_Cons assms list.exhaust not_Nil_step1)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   304
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   305
lemma append_step1I:
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   306
  "step1 r ys xs \<and> vs = us \<or> ys = xs \<and> step1 r vs us
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   307
    \<Longrightarrow> step1 r (ys @ vs) (xs @ us)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   308
  by (smt (verit) append_Cons append_assoc step1_def)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   309
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   310
lemma Snoc_step1_SnocD:
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   311
  assumes "step1 r (ys @ [y]) (xs @ [x])"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   312
  shows "(step1 r ys xs \<and> y = x \<or> ys = xs \<and> r y x)"
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   313
  using assms
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   314
    apply (clarsimp simp: step1_def)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   315
  by (metis butlast.simps(2) butlast_append butlast_snoc last.simps last_appendR list.simps(3))
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   316
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   317
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   318
subsection \<open>Lifting beta-reduction to lists\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   319
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   320
abbreviation
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   321
  list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool"  (infixl "[\<rightarrow>\<^sub>\<beta>]\<^sub>1" 50) where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   322
  "rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<equiv> step1 beta rs ss"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   323
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   324
lemma head_Var_reduction:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   325
  "Var n \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> v \<Longrightarrow> \<exists>ss. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<and> v = Var n \<degree>\<degree> ss"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   326
proof (induct u \<equiv> "Var n \<degree>\<degree> rs" v arbitrary: rs set: beta)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   327
  case (appL s t u)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   328
  then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   329
    by (smt (verit, best) App_eq_foldl_conv app_last append_step1I lam.distinct(1))
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   330
next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   331
  case (appR s t u)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   332
  then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   333
    by (smt (verit, ccfv_SIG) App_eq_foldl_conv Cons_step1_Cons app_last append_step1I lam.distinct(1))
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   334
qed auto
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   335
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   336
lemma apps_betasE [case_names appL appR beta, consumes 1]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   337
  assumes major: "r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> s"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   338
    and cases: "\<And>r'. r \<rightarrow>\<^sub>\<beta> r' \<Longrightarrow> s = r' \<degree>\<degree> rs \<Longrightarrow> R"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   339
      "\<And>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<Longrightarrow> s = r \<degree>\<degree> rs' \<Longrightarrow> R"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   340
      "\<And>t u us. (x \<sharp> r \<Longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us) \<Longrightarrow> R"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   341
  shows R
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   342
proof -
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   343
  note [[simproc del: defined_all]]
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   344
  from major have
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   345
   "(\<exists>r'. r \<rightarrow>\<^sub>\<beta> r' \<and> s = r' \<degree>\<degree> rs) \<or>
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   346
    (\<exists>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<and> s = r \<degree>\<degree> rs') \<or>
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   347
    (\<exists>t u us. x \<sharp> r \<longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us)"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   348
  proof (nominal_induct u \<equiv> "r \<degree>\<degree> rs" s avoiding: x r rs rule: beta.strong_induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   349
    case (beta y t s)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   350
    then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   351
      apply (simp add: App_eq_foldl_conv split: if_split_asm)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   352
       apply blast
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   353
      by (metis (no_types, lifting) abs_fresh(1) alpha' lam.fresh(3) lam.inject(3) subst_rename)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   354
  next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   355
    case (appL s t u)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   356
    then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   357
      apply (simp add: App_eq_foldl_conv split: if_split_asm)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   358
      apply blast
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   359
      by (smt (verit) append_Cons append_step1I snoc_eq_iff_butlast)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   360
  next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   361
    case (appR s t u)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   362
    then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   363
      apply (simp add: App_eq_foldl_conv step1_def split: if_split_asm)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   364
      apply force
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   365
      by (metis snoc_eq_iff_butlast)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   366
  next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   367
    case (abs s t x)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   368
    then show ?case
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   369
      by blast
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   370
  qed
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   371
  with cases show ?thesis by blast
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   372
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   373
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   374
lemma apps_preserves_betas [simp]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   375
    "rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> r \<degree>\<degree> ss"
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   376
proof (induct rs arbitrary: ss rule: rev_induct)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   377
  case Nil
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   378
  then show ?case by simp
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   379
next
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   380
  case (snoc x ts)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   381
  then show ?case 
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   382
    apply (simp add: step1_def)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   383
    by (smt (verit) appR app_last apps_preserves_beta foldl_Cons foldl_append)
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   384
qed
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   385
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   386
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   387
subsection \<open>Standard reduction relation\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   388
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   389
text \<open>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   390
Based on lecture notes by Ralph Matthes,
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   391
original proof idea due to Ralph Loader.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   392
\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   393
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   394
declare listrel_mono [mono_set]
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   395
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   396
lemma listrelp_eqvt [eqvt]:
46317
80dccedd6c14 generalize type of List.listrel
huffman
parents: 46129
diff changeset
   397
  fixes f :: "'a::pt_name \<Rightarrow> 'b::pt_name \<Rightarrow> bool"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   398
  assumes xy: "listrelp f (x::'a::pt_name list) y"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   399
  shows "listrelp ((pi::name prm) \<bullet> f) (pi \<bullet> x) (pi \<bullet> y)" using xy
45966
03ce2b2a29a2 tuned proofs
haftmann
parents: 44890
diff changeset
   400
  by induct (simp_all add: listrelp.intros perm_app [symmetric])
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   401
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   402
inductive
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   403
  sred :: "lam \<Rightarrow> lam \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>s" 50)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   404
  and sredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool"  (infixl "[\<rightarrow>\<^sub>s]" 50)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   405
where
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   406
  "s [\<rightarrow>\<^sub>s] t \<equiv> listrelp (\<rightarrow>\<^sub>s) s t"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   407
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   408
| Abs: "x \<sharp> (ss, ss') \<Longrightarrow> r \<rightarrow>\<^sub>s r' \<Longrightarrow> ss [\<rightarrow>\<^sub>s] ss' \<Longrightarrow> (Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   409
| Beta: "x \<sharp> (s, ss, t) \<Longrightarrow> r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   410
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   411
equivariance sred
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   412
nominal_inductive sred
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   413
  by (simp add: abs_fresh)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   414
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   415
lemma better_sred_Abs:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   416
  assumes H1: "r \<rightarrow>\<^sub>s r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   417
  and H2: "ss [\<rightarrow>\<^sub>s] ss'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   418
  shows "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   419
proof -
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   420
  obtain y::name where y: "y \<sharp> (x, r, r', ss, ss')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   421
    by (rule exists_fresh) (rule fin_supp)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   422
  then have "y \<sharp> (ss, ss')" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   423
  moreover from H1 have "[(y, x)] \<bullet> (r \<rightarrow>\<^sub>s r')" by (rule perm_boolI)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   424
  then have "([(y, x)] \<bullet> r) \<rightarrow>\<^sub>s ([(y, x)] \<bullet> r')" by (simp add: eqvts)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   425
  ultimately have "(Lam [y]. [(y, x)] \<bullet> r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [y]. [(y, x)] \<bullet> r') \<degree>\<degree> ss'" using H2
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   426
    by (rule sred.Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   427
  moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   428
    by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   429
  moreover from y have "(Lam [x].r') = (Lam [y]. [(y, x)] \<bullet> r')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   430
    by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   431
  ultimately show ?thesis by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   432
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   433
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   434
lemma better_sred_Beta:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   435
  assumes H: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   436
  shows "(Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   437
proof -
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   438
  obtain y::name where y: "y \<sharp> (x, r, s, ss, t)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   439
    by (rule exists_fresh) (rule fin_supp)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   440
  then have "y \<sharp> (s, ss, t)" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   441
  moreover from y H have "([(y, x)] \<bullet> r)[y::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   442
    by (simp add: subst_rename)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   443
  ultimately have "(Lam [y].[(y, x)] \<bullet> r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   444
    by (rule sred.Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   445
  moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   446
    by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   447
  ultimately show ?thesis by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   448
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   449
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   450
lemmas better_sred_intros = sred.Var better_sred_Abs better_sred_Beta
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   451
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   452
lemma refl_listrelp: "\<forall>x\<in>set xs. R x x \<Longrightarrow> listrelp R xs xs"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   453
  by (induct xs) (auto intro: listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   454
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   455
lemma refl_sred: "t \<rightarrow>\<^sub>s t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   456
  by (nominal_induct t rule: Apps_lam_induct) (auto intro: refl_listrelp better_sred_intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   457
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   458
lemma listrelp_conj1: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp R x y"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   459
  by (erule listrelp.induct) (auto intro: listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   460
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   461
lemma listrelp_conj2: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp S x y"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   462
  by (erule listrelp.induct) (auto intro: listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   463
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   464
lemma listrelp_app:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   465
  assumes xsys: "listrelp R xs ys"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   466
  shows "listrelp R xs' ys' \<Longrightarrow> listrelp R (xs @ xs') (ys @ ys')" using xsys
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   467
  by (induct arbitrary: xs' ys') (auto intro: listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   468
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   469
lemma lemma1:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   470
  assumes r: "r \<rightarrow>\<^sub>s r'" and s: "s \<rightarrow>\<^sub>s s'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   471
  shows "r \<degree> s \<rightarrow>\<^sub>s r' \<degree> s'" using r
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   472
proof induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   473
  case (Var rs rs' x)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   474
  then have "rs [\<rightarrow>\<^sub>s] rs'" by (rule listrelp_conj1)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   475
  moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   476
  ultimately have "rs @ [s] [\<rightarrow>\<^sub>s] rs' @ [s']" by (rule listrelp_app)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   477
  hence "Var x \<degree>\<degree> (rs @ [s]) \<rightarrow>\<^sub>s Var x \<degree>\<degree> (rs' @ [s'])" by (rule sred.Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   478
  thus ?case by (simp only: app_last)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   479
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   480
  case (Abs x ss ss' r r')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   481
  from Abs(4) have "ss [\<rightarrow>\<^sub>s] ss'" by (rule listrelp_conj1)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   482
  moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   483
  ultimately have "ss @ [s] [\<rightarrow>\<^sub>s] ss' @ [s']" by (rule listrelp_app)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   484
  with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> (ss' @ [s'])"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   485
    by (rule better_sred_Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   486
  thus ?case by (simp only: app_last)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   487
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   488
  case (Beta x u ss t r)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   489
  hence "r[x::=u] \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (simp only: app_last)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   490
  hence "(Lam [x].r) \<degree> u \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (rule better_sred_Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   491
  thus ?case by (simp only: app_last)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   492
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   493
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   494
lemma lemma1':
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   495
  assumes ts: "ts [\<rightarrow>\<^sub>s] ts'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   496
  shows "r \<rightarrow>\<^sub>s r' \<Longrightarrow> r \<degree>\<degree> ts \<rightarrow>\<^sub>s r' \<degree>\<degree> ts'" using ts
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   497
  by (induct arbitrary: r r') (auto intro: lemma1)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   498
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   499
lemma listrelp_betas:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   500
  assumes ts: "listrelp (\<rightarrow>\<^sub>\<beta>\<^sup>*) ts ts'"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   501
  shows "\<And>t t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<degree>\<degree> ts'" using ts
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   502
  by induct auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   503
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   504
lemma lemma2:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   505
  assumes t: "t \<rightarrow>\<^sub>s u"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   506
  shows "t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using t
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   507
  by induct (auto dest: listrelp_conj2
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   508
    intro: listrelp_betas apps_preserves_beta converse_rtranclp_into_rtranclp)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   509
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   510
lemma lemma3:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   511
  assumes r: "r \<rightarrow>\<^sub>s r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   512
  shows "s \<rightarrow>\<^sub>s s' \<Longrightarrow> r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" using r
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   513
proof (nominal_induct avoiding: x s s' rule: sred.strong_induct)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   514
  case (Var rs rs' y)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   515
  hence "map (\<lambda>t. t[x::=s]) rs [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   516
    by induct (auto intro: listrelp.intros Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   517
  moreover have "Var y[x::=s] \<rightarrow>\<^sub>s Var y[x::=s']"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   518
    by (cases "y = x") (auto simp add: Var intro: refl_sred)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   519
  ultimately show ?case by simp (rule lemma1')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   520
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   521
  case (Abs y ss ss' r r')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   522
  then have "r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" by fast
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   523
  moreover from Abs(8) \<open>s \<rightarrow>\<^sub>s s'\<close> have "map (\<lambda>t. t[x::=s]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) ss'"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   524
    by induct (auto intro: listrelp.intros Abs)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   525
  ultimately show ?case using Abs(6) \<open>y \<sharp> x\<close> \<open>y \<sharp> s\<close> \<open>y \<sharp> s'\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   526
    by simp (rule better_sred_Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   527
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   528
  case (Beta y u ss t r)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   529
  thus ?case by (auto simp add: subst_subst fresh_atm intro: better_sred_Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   530
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   531
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   532
lemma lemma4_aux:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   533
  assumes rs: "listrelp (\<lambda>t u. t \<rightarrow>\<^sub>s u \<and> (\<forall>r. u \<rightarrow>\<^sub>\<beta> r \<longrightarrow> t \<rightarrow>\<^sub>s r)) rs rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   534
  shows "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> rs [\<rightarrow>\<^sub>s] ss" using rs
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   535
proof (induct arbitrary: ss)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   536
  case Nil
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   537
  thus ?case by cases (auto intro: listrelp.Nil)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   538
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   539
  case (Cons x y xs ys)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   540
  note Cons' = Cons
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   541
  show ?case
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   542
  proof (cases ss)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   543
    case Nil with Cons show ?thesis by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   544
  next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   545
    case (Cons y' ys')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   546
    hence ss: "ss = y' # ys'" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   547
    from Cons Cons' have "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys \<or> y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   548
    hence "x # xs [\<rightarrow>\<^sub>s] y' # ys'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   549
    proof
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   550
      assume H: "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   551
      with Cons' have "x \<rightarrow>\<^sub>s y'" by blast
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   552
      moreover from Cons' have "xs [\<rightarrow>\<^sub>s] ys" by (iprover dest: listrelp_conj1)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   553
      ultimately have "x # xs [\<rightarrow>\<^sub>s] y' # ys" by (rule listrelp.Cons)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   554
      with H show ?thesis by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   555
    next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   556
      assume H: "y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   557
      with Cons' have "x \<rightarrow>\<^sub>s y'" by blast
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   558
      moreover from H have "xs [\<rightarrow>\<^sub>s] ys'" by (blast intro: Cons')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   559
      ultimately show ?thesis by (rule listrelp.Cons)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   560
    qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   561
    with ss show ?thesis by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   562
  qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   563
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   564
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   565
lemma lemma4:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   566
  assumes r: "r \<rightarrow>\<^sub>s r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   567
  shows "r' \<rightarrow>\<^sub>\<beta> r'' \<Longrightarrow> r \<rightarrow>\<^sub>s r''" using r
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   568
proof (nominal_induct avoiding: r'' rule: sred.strong_induct)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   569
  case (Var rs rs' x)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   570
  then obtain ss where rs: "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss" and r'': "r'' = Var x \<degree>\<degree> ss"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   571
    by (blast dest: head_Var_reduction)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   572
  from Var(1) [simplified] rs have "rs [\<rightarrow>\<^sub>s] ss" by (rule lemma4_aux)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   573
  hence "Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> ss" by (rule sred.Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   574
  with r'' show ?case by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   575
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   576
  case (Abs x ss ss' r r')
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   577
  from \<open>(Lam [x].r') \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''\<close> show ?case
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   578
  proof (cases rule: apps_betasE [where x=x])
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   579
    case (appL s)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   580
    then obtain r''' where s: "s = (Lam [x].r''')" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" using \<open>x \<sharp> r''\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   581
      by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   582
    from r''' have "r \<rightarrow>\<^sub>s r'''" by (blast intro: Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   583
    moreover from Abs have "ss [\<rightarrow>\<^sub>s] ss'" by (iprover dest: listrelp_conj1)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   584
    ultimately have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r''') \<degree>\<degree> ss'" by (rule better_sred_Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   585
    with appL s show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   586
  next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   587
    case (appR rs')
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   588
    from Abs(6) [simplified] \<open>ss' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs'\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   589
    have "ss [\<rightarrow>\<^sub>s] rs'" by (rule lemma4_aux)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   590
    with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> rs'" by (rule better_sred_Abs)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   591
    with appR show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   592
  next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   593
    case (beta t u' us')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   594
    then have Lam_eq: "(Lam [x].r') = (Lam [x].t)" and ss': "ss' = u' # us'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   595
      and r'': "r'' = t[x::=u'] \<degree>\<degree> us'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   596
      by (simp_all add: abs_fresh)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   597
    from Abs(6) ss' obtain u us where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   598
      ss: "ss = u # us" and u: "u \<rightarrow>\<^sub>s u'" and us: "us [\<rightarrow>\<^sub>s] us'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   599
      by cases (auto dest!: listrelp_conj1)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   600
    have "r[x::=u] \<rightarrow>\<^sub>s r'[x::=u']" using \<open>r \<rightarrow>\<^sub>s r'\<close> and u by (rule lemma3)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   601
    with us have "r[x::=u] \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule lemma1')
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   602
    hence "(Lam [x].r) \<degree> u \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule better_sred_Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   603
    with ss r'' Lam_eq show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by (simp add: lam.inject alpha)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   604
  qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   605
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   606
  case (Beta x s ss t r)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   607
  show ?case
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   608
    by (rule better_sred_Beta) (rule Beta)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   609
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   610
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   611
lemma rtrancl_beta_sred:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   612
  assumes r: "r \<rightarrow>\<^sub>\<beta>\<^sup>* r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   613
  shows "r \<rightarrow>\<^sub>s r'" using r
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   614
  by induct (iprover intro: refl_sred lemma4)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   615
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   616
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   617
subsection \<open>Terms in normal form\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   618
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   619
lemma listsp_eqvt [eqvt]:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   620
  assumes xs: "listsp p (xs::'a::pt_name list)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   621
  shows "listsp ((pi::name prm) \<bullet> p) (pi \<bullet> xs)" using xs
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   622
by induction (use perm_app in force)+
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   624
inductive NF :: "lam \<Rightarrow> bool"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   625
where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   626
  App: "listsp NF ts \<Longrightarrow> NF (Var x \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   627
| Abs: "NF t \<Longrightarrow> NF (Lam [x].t)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   628
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   629
equivariance NF
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   630
nominal_inductive NF
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   631
  by (simp add: abs_fresh)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   632
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   633
lemma Abs_NF:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   634
  assumes NF: "NF ((Lam [x].t) \<degree>\<degree> ts)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   635
  shows "ts = []" using NF
80141
022a9c26b14f Tidied up another messy theory
paulson <lp15@cam.ac.uk>
parents: 71989
diff changeset
   636
  by (metis Abs_eq_apps_conv NF.cases Var_apps_neq_Abs_apps)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   637
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   638
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   639
\<^term>\<open>NF\<close> characterizes exactly the terms that are in normal form.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   640
\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   641
  
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   642
lemma NF_eq: "NF t = (\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   643
proof
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   644
  assume H: "NF t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   645
  show "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   646
  proof
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   647
    fix t'
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   648
    from H show "\<not> t \<rightarrow>\<^sub>\<beta> t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   649
    proof (nominal_induct avoiding: t' rule: NF.strong_induct)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   650
      case (App ts t)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   651
      show ?case
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   652
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   653
        assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   654
        then obtain rs where "ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   655
          by (iprover dest: head_Var_reduction)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   656
        with App show False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   657
          by (induct rs arbitrary: ts) (auto del: in_listspD)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   658
      qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   659
    next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   660
      case (Abs t x)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   661
      show ?case
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   662
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   663
        assume "(Lam [x].t) \<rightarrow>\<^sub>\<beta> t'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   664
        then show False using Abs
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   665
          by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   666
      qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   667
    qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   668
  qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   669
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   670
  assume H: "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   671
  then show "NF t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   672
  proof (nominal_induct t rule: Apps_lam_induct)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   673
    case (1 n ts)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   674
    then have "\<forall>ts'. \<not> ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ts'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   675
      by (iprover intro: apps_preserves_betas)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   676
    with 1(1) have "listsp NF ts"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   677
      by (induct ts) (auto simp add: in_listsp_conv_set)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   678
    then show ?case by (rule NF.App)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   679
  next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   680
    case (2 x u ts)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   681
    show ?case
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   682
    proof (cases ts)
39459
7753083c00e6 tidied a few proofs
paulson
parents: 32960
diff changeset
   683
      case Nil thus ?thesis by (metis 2 NF.Abs abs foldl_Nil)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   684
    next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   685
      case (Cons r rs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   686
      have "(Lam [x].u) \<degree> r \<rightarrow>\<^sub>\<beta> u[x::=r]" ..
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   687
      then have "(Lam [x].u) \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   688
        by (rule apps_preserves_beta)
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   689
      with Cons have "(Lam [x].u) \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29097
diff changeset
   690
        by simp
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   691
      with 2 show ?thesis by iprover
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   692
    qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   693
  qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   694
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   695
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   696
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   697
subsection \<open>Leftmost reduction and weakly normalizing terms\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   698
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   699
inductive
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   700
  lred :: "lam \<Rightarrow> lam \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>l" 50)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   701
  and lredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool"  (infixl "[\<rightarrow>\<^sub>l]" 50)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   702
where
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   703
  "s [\<rightarrow>\<^sub>l] t \<equiv> listrelp (\<rightarrow>\<^sub>l) s t"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   704
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   705
| Abs: "r \<rightarrow>\<^sub>l r' \<Longrightarrow> (Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   706
| Beta: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>l t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   707
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   708
lemma lred_imp_sred:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   709
  assumes lred: "s \<rightarrow>\<^sub>l t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   710
  shows "s \<rightarrow>\<^sub>s t" using lred
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   711
proof induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   712
  case (Var rs rs' x)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   713
  then have "rs [\<rightarrow>\<^sub>s] rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   714
    by induct (iprover intro: listrelp.intros)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   715
  then show ?case by (rule sred.Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   716
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   717
  case (Abs r r' x)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   718
  from \<open>r \<rightarrow>\<^sub>s r'\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   719
  have "(Lam [x].r) \<degree>\<degree> [] \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> []" using listrelp.Nil
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   720
    by (rule better_sred_Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   721
  then show ?case by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   722
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   723
  case (Beta r x s ss t)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   724
  from \<open>r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   725
  show ?case by (rule better_sred_Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   726
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   727
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   728
inductive WN :: "lam \<Rightarrow> bool"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   729
  where
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   730
    Var: "listsp WN rs \<Longrightarrow> WN (Var n \<degree>\<degree> rs)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   731
  | Lambda: "WN r \<Longrightarrow> WN (Lam [x].r)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   732
  | Beta: "WN ((r[x::=s]) \<degree>\<degree> ss) \<Longrightarrow> WN (((Lam [x].r) \<degree> s) \<degree>\<degree> ss)"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   733
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   734
lemma listrelp_imp_listsp1:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   735
  assumes H: "listrelp (\<lambda>x y. P x) xs ys"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   736
  shows "listsp P xs" using H
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   737
  by induct auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   738
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   739
lemma listrelp_imp_listsp2:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   740
  assumes H: "listrelp (\<lambda>x y. P y) xs ys"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   741
  shows "listsp P ys" using H
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   742
  by induct auto
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   743
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   744
lemma lemma5:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   745
  assumes lred: "r \<rightarrow>\<^sub>l r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   746
  shows "WN r" and "NF r'" using lred
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   747
  by induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   748
    (iprover dest: listrelp_conj1 listrelp_conj2
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   749
     listrelp_imp_listsp1 listrelp_imp_listsp2 intro: WN.intros
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   750
     NF.intros)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   751
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   752
lemma lemma6:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   753
  assumes wn: "WN r"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   754
  shows "\<exists>r'. r \<rightarrow>\<^sub>l r'" using wn
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   755
proof induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   756
  case (Var rs n)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   757
  then have "\<exists>rs'. rs [\<rightarrow>\<^sub>l] rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   758
    by induct (iprover intro: listrelp.intros)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   759
  then show ?case by (iprover intro: lred.Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   760
qed (iprover intro: lred.intros)+
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   761
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   762
lemma lemma7:
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   763
  assumes r: "r \<rightarrow>\<^sub>s r'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   764
  shows "NF r' \<Longrightarrow> r \<rightarrow>\<^sub>l r'" using r
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   765
proof induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   766
  case (Var rs rs' x)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   767
  from \<open>NF (Var x \<degree>\<degree> rs')\<close> have "listsp NF rs'"
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   768
    by cases simp_all
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   769
  with Var(1) have "rs [\<rightarrow>\<^sub>l] rs'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   770
  proof induct
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   771
    case Nil
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   772
    show ?case by (rule listrelp.Nil)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   773
  next
39459
7753083c00e6 tidied a few proofs
paulson
parents: 32960
diff changeset
   774
    case (Cons x y xs ys) 
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   775
    hence "x \<rightarrow>\<^sub>l y" and "xs [\<rightarrow>\<^sub>l] ys" by (auto del: in_listspD)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   776
    thus ?case by (rule listrelp.Cons)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   777
  qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   778
  thus ?case by (rule lred.Var)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   779
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   780
  case (Abs x ss ss' r r')
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   781
  from \<open>NF ((Lam [x].r') \<degree>\<degree> ss')\<close>
27623
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   782
  have ss': "ss' = []" by (rule Abs_NF)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   783
  from Abs(4) have ss: "ss = []" using ss'
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   784
    by cases simp_all
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   785
  from ss' Abs have "NF (Lam [x].r')" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   786
  hence "NF r'" by (cases rule: NF.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   787
  with Abs have "r \<rightarrow>\<^sub>l r'" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   788
  hence "(Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')" by (rule lred.Abs)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   789
  with ss ss' show ?case by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   790
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   791
  case (Beta x s ss t r)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   792
  hence "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t" by simp
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   793
  thus ?case by (rule lred.Beta)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   794
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   795
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   796
lemma WN_eq: "WN t = (\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   797
proof
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   798
  assume "WN t"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   799
  then have "\<exists>t'. t \<rightarrow>\<^sub>l t'" by (rule lemma6)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   800
  then obtain t' where t': "t \<rightarrow>\<^sub>l t'" ..
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   801
  then have NF: "NF t'" by (rule lemma5)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   802
  from t' have "t \<rightarrow>\<^sub>s t'" by (rule lred_imp_sred)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   803
  then have "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" by (rule lemma2)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   804
  with NF show "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by iprover
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   805
next
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   806
  assume "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   807
  then obtain t' where t': "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and NF: "NF t'"
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   808
    by iprover
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   809
  from t' have "t \<rightarrow>\<^sub>s t'" by (rule rtrancl_beta_sred)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   810
  then have "t \<rightarrow>\<^sub>l t'" using NF by (rule lemma7)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   811
  then show "WN t" by (rule lemma5)
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   812
qed
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   813
8e9c19529a4e Added Standardization theory.
berghofe
parents:
diff changeset
   814
end
47397
d654c73e4b12 no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
haftmann
parents: 46317
diff changeset
   815