diff -r f2094906e491 -r e44d86131648 src/ZF/ex/misc.thy --- a/src/ZF/ex/misc.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/misc.thy Tue Sep 27 16:51:35 2022 +0100 @@ -31,7 +31,7 @@ text\A weird property of ordered pairs.\ -lemma "b\c ==> \ = " +lemma "b\c \ \ = " by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) text\These two are cited in Benzmueller and Kohlhase's system description of @@ -44,15 +44,15 @@ by (blast intro!: equalityI) text\trivial example of term synthesis: apparently hard for some provers!\ -schematic_goal "a \ b ==> a:?X & b \ ?X" +schematic_goal "a \ b \ a:?X & b \ ?X" by blast text\Nice blast benchmark. Proved in 0.3s; old tactics can't manage it!\ -lemma "\x \ S. \y \ S. x \ y ==> \z. S \ {z}" +lemma "\x \ S. \y \ S. x \ y \ \z. S \ {z}" by blast text\variant of the benchmark above\ -lemma "\x \ S. \(S) \ x ==> \z. S \ {z}" +lemma "\x \ S. \(S) \ x \ \z. S \ {z}" by blast (*Example 12 (credited to Peter Andrews) from @@ -87,10 +87,10 @@ by force text\Another version, with meta-level rewriting\ -lemma "(!! A f B g. hom(A,f,B,g) == +lemma "(\A f B g. hom(A,f,B,g) \ {H \ A->B. f \ A*A->A & g \ B*B->B & (\x \ A. \y \ A. H`(f`) = g`)}) - ==> J \ hom(A,f,B,g) & K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" + \ J \ hom(A,f,B,g) & K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" by force @@ -105,38 +105,38 @@ comp_mem_injD2 comp_mem_surjD2 lemma pastre1: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ surj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre3: - "[| (h O g O f) \ surj(A,A); + "\(h O g O f) \ surj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre4: - "[| (h O g O f) \ surj(A,A); + "\(h O g O f) \ surj(A,A); (f O h O g) \ inj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre5: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre6: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ inj(B,B); (g O f O h) \ surj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast)