# HG changeset patch # User Christian Urban # Date 1294320520 0 # Node ID 0579d3a48304bb5faa40f2702fc18870768d03df # Parent f338b455314c5b53c27370233fe3c8a5d1c7a2ab cleaned up weakening proof and added a version with finit sets diff -r f338b455314c -r 0579d3a48304 Nominal/Ex/Typing.thy --- a/Nominal/Ex/Typing.thy Thu Jan 06 13:28:19 2011 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,149 +0,0 @@ -theory Lambda -imports "../Nominal2" -begin - - -atom_decl name - -nominal_datatype lam = - Var "name" -| App "lam" "lam" -| Lam x::"name" l::"lam" bind x in l - -thm lam.distinct -thm lam.induct -thm lam.exhaust lam.strong_exhaust -thm lam.fv_defs -thm lam.bn_defs -thm lam.perm_simps -thm lam.eq_iff -thm lam.fv_bn_eqvt -thm lam.size_eqvt - - -section {* Typing *} - -nominal_datatype ty = - TVar string -| TFun ty ty ("_ \ _") - -lemma ty_fresh: - fixes x::"name" - and T::"ty" - shows "atom x \ T" -apply (nominal_induct T rule: ty.strong_induct) -apply (simp_all add: ty.fresh pure_fresh) -done - -inductive - valid :: "(name \ ty) list \ bool" -where - v_Nil[intro]: "valid []" -| v_Cons[intro]: "\atom x \ Gamma; valid Gamma\ \ valid ((x, T)#Gamma)" - -inductive - typing :: "(name\ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) -where - t_Var[intro]: "\valid \; (x, T) \ set \\ \ \ \ Var x : T" - | t_App[intro]: "\\ \ t1 : T1 \ T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" - | t_Lam[intro]: "\atom x \ \; (x, T1) # \ \ t : T2\ \ \ \ Lam x t : T1 \ T2" - -thm typing.intros -thm typing.induct - -equivariance valid -equivariance typing - -nominal_inductive typing - avoids t_Lam: "x" - by (simp_all add: fresh_star_def ty_fresh lam.fresh) - - -thm typing.strong_induct - -abbreviation - "sub_context" :: "(name \ ty) list \ (name \ ty) list \ bool" ("_ \ _" [60,60] 60) -where - "\1 \ \2 \ \x T. (x, T) \ set \1 \ (x, T) \ set \2" - -text {* Now it comes: The Weakening Lemma *} - -text {* - The first version is, after setting up the induction, - completely automatic except for use of atomize. *} - -lemma weakening_version2: - fixes \1 \2::"(name \ ty) list" - and t ::"lam" - and \ ::"ty" - assumes a: "\1 \ t : T" - and b: "valid \2" - and c: "\1 \ \2" - shows "\2 \ t : T" -using a b c -proof (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) - case (t_Var \1 x T) (* variable case *) - have "\1 \ \2" by fact - moreover - have "valid \2" by fact - moreover - have "(x,T)\ set \1" by fact - ultimately show "\2 \ Var x : T" by auto -next - case (t_Lam x \1 T1 t T2) (* lambda case *) - have vc: "atom x \ \2" by fact (* variable convention *) - have ih: "\valid ((x, T1) # \2); (x, T1) # \1 \ (x, T1) # \2\ \ (x, T1) # \2 \ t : T2" by fact - have "\1 \ \2" by fact - then have "(x, T1) # \1 \ (x, T1) # \2" by simp - moreover - have "valid \2" by fact - then have "valid ((x, T1) # \2)" using vc by (simp add: v_Cons) - ultimately have "(x, T1) # \2 \ t : T2" using ih by simp - with vc show "\2 \ Lam x t : T1 \ T2" by auto -qed (auto) (* app case *) - -lemma weakening_version1: - fixes \1 \2::"(name \ ty) list" - assumes a: "\1 \ t : T" - and b: "valid \2" - and c: "\1 \ \2" - shows "\2 \ t : T" -using a b c -apply (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) -apply (auto | atomize)+ -done - - -inductive - Acc :: "('a::pt \ 'a \ bool) \ 'a \ bool" -where - AccI: "(\y. R y x \ Acc R y) \ Acc R x" - - -lemma Acc_eqvt [eqvt]: - fixes p::"perm" - assumes a: "Acc R x" - shows "Acc (p \ R) (p \ x)" -using a -apply(induct) -apply(rule AccI) -apply(rotate_tac 1) -apply(drule_tac x="-p \ y" in meta_spec) -apply(simp) -apply(drule meta_mp) -apply(rule_tac p="p" in permute_boolE) -apply(perm_simp add: permute_minus_cancel) -apply(assumption) -apply(assumption) -done - - -nominal_inductive Acc . - -thm Acc.strong_induct - - -end - - - diff -r f338b455314c -r 0579d3a48304 Nominal/Ex/Weakening.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/Weakening.thy Thu Jan 06 13:28:40 2011 +0000 @@ -0,0 +1,158 @@ +theory Lambda +imports "../Nominal2" +begin + +section {* The Weakening property in the simply-typed lambda-calculus *} + +atom_decl name + +nominal_datatype lam = + Var "name" +| App "lam" "lam" +| Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100,100] 100) + +text {* Typing *} + +nominal_datatype ty = + TVar string +| TFun ty ty ("_ \ _" [100,100] 100) + +lemma fresh_ty: + fixes x::"name" + and T::"ty" + shows "atom x \ T" + by (nominal_induct T rule: ty.strong_induct) + (simp_all add: ty.fresh pure_fresh) + +text {* Valid typing contexts *} + +inductive + valid :: "(name \ ty) list \ bool" +where + v_Nil[intro]: "valid []" +| v_Cons[intro]: "\atom x \ Gamma; valid Gamma\ \ valid ((x, T) # Gamma)" + +equivariance valid + +text {* Typing judgements *} + +inductive + typing :: "(name \ ty) list \ lam \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) +where + t_Var[intro]: "\valid \; (x, T) \ set \\ \ \ \ Var x : T" + | t_App[intro]: "\\ \ t1 : T1 \ T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" + | t_Lam[intro]: "\atom x \ \; (x, T1) # \ \ t : T2\ \ \ \ Lam [x]. t : T1 \ T2" + +equivariance typing + +text {* Strong induction principle for typing judgements *} + +nominal_inductive typing + avoids t_Lam: "x" + by (simp_all add: fresh_star_def fresh_ty lam.fresh) + + +abbreviation + "sub_context" :: "(name \ ty) list \ (name \ ty) list \ bool" (infixr "\" 60) +where + "\1 \ \2 \ \x T. (x, T) \ set \1 \ (x, T) \ set \2" + + +text {* The proof *} + +lemma weakening_version1: + fixes \1 \2::"(name \ ty) list" + and t ::"lam" + and \ ::"ty" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +proof (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) + case (t_Var \1 x T) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(x, T) \ set \1" by fact + ultimately show "\2 \ Var x : T" by auto +next + case (t_Lam x \1 T1 t T2) (* lambda case *) + have vc: "atom x \ \2" by fact (* variable convention *) + have ih: "\valid ((x, T1) # \2); (x, T1) # \1 \ (x, T1) # \2\ \ (x, T1) # \2 \ t : T2" by fact + have "\1 \ \2" by fact + then have "(x, T1) # \1 \ (x, T1) # \2" by simp + moreover + have "valid \2" by fact + then have "valid ((x, T1) # \2)" using vc by auto + ultimately have "(x, T1) # \2 \ t : T2" using ih by simp + with vc show "\2 \ Lam [x]. t : T1 \ T2" by auto +qed (auto) (* app case *) + +lemma weakening_version2: + fixes \1 \2::"(name \ ty) list" + assumes a: "\1 \ t : T" + and b: "valid \2" + and c: "\1 \ \2" + shows "\2 \ t : T" +using a b c +by (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) + (auto | atomize)+ + + +text {* A version with finite sets as typing contexts *} + +inductive + valid2 :: "(name \ ty) fset \ bool" +where + v2_Empty[intro]: "valid2 {||}" +| v2_Insert[intro]: "\(x, T) |\| Gamma; valid2 Gamma\ \ valid2 (insert_fset (x, T) Gamma)" + +equivariance valid2 + +inductive + typing2 :: "(name \ ty) fset \ lam \ ty \ bool" ("_ 2\ _ : _" [60,60,60] 60) +where + t2_Var[intro]: "\valid2 \; (x, T) |\| \\ \ \ 2\ Var x : T" + | t2_App[intro]: "\\ 2\ t1 : T1 \ T2; \ 2\ t2 : T1\ \ \ 2\ App t1 t2 : T2" + | t2_Lam[intro]: "\atom x \ \; insert_fset (x, T1) \ 2\ t : T2\ \ \ 2\ Lam [x]. t : T1 \ T2" + +equivariance typing2 + +nominal_inductive typing2 + avoids t2_Lam: "x" + by (simp_all add: fresh_star_def fresh_ty lam.fresh) + +lemma weakening_version3: + fixes \::"(name \ ty) fset" + assumes a: "\ 2\ t : T" + and b: "(x, T') |\| \" + shows "{|(x, T')|} |\| \ 2\ t : T" +using a b +apply(nominal_induct \ t T avoiding: x rule: typing2.strong_induct) +apply(auto)[2] +apply(rule t2_Lam) +apply(simp add: fresh_insert_fset fresh_Pair fresh_ty) +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule meta_mp) +apply(simp add: fresh_at_base) +apply(simp add: insert_fset_left_comm) +done + +lemma weakening_version4: + fixes \::"(name \ ty) fset" + assumes a: "\ 2\ t : T" + and b: "(x, T') |\| \" + shows "{|(x, T')|} |\| \ 2\ t : T" +using a b +apply(induct \ t T arbitrary: x) +apply(auto)[2] +apply(rule t2_Lam) +oops + +end + + + diff -r f338b455314c -r 0579d3a48304 Nominal/Nominal2.thy --- a/Nominal/Nominal2.thy Thu Jan 06 13:28:19 2011 +0000 +++ b/Nominal/Nominal2.thy Thu Jan 06 13:28:40 2011 +0000 @@ -154,8 +154,6 @@ val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env bclauses - val _ = tracing ("raw_dt info:\n" ^ @{make_string} raw_dts) - val (raw_dt_full_names, thy1) = Datatype.add_datatype Datatype.default_config raw_dt_names raw_dts thy