# HG changeset patch # User Christian Urban # Date 1295645024 -3600 # Node ID abb6c3ac2df2080527e94b636a9491da6f922ebb # Parent f325eefe803e2b2e72892f59b04bc2be7be557a9 separated type preservation and progress into a separate file diff -r f325eefe803e -r abb6c3ac2df2 Tutorial/Tutorial3.thy --- a/Tutorial/Tutorial3.thy Fri Jan 21 22:02:34 2011 +0100 +++ b/Tutorial/Tutorial3.thy Fri Jan 21 22:23:44 2011 +0100 @@ -6,12 +6,12 @@ text {* Barendregt's proof needs in the variable case a case distinction. - One way to do this in Isar is to use blocks. A block is some sequent - or reasoning steps enclosed in curly braces + One way to do this in Isar is to use blocks. A block consist of some + assumptions and reasoning steps enclosed in curly braces, like { \ - have "statement" + have "last_statement_in_the_block" } Such a block can contain local assumptions like @@ -25,7 +25,7 @@ Where "C" is the last have-statement in this block. The behaviour of such a block to the 'outside' is the implication - \A; B\ \ "C" + A \ B \ C Now if we want to prove a property "smth" using the case-distinctions P1, P2 and P3 then we can use the following reasoning: @@ -52,17 +52,12 @@ P2 \ smth P3 \ smth - If we know that P1, P2 and P3 cover all the cases, that is P1 \ P2 \ P3 is - true, then we have 'ultimately' established the property "smth" + If we know that P1, P2 and P3 cover all the cases, that is P1 \ P2 \ P3 + holds, then we have 'ultimately' established the property "smth" *} -section {* EXERCISE 7 *} - -text {* - Fill in the cases 1.2 and 1.3 and the equational reasoning - in the lambda-case. -*} +subsection {* Two preliminary facts *} lemma forget: shows "atom x \ t \ t[x ::= s] = t" @@ -78,6 +73,14 @@ (auto simp add: lam.fresh fresh_at_base) + +section {* EXERCISE 7 *} + +text {* + Fill in the cases 1.2 and 1.3 and the equational reasoning + in the lambda-case. +*} + lemma assumes a: "x \ y" and b: "atom x \ L" @@ -111,12 +114,12 @@ qed next case (Lam z M1) -- {* case 2: lambdas *} - have ih: "\x \ y; atom x \ L\ \ M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact + have ih: "\x \ y; atom x \ L\ \ M1[x ::= N][y ::= L] = M1[y ::= L][x ::= N[y ::= L]]" by fact have a1: "x \ y" by fact have a2: "atom x \ L" by fact - have fs: "atom z \ x" "atom z \ y" "atom z \ N" "atom z \ L" by fact+ + have fs: "atom z \ x" "atom z \ y" "atom z \ N" "atom z \ L" by fact+ -- {* !! *} then have b: "atom z \ N[y::=L]" by (simp add: fresh_fact) - show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") + show "(Lam [z].M1)[x ::= N][y ::= L] = (Lam [z].M1)[y ::= L][x ::= N[y ::= L]]" (is "?LHS=?RHS") proof - have "?LHS = \" sorry @@ -130,7 +133,7 @@ text {* Again the strong induction principle enables Isabelle to find - the proof of the substitution lemma automatically. + the proof of the substitution lemma completely automatically. *} lemma substitution_lemma_version: diff -r f325eefe803e -r abb6c3ac2df2 Tutorial/Tutorial4.thy --- a/Tutorial/Tutorial4.thy Fri Jan 21 22:02:34 2011 +0100 +++ b/Tutorial/Tutorial4.thy Fri Jan 21 22:23:44 2011 +0100 @@ -1,5 +1,6 @@ + theory Tutorial4 -imports Tutorial1 Tutorial2 +imports Tutorial1 Tutorial2 Tutorial3 begin section {* The CBV Reduction Relation (Small-Step Semantics) *} @@ -163,6 +164,7 @@ shows "t \ t" using a by (induct) (auto) + lemma e_App_elim: assumes a: "App t1 t2 \ v" obtains x t v' where "t1 \ Lam [x].t" "t2 \ v'" "t[x::=v'] \ v" @@ -234,137 +236,8 @@ then show "t1 \ t2" using b cbvs_implies_eval by simp qed -lemma valid_elim: - assumes a: "valid ((x, T) # \)" - shows "atom x \ \ \ valid \" -using a by (cases) (auto) - -lemma valid_insert: - assumes a: "valid (\ @ [(x, T)] @ \)" - shows "valid (\ @ \)" -using a -by (induct \) - (auto simp add: fresh_append fresh_Cons dest!: valid_elim) - -lemma fresh_list: - shows "atom y \ xs = (\x \ set xs. atom y \ x)" -by (induct xs) (simp_all add: fresh_Nil fresh_Cons) - -lemma context_unique: - assumes a1: "valid \" - and a2: "(x, T) \ set \" - and a3: "(x, U) \ set \" - shows "T = U" -using a1 a2 a3 -by (induct) (auto simp add: fresh_list fresh_Pair fresh_at_base) - -lemma type_substitution_aux: - assumes a: "\ @ [(x, T')] @ \ \ e : T" - and b: "\ \ e' : T'" - shows "\ @ \ \ e[x ::= e'] : T" -using a b -proof (nominal_induct \'\"\ @ [(x, T')] @ \" e T avoiding: x e' \ rule: typing.strong_induct) - case (t_Var y T x e' \) - have a1: "valid (\ @ [(x, T')] @ \)" by fact - have a2: "(y,T) \ set (\ @ [(x, T')] @ \)" by fact - have a3: "\ \ e' : T'" by fact - from a1 have a4: "valid (\ @ \)" by (rule valid_insert) - { assume eq: "x = y" - from a1 a2 have "T = T'" using eq by (auto intro: context_unique) - with a3 have "\ @ \ \ Var y[x ::= e'] : T" using eq a4 by (auto intro: weakening) - } - moreover - { assume ineq: "x \ y" - from a2 have "(y, T) \ set (\ @ \)" using ineq by simp - then have "\ @ \ \ Var y[x ::= e'] : T" using ineq a4 by auto - } - ultimately show "\ @ \ \ Var y[x::=e'] : T" by blast -qed (force simp add: fresh_append fresh_Cons)+ - -corollary type_substitution: - assumes a: "(x, T') # \ \ e : T" - and b: "\ \ e' : T'" - shows "\ \ e[x ::= e'] : T" -using a b type_substitution_aux[where \="[]"] -by auto - -lemma t_App_elim: - assumes a: "\ \ App t1 t2 : T" - obtains T' where "\ \ t1 : T' \ T" "\ \ t2 : T'" -using a -by (cases) (auto simp add: lam.eq_iff lam.distinct) -text {* we have not yet generated strong elimination rules *} -lemma t_Lam_elim: - assumes ty: "\ \ Lam [x].t : T" - and fc: "atom x \ \" - obtains T1 T2 where "T = T1 \ T2" "(x, T1) # \ \ t : T2" -using ty fc -apply(cases) -apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff) -apply(auto simp add: Abs1_eq_iff) -apply(rotate_tac 3) -apply(drule_tac p="(x \ xa)" in permute_boolI) -apply(perm_simp) -apply(auto simp add: flip_def swap_fresh_fresh ty_fresh) -done -theorem cbv_type_preservation: - assumes a: "t \cbv t'" - and b: "\ \ t : T" - shows "\ \ t' : T" -using a b -by (nominal_induct avoiding: \ T rule: cbv.strong_induct) - (auto elim!: t_Lam_elim t_App_elim simp add: type_substitution ty.eq_iff) - -corollary cbvs_type_preservation: - assumes a: "t \cbv* t'" - and b: "\ \ t : T" - shows "\ \ t' : T" -using a b -by (induct) (auto intro: cbv_type_preservation) - -text {* - The type-preservation property for the machine and - evaluation relation. -*} - -theorem machine_type_preservation: - assumes a: " \* " - and b: "\ \ t : T" - shows "\ \ t' : T" -proof - - have "t \cbv* t'" using a machines_implies_cbvs by simp - then show "\ \ t' : T" using b cbvs_type_preservation by simp -qed - -theorem eval_type_preservation: - assumes a: "t \ t'" - and b: "\ \ t : T" - shows "\ \ t' : T" -proof - - have " \* " using a eval_implies_machines by simp - then show "\ \ t' : T" using b machine_type_preservation by simp -qed - -text {* The Progress Property *} - -lemma canonical_tArr: - assumes a: "[] \ t : T1 \ T2" - and b: "val t" - obtains x t' where "t = Lam [x].t'" -using b a by (induct) (auto) - -theorem progress: - assumes a: "[] \ t : T" - shows "(\t'. t \cbv t') \ (val t)" -using a -by (induct \\"[]::ty_ctx" t T) - (auto elim: canonical_tArr) - -text {* - Done! -*} end diff -r f325eefe803e -r abb6c3ac2df2 Tutorial/Tutorial5.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Tutorial/Tutorial5.thy Fri Jan 21 22:23:44 2011 +0100 @@ -0,0 +1,142 @@ +theory Tutorial5 +imports Tutorial4 +begin + + +section {* Type Preservation (fixme separate file) *} + + +lemma valid_elim: + assumes a: "valid ((x, T) # \)" + shows "atom x \ \ \ valid \" +using a by (cases) (auto) + +lemma valid_insert: + assumes a: "valid (\ @ [(x, T)] @ \)" + shows "valid (\ @ \)" +using a +by (induct \) + (auto simp add: fresh_append fresh_Cons dest!: valid_elim) + +lemma fresh_list: + shows "atom y \ xs = (\x \ set xs. atom y \ x)" +by (induct xs) (simp_all add: fresh_Nil fresh_Cons) + +lemma context_unique: + assumes a1: "valid \" + and a2: "(x, T) \ set \" + and a3: "(x, U) \ set \" + shows "T = U" +using a1 a2 a3 +by (induct) (auto simp add: fresh_list fresh_Pair fresh_at_base) + +lemma type_substitution_aux: + assumes a: "\ @ [(x, T')] @ \ \ e : T" + and b: "\ \ e' : T'" + shows "\ @ \ \ e[x ::= e'] : T" +using a b +proof (nominal_induct \'\"\ @ [(x, T')] @ \" e T avoiding: x e' \ rule: typing.strong_induct) + case (t_Var y T x e' \) + have a1: "valid (\ @ [(x, T')] @ \)" by fact + have a2: "(y,T) \ set (\ @ [(x, T')] @ \)" by fact + have a3: "\ \ e' : T'" by fact + from a1 have a4: "valid (\ @ \)" by (rule valid_insert) + { assume eq: "x = y" + from a1 a2 have "T = T'" using eq by (auto intro: context_unique) + with a3 have "\ @ \ \ Var y[x ::= e'] : T" using eq a4 by (auto intro: weakening) + } + moreover + { assume ineq: "x \ y" + from a2 have "(y, T) \ set (\ @ \)" using ineq by simp + then have "\ @ \ \ Var y[x ::= e'] : T" using ineq a4 by auto + } + ultimately show "\ @ \ \ Var y[x::=e'] : T" by blast +qed (force simp add: fresh_append fresh_Cons)+ + +corollary type_substitution: + assumes a: "(x, T') # \ \ e : T" + and b: "\ \ e' : T'" + shows "\ \ e[x ::= e'] : T" +using a b type_substitution_aux[where \="[]"] +by auto + +lemma t_App_elim: + assumes a: "\ \ App t1 t2 : T" + obtains T' where "\ \ t1 : T' \ T" "\ \ t2 : T'" +using a +by (cases) (auto simp add: lam.eq_iff lam.distinct) + +text {* we have not yet generated strong elimination rules *} +lemma t_Lam_elim: + assumes ty: "\ \ Lam [x].t : T" + and fc: "atom x \ \" + obtains T1 T2 where "T = T1 \ T2" "(x, T1) # \ \ t : T2" +using ty fc +apply(cases) +apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff) +apply(auto simp add: Abs1_eq_iff) +apply(rotate_tac 3) +apply(drule_tac p="(x \ xa)" in permute_boolI) +apply(perm_simp) +apply(auto simp add: flip_def swap_fresh_fresh ty_fresh) +done + +theorem cbv_type_preservation: + assumes a: "t \cbv t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +by (nominal_induct avoiding: \ T rule: cbv.strong_induct) + (auto elim!: t_Lam_elim t_App_elim simp add: type_substitution ty.eq_iff) + +corollary cbvs_type_preservation: + assumes a: "t \cbv* t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +using a b +by (induct) (auto intro: cbv_type_preservation) + +text {* + The type-preservation property for the machine and + evaluation relation. +*} + +theorem machine_type_preservation: + assumes a: " \* " + and b: "\ \ t : T" + shows "\ \ t' : T" +proof - + have "t \cbv* t'" using a machines_implies_cbvs by simp + then show "\ \ t' : T" using b cbvs_type_preservation by simp +qed + +theorem eval_type_preservation: + assumes a: "t \ t'" + and b: "\ \ t : T" + shows "\ \ t' : T" +proof - + have " \* " using a eval_implies_machines by simp + then show "\ \ t' : T" using b machine_type_preservation by simp +qed + +text {* The Progress Property *} + +lemma canonical_tArr: + assumes a: "[] \ t : T1 \ T2" + and b: "val t" + obtains x t' where "t = Lam [x].t'" +using b a by (induct) (auto) + +theorem progress: + assumes a: "[] \ t : T" + shows "(\t'. t \cbv t') \ (val t)" +using a +by (induct \\"[]::ty_ctx" t T) + (auto elim: canonical_tArr) + +text {* + Done! Congratulations! +*} + +end +