author | haftmann |
Fri, 12 Oct 2007 08:25:48 +0200 | |
changeset 24996 | ebd5f4cc7118 |
parent 24538 | 452c4e02a684 |
child 25107 | dbf09ca6a80e |
permissions | -rw-r--r-- |
24538
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
1 |
(* Title: HOL/Lambda/Standardization.thy |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
2 |
ID: $Id$ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
3 |
Author: Stefan Berghofer |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
4 |
Copyright 2005 TU Muenchen |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
5 |
*) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
6 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
7 |
header {* Standardization *} |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
8 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
9 |
theory Standardization |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
10 |
imports NormalForm |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
11 |
begin |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
12 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
13 |
text {* |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
14 |
Based on lecture notes by Ralph Matthes \cite{Matthes-ESSLLI2000}, |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
15 |
original proof idea due to Ralph Loader \cite{Loader1998}. |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
16 |
*} |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
17 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
18 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
19 |
subsection {* Standard reduction relation *} |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
20 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
21 |
declare listrel_mono [mono_set] |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
22 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
23 |
inductive |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
24 |
sred :: "dB \<Rightarrow> dB \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>s" 50) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
25 |
and sredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>s]" 50) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
26 |
where |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
27 |
"s [\<rightarrow>\<^sub>s] t \<equiv> listrelp op \<rightarrow>\<^sub>s s t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
28 |
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
29 |
| Abs: "r \<rightarrow>\<^sub>s r' \<Longrightarrow> ss [\<rightarrow>\<^sub>s] ss' \<Longrightarrow> Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s Abs r' \<degree>\<degree> ss'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
30 |
| Beta: "r[s/0] \<degree>\<degree> ss \<rightarrow>\<^sub>s t \<Longrightarrow> Abs r \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
31 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
32 |
lemma refl_listrelp: "\<forall>x\<in>set xs. R x x \<Longrightarrow> listrelp R xs xs" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
33 |
by (induct xs) (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
34 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
35 |
lemma refl_sred: "t \<rightarrow>\<^sub>s t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
36 |
by (induct t rule: Apps_dB_induct) (auto intro: refl_listrelp sred.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
37 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
38 |
lemma refl_sreds: "ts [\<rightarrow>\<^sub>s] ts" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
39 |
by (simp add: refl_sred refl_listrelp) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
40 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
41 |
lemma listrelp_conj1: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp R x y" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
42 |
by (erule listrelp.induct) (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
43 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
44 |
lemma listrelp_conj2: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp S x y" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
45 |
by (erule listrelp.induct) (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
46 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
47 |
lemma listrelp_app: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
48 |
assumes xsys: "listrelp R xs ys" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
49 |
shows "listrelp R xs' ys' \<Longrightarrow> listrelp R (xs @ xs') (ys @ ys')" using xsys |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
50 |
by (induct arbitrary: xs' ys') (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
51 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
52 |
lemma lemma1: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
53 |
assumes r: "r \<rightarrow>\<^sub>s r'" and s: "s \<rightarrow>\<^sub>s s'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
54 |
shows "r \<degree> s \<rightarrow>\<^sub>s r' \<degree> s'" using r |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
55 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
56 |
case (Var rs rs' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
57 |
then have "rs [\<rightarrow>\<^sub>s] rs'" by (rule listrelp_conj1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
58 |
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
59 |
ultimately have "rs @ [s] [\<rightarrow>\<^sub>s] rs' @ [s']" by (rule listrelp_app) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
60 |
hence "Var x \<degree>\<degree> (rs @ [s]) \<rightarrow>\<^sub>s Var x \<degree>\<degree> (rs' @ [s'])" by (rule sred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
61 |
thus ?case by (simp only: app_last) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
62 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
63 |
case (Abs r r' ss ss') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
64 |
from Abs(3) have "ss [\<rightarrow>\<^sub>s] ss'" by (rule listrelp_conj1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
65 |
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
66 |
ultimately have "ss @ [s] [\<rightarrow>\<^sub>s] ss' @ [s']" by (rule listrelp_app) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
67 |
with `r \<rightarrow>\<^sub>s r'` have "Abs r \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s Abs r' \<degree>\<degree> (ss' @ [s'])" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
68 |
by (rule sred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
69 |
thus ?case by (simp only: app_last) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
70 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
71 |
case (Beta r u ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
72 |
hence "r[u/0] \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (simp only: app_last) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
73 |
hence "Abs r \<degree> u \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (rule sred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
74 |
thus ?case by (simp only: app_last) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
75 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
76 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
77 |
lemma lemma1': |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
78 |
assumes ts: "ts [\<rightarrow>\<^sub>s] ts'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
79 |
shows "r \<rightarrow>\<^sub>s r' \<Longrightarrow> r \<degree>\<degree> ts \<rightarrow>\<^sub>s r' \<degree>\<degree> ts'" using ts |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
80 |
by (induct arbitrary: r r') (auto intro: lemma1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
81 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
82 |
lemma lemma2_1: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
83 |
assumes beta: "t \<rightarrow>\<^sub>\<beta> u" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
84 |
shows "t \<rightarrow>\<^sub>s u" using beta |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
85 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
86 |
case (beta s t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
87 |
have "Abs s \<degree> t \<degree>\<degree> [] \<rightarrow>\<^sub>s s[t/0] \<degree>\<degree> []" by (iprover intro: sred.Beta refl_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
88 |
thus ?case by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
89 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
90 |
case (appL s t u) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
91 |
thus ?case by (iprover intro: lemma1 refl_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
92 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
93 |
case (appR s t u) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
94 |
thus ?case by (iprover intro: lemma1 refl_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
95 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
96 |
case (abs s t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
97 |
hence "Abs s \<degree>\<degree> [] \<rightarrow>\<^sub>s Abs t \<degree>\<degree> []" by (iprover intro: sred.Abs listrelp.Nil) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
98 |
thus ?case by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
99 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
100 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
101 |
lemma listrelp_betas: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
102 |
assumes ts: "listrelp op \<rightarrow>\<^sub>\<beta>\<^sup>* ts ts'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
103 |
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 |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
104 |
by induct auto |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
105 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
106 |
lemma lemma2_2: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
107 |
assumes t: "t \<rightarrow>\<^sub>s u" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
108 |
shows "t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using t |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
109 |
by induct (auto dest: listrelp_conj2 |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
110 |
intro: listrelp_betas apps_preserves_beta converse_rtranclp_into_rtranclp) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
111 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
112 |
lemma sred_lift: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
113 |
assumes s: "s \<rightarrow>\<^sub>s t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
114 |
shows "lift s i \<rightarrow>\<^sub>s lift t i" using s |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
115 |
proof (induct arbitrary: i) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
116 |
case (Var rs rs' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
117 |
hence "map (\<lambda>t. lift t i) rs [\<rightarrow>\<^sub>s] map (\<lambda>t. lift t i) rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
118 |
by induct (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
119 |
thus ?case by (cases "x < i") (auto intro: sred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
120 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
121 |
case (Abs r r' ss ss') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
122 |
from Abs(3) have "map (\<lambda>t. lift t i) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. lift t i) ss'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
123 |
by induct (auto intro: listrelp.intros) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
124 |
thus ?case by (auto intro: sred.Abs Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
125 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
126 |
case (Beta r s ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
127 |
thus ?case by (auto intro: sred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
128 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
129 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
130 |
lemma lemma3: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
131 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
132 |
shows "s \<rightarrow>\<^sub>s s' \<Longrightarrow> r[s/x] \<rightarrow>\<^sub>s r'[s'/x]" using r |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
133 |
proof (induct arbitrary: s s' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
134 |
case (Var rs rs' y) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
135 |
hence "map (\<lambda>t. t[s/x]) rs [\<rightarrow>\<^sub>s] map (\<lambda>t. t[s'/x]) rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
136 |
by induct (auto intro: listrelp.intros Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
137 |
moreover have "Var y[s/x] \<rightarrow>\<^sub>s Var y[s'/x]" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
138 |
proof (cases "y < x") |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
139 |
case True thus ?thesis by simp (rule refl_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
140 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
141 |
case False |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
142 |
thus ?thesis |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
143 |
by (cases "y = x") (auto simp add: Var intro: refl_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
144 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
145 |
ultimately show ?case by simp (rule lemma1') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
146 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
147 |
case (Abs r r' ss ss') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
148 |
have "lift s 0 \<rightarrow>\<^sub>s lift s' 0" by (rule sred_lift) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
149 |
hence "r[lift s 0/Suc x] \<rightarrow>\<^sub>s r'[lift s' 0/Suc x]" by (fast intro: Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
150 |
moreover from Abs(3) have "map (\<lambda>t. t[s/x]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[s'/x]) ss'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
151 |
by induct (auto intro: listrelp.intros Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
152 |
ultimately show ?case by simp (rule sred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
153 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
154 |
case (Beta r u ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
155 |
thus ?case by (auto simp add: subst_subst intro: sred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
156 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
157 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
158 |
lemma lemma4_aux: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
159 |
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'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
160 |
shows "rs' => ss \<Longrightarrow> rs [\<rightarrow>\<^sub>s] ss" using rs |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
161 |
proof (induct arbitrary: ss) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
162 |
case Nil |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
163 |
thus ?case by cases (auto intro: listrelp.Nil) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
164 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
165 |
case (Cons x y xs ys) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
166 |
note Cons' = Cons |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
167 |
show ?case |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
168 |
proof (cases ss) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
169 |
case Nil with Cons show ?thesis by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
170 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
171 |
case (Cons y' ys') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
172 |
hence ss: "ss = y' # ys'" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
173 |
from Cons Cons' have "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys \<or> y' = y \<and> ys => ys'" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
174 |
hence "x # xs [\<rightarrow>\<^sub>s] y' # ys'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
175 |
proof |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
176 |
assume H: "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
177 |
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
178 |
moreover from Cons' have "xs [\<rightarrow>\<^sub>s] ys" by (iprover dest: listrelp_conj1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
179 |
ultimately have "x # xs [\<rightarrow>\<^sub>s] y' # ys" by (rule listrelp.Cons) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
180 |
with H show ?thesis by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
181 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
182 |
assume H: "y' = y \<and> ys => ys'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
183 |
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
184 |
moreover from H have "xs [\<rightarrow>\<^sub>s] ys'" by (blast intro: Cons') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
185 |
ultimately show ?thesis by (rule listrelp.Cons) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
186 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
187 |
with ss show ?thesis by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
188 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
189 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
190 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
191 |
lemma lemma4: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
192 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
193 |
shows "r' \<rightarrow>\<^sub>\<beta> r'' \<Longrightarrow> r \<rightarrow>\<^sub>s r''" using r |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
194 |
proof (induct arbitrary: r'') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
195 |
case (Var rs rs' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
196 |
then obtain ss where rs: "rs' => ss" and r'': "r'' = Var x \<degree>\<degree> ss" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
197 |
by (blast dest: head_Var_reduction) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
198 |
from Var(1) rs have "rs [\<rightarrow>\<^sub>s] ss" by (rule lemma4_aux) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
199 |
hence "Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> ss" by (rule sred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
200 |
with r'' show ?case by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
201 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
202 |
case (Abs r r' ss ss') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
203 |
from `Abs r' \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''` show ?case |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
204 |
proof |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
205 |
fix s |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
206 |
assume r'': "r'' = s \<degree>\<degree> ss'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
207 |
assume "Abs r' \<rightarrow>\<^sub>\<beta> s" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
208 |
then obtain r''' where s: "s = Abs r'''" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" by cases auto |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
209 |
from r''' have "r \<rightarrow>\<^sub>s r'''" by (blast intro: Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
210 |
moreover from Abs have "ss [\<rightarrow>\<^sub>s] ss'" by (iprover dest: listrelp_conj1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
211 |
ultimately have "Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s Abs r''' \<degree>\<degree> ss'" by (rule sred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
212 |
with r'' s show "Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
213 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
214 |
fix rs' |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
215 |
assume "ss' => rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
216 |
with Abs(3) have "ss [\<rightarrow>\<^sub>s] rs'" by (rule lemma4_aux) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
217 |
with `r \<rightarrow>\<^sub>s r'` have "Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s Abs r' \<degree>\<degree> rs'" by (rule sred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
218 |
moreover assume "r'' = Abs r' \<degree>\<degree> rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
219 |
ultimately show "Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
220 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
221 |
fix t u' us' |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
222 |
assume "ss' = u' # us'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
223 |
with Abs(3) obtain u us where |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
224 |
ss: "ss = u # us" and u: "u \<rightarrow>\<^sub>s u'" and us: "us [\<rightarrow>\<^sub>s] us'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
225 |
by cases (auto dest!: listrelp_conj1) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
226 |
have "r[u/0] \<rightarrow>\<^sub>s r'[u'/0]" by (rule lemma3) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
227 |
with us have "r[u/0] \<degree>\<degree> us \<rightarrow>\<^sub>s r'[u'/0] \<degree>\<degree> us'" by (rule lemma1') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
228 |
hence "Abs r \<degree> u \<degree>\<degree> us \<rightarrow>\<^sub>s r'[u'/0] \<degree>\<degree> us'" by (rule sred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
229 |
moreover assume "Abs r' = Abs t" and "r'' = t[u'/0] \<degree>\<degree> us'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
230 |
ultimately show "Abs r \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" using ss by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
231 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
232 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
233 |
case (Beta r s ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
234 |
show ?case |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
235 |
by (rule sred.Beta) (rule Beta)+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
236 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
237 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
238 |
lemma rtrancl_beta_sred: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
239 |
assumes r: "r \<rightarrow>\<^sub>\<beta>\<^sup>* r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
240 |
shows "r \<rightarrow>\<^sub>s r'" using r |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
241 |
by induct (iprover intro: refl_sred lemma4)+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
242 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
243 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
244 |
subsection {* Leftmost reduction and weakly normalizing terms *} |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
245 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
246 |
inductive |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
247 |
lred :: "dB \<Rightarrow> dB \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
248 |
and lredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>l]" 50) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
249 |
where |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
250 |
"s [\<rightarrow>\<^sub>l] t \<equiv> listrelp op \<rightarrow>\<^sub>l s t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
251 |
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
252 |
| Abs: "r \<rightarrow>\<^sub>l r' \<Longrightarrow> Abs r \<rightarrow>\<^sub>l Abs r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
253 |
| Beta: "r[s/0] \<degree>\<degree> ss \<rightarrow>\<^sub>l t \<Longrightarrow> Abs r \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>l t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
254 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
255 |
lemma lred_imp_sred: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
256 |
assumes lred: "s \<rightarrow>\<^sub>l t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
257 |
shows "s \<rightarrow>\<^sub>s t" using lred |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
258 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
259 |
case (Var rs rs' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
260 |
then have "rs [\<rightarrow>\<^sub>s] rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
261 |
by induct (iprover intro: listrelp.intros)+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
262 |
then show ?case by (rule sred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
263 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
264 |
case (Abs r r') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
265 |
from `r \<rightarrow>\<^sub>s r'` |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
266 |
have "Abs r \<degree>\<degree> [] \<rightarrow>\<^sub>s Abs r' \<degree>\<degree> []" using listrelp.Nil |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
267 |
by (rule sred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
268 |
then show ?case by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
269 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
270 |
case (Beta r s ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
271 |
from `r[s/0] \<degree>\<degree> ss \<rightarrow>\<^sub>s t` |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
272 |
show ?case by (rule sred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
273 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
274 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
275 |
inductive WN :: "dB => bool" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
276 |
where |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
277 |
Var: "listsp WN rs \<Longrightarrow> WN (Var n \<degree>\<degree> rs)" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
278 |
| Lambda: "WN r \<Longrightarrow> WN (Abs r)" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
279 |
| Beta: "WN ((r[s/0]) \<degree>\<degree> ss) \<Longrightarrow> WN ((Abs r \<degree> s) \<degree>\<degree> ss)" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
280 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
281 |
lemma listrelp_imp_listsp1: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
282 |
assumes H: "listrelp (\<lambda>x y. P x) xs ys" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
283 |
shows "listsp P xs" using H |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
284 |
by induct auto |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
285 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
286 |
lemma listrelp_imp_listsp2: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
287 |
assumes H: "listrelp (\<lambda>x y. P y) xs ys" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
288 |
shows "listsp P ys" using H |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
289 |
by induct auto |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
290 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
291 |
lemma lemma5: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
292 |
assumes lred: "r \<rightarrow>\<^sub>l r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
293 |
shows "WN r" and "NF r'" using lred |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
294 |
by induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
295 |
(iprover dest: listrelp_conj1 listrelp_conj2 |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
296 |
listrelp_imp_listsp1 listrelp_imp_listsp2 intro: WN.intros |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
297 |
NF.intros [simplified listall_listsp_eq])+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
298 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
299 |
lemma lemma6: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
300 |
assumes wn: "WN r" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
301 |
shows "\<exists>r'. r \<rightarrow>\<^sub>l r'" using wn |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
302 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
303 |
case (Var rs n) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
304 |
then have "\<exists>rs'. rs [\<rightarrow>\<^sub>l] rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
305 |
by induct (iprover intro: listrelp.intros)+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
306 |
then show ?case by (iprover intro: lred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
307 |
qed (iprover intro: lred.intros)+ |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
308 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
309 |
lemma lemma7: |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
310 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
311 |
shows "NF r' \<Longrightarrow> r \<rightarrow>\<^sub>l r'" using r |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
312 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
313 |
case (Var rs rs' x) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
314 |
from `NF (Var x \<degree>\<degree> rs')` have "listall NF rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
315 |
by cases simp_all |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
316 |
with Var(1) have "rs [\<rightarrow>\<^sub>l] rs'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
317 |
proof induct |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
318 |
case Nil |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
319 |
show ?case by (rule listrelp.Nil) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
320 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
321 |
case (Cons x y xs ys) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
322 |
hence "x \<rightarrow>\<^sub>l y" and "xs [\<rightarrow>\<^sub>l] ys" by simp_all |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
323 |
thus ?case by (rule listrelp.Cons) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
324 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
325 |
thus ?case by (rule lred.Var) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
326 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
327 |
case (Abs r r' ss ss') |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
328 |
from `NF (Abs r' \<degree>\<degree> ss')` |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
329 |
have ss': "ss' = []" by (rule Abs_NF) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
330 |
from Abs(3) have ss: "ss = []" using ss' |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
331 |
by cases simp_all |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
332 |
from ss' Abs have "NF (Abs r')" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
333 |
hence "NF r'" by cases simp_all |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
334 |
with Abs have "r \<rightarrow>\<^sub>l r'" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
335 |
hence "Abs r \<rightarrow>\<^sub>l Abs r'" by (rule lred.Abs) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
336 |
with ss ss' show ?case by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
337 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
338 |
case (Beta r s ss t) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
339 |
hence "r[s/0] \<degree>\<degree> ss \<rightarrow>\<^sub>l t" by simp |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
340 |
thus ?case by (rule lred.Beta) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
341 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
342 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
343 |
lemma WN_eq: "WN t = (\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
344 |
proof |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
345 |
assume "WN t" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
346 |
then have "\<exists>t'. t \<rightarrow>\<^sub>l t'" by (rule lemma6) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
347 |
then obtain t' where t': "t \<rightarrow>\<^sub>l t'" .. |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
348 |
then have NF: "NF t'" by (rule lemma5) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
349 |
from t' have "t \<rightarrow>\<^sub>s t'" by (rule lred_imp_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
350 |
then have "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" by (rule lemma2_2) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
351 |
with NF show "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by iprover |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
352 |
next |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
353 |
assume "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
354 |
then obtain t' where t': "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and NF: "NF t'" |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
355 |
by iprover |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
356 |
from t' have "t \<rightarrow>\<^sub>s t'" by (rule rtrancl_beta_sred) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
357 |
then have "t \<rightarrow>\<^sub>l t'" using NF by (rule lemma7) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
358 |
then show "WN t" by (rule lemma5) |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
359 |
qed |
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
360 |
|
452c4e02a684
New proof of standardization theorem (inspired by Ralph Matthes).
berghofe
parents:
diff
changeset
|
361 |
end |