| author | nipkow | 
| Wed, 22 Jul 2009 10:49:26 +0200 | |
| changeset 32130 | 2a0645733185 | 
| parent 32010 | cb1a1c94b4cd | 
| child 32356 | e11cd88e6ade | 
| permissions | -rw-r--r-- | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Lambda/WeakNorm.thy  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
2  | 
Author: Stefan Berghofer  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
3  | 
Copyright 2003 TU Muenchen  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
6  | 
header {* Weak normalization for simply-typed lambda calculus *}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
7  | 
|
| 22512 | 8  | 
theory WeakNorm  | 
| 
24994
 
c385c4eabb3b
consolidated naming conventions for code generator theories
 
haftmann 
parents: 
24715 
diff
changeset
 | 
9  | 
imports Type NormalForm Code_Integer  | 
| 22512 | 10  | 
begin  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
12  | 
text {*
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
13  | 
Formalization by Stefan Berghofer. Partly based on a paper proof by  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
14  | 
Felix Joachimski and Ralph Matthes \cite{Matthes-Joachimski-AML}.
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
15  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
18  | 
subsection {* Main theorems *}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
19  | 
|
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
20  | 
lemma norm_list:  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
21  | 
assumes f_compat: "\<And>t t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> f t \<rightarrow>\<^sub>\<beta>\<^sup>* f t'"  | 
| 22271 | 22  | 
and f_NF: "\<And>t. NF t \<Longrightarrow> NF (f t)"  | 
23  | 
and uNF: "NF u" and uT: "e \<turnstile> u : T"  | 
|
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
24  | 
shows "\<And>Us. e\<langle>i:T\<rangle> \<tturnstile> as : Us \<Longrightarrow>  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
25  | 
listall (\<lambda>t. \<forall>e T' u i. e\<langle>i:T\<rangle> \<turnstile> t : T' \<longrightarrow>  | 
| 22271 | 26  | 
NF u \<longrightarrow> e \<turnstile> u : T \<longrightarrow> (\<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')) as \<Longrightarrow>  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
27  | 
\<exists>as'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. f (t[u/i])) as \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 22271 | 28  | 
Var j \<degree>\<degree> map f as' \<and> NF (Var j \<degree>\<degree> map f as')"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
29  | 
(is "\<And>Us. _ \<Longrightarrow> listall ?R as \<Longrightarrow> \<exists>as'. ?ex Us as as'")  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
30  | 
proof (induct as rule: rev_induct)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
31  | 
case (Nil Us)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
32  | 
with Var_NF have "?ex Us [] []" by simp  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
33  | 
thus ?case ..  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
34  | 
next  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
35  | 
case (snoc b bs Us)  | 
| 23464 | 36  | 
have "e\<langle>i:T\<rangle> \<tturnstile> bs @ [b] : Us" by fact  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
37  | 
then obtain Vs W where Us: "Us = Vs @ [W]"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
38  | 
and bs: "e\<langle>i:T\<rangle> \<tturnstile> bs : Vs" and bT: "e\<langle>i:T\<rangle> \<turnstile> b : W"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
39  | 
by (rule types_snocE)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
40  | 
from snoc have "listall ?R bs" by simp  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
41  | 
with bs have "\<exists>bs'. ?ex Vs bs bs'" by (rule snoc)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
42  | 
then obtain bs' where  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
43  | 
bsred: "\<And>j. Var j \<degree>\<degree> map (\<lambda>t. f (t[u/i])) bs \<rightarrow>\<^sub>\<beta>\<^sup>* Var j \<degree>\<degree> map f bs'"  | 
| 22271 | 44  | 
and bsNF: "\<And>j. NF (Var j \<degree>\<degree> map f bs')" by iprover  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
45  | 
from snoc have "?R b" by simp  | 
| 22271 | 46  | 
with bT and uNF and uT have "\<exists>b'. b[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* b' \<and> NF b'"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
47  | 
by iprover  | 
| 22271 | 48  | 
then obtain b' where bred: "b[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* b'" and bNF: "NF b'"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
49  | 
by iprover  | 
| 22271 | 50  | 
from bsNF [of 0] have "listall NF (map f bs')"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
51  | 
by (rule App_NF_D)  | 
| 23464 | 52  | 
moreover have "NF (f b')" using bNF by (rule f_NF)  | 
| 22271 | 53  | 
ultimately have "listall NF (map f (bs' @ [b']))"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
54  | 
by simp  | 
| 22271 | 55  | 
hence "\<And>j. NF (Var j \<degree>\<degree> map f (bs' @ [b']))" by (rule NF.App)  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
56  | 
moreover from bred have "f (b[u/i]) \<rightarrow>\<^sub>\<beta>\<^sup>* f b'"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
57  | 
by (rule f_compat)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
58  | 
with bsred have  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
59  | 
"\<And>j. (Var j \<degree>\<degree> map (\<lambda>t. f (t[u/i])) bs) \<degree> f (b[u/i]) \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
60  | 
(Var j \<degree>\<degree> map f bs') \<degree> f b'" by (rule rtrancl_beta_App)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
61  | 
ultimately have "?ex Us (bs @ [b]) (bs' @ [b'])" by simp  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
62  | 
thus ?case ..  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
63  | 
qed  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
64  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
65  | 
lemma subst_type_NF:  | 
| 22271 | 66  | 
"\<And>t e T u i. NF t \<Longrightarrow> e\<langle>i:U\<rangle> \<turnstile> t : T \<Longrightarrow> NF u \<Longrightarrow> e \<turnstile> u : U \<Longrightarrow> \<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
67  | 
(is "PROP ?P U" is "\<And>t e T u i. _ \<Longrightarrow> PROP ?Q t e T u i U")  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
68  | 
proof (induct U)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
69  | 
fix T t  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
70  | 
let ?R = "\<lambda>t. \<forall>e T' u i.  | 
| 22271 | 71  | 
e\<langle>i:T\<rangle> \<turnstile> t : T' \<longrightarrow> NF u \<longrightarrow> e \<turnstile> u : T \<longrightarrow> (\<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
72  | 
assume MI1: "\<And>T1 T2. T = T1 \<Rightarrow> T2 \<Longrightarrow> PROP ?P T1"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
73  | 
assume MI2: "\<And>T1 T2. T = T1 \<Rightarrow> T2 \<Longrightarrow> PROP ?P T2"  | 
| 22271 | 74  | 
assume "NF t"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
75  | 
thus "\<And>e T' u i. PROP ?Q t e T' u i T"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
76  | 
proof induct  | 
| 22271 | 77  | 
fix e T' u i assume uNF: "NF u" and uT: "e \<turnstile> u : T"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
78  | 
    {
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
79  | 
case (App ts x e_ T'_ u_ i_)  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
80  | 
assume "e\<langle>i:T\<rangle> \<turnstile> Var x \<degree>\<degree> ts : T'"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
81  | 
then obtain Us  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
82  | 
where varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : Us \<Rrightarrow> T'"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
83  | 
and argsT: "e\<langle>i:T\<rangle> \<tturnstile> ts : Us"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
84  | 
by (rule var_app_typesE)  | 
| 22271 | 85  | 
from nat_eq_dec show "\<exists>t'. (Var x \<degree>\<degree> ts)[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
86  | 
proof  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
87  | 
assume eq: "x = i"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
88  | 
show ?thesis  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
89  | 
proof (cases ts)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
90  | 
case Nil  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
91  | 
with eq have "(Var x \<degree>\<degree> [])[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* u" by simp  | 
| 17589 | 92  | 
with Nil and uNF show ?thesis by simp iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
93  | 
next  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
94  | 
case (Cons a as)  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
95  | 
with argsT obtain T'' Ts where Us: "Us = T'' # Ts"  | 
| 23464 | 96  | 
by (cases Us) (rule FalseE, simp+, erule that)  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
97  | 
from varT and Us have varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : T'' \<Rightarrow> Ts \<Rrightarrow> T'"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
98  | 
by simp  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
99  | 
from varT eq have T: "T = T'' \<Rightarrow> Ts \<Rrightarrow> T'" by cases auto  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
100  | 
with uT have uT': "e \<turnstile> u : T'' \<Rightarrow> Ts \<Rrightarrow> T'" by simp  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
101  | 
from argsT Us Cons have argsT': "e\<langle>i:T\<rangle> \<tturnstile> as : Ts" by simp  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
102  | 
from argsT Us Cons have argT: "e\<langle>i:T\<rangle> \<turnstile> a : T''" by simp  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
103  | 
from argT uT refl have aT: "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)  | 
| 17589 | 104  | 
from App and Cons have "listall ?R as" by simp (iprover dest: listall_conj2)  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
105  | 
with lift_preserves_beta' lift_NF uNF uT argsT'  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
106  | 
have "\<exists>as'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
107  | 
Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as' \<and>  | 
| 22271 | 108  | 
NF (Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by (rule norm_list)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
109  | 
then obtain as' where  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
110  | 
asred: "Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
111  | 
Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as'"  | 
| 22271 | 112  | 
and asNF: "NF (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')" by iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
113  | 
from App and Cons have "?R a" by simp  | 
| 22271 | 114  | 
with argT and uNF and uT have "\<exists>a'. a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a' \<and> NF a'"  | 
| 17589 | 115  | 
by iprover  | 
| 22271 | 116  | 
then obtain a' where ared: "a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a'" and aNF: "NF a'" by iprover  | 
117  | 
from uNF have "NF (lift u 0)" by (rule lift_NF)  | 
|
118  | 
hence "\<exists>u'. lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u' \<and> NF u'" by (rule app_Var_NF)  | 
|
119  | 
then obtain u' where ured: "lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u'" and u'NF: "NF u'"  | 
|
| 17589 | 120  | 
by iprover  | 
| 22271 | 121  | 
from T and u'NF have "\<exists>ua. u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua \<and> NF ua"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
122  | 
proof (rule MI1)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
123  | 
have "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 \<degree> Var 0 : Ts \<Rrightarrow> T'"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
124  | 
proof (rule typing.App)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
125  | 
from uT' show "e\<langle>0:T''\<rangle> \<turnstile> lift u 0 : T'' \<Rightarrow> Ts \<Rrightarrow> T'" by (rule lift_type)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
126  | 
show "e\<langle>0:T''\<rangle> \<turnstile> Var 0 : T''" by (rule typing.Var) simp  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
127  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
128  | 
with ured show "e\<langle>0:T''\<rangle> \<turnstile> u' : Ts \<Rrightarrow> T'" by (rule subject_reduction')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
129  | 
from ared aT show "e \<turnstile> a' : T''" by (rule subject_reduction')  | 
| 23464 | 130  | 
show "NF a'" by fact  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
131  | 
qed  | 
| 22271 | 132  | 
then obtain ua where uared: "u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" and uaNF: "NF ua"  | 
| 17589 | 133  | 
by iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
134  | 
from ared have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* (lift u 0 \<degree> Var 0)[a'/0]"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
135  | 
by (rule subst_preserves_beta2')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
136  | 
also from ured have "(lift u 0 \<degree> Var 0)[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* u'[a'/0]"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
137  | 
by (rule subst_preserves_beta')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
138  | 
also note uared  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
139  | 
finally have "(lift u 0 \<degree> Var 0)[a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" .  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
140  | 
hence uared': "u \<degree> a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" by simp  | 
| 23464 | 141  | 
from T asNF _ uaNF have "\<exists>r. (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r \<and> NF r"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
142  | 
proof (rule MI2)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
143  | 
have "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as : T'"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
144  | 
proof (rule list_app_typeI)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
145  | 
show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 : Ts \<Rrightarrow> T'" by (rule typing.Var) simp  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
146  | 
from uT argsT' have "e \<tturnstile> map (\<lambda>t. t[u/i]) as : Ts"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
147  | 
by (rule substs_lemma)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
148  | 
hence "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift t 0) (map (\<lambda>t. t[u/i]) as) : Ts"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
149  | 
by (rule lift_types)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
150  | 
thus "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<tturnstile> map (\<lambda>t. lift (t[u/i]) 0) as : Ts"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
151  | 
by (simp_all add: map_compose [symmetric] o_def)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
152  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
153  | 
with asred show "e\<langle>0:Ts \<Rrightarrow> T'\<rangle> \<turnstile> Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as' : T'"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
154  | 
by (rule subject_reduction')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
155  | 
from argT uT refl have "e \<turnstile> a[u/i] : T''" by (rule subst_lemma)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
156  | 
with uT' have "e \<turnstile> u \<degree> a[u/i] : Ts \<Rrightarrow> T'" by (rule typing.App)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
157  | 
with uared' show "e \<turnstile> ua : Ts \<Rrightarrow> T'" by (rule subject_reduction')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
158  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
159  | 
then obtain r where rred: "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r"  | 
| 22271 | 160  | 
and rnf: "NF r" by iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
161  | 
from asred have  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
162  | 
"(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
163  | 
(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0]"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
164  | 
by (rule subst_preserves_beta')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
165  | 
also from uared' have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>*  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
166  | 
(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0]" by (rule subst_preserves_beta2')  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
167  | 
also note rred  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
168  | 
finally have "(Var 0 \<degree>\<degree> map (\<lambda>t. lift (t[u/i]) 0) as)[u \<degree> a[u/i]/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r" .  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
169  | 
with rnf Cons eq show ?thesis  | 
| 17589 | 170  | 
by (simp add: map_compose [symmetric] o_def) iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
171  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
172  | 
next  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
173  | 
assume neq: "x \<noteq> i"  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
174  | 
from App have "listall ?R ts" by (iprover dest: listall_conj2)  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
175  | 
with TrueI TrueI uNF uT argsT  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
176  | 
have "\<exists>ts'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. t[u/i]) ts \<rightarrow>\<^sub>\<beta>\<^sup>* Var j \<degree>\<degree> ts' \<and>  | 
| 22271 | 177  | 
NF (Var j \<degree>\<degree> ts')" (is "\<exists>ts'. ?ex ts'")  | 
| 
18331
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
178  | 
by (rule norm_list [of "\<lambda>t. t", simplified])  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
179  | 
then obtain ts' where NF: "?ex ts'" ..  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
180  | 
from nat_le_dec show ?thesis  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
181  | 
proof  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
182  | 
assume "i < x"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
183  | 
with NF show ?thesis by simp iprover  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
184  | 
next  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
185  | 
assume "\<not> (i < x)"  | 
| 
 
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
 
berghofe 
parents: 
18257 
diff
changeset
 | 
186  | 
with NF neq show ?thesis by (simp add: subst_Var) iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
187  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
188  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
189  | 
next  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
190  | 
case (Abs r e_ T'_ u_ i_)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
191  | 
assume absT: "e\<langle>i:T\<rangle> \<turnstile> Abs r : T'"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
192  | 
then obtain R S where "e\<langle>0:R\<rangle>\<langle>Suc i:T\<rangle> \<turnstile> r : S" by (rule abs_typeE) simp  | 
| 23464 | 193  | 
moreover have "NF (lift u 0)" using `NF u` by (rule lift_NF)  | 
194  | 
moreover have "e\<langle>0:R\<rangle> \<turnstile> lift u 0 : T" using uT by (rule lift_type)  | 
|
| 22271 | 195  | 
ultimately have "\<exists>t'. r[lift u 0/Suc i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by (rule Abs)  | 
196  | 
thus "\<exists>t'. Abs r[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"  | 
|
| 17589 | 197  | 
by simp (iprover intro: rtrancl_beta_Abs NF.Abs)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
198  | 
}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
199  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
200  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
202  | 
|
| 22271 | 203  | 
-- {* A computationally relevant copy of @{term "e \<turnstile> t : T"} *}
 | 
| 23750 | 204  | 
inductive rtyping :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile>\<^sub>R _ : _" [50, 50, 50] 50)
 | 
| 22271 | 205  | 
where  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
206  | 
Var: "e x = T \<Longrightarrow> e \<turnstile>\<^sub>R Var x : T"  | 
| 22271 | 207  | 
| Abs: "e\<langle>0:T\<rangle> \<turnstile>\<^sub>R t : U \<Longrightarrow> e \<turnstile>\<^sub>R Abs t : (T \<Rightarrow> U)"  | 
208  | 
| App: "e \<turnstile>\<^sub>R s : T \<Rightarrow> U \<Longrightarrow> e \<turnstile>\<^sub>R t : T \<Longrightarrow> e \<turnstile>\<^sub>R (s \<degree> t) : U"  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
209  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
210  | 
lemma rtyping_imp_typing: "e \<turnstile>\<^sub>R t : T \<Longrightarrow> e \<turnstile> t : T"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
211  | 
apply (induct set: rtyping)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
212  | 
apply (erule typing.Var)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
213  | 
apply (erule typing.Abs)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
214  | 
apply (erule typing.App)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
215  | 
apply assumption  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
216  | 
done  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
217  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
218  | 
|
| 18513 | 219  | 
theorem type_NF:  | 
220  | 
assumes "e \<turnstile>\<^sub>R t : T"  | 
|
| 23464 | 221  | 
shows "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" using assms  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
222  | 
proof induct  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
223  | 
case Var  | 
| 17589 | 224  | 
show ?case by (iprover intro: Var_NF)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
225  | 
next  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
226  | 
case Abs  | 
| 17589 | 227  | 
thus ?case by (iprover intro: rtrancl_beta_Abs NF.Abs)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
228  | 
next  | 
| 22271 | 229  | 
case (App e s T U t)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
230  | 
from App obtain s' t' where  | 
| 23464 | 231  | 
sred: "s \<rightarrow>\<^sub>\<beta>\<^sup>* s'" and "NF s'"  | 
| 22271 | 232  | 
and tred: "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and tNF: "NF t'" by iprover  | 
233  | 
have "\<exists>u. (Var 0 \<degree> lift t' 0)[s'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* u \<and> NF u"  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
234  | 
proof (rule subst_type_NF)  | 
| 23464 | 235  | 
have "NF (lift t' 0)" using tNF by (rule lift_NF)  | 
| 22271 | 236  | 
hence "listall NF [lift t' 0]" by (rule listall_cons) (rule listall_nil)  | 
237  | 
hence "NF (Var 0 \<degree>\<degree> [lift t' 0])" by (rule NF.App)  | 
|
238  | 
thus "NF (Var 0 \<degree> lift t' 0)" by simp  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
239  | 
show "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> Var 0 \<degree> lift t' 0 : U"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
240  | 
proof (rule typing.App)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
241  | 
show "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> Var 0 : T \<Rightarrow> U"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
242  | 
by (rule typing.Var) simp  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
243  | 
from tred have "e \<turnstile> t' : T"  | 
| 23464 | 244  | 
by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
245  | 
thus "e\<langle>0:T \<Rightarrow> U\<rangle> \<turnstile> lift t' 0 : T"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
246  | 
by (rule lift_type)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
247  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
248  | 
from sred show "e \<turnstile> s' : T \<Rightarrow> U"  | 
| 23464 | 249  | 
by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps)  | 
250  | 
show "NF s'" by fact  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
251  | 
qed  | 
| 22271 | 252  | 
then obtain u where ured: "s' \<degree> t' \<rightarrow>\<^sub>\<beta>\<^sup>* u" and unf: "NF u" by simp iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
253  | 
from sred tred have "s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'" by (rule rtrancl_beta_App)  | 
| 23750 | 254  | 
hence "s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using ured by (rule rtranclp_trans)  | 
| 17589 | 255  | 
with unf show ?case by iprover  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
256  | 
qed  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
258  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
259  | 
subsection {* Extracting the program *}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
261  | 
declare NF.induct [ind_realizer]  | 
| 23750 | 262  | 
declare rtranclp.induct [ind_realizer irrelevant]  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
263  | 
declare rtyping.induct [ind_realizer]  | 
| 22271 | 264  | 
lemmas [extraction_expand] = conj_assoc listall_cons_eq  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
266  | 
extract type_NF  | 
| 23750 | 267  | 
lemma rtranclR_rtrancl_eq: "rtranclpR r a b = r\<^sup>*\<^sup>* a b"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
268  | 
apply (rule iffI)  | 
| 23750 | 269  | 
apply (erule rtranclpR.induct)  | 
270  | 
apply (rule rtranclp.rtrancl_refl)  | 
|
271  | 
apply (erule rtranclp.rtrancl_into_rtrancl)  | 
|
| 22271 | 272  | 
apply assumption  | 
| 23750 | 273  | 
apply (erule rtranclp.induct)  | 
274  | 
apply (rule rtranclpR.rtrancl_refl)  | 
|
275  | 
apply (erule rtranclpR.rtrancl_into_rtrancl)  | 
|
| 22271 | 276  | 
apply assumption  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
277  | 
done  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
278  | 
|
| 22271 | 279  | 
lemma NFR_imp_NF: "NFR nf t \<Longrightarrow> NF t"  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
280  | 
apply (erule NFR.induct)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
281  | 
apply (rule NF.intros)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
282  | 
apply (simp add: listall_def)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
283  | 
apply (erule NF.intros)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
284  | 
done  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
285  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
286  | 
text_raw {*
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
287  | 
\begin{figure}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
288  | 
\renewcommand{\isastyle}{\scriptsize\it}%
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
289  | 
@{thm [display,eta_contract=false,margin=100] subst_type_NF_def}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
290  | 
\renewcommand{\isastyle}{\small\it}%
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
291  | 
\caption{Program extracted from @{text subst_type_NF}}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
292  | 
\label{fig:extr-subst-type-nf}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
293  | 
\end{figure}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
294  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
295  | 
\begin{figure}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
296  | 
\renewcommand{\isastyle}{\scriptsize\it}%
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
297  | 
@{thm [display,margin=100] subst_Var_NF_def}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
298  | 
@{thm [display,margin=100] app_Var_NF_def}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
299  | 
@{thm [display,margin=100] lift_NF_def}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
300  | 
@{thm [display,eta_contract=false,margin=100] type_NF_def}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
301  | 
\renewcommand{\isastyle}{\small\it}%
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
302  | 
\caption{Program extracted from lemmas and main theorem}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
303  | 
\label{fig:extr-type-nf}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
304  | 
\end{figure}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
305  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
307  | 
text {*
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
308  | 
The program corresponding to the proof of the central lemma, which  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
309  | 
performs substitution and normalization, is shown in Figure  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
310  | 
\ref{fig:extr-subst-type-nf}. The correctness
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
311  | 
theorem corresponding to the program @{text "subst_type_NF"} is
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
312  | 
@{thm [display,margin=100] subst_type_NF_correctness
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
313  | 
[simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
314  | 
where @{text NFR} is the realizability predicate corresponding to
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
315  | 
the datatype @{text NFT}, which is inductively defined by the rules
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
316  | 
\pagebreak  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
317  | 
@{thm [display,margin=90] NFR.App [of ts nfs x] NFR.Abs [of nf t]}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
318  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
319  | 
The programs corresponding to the main theorem @{text "type_NF"}, as
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
320  | 
well as to some lemmas, are shown in Figure \ref{fig:extr-type-nf}.
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
321  | 
The correctness statement for the main function @{text "type_NF"} is
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
322  | 
@{thm [display,margin=100] type_NF_correctness
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
323  | 
[simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
324  | 
where the realizability predicate @{text "rtypingR"} corresponding to the
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
325  | 
computationally relevant version of the typing judgement is inductively  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
326  | 
defined by the rules  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
327  | 
@{thm [display,margin=100] rtypingR.Var [no_vars]
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
328  | 
rtypingR.Abs [of ty, no_vars] rtypingR.App [of ty e s T U ty' t]}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
329  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
330  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
331  | 
subsection {* Generating executable code *}
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
332  | 
|
| 27982 | 333  | 
instantiation NFT :: default  | 
334  | 
begin  | 
|
335  | 
||
336  | 
definition "default = Dummy ()"  | 
|
337  | 
||
338  | 
instance ..  | 
|
339  | 
||
340  | 
end  | 
|
341  | 
||
342  | 
instantiation dB :: default  | 
|
343  | 
begin  | 
|
344  | 
||
345  | 
definition "default = dB.Var 0"  | 
|
346  | 
||
347  | 
instance ..  | 
|
348  | 
||
349  | 
end  | 
|
350  | 
||
351  | 
instantiation * :: (default, default) default  | 
|
352  | 
begin  | 
|
353  | 
||
354  | 
definition "default = (default, default)"  | 
|
355  | 
||
356  | 
instance ..  | 
|
357  | 
||
358  | 
end  | 
|
359  | 
||
360  | 
instantiation list :: (type) default  | 
|
361  | 
begin  | 
|
362  | 
||
363  | 
definition "default = []"  | 
|
364  | 
||
365  | 
instance ..  | 
|
366  | 
||
367  | 
end  | 
|
368  | 
||
369  | 
instantiation "fun" :: (type, default) default  | 
|
370  | 
begin  | 
|
371  | 
||
372  | 
definition "default = (\<lambda>x. default)"  | 
|
373  | 
||
374  | 
instance ..  | 
|
375  | 
||
376  | 
end  | 
|
377  | 
||
378  | 
definition int_of_nat :: "nat \<Rightarrow> int" where  | 
|
379  | 
"int_of_nat = of_nat"  | 
|
380  | 
||
381  | 
text {*
 | 
|
382  | 
  The following functions convert between Isabelle's built-in {\tt term}
 | 
|
383  | 
  datatype and the generated {\tt dB} datatype. This allows to
 | 
|
384  | 
generate example terms using Isabelle's parser and inspect  | 
|
385  | 
normalized terms using Isabelle's pretty printer.  | 
|
386  | 
*}  | 
|
387  | 
||
388  | 
ML {*
 | 
|
389  | 
fun dBtype_of_typ (Type ("fun", [T, U])) =
 | 
|
390  | 
      @{code Fun} (dBtype_of_typ T, dBtype_of_typ U)
 | 
|
391  | 
| dBtype_of_typ (TFree (s, _)) = (case explode s of  | 
|
392  | 
        ["'", a] => @{code Atom} (@{code nat} (ord a - 97))
 | 
|
393  | 
| _ => error "dBtype_of_typ: variable name")  | 
|
394  | 
| dBtype_of_typ _ = error "dBtype_of_typ: bad type";  | 
|
395  | 
||
396  | 
fun dB_of_term (Bound i) = @{code dB.Var} (@{code nat} i)
 | 
|
397  | 
  | dB_of_term (t $ u) = @{code dB.App} (dB_of_term t, dB_of_term u)
 | 
|
398  | 
  | dB_of_term (Abs (_, _, t)) = @{code dB.Abs} (dB_of_term t)
 | 
|
399  | 
| dB_of_term _ = error "dB_of_term: bad term";  | 
|
400  | 
||
401  | 
fun term_of_dB Ts (Type ("fun", [T, U])) (@{code dB.Abs} dBt) =
 | 
|
402  | 
      Abs ("x", T, term_of_dB (T :: Ts) U dBt)
 | 
|
403  | 
| term_of_dB Ts _ dBt = term_of_dB' Ts dBt  | 
|
404  | 
and term_of_dB' Ts (@{code dB.Var} n) = Bound (@{code int_of_nat} n)
 | 
|
405  | 
  | term_of_dB' Ts (@{code dB.App} (dBt, dBu)) =
 | 
|
406  | 
let val t = term_of_dB' Ts dBt  | 
|
407  | 
in case fastype_of1 (Ts, t) of  | 
|
408  | 
          Type ("fun", [T, U]) => t $ term_of_dB Ts T dBu
 | 
|
409  | 
| _ => error "term_of_dB: function type expected"  | 
|
410  | 
end  | 
|
411  | 
| term_of_dB' _ _ = error "term_of_dB: term not in normal form";  | 
|
412  | 
||
413  | 
fun typing_of_term Ts e (Bound i) =  | 
|
414  | 
      @{code Var} (e, @{code nat} i, dBtype_of_typ (nth Ts i))
 | 
|
415  | 
| typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of  | 
|
416  | 
        Type ("fun", [T, U]) => @{code App} (e, dB_of_term t,
 | 
|
417  | 
dBtype_of_typ T, dBtype_of_typ U, dB_of_term u,  | 
|
418  | 
typing_of_term Ts e t, typing_of_term Ts e u)  | 
|
419  | 
| _ => error "typing_of_term: function type expected")  | 
|
420  | 
| typing_of_term Ts e (Abs (s, T, t)) =  | 
|
421  | 
let val dBT = dBtype_of_typ T  | 
|
422  | 
      in @{code Abs} (e, dBT, dB_of_term t,
 | 
|
423  | 
dBtype_of_typ (fastype_of1 (T :: Ts, t)),  | 
|
424  | 
        typing_of_term (T :: Ts) (@{code shift} e @{code "0::nat"} dBT) t)
 | 
|
425  | 
end  | 
|
426  | 
| typing_of_term _ _ _ = error "typing_of_term: bad term";  | 
|
427  | 
||
428  | 
fun dummyf _ = error "dummy";  | 
|
429  | 
||
430  | 
val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"};
 | 
|
431  | 
val (dB1, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct1));
 | 
|
| 32010 | 432  | 
val ct1' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct1)) dB1);
 | 
| 27982 | 433  | 
|
434  | 
val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"};
 | 
|
435  | 
val (dB2, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct2));
 | 
|
| 32010 | 436  | 
val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2);
 | 
| 27982 | 437  | 
*}  | 
438  | 
||
439  | 
text {*
 | 
|
440  | 
The same story again for the SML code generator.  | 
|
441  | 
*}  | 
|
442  | 
||
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
443  | 
consts_code  | 
| 27982 | 444  | 
  "default" ("(error \"default\")")
 | 
445  | 
  "default :: 'a \<Rightarrow> 'b::default" ("(fn '_ => error \"default\")")
 | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
446  | 
|
| 17145 | 447  | 
code_module Norm  | 
448  | 
contains  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
449  | 
test = "type_NF"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
450  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
451  | 
ML {*
 | 
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20503 
diff
changeset
 | 
452  | 
fun nat_of_int 0 = Norm.zero  | 
| 17145 | 453  | 
| nat_of_int n = Norm.Suc (nat_of_int (n-1));  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
454  | 
|
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20503 
diff
changeset
 | 
455  | 
fun int_of_nat Norm.zero = 0  | 
| 17145 | 456  | 
| int_of_nat (Norm.Suc n) = 1 + int_of_nat n;  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
458  | 
fun dBtype_of_typ (Type ("fun", [T, U])) =
 | 
| 17145 | 459  | 
Norm.Fun (dBtype_of_typ T, dBtype_of_typ U)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
460  | 
| dBtype_of_typ (TFree (s, _)) = (case explode s of  | 
| 17145 | 461  | 
["'", a] => Norm.Atom (nat_of_int (ord a - 97))  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
462  | 
| _ => error "dBtype_of_typ: variable name")  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
463  | 
| dBtype_of_typ _ = error "dBtype_of_typ: bad type";  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
464  | 
|
| 17145 | 465  | 
fun dB_of_term (Bound i) = Norm.dB_Var (nat_of_int i)  | 
466  | 
| dB_of_term (t $ u) = Norm.App (dB_of_term t, dB_of_term u)  | 
|
467  | 
| dB_of_term (Abs (_, _, t)) = Norm.Abs (dB_of_term t)  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
468  | 
| dB_of_term _ = error "dB_of_term: bad term";  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
469  | 
|
| 17145 | 470  | 
fun term_of_dB Ts (Type ("fun", [T, U])) (Norm.Abs dBt) =
 | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
471  | 
      Abs ("x", T, term_of_dB (T :: Ts) U dBt)
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
472  | 
| term_of_dB Ts _ dBt = term_of_dB' Ts dBt  | 
| 17145 | 473  | 
and term_of_dB' Ts (Norm.dB_Var n) = Bound (int_of_nat n)  | 
474  | 
| term_of_dB' Ts (Norm.App (dBt, dBu)) =  | 
|
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
475  | 
let val t = term_of_dB' Ts dBt  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
476  | 
in case fastype_of1 (Ts, t) of  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
477  | 
          Type ("fun", [T, U]) => t $ term_of_dB Ts T dBu
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
478  | 
| _ => error "term_of_dB: function type expected"  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
479  | 
end  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
480  | 
| term_of_dB' _ _ = error "term_of_dB: term not in normal form";  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
481  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
482  | 
fun typing_of_term Ts e (Bound i) =  | 
| 17145 | 483  | 
Norm.Var (e, nat_of_int i, dBtype_of_typ (List.nth (Ts, i)))  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
484  | 
| typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of  | 
| 17145 | 485  | 
        Type ("fun", [T, U]) => Norm.rtypingT_App (e, dB_of_term t,
 | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
486  | 
dBtype_of_typ T, dBtype_of_typ U, dB_of_term u,  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
487  | 
typing_of_term Ts e t, typing_of_term Ts e u)  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
488  | 
| _ => error "typing_of_term: function type expected")  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
489  | 
| typing_of_term Ts e (Abs (s, T, t)) =  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
490  | 
let val dBT = dBtype_of_typ T  | 
| 17145 | 491  | 
in Norm.rtypingT_Abs (e, dBT, dB_of_term t,  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
492  | 
dBtype_of_typ (fastype_of1 (T :: Ts, t)),  | 
| 
20713
 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 
haftmann 
parents: 
20503 
diff
changeset
 | 
493  | 
typing_of_term (T :: Ts) (Norm.shift e Norm.zero dBT) t)  | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
494  | 
end  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
495  | 
| typing_of_term _ _ _ = error "typing_of_term: bad term";  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
496  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
497  | 
fun dummyf _ = error "dummy";  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
498  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
499  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
500  | 
text {*
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
501  | 
We now try out the extracted program @{text "type_NF"} on some example terms.
 | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
502  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
503  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
504  | 
ML {*
 | 
| 22512 | 505  | 
val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"};
 | 
| 17145 | 506  | 
val (dB1, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct1));  | 
| 32010 | 507  | 
val ct1' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct1)) dB1);
 | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
508  | 
|
| 22512 | 509  | 
val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"};
 | 
| 17145 | 510  | 
val (dB2, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct2));  | 
| 32010 | 511  | 
val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2);
 | 
| 
14063
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
512  | 
*}  | 
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
513  | 
|
| 
 
e61a310cde02
New proof of weak normalization with program extraction.
 
berghofe 
parents:  
diff
changeset
 | 
514  | 
end  |