diff -r 0b233f430076 -r 336b0bcbd27c src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Wed May 08 09:14:56 2002 +0200 +++ b/src/ZF/CardinalArith.thy Wed May 08 10:12:57 2002 +0200 @@ -40,6 +40,42 @@ "op |*|" :: "[i,i] => i" (infixl "\" 70) +(*** The following really belong early in the development ***) + +lemma relation_converse_converse [simp]: + "relation(r) ==> converse(converse(r)) = r" +by (simp add: relation_def, blast) + +lemma relation_restrict [simp]: "relation(restrict(r,A))" +by (simp add: restrict_def relation_def, blast) + +(*** The following really belong in Order ***) + +lemma subset_ord_iso_Memrel: + "\f: ord_iso(A,Memrel(B),C,r); A<=B\ \ f: ord_iso(A,Memrel(A),C,r)" +apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN fun_is_rel]) +apply (frule ord_iso_trans [OF id_ord_iso_Memrel], assumption) +apply (simp add: right_comp_id) +done + +lemma restrict_ord_iso: + "\f \ ord_iso(i, Memrel(i), Order.pred(A,a,r), r); a \ A; j < i; + trans[A](r)\ + \ restrict(f,j) \ ord_iso(j, Memrel(j), Order.pred(A,f`j,r), r)" +apply (frule ltD) +apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption) +apply (frule ord_iso_restrict_pred, assumption) +apply (simp add: pred_iff trans_pred_pred_eq lt_pred_Memrel) +apply (blast intro!: subset_ord_iso_Memrel le_imp_subset [OF leI]) +done + +lemma restrict_ord_iso2: + "\f \ ord_iso(Order.pred(A,a,r), r, i, Memrel(i)); a \ A; + j < i; trans[A](r)\ + \ converse(restrict(converse(f), j)) + \ ord_iso(Order.pred(A, converse(f)`j, r), r, j, Memrel(j))" +by (blast intro: restrict_ord_iso ord_iso_sym ltI) + (*** The following really belong in OrderType ***) lemma oadd_eq_0_iff: "\Ord(i); Ord(j)\ \ (i ++ j) = 0 <-> i=0 & j=0"