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