author | wenzelm |
Thu, 16 Feb 2006 21:12:00 +0100 | |
changeset 19086 | 1b3780be6cc2 |
parent 18513 | 791b53bf4073 |
child 19363 | 667b5ea637dd |
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 |
|
16417 | 9 |
theory WeakNorm imports Type begin |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
10 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
11 |
text {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
12 |
Formalization by Stefan Berghofer. Partly based on a paper proof by |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
13 |
Felix Joachimski and Ralph Matthes \cite{Matthes-Joachimski-AML}. |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
14 |
*} |
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 |
subsection {* Terms in normal form *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
18 |
|
19086 | 19 |
definition |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
20 |
listall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" |
19086 | 21 |
"listall P xs == (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))" |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
22 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
23 |
declare listall_def [extraction_expand] |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
24 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
25 |
theorem listall_nil: "listall P []" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
26 |
by (simp add: listall_def) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
27 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
28 |
theorem listall_nil_eq [simp]: "listall P [] = True" |
17589 | 29 |
by (iprover intro: listall_nil) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
30 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
31 |
theorem listall_cons: "P x \<Longrightarrow> listall P xs \<Longrightarrow> listall P (x # xs)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
32 |
apply (simp add: listall_def) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
33 |
apply (rule allI impI)+ |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
34 |
apply (case_tac i) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
35 |
apply simp+ |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
36 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
37 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
38 |
theorem listall_cons_eq [simp]: "listall P (x # xs) = (P x \<and> listall P xs)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
39 |
apply (rule iffI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
40 |
prefer 2 |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
41 |
apply (erule conjE) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
42 |
apply (erule listall_cons) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
43 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
44 |
apply (unfold listall_def) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
45 |
apply (rule conjI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
46 |
apply (erule_tac x=0 in allE) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
47 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
48 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
49 |
apply (rule allI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
50 |
apply (erule_tac x="Suc i" in allE) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
51 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
52 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
53 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
54 |
lemma listall_conj1: "listall (\<lambda>x. P x \<and> Q x) xs \<Longrightarrow> listall P xs" |
18241 | 55 |
by (induct xs) simp_all |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
56 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
57 |
lemma listall_conj2: "listall (\<lambda>x. P x \<and> Q x) xs \<Longrightarrow> listall Q xs" |
18241 | 58 |
by (induct xs) simp_all |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
59 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
60 |
lemma listall_app: "listall P (xs @ ys) = (listall P xs \<and> listall P ys)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
61 |
apply (induct xs) |
18241 | 62 |
apply (rule iffI, simp, simp) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
63 |
apply (rule iffI, simp, simp) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
64 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
65 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
66 |
lemma listall_snoc [simp]: "listall P (xs @ [x]) = (listall P xs \<and> P x)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
67 |
apply (rule iffI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
68 |
apply (simp add: listall_app)+ |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
69 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
70 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
71 |
lemma listall_cong [cong, extraction_expand]: |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
72 |
"xs = ys \<Longrightarrow> listall P xs = listall P ys" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
73 |
-- {* Currently needed for strange technical reasons *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
74 |
by (unfold listall_def) simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
75 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
76 |
consts NF :: "dB set" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
77 |
inductive NF |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
78 |
intros |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
79 |
App: "listall (\<lambda>t. t \<in> NF) ts \<Longrightarrow> Var x \<degree>\<degree> ts \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
80 |
Abs: "t \<in> NF \<Longrightarrow> Abs t \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
81 |
monos listall_def |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
82 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
83 |
lemma nat_eq_dec: "\<And>n::nat. m = n \<or> m \<noteq> n" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
84 |
apply (induct m) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
85 |
apply (case_tac n) |
15236
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
nipkow
parents:
14860
diff
changeset
|
86 |
apply (case_tac [3] n) |
17589 | 87 |
apply (simp only: nat.simps, iprover?)+ |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
88 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
89 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
90 |
lemma nat_le_dec: "\<And>n::nat. m < n \<or> \<not> (m < n)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
91 |
apply (induct m) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
92 |
apply (case_tac n) |
15236
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
nipkow
parents:
14860
diff
changeset
|
93 |
apply (case_tac [3] n) |
17589 | 94 |
apply (simp del: simp_thms, iprover?)+ |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
95 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
96 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
97 |
lemma App_NF_D: assumes NF: "Var n \<degree>\<degree> ts \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
98 |
shows "listall (\<lambda>t. t \<in> NF) ts" using NF |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
99 |
by cases simp_all |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
100 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
101 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
102 |
subsection {* Properties of @{text NF} *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
103 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
104 |
lemma Var_NF: "Var n \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
105 |
apply (subgoal_tac "Var n \<degree>\<degree> [] \<in> NF") |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
106 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
107 |
apply (rule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
108 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
109 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
110 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
111 |
lemma subst_terms_NF: "listall (\<lambda>t. t \<in> NF) ts \<Longrightarrow> |
18513 | 112 |
listall (\<lambda>t. \<forall>i j. t[Var i/j] \<in> NF) ts \<Longrightarrow> |
113 |
listall (\<lambda>t. t \<in> NF) (map (\<lambda>t. t[Var i/j]) ts)" |
|
114 |
by (induct ts) simp_all |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
115 |
|
18257 | 116 |
lemma subst_Var_NF: "t \<in> NF \<Longrightarrow> t[Var i/j] \<in> NF" |
117 |
apply (induct fixing: i j set: NF) |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
118 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
119 |
apply (frule listall_conj1) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
120 |
apply (drule listall_conj2) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
121 |
apply (drule_tac i=i and j=j in subst_terms_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
122 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
123 |
apply (rule_tac m=x and n=j in nat_eq_dec [THEN disjE, standard]) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
124 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
125 |
apply (erule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
126 |
apply (rule_tac m=j and n=x in nat_le_dec [THEN disjE, standard]) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
127 |
apply simp |
17589 | 128 |
apply (iprover intro: NF.App) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
129 |
apply simp |
17589 | 130 |
apply (iprover intro: NF.App) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
131 |
apply simp |
17589 | 132 |
apply (iprover intro: NF.Abs) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
133 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
134 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
135 |
lemma app_Var_NF: "t \<in> NF \<Longrightarrow> \<exists>t'. t \<degree> Var i \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
136 |
apply (induct set: NF) |
15948 | 137 |
apply (simplesubst app_last) --{*Using @{text subst} makes extraction fail*} |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
138 |
apply (rule exI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
139 |
apply (rule conjI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
140 |
apply (rule rtrancl_refl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
141 |
apply (rule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
142 |
apply (drule listall_conj1) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
143 |
apply (simp add: listall_app) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
144 |
apply (rule Var_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
145 |
apply (rule exI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
146 |
apply (rule conjI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
147 |
apply (rule rtrancl_into_rtrancl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
148 |
apply (rule rtrancl_refl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
149 |
apply (rule beta) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
150 |
apply (erule subst_Var_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
151 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
152 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
153 |
lemma lift_terms_NF: "listall (\<lambda>t. t \<in> NF) ts \<Longrightarrow> |
18513 | 154 |
listall (\<lambda>t. \<forall>i. lift t i \<in> NF) ts \<Longrightarrow> |
155 |
listall (\<lambda>t. t \<in> NF) (map (\<lambda>t. lift t i) ts)" |
|
18257 | 156 |
by (induct ts) simp_all |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
157 |
|
18257 | 158 |
lemma lift_NF: "t \<in> NF \<Longrightarrow> lift t i \<in> NF" |
159 |
apply (induct fixing: i set: NF) |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
160 |
apply (frule listall_conj1) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
161 |
apply (drule listall_conj2) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
162 |
apply (drule_tac i=i in lift_terms_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
163 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
164 |
apply (rule_tac m=x and n=i in nat_le_dec [THEN disjE, standard]) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
165 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
166 |
apply (rule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
167 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
168 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
169 |
apply (rule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
170 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
171 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
172 |
apply (rule NF.Abs) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
173 |
apply simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
174 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
175 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
176 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
177 |
subsection {* Main theorems *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
178 |
|
18331
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
179 |
lemma norm_list: |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
180 |
assumes f_compat: "\<And>t t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> f t \<rightarrow>\<^sub>\<beta>\<^sup>* f t'" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
181 |
and f_NF: "\<And>t. t \<in> NF \<Longrightarrow> f t \<in> NF" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
182 |
and uNF: "u \<in> NF" and uT: "e \<turnstile> u : T" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
183 |
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
|
184 |
listall (\<lambda>t. \<forall>e T' u i. e\<langle>i:T\<rangle> \<turnstile> t : T' \<longrightarrow> |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
185 |
u \<in> NF \<longrightarrow> e \<turnstile> u : T \<longrightarrow> (\<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF)) as \<Longrightarrow> |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
186 |
\<exists>as'. \<forall>j. Var j \<degree>\<degree> map (\<lambda>t. f (t[u/i])) as \<rightarrow>\<^sub>\<beta>\<^sup>* |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
187 |
Var j \<degree>\<degree> map f as' \<and> Var j \<degree>\<degree> map f as' \<in> NF" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
188 |
(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
|
189 |
proof (induct as rule: rev_induct) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
190 |
case (Nil Us) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
191 |
with Var_NF have "?ex Us [] []" by simp |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
192 |
thus ?case .. |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
193 |
next |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
194 |
case (snoc b bs Us) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
195 |
have "e\<langle>i:T\<rangle> \<tturnstile> bs @ [b] : Us" . |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
196 |
then obtain Vs W where Us: "Us = Vs @ [W]" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
197 |
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
|
198 |
by (rule types_snocE) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
199 |
from snoc have "listall ?R bs" by simp |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
200 |
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
|
201 |
then obtain bs' where |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
202 |
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'" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
203 |
and bsNF: "\<And>j. Var j \<degree>\<degree> map f bs' \<in> NF" by iprover |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
204 |
from snoc have "?R b" by simp |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
205 |
with bT and uNF and uT have "\<exists>b'. b[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* b' \<and> b' \<in> NF" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
206 |
by iprover |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
207 |
then obtain b' where bred: "b[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* b'" and bNF: "b' \<in> NF" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
208 |
by iprover |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
209 |
from bsNF [of 0] have "listall (\<lambda>t. t \<in> NF) (map f bs')" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
210 |
by (rule App_NF_D) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
211 |
moreover have "f b' \<in> NF" by (rule f_NF) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
212 |
ultimately have "listall (\<lambda>t. t \<in> NF) (map f (bs' @ [b']))" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
213 |
by simp |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
214 |
hence "\<And>j. Var j \<degree>\<degree> map f (bs' @ [b']) \<in> NF" by (rule NF.App) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
215 |
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
|
216 |
by (rule f_compat) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
217 |
with bsred have |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
218 |
"\<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
|
219 |
(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
|
220 |
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
|
221 |
thus ?case .. |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
222 |
qed |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
223 |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
224 |
lemma subst_type_NF: |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
225 |
"\<And>t e T u i. t \<in> NF \<Longrightarrow> e\<langle>i:U\<rangle> \<turnstile> t : T \<Longrightarrow> u \<in> NF \<Longrightarrow> e \<turnstile> u : U \<Longrightarrow> \<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
226 |
(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
|
227 |
proof (induct U) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
228 |
fix T t |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
229 |
let ?R = "\<lambda>t. \<forall>e T' u i. |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
230 |
e\<langle>i:T\<rangle> \<turnstile> t : T' \<longrightarrow> u \<in> NF \<longrightarrow> e \<turnstile> u : T \<longrightarrow> (\<exists>t'. t[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
231 |
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
|
232 |
assume MI2: "\<And>T1 T2. T = T1 \<Rightarrow> T2 \<Longrightarrow> PROP ?P T2" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
233 |
assume "t \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
234 |
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
|
235 |
proof induct |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
236 |
fix e T' u i assume uNF: "u \<in> NF" and uT: "e \<turnstile> u : T" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
237 |
{ |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
238 |
case (App ts x e_ T'_ u_ i_) |
18331
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
239 |
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
|
240 |
then obtain Us |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
241 |
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
|
242 |
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
|
243 |
by (rule var_app_typesE) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
244 |
from nat_eq_dec show "\<exists>t'. (Var x \<degree>\<degree> ts)[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
245 |
proof |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
246 |
assume eq: "x = i" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
247 |
show ?thesis |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
248 |
proof (cases ts) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
249 |
case Nil |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
250 |
with eq have "(Var x \<degree>\<degree> [])[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* u" by simp |
17589 | 251 |
with Nil and uNF show ?thesis by simp iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
252 |
next |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
253 |
case (Cons a as) |
18331
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
254 |
with argsT obtain T'' Ts where Us: "Us = T'' # Ts" |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
255 |
by (cases Us) (rule FalseE, simp+) |
18331
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
256 |
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
|
257 |
by simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
258 |
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
|
259 |
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
|
260 |
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
|
261 |
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
|
262 |
from argT uT refl have aT: "e \<turnstile> a[u/i] : T''" by (rule subst_lemma) |
17589 | 263 |
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
|
264 |
with lift_preserves_beta' lift_NF uNF uT argsT' |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
265 |
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
|
266 |
Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as' \<and> |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
267 |
Var j \<degree>\<degree> map (\<lambda>t. lift t 0) as' \<in> NF" by (rule norm_list) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
268 |
then obtain as' where |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
269 |
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
|
270 |
Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as'" |
17589 | 271 |
and asNF: "Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as' \<in> NF" by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
272 |
from App and Cons have "?R a" by simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
273 |
with argT and uNF and uT have "\<exists>a'. a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a' \<and> a' \<in> NF" |
17589 | 274 |
by iprover |
275 |
then obtain a' where ared: "a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* a'" and aNF: "a' \<in> NF" by iprover |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
276 |
from uNF have "lift u 0 \<in> NF" by (rule lift_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
277 |
hence "\<exists>u'. lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u' \<and> u' \<in> NF" by (rule app_Var_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
278 |
then obtain u' where ured: "lift u 0 \<degree> Var 0 \<rightarrow>\<^sub>\<beta>\<^sup>* u'" and u'NF: "u' \<in> NF" |
17589 | 279 |
by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
280 |
from T and u'NF have "\<exists>ua. u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua \<and> ua \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
281 |
proof (rule MI1) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
282 |
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
|
283 |
proof (rule typing.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
284 |
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
|
285 |
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
|
286 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
287 |
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
|
288 |
from ared aT show "e \<turnstile> a' : T''" by (rule subject_reduction') |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
289 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
290 |
then obtain ua where uared: "u'[a'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" and uaNF: "ua \<in> NF" |
17589 | 291 |
by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
292 |
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
|
293 |
by (rule subst_preserves_beta2') |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
294 |
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
|
295 |
by (rule subst_preserves_beta') |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
296 |
also note uared |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
297 |
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
|
298 |
hence uared': "u \<degree> a[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* ua" by simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
299 |
from T have "\<exists>r. (Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r \<and> r \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
300 |
proof (rule MI2) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
301 |
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
|
302 |
proof (rule list_app_typeI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
by (rule substs_lemma) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
306 |
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
|
307 |
by (rule lift_types) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
308 |
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
|
309 |
by (simp_all add: map_compose [symmetric] o_def) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
310 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
311 |
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
|
312 |
by (rule subject_reduction') |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
313 |
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
|
314 |
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
|
315 |
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
|
316 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
317 |
then obtain r where rred: "(Var 0 \<degree>\<degree> map (\<lambda>t. lift t 0) as')[ua/0] \<rightarrow>\<^sub>\<beta>\<^sup>* r" |
17589 | 318 |
and rnf: "r \<in> NF" by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
319 |
from asred have |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
320 |
"(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
|
321 |
(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
|
322 |
by (rule subst_preserves_beta') |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
323 |
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
|
324 |
(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
|
325 |
also note rred |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
326 |
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
|
327 |
with rnf Cons eq show ?thesis |
17589 | 328 |
by (simp add: map_compose [symmetric] o_def) iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
329 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
330 |
next |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
331 |
assume neq: "x \<noteq> i" |
18331
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
332 |
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
|
333 |
with TrueI TrueI uNF uT argsT |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
334 |
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> |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
335 |
Var j \<degree>\<degree> ts' \<in> NF" (is "\<exists>ts'. ?ex ts'") |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
336 |
by (rule norm_list [of "\<lambda>t. t", simplified]) |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
337 |
then obtain ts' where NF: "?ex ts'" .. |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
338 |
from nat_le_dec show ?thesis |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
339 |
proof |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
340 |
assume "i < x" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
341 |
with NF show ?thesis by simp iprover |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
342 |
next |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
343 |
assume "\<not> (i < x)" |
eb3a7d3d874b
Factored out proof for normalization of applications (norm_list).
berghofe
parents:
18257
diff
changeset
|
344 |
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
|
345 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
346 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
347 |
next |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
348 |
case (Abs r e_ T'_ u_ i_) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
349 |
assume absT: "e\<langle>i:T\<rangle> \<turnstile> Abs r : T'" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
350 |
then obtain R S where "e\<langle>0:R\<rangle>\<langle>Suc i:T\<rangle> \<turnstile> r : S" by (rule abs_typeE) simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
351 |
moreover have "lift u 0 \<in> NF" by (rule lift_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
352 |
moreover have "e\<langle>0:R\<rangle> \<turnstile> lift u 0 : T" by (rule lift_type) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
353 |
ultimately have "\<exists>t'. r[lift u 0/Suc i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" by (rule Abs) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
354 |
thus "\<exists>t'. Abs r[u/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" |
17589 | 355 |
by simp (iprover intro: rtrancl_beta_Abs NF.Abs) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
356 |
} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
357 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
358 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
359 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
360 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
361 |
consts -- {* A computationally relevant copy of @{term "e \<turnstile> t : T"} *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
362 |
rtyping :: "((nat \<Rightarrow> type) \<times> dB \<times> type) set" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
363 |
|
19086 | 364 |
abbreviation (output) |
365 |
rtyping_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ |-\<^sub>R _ : _" [50, 50, 50] 50) |
|
366 |
"(e |-\<^sub>R t : T) = ((e, t, T) \<in> rtyping)" |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
367 |
syntax (xsymbols) |
19086 | 368 |
rtyping_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile>\<^sub>R _ : _" [50, 50, 50] 50) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
369 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
370 |
inductive rtyping |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
371 |
intros |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
372 |
Var: "e x = T \<Longrightarrow> e \<turnstile>\<^sub>R Var x : T" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
373 |
Abs: "e\<langle>0:T\<rangle> \<turnstile>\<^sub>R t : U \<Longrightarrow> e \<turnstile>\<^sub>R Abs t : (T \<Rightarrow> U)" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
374 |
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" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
375 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
376 |
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
|
377 |
apply (induct set: rtyping) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
378 |
apply (erule typing.Var) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
379 |
apply (erule typing.Abs) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
380 |
apply (erule typing.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
381 |
apply assumption |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
382 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
383 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
384 |
|
18513 | 385 |
theorem type_NF: |
386 |
assumes "e \<turnstile>\<^sub>R t : T" |
|
387 |
shows "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> t' \<in> NF" using prems |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
388 |
proof induct |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
389 |
case Var |
17589 | 390 |
show ?case by (iprover intro: Var_NF) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
391 |
next |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
392 |
case Abs |
17589 | 393 |
thus ?case by (iprover intro: rtrancl_beta_Abs NF.Abs) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
394 |
next |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
395 |
case (App T U e s t) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
396 |
from App obtain s' t' where |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
397 |
sred: "s \<rightarrow>\<^sub>\<beta>\<^sup>* s'" and sNF: "s' \<in> NF" |
17589 | 398 |
and tred: "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and tNF: "t' \<in> NF" by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
399 |
have "\<exists>u. (Var 0 \<degree> lift t' 0)[s'/0] \<rightarrow>\<^sub>\<beta>\<^sup>* u \<and> u \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
400 |
proof (rule subst_type_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
401 |
have "lift t' 0 \<in> NF" by (rule lift_NF) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
402 |
hence "listall (\<lambda>t. t \<in> NF) [lift t' 0]" by (rule listall_cons) (rule listall_nil) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
403 |
hence "Var 0 \<degree>\<degree> [lift t' 0] \<in> NF" by (rule NF.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
404 |
thus "Var 0 \<degree> lift t' 0 \<in> NF" by simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
405 |
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
|
406 |
proof (rule typing.App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
407 |
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
|
408 |
by (rule typing.Var) simp |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
409 |
from tred have "e \<turnstile> t' : T" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
410 |
by (rule subject_reduction') (rule rtyping_imp_typing) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
411 |
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
|
412 |
by (rule lift_type) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
413 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
414 |
from sred show "e \<turnstile> s' : T \<Rightarrow> U" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
415 |
by (rule subject_reduction') (rule rtyping_imp_typing) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
416 |
qed |
17589 | 417 |
then obtain u where ured: "s' \<degree> t' \<rightarrow>\<^sub>\<beta>\<^sup>* u" and unf: "u \<in> NF" by simp iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
418 |
from sred tred have "s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'" by (rule rtrancl_beta_App) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
419 |
hence "s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using ured by (rule rtrancl_trans) |
17589 | 420 |
with unf show ?case by iprover |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
421 |
qed |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
422 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
423 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
424 |
subsection {* Extracting the program *} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
425 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
426 |
declare NF.induct [ind_realizer] |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
427 |
declare rtrancl.induct [ind_realizer irrelevant] |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
428 |
declare rtyping.induct [ind_realizer] |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
429 |
lemmas [extraction_expand] = trans_def conj_assoc listall_cons_eq |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
430 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
431 |
extract type_NF |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
432 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
433 |
lemma rtranclR_rtrancl_eq: "((a, b) \<in> rtranclR r) = ((a, b) \<in> rtrancl (Collect r))" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
434 |
apply (rule iffI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
435 |
apply (erule rtranclR.induct) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
436 |
apply (rule rtrancl_refl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
437 |
apply (erule rtrancl_into_rtrancl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
438 |
apply (erule CollectI) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
439 |
apply (erule rtrancl.induct) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
440 |
apply (rule rtranclR.rtrancl_refl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
441 |
apply (erule rtranclR.rtrancl_into_rtrancl) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
442 |
apply (erule CollectD) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
443 |
done |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
444 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
445 |
lemma NFR_imp_NF: "(nf, t) \<in> NFR \<Longrightarrow> t \<in> NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
446 |
apply (erule NFR.induct) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
447 |
apply (rule NF.intros) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
448 |
apply (simp add: listall_def) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
449 |
apply (erule NF.intros) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
450 |
done |
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 |
text_raw {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
453 |
\begin{figure} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
454 |
\renewcommand{\isastyle}{\scriptsize\it}% |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
455 |
@{thm [display,eta_contract=false,margin=100] subst_type_NF_def} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
456 |
\renewcommand{\isastyle}{\small\it}% |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
457 |
\caption{Program extracted from @{text subst_type_NF}} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
458 |
\label{fig:extr-subst-type-nf} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
459 |
\end{figure} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
460 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
461 |
\begin{figure} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
462 |
\renewcommand{\isastyle}{\scriptsize\it}% |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
463 |
@{thm [display,margin=100] subst_Var_NF_def} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
464 |
@{thm [display,margin=100] app_Var_NF_def} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
465 |
@{thm [display,margin=100] lift_NF_def} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
466 |
@{thm [display,eta_contract=false,margin=100] type_NF_def} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
467 |
\renewcommand{\isastyle}{\small\it}% |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
468 |
\caption{Program extracted from lemmas and main theorem} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
469 |
\label{fig:extr-type-nf} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
470 |
\end{figure} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
471 |
*} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
472 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
473 |
text {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
474 |
The program corresponding to the proof of the central lemma, which |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
475 |
performs substitution and normalization, is shown in Figure |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
476 |
\ref{fig:extr-subst-type-nf}. The correctness |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
477 |
theorem corresponding to the program @{text "subst_type_NF"} is |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
478 |
@{thm [display,margin=100] subst_type_NF_correctness |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
479 |
[simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
480 |
where @{text NFR} is the realizability predicate corresponding to |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
481 |
the datatype @{text NFT}, which is inductively defined by the rules |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
482 |
\pagebreak |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
483 |
@{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
|
484 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
485 |
The programs corresponding to the main theorem @{text "type_NF"}, as |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
486 |
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
|
487 |
The correctness statement for the main function @{text "type_NF"} is |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
488 |
@{thm [display,margin=100] type_NF_correctness |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
489 |
[simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
490 |
where the realizability predicate @{text "rtypingR"} corresponding to the |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
491 |
computationally relevant version of the typing judgement is inductively |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
492 |
defined by the rules |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
493 |
@{thm [display,margin=100] rtypingR.Var [no_vars] |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
494 |
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
|
495 |
*} |
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 |
subsection {* Generating executable code *} |
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 |
consts_code |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
500 |
arbitrary :: "'a" ("(error \"arbitrary\")") |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
501 |
arbitrary :: "'a \<Rightarrow> 'b" ("(fn '_ => error \"arbitrary\")") |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
502 |
|
17145 | 503 |
code_module Norm |
504 |
contains |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
505 |
test = "type_NF" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
506 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
507 |
text {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
508 |
The following functions convert between Isabelle's built-in {\tt term} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
509 |
datatype and the generated {\tt dB} datatype. This allows to |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
510 |
generate example terms using Isabelle's parser and inspect |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
511 |
normalized terms using Isabelle's pretty printer. |
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 |
ML {* |
17145 | 515 |
fun nat_of_int 0 = Norm.id_0 |
516 |
| 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
|
517 |
|
17145 | 518 |
fun int_of_nat Norm.id_0 = 0 |
519 |
| 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
|
520 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
521 |
fun dBtype_of_typ (Type ("fun", [T, U])) = |
17145 | 522 |
Norm.Fun (dBtype_of_typ T, dBtype_of_typ U) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
523 |
| dBtype_of_typ (TFree (s, _)) = (case explode s of |
17145 | 524 |
["'", a] => Norm.Atom (nat_of_int (ord a - 97)) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
525 |
| _ => error "dBtype_of_typ: variable name") |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
526 |
| dBtype_of_typ _ = error "dBtype_of_typ: bad type"; |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
527 |
|
17145 | 528 |
fun dB_of_term (Bound i) = Norm.dB_Var (nat_of_int i) |
529 |
| dB_of_term (t $ u) = Norm.App (dB_of_term t, dB_of_term u) |
|
530 |
| 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
|
531 |
| dB_of_term _ = error "dB_of_term: bad term"; |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
532 |
|
17145 | 533 |
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
|
534 |
Abs ("x", T, term_of_dB (T :: Ts) U dBt) |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
535 |
| term_of_dB Ts _ dBt = term_of_dB' Ts dBt |
17145 | 536 |
and term_of_dB' Ts (Norm.dB_Var n) = Bound (int_of_nat n) |
537 |
| term_of_dB' Ts (Norm.App (dBt, dBu)) = |
|
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
538 |
let val t = term_of_dB' Ts dBt |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
539 |
in case fastype_of1 (Ts, t) of |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
540 |
Type ("fun", [T, U]) => t $ term_of_dB Ts T dBu |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
541 |
| _ => error "term_of_dB: function type expected" |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
542 |
end |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
543 |
| 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
|
544 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
545 |
fun typing_of_term Ts e (Bound i) = |
17145 | 546 |
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
|
547 |
| typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of |
17145 | 548 |
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
|
549 |
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
|
550 |
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
|
551 |
| _ => error "typing_of_term: function type expected") |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
552 |
| typing_of_term Ts e (Abs (s, T, t)) = |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
553 |
let val dBT = dBtype_of_typ T |
17145 | 554 |
in Norm.rtypingT_Abs (e, dBT, dB_of_term t, |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
555 |
dBtype_of_typ (fastype_of1 (T :: Ts, t)), |
17145 | 556 |
typing_of_term (T :: Ts) (Norm.shift e Norm.id_0 dBT) t) |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
557 |
end |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
558 |
| typing_of_term _ _ _ = error "typing_of_term: bad term"; |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
559 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
560 |
fun dummyf _ = error "dummy"; |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
561 |
*} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
562 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
563 |
text {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
564 |
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
|
565 |
*} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
566 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
567 |
ML {* |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
568 |
val sg = sign_of (the_context()); |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
569 |
fun rd s = read_cterm sg (s, TypeInfer.logicT); |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
570 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
571 |
val ct1 = rd "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"; |
17145 | 572 |
val (dB1, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct1)); |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
573 |
val ct1' = cterm_of sg (term_of_dB [] (#T (rep_cterm ct1)) dB1); |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
574 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
575 |
val ct2 = rd |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
576 |
"%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 | 577 |
val (dB2, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct2)); |
14063
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
578 |
val ct2' = cterm_of sg (term_of_dB [] (#T (rep_cterm ct2)) dB2); |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
579 |
*} |
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
580 |
|
e61a310cde02
New proof of weak normalization with program extraction.
berghofe
parents:
diff
changeset
|
581 |
end |