# HG changeset patch # User urbanc # Date 1180956142 -7200 # Node ID 76462538f3495cc560518832d7532601e97a3ff2 # Parent 861ab9c18e1818f80009eca5e40fed83714661df added a few comments to the proofs diff -r 861ab9c18e18 -r 76462538f349 src/HOL/Nominal/Examples/LocalWeakening.thy --- a/src/HOL/Nominal/Examples/LocalWeakening.thy Mon Jun 04 11:39:19 2007 +0200 +++ b/src/HOL/Nominal/Examples/LocalWeakening.thy Mon Jun 04 13:22:22 2007 +0200 @@ -11,14 +11,17 @@ atom_decl name -(* Curry-style lambda terms in locally nameless *) -(* representation without any binders *) +text {* Curry-style lambda terms in locally nameless + representation without any binders *} + nominal_datatype llam = lPar "name" | lVar "nat" | lApp "llam" "llam" | lLam "llam" +text {* definition of vsub - at the moment a bit cumbersome *} + lemma llam_cases: "(\a. t = lPar a) \ (\n. t = lVar n) \ (\t1 t2. t = lApp t1 t2) \ (\t1. t = lLam t1)" @@ -69,6 +72,8 @@ shows "pi\(freshen t p) = freshen (pi\t) (pi\p)" by (simp add: vsub_eqvt freshen_def perm_nat_def) +text {* types *} + nominal_datatype ty = TVar "nat" | TArr "ty" "ty" (infix "\" 200) @@ -88,6 +93,7 @@ (simp_all add: fresh_nat) text {* valid contexts *} + types cxt = "(name\ty) list" inductive2 @@ -104,12 +110,14 @@ using valid.cases[OF a] by (auto) +text {* "weak" typing relation *} + inductive2 - typing :: "cxt\llam\ty\bool" (" _ \ _ : _ " [80,80,80] 80) + typing :: "cxt\llam\ty\bool" (" _ \ _ : _ " [60,60,60] 60) where - t_lPar[intro]: "\valid \; (a,T)\set \\\ \ \ lPar a : T" + t_lPar[intro]: "\valid \; (x,T)\set \\\ \ \ lPar x : T" | t_lApp[intro]: "\\ \ t1 : T1\T2; \ \ t2 : T1\\ \ \ lApp t1 t2 : T2" - | t_lLam[intro]: "\a\t; ((a,T1)#\) \ (freshen t a) : T2\\ \ \ lLam t : T1\T2" + | t_lLam[intro]: "\x\t; (x,T1)#\ \ freshen t x : T2\\ \ \ lLam t : T1\T2" equivariance typing @@ -119,32 +127,34 @@ using a by (induct) (auto dest: v2_elim) +text {* we explicitly have to say to strengthen over the variable x *} nominal_inductive typing - avoids t_lLam: a + avoids t_lLam: x by (auto simp add: fresh_prod dest: v2_elim typing_implies_valid) -(* strong induction principle for typing *) +text {* strong induction principle for typing *} thm typing.strong_induct +text {* sub-contexts *} + abbreviation "sub_context" :: "cxt \ cxt \ bool" ("_ \ _" [60,60] 60) where "\1 \ \2 \ \x T. (x,T)\set \1 \ (x,T)\set \2" -lemma weakening_automatic: +lemma weakening_almost_automatic: assumes a: "\1 \ t : T" and b: "\1 \ \2" and c: "valid \2" shows "\2 \ t : T" using a b c apply(nominal_induct \1 t T avoiding: \2 rule: typing.strong_induct) -apply(auto)[1] -apply(auto)[1] -apply(drule_tac x="(a,T1)#\2" in meta_spec) +apply(auto) +apply(drule_tac x="(x,T1)#\2" in meta_spec) apply(auto intro!: t_lLam) done -lemma weakening: +lemma weakening_in_more_detail: assumes a: "\1 \ t : T" and b: "\1 \ \2" and c: "valid \2" @@ -161,14 +171,14 @@ next case (t_lLam x t T1 \1 T2 \2) (* lambda case *) have vc: "x\\2" "x\t" by fact (* variable convention *) - have ih: "\(x,T1)#\1 \ (x,T1)#\2; valid ((x,T1)#\2)\ \ ((x,T1)#\2) \ freshen t x : T2" by fact + have ih: "\(x,T1)#\1 \ (x,T1)#\2; valid ((x,T1)#\2)\ \ (x,T1)#\2 \ freshen t x : 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: v2) - ultimately have "((x,T1)#\2) \ freshen t x : T2" using ih by simp - with vc show "\2 \ lLam t : T1\T2" by (auto simp add: fresh_prod) + ultimately have "(x,T1)#\2 \ freshen t x : T2" using ih by simp + with vc show "\2 \ lLam t : T1\T2" by auto qed (auto) (* app case *) end