diff -r f075640b8868 -r 3abf6a722518 src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Tue Jan 16 09:12:16 2018 +0100 +++ b/src/ZF/Constructible/Reflection.thy Tue Jan 16 09:30:00 2018 +0100 @@ -38,9 +38,9 @@ defines "M(x) == \a. Ord(a) & x \ Mset(a)" and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) & (\a. Cl(a) \ (\x\Mset(a). P(x) \ Q(a,x)))" - fixes F0 \\ordinal for a specific value @{term y}\ - fixes FF \\sup over the whole level, @{term "y\Mset(a)"}\ - fixes ClEx \\Reflecting ordinals for the formula @{term "\z. P"}\ + fixes F0 \ \ordinal for a specific value @{term y}\ + fixes FF \ \sup over the whole level, @{term "y\Mset(a)"}\ + fixes ClEx \ \Reflecting ordinals for the formula @{term "\z. P"}\ defines "F0(P,y) == \ b. (\z. M(z) & P()) \ (\z\Mset(b). P())" and "FF(P) == \a. \y\Mset(a). F0(P,y)" @@ -136,9 +136,9 @@ text\Locale for the induction hypothesis\ locale ex_reflection = reflection + - fixes P \"the original formula" - fixes Q \"the reflected formula" - fixes Cl \"the class of reflecting ordinals" + fixes P \ \the original formula\ + fixes Q \ \the reflected formula\ + fixes Cl \ \the class of reflecting ordinals\ assumes Cl_reflects: "[| Cl(a); Ord(a) |] ==> \x\Mset(a). P(x) \ Q(a,x)" lemma (in ex_reflection) ClEx_downward: