# HG changeset patch # User urbanc # Date 1191821893 -7200 # Node ID 08865bb8709817ca234e7e66bcb7adf12eab51e6 # Parent 799ca514244fc4cbea299368f55bec2f97aaf8ec Isar-fied many proofs diff -r 799ca514244f -r 08865bb87098 src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Mon Oct 08 07:09:29 2007 +0200 +++ b/src/HOL/Nominal/Examples/SN.thy Mon Oct 08 07:38:13 2007 +0200 @@ -211,10 +211,8 @@ assumes a: "SN a" and b: "SN b" and hyp: "\x z. - (\y. x \\<^isub>\ y \ SN y) \ - (\y. x \\<^isub>\ y \ P y z) \ - (\u. z \\<^isub>\ u \ SN u) \ - (\u. z \\<^isub>\ u \ P x u) \ P x z" + \\y. x \\<^isub>\ y \ SN y; \y. x \\<^isub>\ y \ P y z; + \u. z \\<^isub>\ u \ SN u; \u. z \\<^isub>\ u \ P x u\ \ P x z" shows "P a b" proof - from a @@ -259,6 +257,7 @@ "RED (\\\) = {t. \u. (u\RED \ \ (App t u)\RED \)}" by (rule TrueI)+ +(* neutral terms *) constdefs NEUT :: "lam \ bool" "NEUT t \ (\a. t = Var a) \ (\t1 t2. t = App t1 t2)" @@ -270,15 +269,32 @@ where fst[intro!]: "(App t s) \ t" +consts + fst_app_aux::"lam\lam option" + +nominal_primrec + "fst_app_aux (Var a) = None" + "fst_app_aux (App t1 t2) = Some t1" + "fst_app_aux (Lam [x].t) = None" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: fresh_none) +apply(fresh_guess)+ +done + +definition + fst_app_def[simp]: "fst_app t = the (fst_app_aux t)" + lemma SN_of_FST_of_App: assumes a: "SN (App t s)" - shows "SN t" + shows "SN (fst_app (App t s))" using a proof - from a have "\z. (App t s \ z) \ SN z" by (induct rule: SN.SN.induct) (blast elim: FST.cases intro: SN_intro) - then show "SN t" by blast + then have "SN t" by blast + then show "SN (fst_app (App t s))" by simp qed section {* Candidates *} @@ -378,7 +394,7 @@ ultimately have "(Var a)\ RED \1" by (simp add: CR4_def) with a have "App t (Var a) \ RED \2" by simp hence "SN (App t (Var a))" using ih_CR1_\2 by (simp add: CR1_def) - thus "SN(t)" by (rule SN_of_FST_of_App) + thus "SN(t)" by (auto dest: SN_of_FST_of_App) qed next case 2 @@ -522,10 +538,8 @@ have ih: "\\. \ closes ((a,\)#\) \ \ \ RED \" by fact have \_cond: "\ closes \" by fact have fresh: "a\\" "a\\" by fact+ - from ih have "\s\RED \. ((a,s)#\) \ RED \" - using fresh \_cond fresh_context by simp - then have "\s\RED \. \[a::=s] \ RED \" - using fresh by (simp add: psubst_subst) + from ih have "\s\RED \. ((a,s)#\) \ RED \" using fresh \_cond fresh_context by simp + then have "\s\RED \. \[a::=s] \ RED \" using fresh by (simp add: psubst_subst) then have "Lam [a].(\) \ RED (\ \ \)" by (simp only: abs_RED) then show "\<(Lam [a].t)> \ RED (\ \ \)" using fresh by simp qed auto