# HG changeset patch # User paulson # Date 1020845577 -7200 # Node ID 336b0bcbd27cbd746099b49029b773eb6d6a1f52 # Parent 0b233f430076fd74e9ad41919c2065f4c7b5e49b new lemmas 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" diff -r 0b233f430076 -r 336b0bcbd27c src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Wed May 08 09:14:56 2002 +0200 +++ b/src/ZF/OrdQuant.thy Wed May 08 10:12:57 2002 +0200 @@ -62,6 +62,33 @@ declare Ord_UN [intro,simp,TC] declare Ord_Union [intro,simp,TC] +(** Now some very basic ZF theorems **) + +lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" +by blast + +lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" +by blast + +lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)" +by (unfold trans_def trans_on_def, blast) + +lemma image_is_UN: "\function(g); x <= domain(g)\ \ g``x = (UN k:x. {g`k})" +by (blast intro: function_apply_equality [THEN sym] function_apply_Pair) + +lemma functionI: + "\!!x y y'. \:r; :r\ \ y=y'\ \ function(r)" +by (simp add: function_def, blast) + +lemma function_lam: "function (lam x:A. b(x))" +by (simp add: function_def lam_def) + +lemma relation_lam: "relation (lam x:A. b(x))" +by (simp add: relation_def lam_def, blast) + +lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" +by (simp add: restrict_def) + (** These mostly belong to theory Ordinal **) lemma Union_upper_le: