# HG changeset patch # User urbanc # Date 1175012902 -7200 # Node ID 62c76754da32598e5bd12fa1b4e88247e7da7376 # Parent 7b9f346ac3662cc09e49ff502486b6159e6571d8 adapted to new nominal_inductive infrastructure diff -r 7b9f346ac366 -r 62c76754da32 src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Tue Mar 27 17:57:42 2007 +0200 +++ b/src/HOL/Nominal/Examples/Weakening.thy Tue Mar 27 18:28:22 2007 +0200 @@ -18,113 +18,43 @@ TVar "nat" | TArr "ty" "ty" (infix "\" 200) -lemma ty_perm[simp]: - fixes pi ::"name prm" - and T ::"ty" - shows "pi\T = T" -by (induct T rule: ty.induct_weak) - (simp_all add: perm_nat_def) +lemma ty_fresh: + fixes x::"name" + and T::"ty" + shows "x\T" +by (nominal_induct T rule: ty.induct) + (auto simp add: fresh_nat) text {* valid contexts *} + inductive2 valid :: "(name\ty) list \ bool" where v1[intro]: "valid []" | v2[intro]: "\valid \;a\\\\ valid ((a,\)#\)" -lemma eqvt_valid[eqvt]: - fixes pi:: "name prm" - assumes a: "valid \" - shows "valid (pi\\)" -using a -by (induct) (auto simp add: fresh_bij) +equivariance valid text{* typing judgements *} inductive2 typing :: "(name\ty) list\lam\ty\bool" (" _ \ _ : _ " [80,80,80] 80) where - t_Var[intro]: "\valid \; (a,T)\set \\\ \ \ Var a : T" + 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]: "\a\\;((a,T1)#\) \ t : T2\ \ \ \ Lam [a].t : T1\T2" - -lemma eqvt_typing[eqvt]: - fixes pi:: "name prm" - assumes a: "\ \ t : T" - shows "(pi\\) \ (pi\t) : (pi\T)" -using a -proof (induct) - case (t_Var \ a T) - have "valid (pi\\)" by (rule eqvt_valid) - moreover - have "(pi\(a,T))\(pi\set \)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) - ultimately show "(pi\\) \ (pi\Var a) : (pi\T)" - using typing.intros by (force simp add: set_eqvt) -next - case (t_Lam a \ T1 t T2) - moreover have "(pi\a)\(pi\\)" by (simp add: fresh_bij) - ultimately show "(pi\\) \ (pi\Lam [a].t) :(pi\T1\T2)" by force -qed (auto) - -text {* the strong induction principle needs to be derived manually *} + | t_Lam[intro]: "\x\\;((x,T1)#\) \ t : T2\ \ \ \ Lam [x].t : T1\T2" -lemma typing_induct[consumes 1, case_names t_Var t_App t_Lam]: - fixes P :: "'a::fs_name\(name\ty) list \ lam \ ty \bool" - and \ :: "(name\ty) list" - and t :: "lam" - and T :: "ty" - and x :: "'a::fs_name" - assumes a: "\ \ t : T" - and a1: "\\ a T x. \valid \; (a,T) \ set \\ \ P x \ (Var a) T" - and a2: "\\ T1 T2 t1 t2 x. \\z. P z \ t1 (T1\T2); \z. P z \ t2 T1\ - \ P x \ (App t1 t2) T2" - and a3: "\a \ T1 T2 t x. \a\x; a\\; \z. P z ((a,T1)#\) t T2\ - \ P x \ (Lam [a].t) (T1\T2)" - shows "P x \ t T" -proof - - from a have "\(pi::name prm) x. P x (pi\\) (pi\t) (pi\T)" - proof (induct) - case (t_Var \ a T) - have "valid \" by fact - then have "valid (pi\\)" by (rule eqvt) - moreover - have "(a,T)\set \" by fact - then have "pi\(a,T)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) - then have "(pi\a,T)\set (pi\\)" by (simp add: set_eqvt) - ultimately show "P x (pi\\) (pi\(Var a)) (pi\T)" using a1 by simp - next - case (t_App \ t1 T1 T2 t2) - thus "P x (pi\\) (pi\(App t1 t2)) (pi\T2)" using a2 - by (simp only: eqvt) (blast) - next - case (t_Lam a \ T1 t T2) - obtain c::"name" where fs: "c\(pi\a,pi\t,pi\\,x)" by (rule exists_fresh[OF fs_name1]) - let ?sw="[(pi\a,c)]" - let ?pi'="?sw@pi" - have f1: "a\\" by fact - have f2: "(pi\a)\(pi\\)" using f1 by (simp add: fresh_bij) - have f3: "c\?pi'\\" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp) - have ih1: "\x. P x (?pi'\((a,T1)#\)) (?pi'\t) (?pi'\T2)" by fact - then have "\x. P x ((c,T1)#(?pi'\\)) (?pi'\t) (?pi'\T2)" by (simp add: calc_atm) - then have "P x (?pi'\\) (Lam [c].(?pi'\t)) (T1\T2)" using a3 f3 fs by simp - then have "P x (?sw\pi\\) (?sw\(Lam [(pi\a)].(pi\t))) (T1\T2)" - by (simp del: append_Cons add: calc_atm pt_name2) - moreover have "(?sw\(pi\\)) = (pi\\)" - by (rule perm_fresh_fresh) (simp_all add: fs f2) - moreover have "(?sw\(Lam [(pi\a)].(pi\t))) = Lam [(pi\a)].(pi\t)" - by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh) - ultimately show "P x (pi\\) (pi\(Lam [a].t)) (pi\T1\T2)" by (simp) - qed - hence "P x (([]::name prm)\\) (([]::name prm)\t) (([]::name prm)\T)" by blast - thus "P x \ t T" by simp -qed +(* automatically deriving the strong induction principle *) +nominal_inductive typing + by (simp_all add: abs_fresh ty_fresh) text {* definition of a subcontext *} abbreviation - "sub" :: "(name\ty) list \ (name\ty) list \ bool" (" _ \ _ " [80,80] 80) where - "\1 \ \2 \ \a \. (a,\)\set \1 \ (a,\)\set \2" + "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" (" _ \ _ " [80,80] 80) +where + "\1 \ \2 \ \x T. (x,T)\set \1 \ (x,T)\set \2" -text {* now it comes: The Weakening Lemma *} +text {* Now it comes: The Weakening Lemma *} lemma weakening_version1: assumes a: "\1 \ t : T" @@ -132,9 +62,8 @@ and c: "\1 \ \2" shows "\2 \ t : T" using a b c -by (nominal_induct \1 t T avoiding: \2 rule: typing_induct) +by (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) (auto | atomize)+ -(* FIXME: meta-quantifiers seem to be not as "automatic" as object-quantifiers *) lemma weakening_version2: fixes \1::"(name\ty) list" @@ -145,25 +74,25 @@ and c: "\1 \ \2" shows "\2 \ t:T" using a b c -proof (nominal_induct \1 t T avoiding: \2 rule: typing_induct) - case (t_Var \1 a T) (* variable case *) +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 "(a,T)\ set \1" by fact - ultimately show "\2 \ Var a : T" by auto + have "(x,T)\ set \1" by fact + ultimately show "\2 \ Var x : T" by auto next - case (t_Lam a \1 T1 T2 t) (* lambda case *) - have vc: "a\\2" by fact (* variable convention *) - have ih: "\\3. \valid \3; ((a,T1)#\1) \ \3\ \ \3 \ t:T2" by fact + case (t_Lam x \1 T1 t T2) (* lambda case *) + have vc: "x\\2" by fact (* variable convention *) + have ih: "\\3. \valid \3; ((x,T1)#\1) \ \3\ \ \3 \ t:T2" by fact have "\1 \ \2" by fact - then have "((a,T1)#\1) \ ((a,T1)#\2)" by simp + then have "((x,T1)#\1) \ ((x,T1)#\2)" by simp moreover have "valid \2" by fact - then have "valid ((a,T1)#\2)" using vc by (simp add: v2) - ultimately have "((a,T1)#\2) \ t:T2" using ih by simp - with vc show "\2 \ (Lam [a].t) : T1\T2" by auto + then have "valid ((x,T1)#\2)" using vc by (simp add: v2) + 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_version3: @@ -172,17 +101,17 @@ and c: "\1 \ \2" shows "\2 \ t : T" using a b c -proof (nominal_induct \1 t T avoiding: \2 rule: typing_induct) - case (t_Lam a \1 T1 T2 t) (* lambda case *) - have vc: "a\\2" by fact (* variable convention *) - have ih: "\\3. \valid \3; ((a,T1)#\1) \ \3\ \ \3 \ t : T2" by fact +proof (nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) + case (t_Lam x \1 T1 t T2) (* lambda case *) + have vc: "x\\2" by fact (* variable convention *) + have ih: "\\3. \valid \3; ((x,T1)#\1) \ \3\ \ \3 \ t : T2" by fact have "\1 \ \2" by fact - then have "((a,T1)#\1) \ ((a,T1)#\2)" by simp + then have "((x,T1)#\1) \ ((x,T1)#\2)" by simp moreover have "valid \2" by fact - then have "valid ((a,T1)#\2)" using vc by (simp add: v2) - ultimately have "((a,T1)#\2) \ t : T2" using ih by simp - with vc show "\2 \ (Lam [a].t) : T1 \ T2" by auto + then have "valid ((x,T1)#\2)" using vc by (simp add: v2) + 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 and var case *) text{* The original induction principle for the typing relation @@ -194,24 +123,24 @@ shows "\2 \ t : T" using a b c proof (induct arbitrary: \2) - case (t_Var \1 a T) (* variable case *) + case (t_Var \1 x T) (* variable case *) have "\1 \ \2" by fact moreover have "valid \2" by fact moreover - have "(a,T) \ (set \1)" by fact - ultimately show "\2 \ Var a : T" by auto + have "(x,T) \ (set \1)" by fact + ultimately show "\2 \ Var x : T" by auto next - case (t_Lam a \1 T1 t T2) (* lambda case *) + case (t_Lam x \1 T1 t T2) (* lambda case *) (* all assumptions available in this case*) - have a0: "a\\1" by fact - have a1: "((a,T1)#\1) \ t : T2" by fact + have a0: "x\\1" by fact + have a1: "((x,T1)#\1) \ t : T2" by fact have a2: "\1 \ \2" by fact have a3: "valid \2" by fact - have ih: "\\3. \valid \3; ((a,T1)#\1) \ \3\ \ \3 \ t : T2" by fact - have "((a,T1)#\1) \ ((a,T1)#\2)" using a2 by simp + have ih: "\\3. \valid \3; ((x,T1)#\1) \ \3\ \ \3 \ t : T2" by fact + have "((x,T1)#\1) \ ((x,T1)#\2)" using a2 by simp moreover - have "valid ((a,T1)#\2)" using v2 (* fails *) + have "valid ((x,T1)#\2)" using v2 (* fails *) oops end \ No newline at end of file