diff -r dd50de393330 -r 3dcc34f18bfa src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Wed Nov 30 12:28:47 2005 +0100 +++ b/src/HOL/Nominal/Examples/Weakening.thy Wed Nov 30 14:27:09 2005 +0100 @@ -96,35 +96,35 @@ lemma typing_induct_weak[THEN spec, case_names t1 t2 t3]: - fixes P :: "(name\ty) list \ lam \ ty \'a\bool" + fixes P :: "'a\(name\ty) list \ lam \ ty \bool" and \ :: "(name\ty) list" and t :: "lam" and \ :: "ty" assumes a: "\ \ t : \" - and a1: "\x \ (a::name) \. valid \ \ (a,\) \ set \ \ P \ (Var a) \ x" - and a2: "\x \ \ \ t1 t2. - \ \ t1 : \\\ \ (\z. P \ t1 (\\\) z) \ \ \ t2 : \ \ (\z. P \ t2 \ z) - \ P \ (App t1 t2) \ x" - and a3: "\x (a::name) \ \ \ t. - a\\ \ ((a,\) # \) \ t : \ \ (\z. P ((a,\)#\) t \ z) - \ P \ (Lam [a].t) (\\\) x" - shows "\x. P \ t \ x" + and a1: "\\ (a::name) \ x. valid \ \ (a,\) \ set \ \ P x \ (Var a) \" + and a2: "\\ \ \ t1 t2 x. + \ \ t1 : \\\ \ (\z. P z \ t1 (\\\)) \ \ \ t2 : \ \ (\z. P z \ t2 \) + \ P x \ (App t1 t2) \" + and a3: "\(a::name) \ \ \ t x. + a\\ \ ((a,\) # \) \ t : \ \ (\z. P z ((a,\)#\) t \) + \ P x \ (Lam [a].t) (\\\)" + shows "\x. P x \ t \" using a by (induct, simp_all add: a1 a2 a3) lemma typing_induct_aux[rule_format]: - fixes P :: "(name\ty) list \ lam \ ty \'a::fs_name\bool" + fixes P :: "'a::fs_name \ (name\ty) list \ lam \ ty \bool" and \ :: "(name\ty) list" and t :: "lam" and \ :: "ty" assumes a: "\ \ t : \" - and a1: "\x \ (a::name) \. valid \ \ (a,\) \ set \ \ P \ (Var a) \ x" - and a2: "\x \ \ \ t1 t2. - \ \ t1 : \\\ \ (\z. P \ t1 (\\\) z) \ \ \ t2 : \ \ (\z. P \ t2 \ z) - \ P \ (App t1 t2) \ x" - and a3: "\x (a::name) \ \ \ t. - a\x \ a\\ \ ((a,\) # \) \ t : \ \ (\z. P ((a,\)#\) t \ z) - \ P \ (Lam [a].t) (\\\) x" - shows "\(pi::name prm) (x::'a::fs_name). P (pi\\) (pi\t) \ x" + and a1: "\\ (a::name) \ x. valid \ \ (a,\) \ set \ \ P x \ (Var a) \" + and a2: "\\ \ \ t1 t2 x. + \ \ t1 : \\\ \ (\z. P z \ t1 (\\\)) \ \ \ t2 : \ \ (\z. P z \ t2 \) + \ P x \ (App t1 t2) \" + and a3: "\(a::name) \ \ \ t x. + a\x \ a\\ \ ((a,\) # \) \ t : \ \ (\z. P z ((a,\)#\) t \) + \ P x \ (Lam [a].t) (\\\)" + shows "\(pi::name prm) (x::'a::fs_name). P x (pi\\) (pi\t) \" using a proof (induct) case (t1 \ \ a) @@ -136,7 +136,7 @@ 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 (pi\\) (Var (pi\a)) \ x" using a1 j3 j4 by force + show "P x (pi\\) (Var (pi\a)) \" using a1 j3 j4 by force qed next case (t2 \ \ \ t1 t2) @@ -145,7 +145,7 @@ case (t3 \ \ \ a t) have k1: "a\\" by fact have k2: "((a,\)#\)\t:\" by fact - have k3: "\(pi::name prm) (x::'a::fs_name). P (pi \ ((a,\)#\)) (pi\t) \ x" by fact + have k3: "\(pi::name prm) (x::'a::fs_name). P x (pi \((a,\)#\)) (pi\t) \" by fact show ?case proof (intro strip, simp) fix pi::"name prm" and x::"'a::fs_name" @@ -159,36 +159,36 @@ pt_rev_pi[OF pt_name_inst, OF at_name_inst]) 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]) - have "\x. P (([(c,pi\a)]@pi)\((a,\)#\)) (([(c,pi\a)]@pi)\t) \ x" using k3 by force - hence l2: "\x. P ((c, \)#(pi\\)) (([(c,pi\a)]@pi)\t) \ x" using f1 l1 + 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] split: if_splits) 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] split: if_splits) - have l4: "P (pi\\) (Lam [c].(([(c,pi\a)]@pi)\t)) (\ \ \) x" using f2 f4 l2 l3 a3 by auto + 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 (pi\\) (Lam [(pi\a)].(pi\t)) (\ \ \) x" using l4 alpha + show "P x (pi\\) (Lam [(pi\a)].(pi\t)) (\ \ \)" using l4 alpha by (simp only: pt2[OF pt_name_inst]) qed qed -lemma typing_induct[case_names t1 t2 t3]: - fixes P :: "(name\ty) list \ lam \ ty \'a::fs_name\bool" +lemma typing_induct[consumes 1, case_names t1 t2 t3]: + 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: "\x \ (a::name) \. valid \ \ (a,\) \ set \ \ P \ (Var a) \ x" - and a2: "\x \ \ \ t1 t2. - \ \ t1 : \\\ \ (\z. P \ t1 (\\\) z) \ \ \ t2 : \ \ (\z. P \ t2 \ z) - \ P \ (App t1 t2) \ x" - and a3: "\x (a::name) \ \ \ t. - a\x \ a\\ \ ((a,\) # \) \ t : \ \ (\z. P ((a,\)#\) t \ z) - \ P \ (Lam [a].t) (\\\) x" - shows "P \ t \ x" -using a a1 a2 a3 typing_induct_aux[of "\" "t" "\" "P" "[]" "x", simplified] by force + and a1: "\\ (a::name) \ x. valid \ \ (a,\) \ set \ \ P x \ (Var a) \" + and a2: "\\ \ \ t1 t2 x. + \ \ t1 : \\\ \ (\z. P z \ t1 (\\\)) \ \ \ t2 : \ \ (\z. P z \ t2 \) + \ P x \ (App t1 t2) \" + and a3: "\(a::name) \ \ \ t x. + a\x \ a\\ \ ((a,\) # \) \ t : \ \ (\z. P z ((a,\)#\) t \) + \ P x \ (Lam [a].t) (\\\)" + shows "P x \ t \" +using a a1 a2 a3 typing_induct_aux[of "\" "t" "\" "P" "x" "[]", simplified] by force (* Now it comes: The Weakening Lemma *) @@ -198,11 +198,15 @@ "\1 \ \2 \ \a \. (a,\)\set \1 \ (a,\)\set \2" lemma weakening_version1[rule_format]: - assumes a: "\1 \ t : \" - shows "valid \2 \ \1 \ \2 \ \2 \ t:\" -using a -apply(nominal_induct \1 t \ rule: typing_induct) + assumes "\1 \ t : \" + and "valid \2" + and "\1 \ \2" + shows "\2 \ t:\" +using prems +apply(nominal_induct \1 t \ fresh: \2 rule: typing_induct) apply(auto simp add: sub_def) +apply(atomize) (* FIXME: earlier this was completely automatic :o( *) +apply(auto) done lemma weakening_version2[rule_format]: @@ -211,19 +215,19 @@ and \ ::"ty" assumes a: "\1 \ t:\" shows "valid \2 \ \1 \ \2 \ \2 \ t:\" -using a -proof (nominal_induct \1 t \ rule: typing_induct, auto) - case (t1 \2 \1 a \) (* variable case *) +using prems +proof (nominal_induct \1 t \ fresh: \2 rule: typing_induct, auto) + case (t1 \1 a \ \2) (* variable case *) assume "\1 \ \2" and "valid \2" and "(a,\)\ set \1" thus "\2 \ Var a : \" by (force simp add: sub_def) next - case (t3 \2 a \1 \ \ t) (* lambda case *) + case (t3 a \1 \ \ t \2) (* lambda case *) assume a1: "\1 \ \2" and a2: "valid \2" and a3: "a\\2" - have i: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact + have i: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact have "((a,\)#\1) \ ((a,\)#\2)" using a1 by (simp add: sub_def) moreover have "valid ((a,\)#\2)" using a2 a3 v2 by force @@ -235,19 +239,19 @@ fixes \1::"(name\ty) list" and t ::"lam" and \ ::"ty" - assumes a: "\1 \ t:\" + assumes "\1 \ t:\" shows "valid \2 \ \1 \ \2 \ \2 \ t:\" -using a -proof (nominal_induct \1 t \ rule: typing_induct) - case (t1 \2 \1 a \) (* variable case *) +using prems +proof (nominal_induct \1 t \ fresh: \2 rule: typing_induct) + case (t1 \1 a \ \2) (* variable case *) thus "valid \2 \ \1 \ \2 \ \2 \ Var a : \" by (force simp add: sub_def) next - case (t2 \2 \1 \ \ t1 t2) (* variable case *) + case (t2 \1 \ \ t1 t2 \2) (* variable case *) thus "valid \2 \ \1 \ \2 \ \2 \ App t1 t2 : \" by force next - case (t3 \2 a \1 \ \ t) (* lambda case *) - have a3: "a\\2" - and i: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact + case (t3 a \1 \ \ t \2) (* lambda case *) + have a3: "a\\2" by fact + have ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact show "valid \2 \ \1 \ \2 \ \2 \ (Lam [a].t) : \ \ \" proof (intro strip) assume a1: "\1 \ \2" @@ -255,19 +259,19 @@ have "((a,\)#\1) \ ((a,\)#\2)" using a1 by (simp add: sub_def) moreover have "valid ((a,\)#\2)" using a2 a3 v2 by force - ultimately have "((a,\)#\2) \ t:\" using i by force + ultimately have "((a,\)#\2) \ t:\" using ih by force with a3 show "\2 \ (Lam [a].t) : \ \ \" by force qed qed lemma weakening_version4[rule_format]: - assumes a: "\1 \ t:\" + assumes "\1 \ t:\" shows "valid \2 \ \1 \ \2 \ \2 \ t:\" -using a -proof (nominal_induct \1 t \ rule: typing_induct) - case (t3 \2 a \1 \ \ t) (* lambda case *) - have fc: "a\\2" - and ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact +using prems +proof (nominal_induct \1 t \ fresh: \2 rule: typing_induct) + case (t3 a \1 \ \ t \2) (* lambda case *) + have fc: "a\\2" by fact + have ih: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\" by fact show "valid \2 \ \1 \ \2 \ \2 \ (Lam [a].t) : \ \ \" proof (intro strip) assume a1: "\1 \ \2" @@ -284,17 +288,17 @@ (* original induction principle is not strong *) (* enough - so the simple proof fails *) lemma weakening_too_weak[rule_format]: - assumes a: "\1 \ t:\" + assumes "\1 \ t:\" shows "valid \2 \ \1 \ \2 \ \2 \ t:\" -using a -proof (nominal_induct \1 t \ rule: typing_induct_weak, auto) - case (t1 \2 \1 a \) (* variable case *) +using prems +proof (induct, auto) + case (t1 \1 \ a) (* variable case *) assume "\1 \ \2" and "valid \2" - and "(a,\)\ set \1" + and "(a,\) \ (set \1)" thus "\2 \ Var a : \" by (force simp add: sub_def) next - case (t3 \2 a \1 \ \ t) (* lambda case *) + case (t3 \1 \ \ a t) (* lambda case *) assume a1: "\1 \ \2" and a2: "valid \2" and i: "\\3. valid \3 \ ((a,\)#\1) \ \3 \ \3 \ t:\"