diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Acc.thy --- a/src/ZF/Induct/Acc.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Acc.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,7 +12,7 @@ \ consts - acc :: "i => i" + acc :: "i \ i" inductive domains "acc(r)" \ "field(r)" @@ -28,15 +28,15 @@ The intended introduction rule: \ -lemma accI: "\\b. :r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" +lemma accI: "\\b. \b,a\:r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" by (blast intro: acc.intros) -lemma acc_downward: "\b \ acc(r); : r\ \ a \ acc(r)" +lemma acc_downward: "\b \ acc(r); \a,b\: r\ \ a \ acc(r)" by (erule acc.cases) blast lemma acc_induct [consumes 1, case_names vimage, induct set: acc]: "\a \ acc(r); - \x. \x \ acc(r); \y. :r \ P(y)\ \ P(x) + \x. \x \ acc(r); \y. \y,x\:r \ P(y)\ \ P(x) \ \ P(a)" by (erule acc.induct) (blast intro: acc.intros)