# HG changeset patch # User urbanc # Date 1161205611 -7200 # Node ID ec5531061ed64276b6a59083fd380e201ca17acd # Parent c49467a9c1e147bd51334f0b5af8446ac269750d adapted to Stefan's new inductive package and cleaning up diff -r c49467a9c1e1 -r ec5531061ed6 src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Wed Oct 18 16:13:03 2006 +0200 +++ b/src/HOL/Nominal/Examples/Weakening.thy Wed Oct 18 23:06:51 2006 +0200 @@ -1,37 +1,33 @@ (* $Id$ *) theory Weakening -imports "../Nominal" +imports "Nominal" begin -(* WEAKENING EXAMPLE*) - -section {* Simply-Typed Lambda-Calculus *} -(*======================================*) +section {* Weakening Example for the Simply-Typed Lambda-Calculus *} +(*================================================================*) atom_decl name -nominal_datatype lam = Var "name" - | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) +nominal_datatype lam = + Var "name" + | App "lam" "lam" + | Lam "\name\lam" ("Lam [_]._" [100,100] 100) nominal_datatype ty = TVar "nat" | TArr "ty" "ty" (infix "\" 200) -lemma perm_ty[simp]: +lemma [simp]: fixes pi ::"name prm" and \ ::"ty" shows "pi\\ = \" -by (induct \ rule: ty.induct_weak, simp_all add: perm_nat_def) +by (induct \ rule: ty.induct_weak) + (simp_all add: perm_nat_def) -(* valid contexts *) -consts - ctxts :: "((name\ty) list) set" +text {* valid contexts *} +inductive2 valid :: "(name\ty) list \ bool" -translations - "valid \" \ "\ \ ctxts" -inductive ctxts intros v1[intro]: "valid []" v2[intro]: "\valid \;a\\\\ valid ((a,\)#\)" @@ -41,108 +37,100 @@ assumes a: "valid \" shows "valid (pi\\)" using a -apply(induct) -apply(auto simp add: fresh_bij) -done +by (induct) + (auto simp add: fresh_bij) -(* typing judgements *) -consts - typing :: "(((name\ty) list)\lam\ty) set" -syntax - "_typing_judge" :: "(name\ty) list\lam\ty\bool" (" _ \ _ : _ " [80,80,80] 80) -translations - "\ \ t : \" \ "(\,t,\) \ typing" - -inductive typing +text{* typing judgements *} +inductive2 + typing :: "(name\ty) list\lam\ty\bool" (" _ \ _ : _ " [80,80,80] 80) intros -t1[intro]: "\valid \; (a,\)\set \\\ \ \ Var a : \" -t2[intro]: "\\ \ t1 : \\\; \ \ t2 : \\\ \ \ App t1 t2 : \" -t3[intro]: "\a\\;((a,\)#\) \ t : \\ \ \ \ Lam [a].t : \\\" +t_Var[intro]: "\valid \; (a,\)\set \\\ \ \ Var a : \" +t_App[intro]: "\\ \ t1 : \\\; \ \ t2 : \\\ \ \ App t1 t2 : \" +t_Lam[intro]: "\a\\;((a,\)#\) \ t : \\ \ \ \ Lam [a].t : \\\" lemma eqvt_typing: - fixes \ :: "(name\ty) list" - and t :: "lam" - and \ :: "ty" - and pi:: "name prm" + fixes pi:: "name prm" assumes a: "\ \ t : \" shows "(pi\\) \ (pi\t) : \" using a proof (induct) - case (t1 \ \ a) + case (t_Var \ a \) have "valid (pi\\)" by (rule eqvt_valid) moreover have "(pi\(a,\))\((pi::name prm)\set \)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) ultimately show "(pi\\) \ ((pi::name prm)\Var a) : \" using typing.intros by (force simp add: pt_list_set_pi[OF pt_name_inst, symmetric]) next - case (t3 \ \ \ a t) + case (t_Lam a \ \ t \) moreover have "(pi\a)\(pi\\)" by (simp add: fresh_bij) ultimately show "(pi\\) \ (pi\Lam [a].t) :\\\" by force qed (auto) -lemma typing_induct[consumes 1, case_names t1 t2 t3]: +text {* the strong induction principle needs to be derived manually *} + +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 \ :: "ty" and x :: "'a::fs_name" assumes a: "\ \ t : \" - and a1: "\\ (a::name) \ x. valid \ \ (a,\) \ set \ \ P x \ (Var a) \" + and a1: "\\ a \ x. \valid \; (a,\) \ set \\ \ P x \ (Var a) \" and a2: "\\ \ \ t1 t2 x. - \ \ t1 : \\\ \ (\z. P z \ t1 (\\\)) \ \ \ t2 : \ \ (\z. P z \ t2 \) + \\ \ t1 : \\\; (\z. P z \ t1 (\\\)); \ \ t2 : \; (\z. P z \ t2 \)\ \ P x \ (App t1 t2) \" - and a3: "\a \ \ \ t x. a\x \ a\\ \ ((a,\) # \) \ t : \ \ (\z. P z ((a,\)#\) t \) + and a3: "\a \ \ \ t x. \a\x; a\\; ((a,\)#\) \ t : \; (\z. P z ((a,\)#\) t \)\ \ P x \ (Lam [a].t) (\\\)" shows "P x \ t \" proof - from a have "\(pi::name prm) x. P x (pi\\) (pi\t) \" proof (induct) - case (t1 \ \ a) - have j1: "valid \" by fact - have j2: "(a,\)\set \" by fact - from j1 have j3: "valid (pi\\)" by (rule eqvt_valid) - from j2 have "pi\(a,\)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) - hence j4: "(pi\a,\)\set (pi\\)" by (simp add: pt_list_set_pi[OF pt_name_inst]) - show "P x (pi\\) (pi\(Var a)) \" using a1 j3 j4 by simp + case (t_Var \ a \) + have "valid \" by fact + then have "valid (pi\\)" by (rule eqvt_valid) + moreover + have "(a,\)\set \" by fact + then have "pi\(a,\)\pi\(set \)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + then have "(pi\a,\)\set (pi\\)" by (simp add: pt_list_set_pi[OF pt_name_inst]) + ultimately show "P x (pi\\) (pi\(Var a)) \" using a1 by simp next - case (t2 \ \ \ t1 t2) - thus ?case using a2 by (simp, blast intro: eqvt_typing) + case (t_App \ t1 \ \ t2) + thus "P x (pi\\) (pi\(App t1 t2)) \" using a2 by (simp, blast intro: eqvt_typing) next - case (t3 \ \ \ a t) + case (t_Lam a \ \ t \) have k1: "a\\" by fact have k2: "((a,\)#\)\t:\" by fact have k3: "\(pi::name prm) (x::'a::fs_name). P x (pi \((a,\)#\)) (pi\t) \" by fact have f: "\c::name. c\(pi\a,pi\t,pi\\,x)" - by (rule at_exists_fresh[OF at_name_inst], simp add: fs_name1) + by (rule exists_fresh, simp add: fs_name1) then obtain c::"name" where f1: "c\(pi\a)" and f2: "c\x" and f3: "c\(pi\t)" and f4: "c\(pi\\)" - by (force simp add: fresh_prod at_fresh[OF at_name_inst]) - from k1 have k1a: "(pi\a)\(pi\\)" - by (simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] - pt_rev_pi[OF pt_name_inst, OF at_name_inst]) + by (force simp add: fresh_prod fresh_atm) + from k1 have k1a: "(pi\a)\(pi\\)" by (simp add: fresh_bij) have l1: "(([(c,pi\a)]@pi)\\) = (pi\\)" using f4 k1a - by (simp only: pt2[OF pt_name_inst], rule pt_fresh_fresh[OF pt_name_inst, OF at_name_inst]) + by (simp only: pt_name2, rule perm_fresh_fresh) have "\x. P x (([(c,pi\a)]@pi)\((a,\)#\)) (([(c,pi\a)]@pi)\t) \" using k3 by force hence l2: "\x. P x ((c, \)#(pi\\)) (([(c,pi\a)]@pi)\t) \" using f1 l1 - by (force simp add: pt2[OF pt_name_inst] at_calc[OF at_name_inst]) + by (force simp add: pt_name2 calc_atm) have "(([(c,pi\a)]@pi)\((a,\)#\)) \ (([(c,pi\a)]@pi)\t) : \" using k2 by (rule eqvt_typing) hence l3: "((c, \)#(pi\\)) \ (([(c,pi\a)]@pi)\t) : \" using l1 f1 - by (force simp add: pt2[OF pt_name_inst] at_calc[OF at_name_inst]) + by (force simp add: pt_name2 calc_atm) have l4: "P x (pi\\) (Lam [c].(([(c,pi\a)]@pi)\t)) (\ \ \)" using f2 f4 l2 l3 a3 by auto have alpha: "(Lam [c].([(c,pi\a)]\(pi\t))) = (Lam [(pi\a)].(pi\t))" using f1 f3 by (simp add: lam.inject alpha) - show "P x (pi\\) (pi\(Lam [a].t)) (\ \ \)" using l4 alpha - by (simp only: pt2[OF pt_name_inst], simp) + show "P x (pi\\) (pi\(Lam [a].t)) (\ \ \)" using l4 alpha by (simp only: pt_name2, simp) qed hence "P x (([]::name prm)\\) (([]::name prm)\t) \" by blast thus "P x \ t \" by simp qed -(* Now it comes: The Weakening Lemma *) +text {* definition of a subcontext *} -constdefs +abbreviation "sub" :: "(name\ty) list \ (name\ty) list \ bool" (" _ \ _ " [80,80] 80) - "\1 \ \2 \ \a \. (a,\)\set \1 \ (a,\)\set \2" + "\1 \ \2 \ \a \. (a,\)\set \1 \ (a,\)\set \2" + +text {* Now it comes: The Weakening Lemma *} lemma weakening_version1: assumes a: "\1 \ t : \" @@ -151,11 +139,8 @@ shows "\2 \ t:\" using a b c apply(nominal_induct \1 t \ avoiding: \2 rule: typing_induct) -apply(auto simp add: sub_def) -(* FIXME: this was completely automatic before the *) -(* change to meta-connectives :o( *) -apply(atomize) -apply(auto) +apply(auto | atomize)+ +(* FIXME: meta-quantifiers seem to not ba as "automatic" as object-quantifiers *) done lemma weakening_version2: @@ -167,24 +152,26 @@ and c: "\1 \ \2" shows "\2 \ t:\" using a b c -proof (nominal_induct \1 t \ avoiding: \2 rule: typing_induct, auto) - case (t1 \1 a \) (* variable case *) - have "\1 \ \2" - and "valid \2" - and "(a,\)\ set \1" by fact+ - thus "\2 \ Var a : \" by (force simp add: sub_def) +proof (nominal_induct \1 t \ avoiding: \2 rule: typing_induct) + case (t_Var \1 a \) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(a,\)\ set \1" by fact + ultimately show "\2 \ Var a : \" by auto next - case (t3 a \1 \ \ t) (* lambda case *) - have a1: "\1 \ \2" by fact - have a2: "valid \2" by fact - have a3: "a\\2" by fact - have ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact - have "((a,\)#\1) \ ((a,\)#\2)" using a1 by (simp add: sub_def) + case (t_Lam a \1 \ \ t) (* lambda case *) + have vc: "a\\2" by fact (* variable convention *) + have ih: "\\3. \valid \3; ((a,\)#\1) \ \3\ \ \3 \ t:\" by fact + have "\1 \ \2" by fact + then have "((a,\)#\1) \ ((a,\)#\2)" by simp moreover - have "valid ((a,\)#\2)" using a2 a3 v2 by force - ultimately have "((a,\)#\2) \ t:\" using ih by force - with a3 show "\2 \ (Lam [a].t) : \ \ \" by force -qed + have "valid \2" by fact + then have "valid ((a,\)#\2)" using vc v2 by simp + ultimately have "((a,\)#\2) \ t:\" using ih by simp + with vc show "\2 \ (Lam [a].t) : \ \ \" by auto +qed (auto) lemma weakening_version3: assumes a: "\1 \ t:\" @@ -193,44 +180,43 @@ shows "\2 \ t:\" using a b c proof (nominal_induct \1 t \ avoiding: \2 rule: typing_induct) - case (t3 a \1 \ \ t) (* lambda case *) - have fc: "a\\2" by fact - have ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact - have a1: "\1 \ \2" by fact - have a2: "valid \2" by fact - have "((a,\)#\1) \ ((a,\)#\2)" using a1 sub_def by simp + case (t_Lam a \1 \ \ t) (* lambda case *) + have vc: "a\\2" by fact (* variable convention *) + have ih: "\\3. \valid \3; ((a,\)#\1) \ \3\ \ \3 \ t:\" by fact + have "\1 \ \2" by fact + then have "((a,\)#\1) \ ((a,\)#\2)" by simp moreover - have "valid ((a,\)#\2)" using a2 fc by force - ultimately have "((a,\)#\2) \ t:\" using ih by simp - with fc show "\2 \ (Lam [a].t) : \ \ \" by force -qed (auto simp add: sub_def) (* app and var case *) + have "valid \2" by fact + then have "valid ((a,\)#\2)" using vc v2 by simp + ultimately have "((a,\)#\2) \ t:\" using ih by simp + with vc show "\2 \ (Lam [a].t) : \ \ \" by auto +qed (auto) (* app and var case *) -text{* The original induction principle for typing - is not strong enough - so the simple proof fails *} +text{* The original induction principle for the typing relation + is not strong enough - even this simple lemma fails *} lemma weakening_too_weak: assumes a: "\1 \ t:\" and b: "valid \2" and c: "\1 \ \2" shows "\2 \ t:\" - -thm typing.induct[no_vars] - using a b c proof (induct arbitrary: \2) - case (t1 \1 \ a) (* variable case *) - have "\1 \ \2" - and "valid \2" - and "(a,\) \ (set \1)" by fact+ - thus "\2 \ Var a : \" by (force simp add: sub_def) + case (t_Var \1 a \) (* variable case *) + have "\1 \ \2" by fact + moreover + have "valid \2" by fact + moreover + have "(a,\) \ (set \1)" by fact + ultimately show "\2 \ Var a : \" by auto next - case (t3 \1 \ \ a t) (* lambda case *) + case (t_Lam a \1 \ t \) (* lambda case *) (* all assumption in this case*) have a0: "a\\1" by fact have a1: "((a,\)#\1) \ t : \" by fact have a2: "\1 \ \2" by fact have a3: "valid \2" by fact - have ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact - have "((a,\)#\1) \ ((a,\)#\2)" using a2 by (simp add: sub_def) + have ih: "\\3. \valid \3; ((a,\)#\1) \ \3\ \ \3 \ t:\" by fact + have "((a,\)#\1) \ ((a,\)#\2)" using a2 by simp moreover have "valid ((a,\)#\2)" using v2 (* fails *) oops