| author | huffman | 
| Wed, 27 Sep 2006 00:52:59 +0200 | |
| changeset 20718 | 4c4869e4ddb7 | 
| parent 20503 | 503ac4c5ef91 | 
| child 21210 | c17fd2df4e9e | 
| permissions | -rw-r--r-- | 
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 1 | (* Title: HOL/Lambda/Type.thy | 
| 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 2 | ID: $Id$ | 
| 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 3 | Author: Stefan Berghofer | 
| 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 4 | Copyright 2000 TU Muenchen | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 5 | *) | 
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 6 | |
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 7 | header {* Simply-typed lambda terms *}
 | 
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 8 | |
| 16417 | 9 | theory Type imports ListApplication begin | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 10 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 11 | |
| 11946 | 12 | subsection {* Environments *}
 | 
| 13 | ||
| 19086 | 14 | definition | 
| 12011 | 15 |   shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"    ("_<_:_>" [90, 0, 0] 91)
 | 
| 19086 | 16 | "e<i:a> = (\<lambda>j. if j < i then e j else if j = i then a else e (j - 1))" | 
| 19363 | 17 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 18 | const_syntax (xsymbols) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 19 |   shift  ("_\<langle>_:_\<rangle>" [90, 0, 0] 91)
 | 
| 19363 | 20 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 21 | const_syntax (HTML output) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 22 |   shift  ("_\<langle>_:_\<rangle>" [90, 0, 0] 91)
 | 
| 11946 | 23 | |
| 24 | lemma shift_eq [simp]: "i = j \<Longrightarrow> (e\<langle>i:T\<rangle>) j = T" | |
| 25 | by (simp add: shift_def) | |
| 26 | ||
| 27 | lemma shift_gt [simp]: "j < i \<Longrightarrow> (e\<langle>i:T\<rangle>) j = e j" | |
| 28 | by (simp add: shift_def) | |
| 29 | ||
| 30 | lemma shift_lt [simp]: "i < j \<Longrightarrow> (e\<langle>i:T\<rangle>) j = e (j - 1)" | |
| 31 | by (simp add: shift_def) | |
| 32 | ||
| 33 | lemma shift_commute [simp]: "e\<langle>i:U\<rangle>\<langle>0:T\<rangle> = e\<langle>0:T\<rangle>\<langle>Suc i:U\<rangle>" | |
| 34 | apply (rule ext) | |
| 35 | apply (case_tac x) | |
| 36 | apply simp | |
| 37 | apply (case_tac nat) | |
| 38 | apply (simp_all add: shift_def) | |
| 39 | done | |
| 40 | ||
| 41 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 42 | subsection {* Types and typing rules *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 43 | |
| 9641 | 44 | datatype type = | 
| 9622 | 45 | Atom nat | 
| 11945 | 46 | | Fun type type (infixr "\<Rightarrow>" 200) | 
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 47 | |
| 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 48 | consts | 
| 11943 | 49 | typing :: "((nat \<Rightarrow> type) \<times> dB \<times> type) set" | 
| 50 | typings :: "(nat \<Rightarrow> type) \<Rightarrow> dB list \<Rightarrow> type list \<Rightarrow> bool" | |
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 51 | |
| 19363 | 52 | abbreviation | 
| 19086 | 53 | funs :: "type list \<Rightarrow> type \<Rightarrow> type" (infixr "=>>" 200) | 
| 19363 | 54 | "Ts =>> T == foldr Fun Ts T" | 
| 19086 | 55 | |
| 56 |   typing_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"    ("_ |- _ : _" [50, 50, 50] 50)
 | |
| 19363 | 57 | "env |- t : T == (env, t, T) \<in> typing" | 
| 19086 | 58 | |
| 59 | typings_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB list \<Rightarrow> type list \<Rightarrow> bool" | |
| 11943 | 60 |     ("_ ||- _ : _" [50, 50, 50] 50)
 | 
| 19363 | 61 | "env ||- ts : Ts == typings env ts Ts" | 
| 19086 | 62 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 63 | const_syntax (xsymbols) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 64 |   typing_rel  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
 | 
| 19363 | 65 | |
| 19656 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 66 | const_syntax (latex) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 67 | funs (infixr "\<Rrightarrow>" 200) | 
| 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 wenzelm parents: 
19380diff
changeset | 68 |   typings_rel  ("_ \<tturnstile> _ : _" [50, 50, 50] 50)
 | 
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 69 | |
| 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 70 | inductive typing | 
| 11638 | 71 | intros | 
| 11943 | 72 | Var [intro!]: "env x = T \<Longrightarrow> env \<turnstile> Var x : T" | 
| 11946 | 73 | Abs [intro!]: "env\<langle>0:T\<rangle> \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs t : (T \<Rightarrow> U)" | 
| 12011 | 74 | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" | 
| 9622 | 75 | |
| 11943 | 76 | inductive_cases typing_elims [elim!]: | 
| 77 | "e \<turnstile> Var i : T" | |
| 12011 | 78 | "e \<turnstile> t \<degree> u : T" | 
| 11943 | 79 | "e \<turnstile> Abs t : T" | 
| 11935 | 80 | |
| 11943 | 81 | primrec | 
| 82 | "(e \<tturnstile> [] : Ts) = (Ts = [])" | |
| 83 | "(e \<tturnstile> (t # ts) : Ts) = | |
| 84 | (case Ts of | |
| 85 | [] \<Rightarrow> False | |
| 86 | | T # Ts \<Rightarrow> e \<turnstile> t : T \<and> e \<tturnstile> ts : Ts)" | |
| 9114 
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
 berghofe parents: diff
changeset | 87 | |
| 9622 | 88 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 89 | subsection {* Some examples *}
 | 
| 9622 | 90 | |
| 12011 | 91 | lemma "e \<turnstile> Abs (Abs (Abs (Var 1 \<degree> (Var 2 \<degree> Var 1 \<degree> Var 0)))) : ?T" | 
| 11935 | 92 | by force | 
| 9622 | 93 | |
| 12011 | 94 | lemma "e \<turnstile> Abs (Abs (Abs (Var 2 \<degree> Var 0 \<degree> (Var 1 \<degree> Var 0)))) : ?T" | 
| 11935 | 95 | by force | 
| 9622 | 96 | |
| 97 | ||
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 98 | subsection {* Lists of types *}
 | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 99 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 100 | lemma lists_typings: | 
| 18241 | 101 |     "e \<tturnstile> ts : Ts \<Longrightarrow> ts \<in> lists {t. \<exists>T. e \<turnstile> t : T}"
 | 
| 20503 | 102 | apply (induct ts arbitrary: Ts) | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 103 | apply (case_tac Ts) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 104 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 105 | apply (rule lists.Nil) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 106 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 107 | apply (case_tac Ts) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 108 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 109 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 110 | apply (rule lists.Cons) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 111 | apply blast | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 112 | apply blast | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 113 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 114 | |
| 18241 | 115 | lemma types_snoc: "e \<tturnstile> ts : Ts \<Longrightarrow> e \<turnstile> t : T \<Longrightarrow> e \<tturnstile> ts @ [t] : Ts @ [T]" | 
| 20503 | 116 | apply (induct ts arbitrary: Ts) | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 117 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 118 | apply (case_tac Ts) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 119 | apply simp+ | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 120 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 121 | |
| 18241 | 122 | lemma types_snoc_eq: "e \<tturnstile> ts @ [t] : Ts @ [T] = | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 123 | (e \<tturnstile> ts : Ts \<and> e \<turnstile> t : T)" | 
| 20503 | 124 | apply (induct ts arbitrary: Ts) | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 125 | apply (case_tac Ts) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 126 | apply simp+ | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 127 | apply (case_tac Ts) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
14565diff
changeset | 128 | apply (case_tac "ts @ [t]") | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 129 | apply simp+ | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 130 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 131 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 132 | lemma rev_exhaust2 [case_names Nil snoc, extraction_expand]: | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 133 | "(xs = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>ys y. xs = ys @ [y] \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 134 |   -- {* Cannot use @{text rev_exhaust} from the @{text List}
 | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 135 | theory, since it is not constructive *} | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 136 | apply (subgoal_tac "\<forall>ys. xs = rev ys \<longrightarrow> P") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 137 | apply (erule_tac x="rev xs" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 138 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 139 | apply (rule allI) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 140 | apply (rule impI) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 141 | apply (case_tac ys) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 142 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 143 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 144 | apply atomize | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 145 | apply (erule allE)+ | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 146 | apply (erule mp, rule conjI) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 147 | apply (rule refl)+ | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 148 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 149 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 150 | lemma types_snocE: "e \<tturnstile> ts @ [t] : Ts \<Longrightarrow> | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 151 | (\<And>Us U. Ts = Us @ [U] \<Longrightarrow> e \<tturnstile> ts : Us \<Longrightarrow> e \<turnstile> t : U \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 152 | apply (cases Ts rule: rev_exhaust2) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 153 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 154 | apply (case_tac "ts @ [t]") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 155 | apply (simp add: types_snoc_eq)+ | 
| 17589 | 156 | apply iprover | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 157 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 158 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 159 | |
| 11950 | 160 | subsection {* n-ary function types *}
 | 
| 9622 | 161 | |
| 11987 | 162 | lemma list_app_typeD: | 
| 18241 | 163 | "e \<turnstile> t \<degree>\<degree> ts : T \<Longrightarrow> \<exists>Ts. e \<turnstile> t : Ts \<Rrightarrow> T \<and> e \<tturnstile> ts : Ts" | 
| 20503 | 164 | apply (induct ts arbitrary: t T) | 
| 9622 | 165 | apply simp | 
| 11987 | 166 | apply atomize | 
| 9622 | 167 | apply simp | 
| 12011 | 168 | apply (erule_tac x = "t \<degree> a" in allE) | 
| 9622 | 169 | apply (erule_tac x = T in allE) | 
| 170 | apply (erule impE) | |
| 171 | apply assumption | |
| 172 | apply (elim exE conjE) | |
| 12011 | 173 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 9622 | 174 | apply (rule_tac x = "Ta # Ts" in exI) | 
| 175 | apply simp | |
| 176 | done | |
| 177 | ||
| 11935 | 178 | lemma list_app_typeE: | 
| 12011 | 179 | "e \<turnstile> t \<degree>\<degree> ts : T \<Longrightarrow> (\<And>Ts. e \<turnstile> t : Ts \<Rrightarrow> T \<Longrightarrow> e \<tturnstile> ts : Ts \<Longrightarrow> C) \<Longrightarrow> C" | 
| 11935 | 180 | by (insert list_app_typeD) fast | 
| 181 | ||
| 11987 | 182 | lemma list_app_typeI: | 
| 18241 | 183 | "e \<turnstile> t : Ts \<Rrightarrow> T \<Longrightarrow> e \<tturnstile> ts : Ts \<Longrightarrow> e \<turnstile> t \<degree>\<degree> ts : T" | 
| 20503 | 184 | apply (induct ts arbitrary: t T Ts) | 
| 9622 | 185 | apply simp | 
| 11987 | 186 | apply atomize | 
| 9622 | 187 | apply (case_tac Ts) | 
| 188 | apply simp | |
| 189 | apply simp | |
| 12011 | 190 | apply (erule_tac x = "t \<degree> a" in allE) | 
| 9622 | 191 | apply (erule_tac x = T in allE) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
14565diff
changeset | 192 | apply (erule_tac x = list in allE) | 
| 9622 | 193 | apply (erule impE) | 
| 194 | apply (erule conjE) | |
| 195 | apply (erule typing.App) | |
| 196 | apply assumption | |
| 197 | apply blast | |
| 198 | done | |
| 199 | ||
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 200 | text {*
 | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 201 | For the specific case where the head of the term is a variable, | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 202 | the following theorems allow to infer the types of the arguments | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 203 | without analyzing the typing derivation. This is crucial | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 204 | for program extraction. | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 205 | *} | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 206 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 207 | theorem var_app_type_eq: | 
| 18241 | 208 | "e \<turnstile> Var i \<degree>\<degree> ts : T \<Longrightarrow> e \<turnstile> Var i \<degree>\<degree> ts : U \<Longrightarrow> T = U" | 
| 20503 | 209 | apply (induct ts arbitrary: T U rule: rev_induct) | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 210 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 211 | apply (ind_cases "e \<turnstile> Var i : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 212 | apply (ind_cases "e \<turnstile> Var i : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 213 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 214 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 215 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 216 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 217 | apply atomize | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 218 | apply (erule_tac x="Ta \<Rightarrow> T" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 219 | apply (erule_tac x="Tb \<Rightarrow> U" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 220 | apply (erule impE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 221 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 222 | apply (erule impE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 223 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 224 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 225 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 226 | |
| 18241 | 227 | lemma var_app_types: "e \<turnstile> Var i \<degree>\<degree> ts \<degree>\<degree> us : T \<Longrightarrow> e \<tturnstile> ts : Ts \<Longrightarrow> | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 228 | e \<turnstile> Var i \<degree>\<degree> ts : U \<Longrightarrow> \<exists>Us. U = Us \<Rrightarrow> T \<and> e \<tturnstile> us : Us" | 
| 20503 | 229 | apply (induct us arbitrary: ts Ts U) | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 230 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 231 | apply (erule var_app_type_eq) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 232 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 233 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 234 | apply atomize | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 235 | apply (case_tac U) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 236 | apply (rule FalseE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 237 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 238 | apply (erule list_app_typeE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 239 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 240 | apply (drule_tac T="Atom nat" and U="Ta \<Rightarrow> Tsa \<Rrightarrow> T" in var_app_type_eq) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 241 | apply assumption | 
| 9622 | 242 | apply simp | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 243 | apply (erule_tac x="ts @ [a]" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 244 | apply (erule_tac x="Ts @ [type1]" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 245 | apply (erule_tac x="type2" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 246 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 247 | apply (erule impE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 248 | apply (rule types_snoc) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 249 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 250 | apply (erule list_app_typeE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 251 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 252 | apply (drule_tac T="type1 \<Rightarrow> type2" and U="Ta \<Rightarrow> Tsa \<Rrightarrow> T" in var_app_type_eq) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 253 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 254 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 255 | apply (erule impE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 256 | apply (rule typing.App) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 257 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 258 | apply (erule list_app_typeE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 259 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 260 | apply (frule_tac T="type1 \<Rightarrow> type2" and U="Ta \<Rightarrow> Tsa \<Rrightarrow> T" in var_app_type_eq) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 261 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 262 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 263 | apply (erule exE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 264 | apply (rule_tac x="type1 # Us" in exI) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 265 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 266 | apply (erule list_app_typeE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 267 | apply (ind_cases "e \<turnstile> t \<degree> u : T") | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 268 | apply (frule_tac T="type1 \<Rightarrow> Us \<Rrightarrow> T" and U="Ta \<Rightarrow> Tsa \<Rrightarrow> T" in var_app_type_eq) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 269 | apply assumption | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 270 | apply simp | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 271 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 272 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 273 | lemma var_app_typesE: "e \<turnstile> Var i \<degree>\<degree> ts : T \<Longrightarrow> | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 274 | (\<And>Ts. e \<turnstile> Var i : Ts \<Rrightarrow> T \<Longrightarrow> e \<tturnstile> ts : Ts \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 275 | apply (drule var_app_types [of _ _ "[]", simplified]) | 
| 17589 | 276 | apply (iprover intro: typing.Var)+ | 
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 277 | done | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 278 | |
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 279 | lemma abs_typeE: "e \<turnstile> Abs t : T \<Longrightarrow> (\<And>U V. e\<langle>0:U\<rangle> \<turnstile> t : V \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 280 | apply (cases T) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 281 | apply (rule FalseE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 282 | apply (erule typing.elims) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 283 | apply simp_all | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 284 | apply atomize | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 285 | apply (erule_tac x="type1" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 286 | apply (erule_tac x="type2" in allE) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 287 | apply (erule mp) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 288 | apply (erule typing.elims) | 
| 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 289 | apply simp_all | 
| 9622 | 290 | done | 
| 291 | ||
| 292 | ||
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 293 | subsection {* Lifting preserves well-typedness *}
 | 
| 9622 | 294 | |
| 18257 | 295 | lemma lift_type [intro!]: "e \<turnstile> t : T \<Longrightarrow> e\<langle>i:U\<rangle> \<turnstile> lift t i : T" | 
| 20503 | 296 | by (induct arbitrary: i U set: typing) auto | 
| 12171 | 297 | |
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 298 | lemma lift_types: | 
| 18241 | 299 | "e \<tturnstile> ts : Ts \<Longrightarrow> e\<langle>i:U\<rangle> \<tturnstile> (map (\<lambda>t. lift t i) ts) : Ts" | 
| 20503 | 300 | apply (induct ts arbitrary: Ts) | 
| 9622 | 301 | apply simp | 
| 302 | apply (case_tac Ts) | |
| 11946 | 303 | apply auto | 
| 9622 | 304 | done | 
| 305 | ||
| 306 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 307 | subsection {* Substitution lemmas *}
 | 
| 9622 | 308 | |
| 11994 | 309 | lemma subst_lemma: | 
| 18257 | 310 | "e \<turnstile> t : T \<Longrightarrow> e' \<turnstile> u : U \<Longrightarrow> e = e'\<langle>i:U\<rangle> \<Longrightarrow> e' \<turnstile> t[u/i] : T" | 
| 20503 | 311 | apply (induct arbitrary: e' i U u set: typing) | 
| 11946 | 312 | apply (rule_tac x = x and y = i in linorder_cases) | 
| 313 | apply auto | |
| 314 | apply blast | |
| 9622 | 315 | done | 
| 316 | ||
| 12011 | 317 | lemma substs_lemma: | 
| 18241 | 318 | "e \<turnstile> u : T \<Longrightarrow> e\<langle>i:T\<rangle> \<tturnstile> ts : Ts \<Longrightarrow> | 
| 11943 | 319 | e \<tturnstile> (map (\<lambda>t. t[u/i]) ts) : Ts" | 
| 20503 | 320 | apply (induct ts arbitrary: Ts) | 
| 9622 | 321 | apply (case_tac Ts) | 
| 322 | apply simp | |
| 323 | apply simp | |
| 12011 | 324 | apply atomize | 
| 9622 | 325 | apply (case_tac Ts) | 
| 326 | apply simp | |
| 327 | apply simp | |
| 328 | apply (erule conjE) | |
| 12011 | 329 | apply (erule (1) subst_lemma) | 
| 11994 | 330 | apply (rule refl) | 
| 331 | done | |
| 332 | ||
| 9622 | 333 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 334 | subsection {* Subject reduction *}
 | 
| 9622 | 335 | |
| 18257 | 336 | lemma subject_reduction: "e \<turnstile> t : T \<Longrightarrow> t -> t' \<Longrightarrow> e \<turnstile> t' : T" | 
| 20503 | 337 | apply (induct arbitrary: t' set: typing) | 
| 9622 | 338 | apply blast | 
| 339 | apply blast | |
| 11994 | 340 | apply atomize | 
| 12011 | 341 | apply (ind_cases "s \<degree> t -> t'") | 
| 9622 | 342 | apply hypsubst | 
| 11945 | 343 | apply (ind_cases "env \<turnstile> Abs t : T \<Rightarrow> U") | 
| 9622 | 344 | apply (rule subst_lemma) | 
| 345 | apply assumption | |
| 346 | apply assumption | |
| 347 | apply (rule ext) | |
| 11935 | 348 | apply (case_tac x) | 
| 11946 | 349 | apply auto | 
| 9622 | 350 | done | 
| 351 | ||
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 352 | theorem subject_reduction': "t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> e \<turnstile> t : T \<Longrightarrow> e \<turnstile> t' : T" | 
| 17589 | 353 | by (induct set: rtrancl) (iprover intro: subject_reduction)+ | 
| 9622 | 354 | |
| 355 | ||
| 14064 
35d36f43ba06
Moved strong normalization proof to StrongNorm.thy
 berghofe parents: 
13596diff
changeset | 356 | subsection {* Alternative induction rule for types *}
 | 
| 9622 | 357 | |
| 11935 | 358 | lemma type_induct [induct type]: | 
| 18241 | 359 | assumes | 
| 11945 | 360 | "(\<And>T. (\<And>T1 T2. T = T1 \<Rightarrow> T2 \<Longrightarrow> P T1) \<Longrightarrow> | 
| 18241 | 361 | (\<And>T1 T2. T = T1 \<Rightarrow> T2 \<Longrightarrow> P T2) \<Longrightarrow> P T)" | 
| 362 | shows "P T" | |
| 363 | proof (induct T) | |
| 364 | case Atom | |
| 365 | show ?case by (rule prems) simp_all | |
| 366 | next | |
| 367 | case Fun | |
| 20369 | 368 | show ?case by (rule prems) (insert Fun, simp_all) | 
| 11935 | 369 | qed | 
| 370 | ||
| 11638 | 371 | end |