diff -r a80d8ec6c998 -r 3dda49e08b9d src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Fri Jan 04 21:49:06 2019 +0100 +++ b/src/ZF/Constructible/Reflection.thy Fri Jan 04 23:22:53 2019 +0100 @@ -22,10 +22,10 @@ text\First part: the cumulative hierarchy defining the class \M\. To avoid handling multiple arguments, we assume that \Mset(l)\ is closed under ordered pairing provided \l\ is limit. Possibly this -could be avoided: the induction hypothesis @{term Cl_reflects} +could be avoided: the induction hypothesis \<^term>\Cl_reflects\ (in locale \ex_reflection\) could be weakened to -@{term "\y\Mset(a). \z\Mset(a). P() \ Q(a,)"}, removing most -uses of @{term Pair_in_Mset}. But there isn't much point in doing so, since +\<^term>\\y\Mset(a). \z\Mset(a). P() \ Q(a,)\, removing most +uses of \<^term>\Pair_in_Mset\. But there isn't much point in doing so, since ultimately the \ex_reflection\ proof is packaged up using the predicate \Reflects\. \ @@ -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)" @@ -115,8 +115,8 @@ apply (simp add: cont_Ord_def FF_def, blast) done -text\Recall that @{term F0} depends upon @{term "y\Mset(a)"}, -while @{term FF} depends only upon @{term a}.\ +text\Recall that \<^term>\F0\ depends upon \<^term>\y\Mset(a)\, +while \<^term>\FF\ depends only upon \<^term>\a\.\ lemma (in reflection) FF_works: "[| M(z); y\Mset(a); P(); Ord(a) |] ==> \z\Mset(FF(P,a)). P()" apply (simp add: FF_def) @@ -275,7 +275,7 @@ by fast text\Problem here: there needs to be a conjunction (class intersection) -in the class of reflecting ordinals. The @{term "Ord(a)"} is redundant, +in the class of reflecting ordinals. The \<^term>\Ord(a)\ is redundant, though harmless.\ lemma (in reflection) "Reflects(\a. Ord(a) & ClEx(\x. fst(x) \ snd(x), a), @@ -328,7 +328,7 @@ done text\Example 3. Warning: the following examples make sense only -if @{term P} is quantifier-free, since it is not being relativized.\ +if \<^term>\P\ is quantifier-free, since it is not being relativized.\ schematic_goal (in reflection) "Reflects(?Cl, \x. \y. M(y) & (\z. M(z) \ z \ y \ z \ x & P(z)),