diff -r d74a16b84e05 -r 75873e94357c src/HOL/Library/SCT_Interpretation.thy --- a/src/HOL/Library/SCT_Interpretation.thy Wed Jul 11 11:25:24 2007 +0200 +++ b/src/HOL/Library/SCT_Interpretation.thy Wed Jul 11 11:27:46 2007 +0200 @@ -13,34 +13,34 @@ "idseq R s x = (s 0 = x \ (\i. R (s (Suc i)) (s i)))" lemma not_acc_smaller: - assumes notacc: "\ acc R x" - shows "\y. R y x \ \ acc R y" + assumes notacc: "\ accp R x" + shows "\y. R y x \ \ accp R y" proof (rule classical) assume "\ ?thesis" - hence "\y. R y x \ acc R y" by blast - with accI have "acc R x" . + hence "\y. R y x \ accp R y" by blast + with accp.accI have "accp R x" . with notacc show ?thesis by contradiction qed lemma non_acc_has_idseq: - assumes "\ acc R x" + assumes "\ accp R x" shows "\s. idseq R s x" proof - - have "\f. \x. \acc R x \ R (f x) x \ \acc R (f x)" + have "\f. \x. \accp R x \ R (f x) x \ \accp R (f x)" by (rule choice, auto simp:not_acc_smaller) then obtain f where - in_R: "\x. \acc R x \ R (f x) x" - and nia: "\x. \acc R x \ \acc R (f x)" + in_R: "\x. \accp R x \ R (f x) x" + and nia: "\x. \accp R x \ \accp R (f x)" by blast let ?s = "\i. (f ^ i) x" { fix i - have "\acc R (?s i)" - by (induct i) (auto simp:nia `\acc R x`) + have "\accp R (?s i)" + by (induct i) (auto simp:nia `\accp R x`) hence "R (f (?s i)) (?s i)" by (rule in_R) } @@ -281,10 +281,10 @@ assumes R: "R = mk_rel RDs" assumes sound: "sound_int \ RDs M" assumes "SCT \" - shows "\x. acc R x" + shows "\x. accp R x" proof (rule, rule classical) fix x - assume "\ acc R x" + assume "\ accp R x" with non_acc_has_idseq have "\s. idseq R s x" . then obtain s where "idseq R s x" .. @@ -409,7 +409,7 @@ qed ultimately have False by (rule no_inf_desc_nat_sequence[of "Suc n"]) simp - thus "acc R x" .. + thus "accp R x" .. qed end