# HG changeset patch # User wenzelm # Date 1004105894 -7200 # Node ID adef41692ab0b97a8ef549fee63558fc62942ea6 # Parent 1b540afebf4d10fc48bbb990b07c62fa0d6eaed5 tuned; diff -r 1b540afebf4d -r adef41692ab0 src/HOL/Lambda/InductTermi.thy --- a/src/HOL/Lambda/InductTermi.thy Fri Oct 26 14:22:33 2001 +0200 +++ b/src/HOL/Lambda/InductTermi.thy Fri Oct 26 16:18:14 2001 +0200 @@ -19,9 +19,9 @@ inductive IT intros - Var: "rs : lists IT ==> Var n $$ rs : IT" - Lambda: "r : IT ==> Abs r : IT" - Beta: "[| (r[s/0]) $$ ss : IT; s : IT |] ==> (Abs r $ s) $$ ss : IT" + Var [intro]: "rs : lists IT ==> Var n $$ rs : IT" + Lambda [intro]: "r : IT ==> Abs r : IT" + Beta [intro]: "(r[s/0]) $$ ss : IT ==> s : IT ==> (Abs r $ s) $$ ss : IT" subsection {* Every term in IT terminates *} @@ -65,7 +65,7 @@ declare Var_apps_neq_Abs_apps [THEN not_sym, simp] -lemma [simp, THEN not_sym, simp]: "Var n $$ ss ~= Abs r $ s $$ ts" +lemma [simp, THEN not_sym, simp]: "Var n $$ ss \ Abs r $ s $$ ts" apply (simp add: foldl_Cons [symmetric] del: foldl_Cons) done @@ -101,7 +101,7 @@ apply (rename_tac u ts) apply (case_tac ts) apply simp - apply (blast intro: IT.intros) + apply blast apply (rename_tac s ss) apply simp apply clarify diff -r 1b540afebf4d -r adef41692ab0 src/HOL/Lambda/Type.thy --- a/src/HOL/Lambda/Type.thy Fri Oct 26 14:22:33 2001 +0200 +++ b/src/HOL/Lambda/Type.thy Fri Oct 26 16:18:14 2001 +0200 @@ -15,6 +15,30 @@ *} +subsection {* Environments *} + +constdefs + shift :: "(nat \ 'a) \ nat \ 'a \ nat \ 'a" ("_\_:_\" [90, 0, 0] 91) + "e\i:a\ \ \j. if j < i then e j else if j = i then a else e (j - 1)" + +lemma shift_eq [simp]: "i = j \ (e\i:T\) j = T" + by (simp add: shift_def) + +lemma shift_gt [simp]: "j < i \ (e\i:T\) j = e j" + by (simp add: shift_def) + +lemma shift_lt [simp]: "i < j \ (e\i:T\) j = e (j - 1)" + by (simp add: shift_def) + +lemma shift_commute [simp]: "e\i:U\\0:T\ = e\0:T\\Suc i:U\" + apply (rule ext) + apply (case_tac x) + apply simp + apply (case_tac nat) + apply (simp_all add: shift_def) + done + + subsection {* Types and typing rules *} datatype type = @@ -30,14 +54,12 @@ "_typing" :: "(nat \ type) \ dB \ type \ bool" ("_ |- _ : _" [50, 50, 50] 50) "_typings" :: "(nat \ type) \ dB list \ type list \ bool" ("_ ||- _ : _" [50, 50, 50] 50) - syntax (symbols) "_typing" :: "(nat \ type) \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) syntax (latex) "_funs" :: "type list \ type \ type" (infixr "\" 200) "_typings" :: "(nat \ type) \ dB list \ type list \ bool" ("_ \ _ : _" [50, 50, 50] 50) - translations "Ts \ T" \ "foldr Fun Ts T" "env \ t : T" \ "(env, t, T) \ typing" @@ -46,7 +68,7 @@ inductive typing intros Var [intro!]: "env x = T \ env \ Var x : T" - Abs [intro!]: "nat_case T env \ t : U \ env \ Abs t : (T \ U)" + Abs [intro!]: "env\0:T\ \ t : U \ env \ Abs t : (T \ U)" App [intro!]: "env \ s : T \ U \ env \ t : T \ env \ (s \<^sub>\ t) : U" inductive_cases typing_elims [elim!]: @@ -61,15 +83,6 @@ [] \ False | T # Ts \ e \ t : T \ e \ ts : Ts)" -inductive_cases lists_elim [elim!]: - "x # xs \ lists S" - -declare IT.intros [intro!] - -constdefs - shift :: "(nat \ 'a) \ nat \ 'a \ nat \ 'a" ("_\_:_\" [90, 0, 0] 91) - "e\i:a\ \ \j. if j < i then e j else if j = i then a else e (j - 1)" - subsection {* Some examples *} @@ -173,45 +186,20 @@ "ts \ lists IT \ map (\t. lift t 0) ts \ lists IT" by (induct ts) auto -lemma shift_env [simp]: - "nat_case T - (\j. if j < i then e j else if j = i then Ua else e (j - 1)) = - (\j. if j < Suc i then nat_case T e j else if j = Suc i then Ua - else nat_case T e (j - 1))" - apply (rule ext) - apply (case_tac j) - apply simp - apply (case_tac nat) - apply simp_all - done - -lemma lift_type': - "e \ t : T \ e\i:U\ \ lift t i : T" +lemma lift_type [intro!]: "e \ t : T \ e\i:U\ \ lift t i : T" proof - assume "e \ t : T" thus "\i U. e\i:U\ \ lift t i : T" - by induct (auto simp add: shift_def) + by induct auto qed -lemma lift_type [intro!]: - "e \ t : T \ nat_case U e \ lift t 0 : T" - apply (subgoal_tac "nat_case U e = e\0:U\") - apply (erule ssubst) - apply (erule lift_type') - apply (rule ext) - apply (case_tac x) - apply (simp_all add: shift_def) - done - lemma lift_typings [rule_format]: "\Ts. e \ ts : Ts \ (e\i:U\) \ (map (\t. lift t i) ts) : Ts" apply (induct_tac ts) apply simp apply (intro strip) apply (case_tac Ts) - apply simp_all - apply (rule lift_type') - apply (erule conjunct1) + apply auto done @@ -219,22 +207,11 @@ lemma subst_lemma [rule_format]: "e \ t : T \ \e' i U u. e' \ u : U \ e = e'\i:U\ \ e' \ t[u/i] : T" - apply (unfold shift_def) apply (erule typing.induct) apply (intro strip) - apply (case_tac "x = i") - apply simp - apply (frule linorder_neq_iff [THEN iffD1]) - apply (erule disjE) - apply simp - apply (rule typing.Var) - apply assumption - apply (frule order_less_not_sym) - apply (simp only: subst_gt split: split_if add: if_False) - apply (rule typing.Var) - apply assumption - apply fastsimp - apply auto + apply (rule_tac x = x and y = i in linorder_cases) + apply auto + apply blast done lemma substs_lemma [rule_format]: @@ -272,7 +249,7 @@ apply assumption apply (rule ext) apply (case_tac x) - apply (auto simp add: shift_def) + apply auto done @@ -392,7 +369,7 @@ and argT: "e\i:T\ \ a : T''" by cases simp_all from varT True have T: "T = T'' \ Ts \ T'" - by cases (auto simp add: shift_def) + by cases auto with uT have uT': "e \ u : T'' \ Ts \ T'" by simp from Var have SI: "?R a" by cases (simp_all add: Cons) from T have "(Var 0 \<^sub>\\<^sub>\ map (\t. lift t 0) @@ -405,14 +382,14 @@ show "e\0:T''\ \ lift u 0 \<^sub>\ Var 0 : Ts \ T'" proof (rule typing.App) show "e\0:T''\ \ lift u 0 : T'' \ Ts \ T'" - by (rule lift_type') (rule uT') + by (rule lift_type) (rule uT') show "e\0:T''\ \ Var 0 : T''" - by (rule typing.Var) (simp add: shift_def) + by (rule typing.Var) simp qed from argT uIT uT show "a[u/i] \ IT" by (rule SI[rule_format]) from argT uT show "e \ a[u/i] : T''" - by (rule subst_lemma) (simp add: shift_def) + by (rule subst_lemma) simp qed thus "u \<^sub>\ a[u/i] \ IT" by simp from Var have "as \ lists {t. ?R t}" @@ -438,7 +415,7 @@ qed thus "Var 0 \<^sub>\\<^sub>\ ?ls as \ IT" by (rule IT.Var) have "e\0:Ts \ T'\ \ Var 0 : Ts \ T'" - by (rule typing.Var) (simp add: shift_def) + by (rule typing.Var) simp moreover from uT argsT have "e \ map (\t. t[u/i]) as : Ts" by (rule substs_lemma) hence "(e\0:Ts \ T'\) \ ?ls as : Ts" @@ -481,7 +458,7 @@ assume "e\i:T\ \ Abs r : T'" and "\e T' u i. PROP ?Q r e T' u i T" with uIT uT show "Abs r[u/i] \ IT" - by (fastsimp simp add: shift_def) + by fastsimp next case (Beta r a as) assume T: "e\i:T\ \ Abs r \<^sub>\ a \<^sub>\\<^sub>\ as : T'" @@ -526,12 +503,12 @@ have "lift t 0 \ IT" by (rule lift_IT) hence "[lift t 0] \ lists IT" by (rule lists.Cons) (rule lists.Nil) hence "Var 0 \<^sub>\\<^sub>\ [lift t 0] \ IT" by (rule IT.Var) - also have "(Var 0 \<^sub>\\<^sub>\ [lift t 0]) = (Var 0 \<^sub>\ lift t 0)" by simp + also have "Var 0 \<^sub>\\<^sub>\ [lift t 0] = Var 0 \<^sub>\ lift t 0" by simp finally show "\ \ IT" . have "e\0:T \ U\ \ Var 0 : T \ U" - by (rule typing.Var) (simp add: shift_def) + by (rule typing.Var) simp moreover have "e\0:T \ U\ \ lift t 0 : T" - by (rule lift_type') + by (rule lift_type) ultimately show "e\0:T \ U\ \ Var 0 \<^sub>\ lift t 0 : U" by (rule typing.App) qed