# HG changeset patch # User paulson # Date 1331829302 0 # Node ID 2b6e55924af31a67534de3a5b1e2f238a324a64c # Parent 5e1bcfdcb1756517c69fe5dbafb9615bc33dce3e replacing ":" by "\" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/AC.thy --- a/src/ZF/AC.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/AC.thy Thu Mar 15 16:35:02 2012 +0000 @@ -9,7 +9,7 @@ text{*This definition comes from Halmos (1960), page 59.*} axiomatization where - AC: "[| a: A; !!x. x:A ==> (\y. y:B(x)) |] ==> \z. z \ Pi(A,B)" + AC: "[| a \ A; !!x. x \ A ==> (\y. y \ B(x)) |] ==> \z. z \ Pi(A,B)" (*The same as AC, but no premise @{term"a \ A"}*) lemma AC_Pi: "[| !!x. x \ A ==> (\y. y \ B(x)) |] ==> \z. z \ Pi(A,B)" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Bin.thy --- a/src/ZF/Bin.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Bin.thy Thu Mar 15 16:35:02 2012 +0000 @@ -24,7 +24,7 @@ datatype "bin" = Pls | Min - | Bit ("w: bin", "b: bool") (infixl "BIT" 90) + | Bit ("w \ bin", "b \ bool") (infixl "BIT" 90) consts integ_of :: "i=>i" @@ -132,26 +132,26 @@ (** Type checking **) -lemma integ_of_type [TC]: "w: bin ==> integ_of(w) \ int" +lemma integ_of_type [TC]: "w \ bin ==> integ_of(w) \ int" apply (induct_tac "w") apply (simp_all add: bool_into_nat) done -lemma NCons_type [TC]: "[| w: bin; b: bool |] ==> NCons(w,b) \ bin" +lemma NCons_type [TC]: "[| w \ bin; b \ bool |] ==> NCons(w,b) \ bin" by (induct_tac "w", auto) -lemma bin_succ_type [TC]: "w: bin ==> bin_succ(w) \ bin" +lemma bin_succ_type [TC]: "w \ bin ==> bin_succ(w) \ bin" by (induct_tac "w", auto) -lemma bin_pred_type [TC]: "w: bin ==> bin_pred(w) \ bin" +lemma bin_pred_type [TC]: "w \ bin ==> bin_pred(w) \ bin" by (induct_tac "w", auto) -lemma bin_minus_type [TC]: "w: bin ==> bin_minus(w) \ bin" +lemma bin_minus_type [TC]: "w \ bin ==> bin_minus(w) \ bin" by (induct_tac "w", auto) (*This proof is complicated by the mutual recursion*) lemma bin_add_type [rule_format,TC]: - "v: bin ==> \w\bin. bin_add(v,w) \ bin" + "v \ bin ==> \w\bin. bin_add(v,w) \ bin" apply (unfold bin_add_def) apply (induct_tac "v") apply (rule_tac [3] ballI) @@ -160,7 +160,7 @@ apply (simp_all add: NCons_type) done -lemma bin_mult_type [TC]: "[| v: bin; w: bin |] ==> bin_mult(v,w) \ bin" +lemma bin_mult_type [TC]: "[| v \ bin; w \ bin |] ==> bin_mult(v,w) \ bin" by (induct_tac "v", auto) @@ -169,19 +169,19 @@ (*NCons preserves the integer value of its argument*) lemma integ_of_NCons [simp]: - "[| w: bin; b: bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)" + "[| w \ bin; b \ bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)" apply (erule bin.cases) apply (auto elim!: boolE) done lemma integ_of_succ [simp]: - "w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)" + "w \ bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac elim!: boolE) done lemma integ_of_pred [simp]: - "w: bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" + "w \ bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac elim!: boolE) done @@ -189,7 +189,7 @@ subsubsection{*@{term bin_minus}: Unary Negation of Binary Integers*} -lemma integ_of_minus: "w: bin ==> integ_of(bin_minus(w)) = $- integ_of(w)" +lemma integ_of_minus: "w \ bin ==> integ_of(bin_minus(w)) = $- integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE) done @@ -197,18 +197,18 @@ subsubsection{*@{term bin_add}: Binary Addition*} -lemma bin_add_Pls [simp]: "w: bin ==> bin_add(Pls,w) = w" +lemma bin_add_Pls [simp]: "w \ bin ==> bin_add(Pls,w) = w" by (unfold bin_add_def, simp) -lemma bin_add_Pls_right: "w: bin ==> bin_add(w,Pls) = w" +lemma bin_add_Pls_right: "w \ bin ==> bin_add(w,Pls) = w" apply (unfold bin_add_def) apply (erule bin.induct, auto) done -lemma bin_add_Min [simp]: "w: bin ==> bin_add(Min,w) = bin_pred(w)" +lemma bin_add_Min [simp]: "w \ bin ==> bin_add(Min,w) = bin_pred(w)" by (unfold bin_add_def, simp) -lemma bin_add_Min_right: "w: bin ==> bin_add(w,Min) = bin_pred(w)" +lemma bin_add_Min_right: "w \ bin ==> bin_add(w,Min) = bin_pred(w)" apply (unfold bin_add_def) apply (erule bin.induct, auto) done @@ -220,13 +220,13 @@ by (unfold bin_add_def, simp) lemma bin_add_BIT_BIT [simp]: - "[| w: bin; y: bool |] + "[| w \ bin; y \ bool |] ==> bin_add(v BIT x, w BIT y) = NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)" by (unfold bin_add_def, simp) lemma integ_of_add [rule_format]: - "v: bin ==> + "v \ bin ==> \w\bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)" apply (erule bin.induct, simp, simp) apply (rule ballI) @@ -236,7 +236,7 @@ (*Subtraction*) lemma diff_integ_of_eq: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" apply (unfold zdiff_def) apply (simp add: integ_of_add integ_of_minus) @@ -246,7 +246,7 @@ subsubsection{*@{term bin_mult}: Binary Multiplication*} lemma integ_of_mult: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)" apply (induct_tac "v", simp) apply (simp add: integ_of_minus) @@ -280,15 +280,15 @@ (** extra rules for bin_add **) -lemma bin_add_BIT_11: "w: bin ==> bin_add(v BIT 1, w BIT 1) = +lemma bin_add_BIT_11: "w \ bin ==> bin_add(v BIT 1, w BIT 1) = NCons(bin_add(v, bin_succ(w)), 0)" by simp -lemma bin_add_BIT_10: "w: bin ==> bin_add(v BIT 1, w BIT 0) = +lemma bin_add_BIT_10: "w \ bin ==> bin_add(v BIT 1, w BIT 0) = NCons(bin_add(v,w), 1)" by simp -lemma bin_add_BIT_0: "[| w: bin; y: bool |] +lemma bin_add_BIT_0: "[| w \ bin; y \ bool |] ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)" by simp @@ -345,7 +345,7 @@ (** Equals (=) **) lemma eq_integ_of_eq: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> ((integ_of(v)) = integ_of(w)) \ iszero (integ_of (bin_add (v, bin_minus(w))))" apply (unfold iszero_def) @@ -362,7 +362,7 @@ done lemma iszero_integ_of_BIT: - "[| w: bin; x: bool |] + "[| w \ bin; x \ bool |] ==> iszero (integ_of (w BIT x)) \ (x=0 & iszero (integ_of(w)))" apply (unfold iszero_def, simp) apply (subgoal_tac "integ_of (w) \ int") @@ -374,10 +374,10 @@ done lemma iszero_integ_of_0: - "w: bin ==> iszero (integ_of (w BIT 0)) \ iszero (integ_of(w))" + "w \ bin ==> iszero (integ_of (w BIT 0)) \ iszero (integ_of(w))" by (simp only: iszero_integ_of_BIT, blast) -lemma iszero_integ_of_1: "w: bin ==> ~ iszero (integ_of (w BIT 1))" +lemma iszero_integ_of_1: "w \ bin ==> ~ iszero (integ_of (w BIT 1))" by (simp only: iszero_integ_of_BIT, blast) @@ -385,7 +385,7 @@ (** Less-than (<) **) lemma less_integ_of_eq_neg: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $< integ_of(w) \ znegative (integ_of (bin_add (v, bin_minus(w))))" apply (unfold zless_def zdiff_def) @@ -399,7 +399,7 @@ by simp lemma neg_integ_of_BIT: - "[| w: bin; x: bool |] + "[| w \ bin; x \ bool |] ==> znegative (integ_of (w BIT x)) \ znegative (integ_of(w))" apply simp apply (subgoal_tac "integ_of (w) \ int") @@ -471,24 +471,24 @@ (** Simplification of arithmetic when nested to the right **) lemma add_integ_of_left [simp]: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)" by (simp add: zadd_assoc [symmetric]) lemma mult_integ_of_left [simp]: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)" by (simp add: zmult_assoc [symmetric]) lemma add_integ_of_diff1 [simp]: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)" apply (unfold zdiff_def) apply (rule add_integ_of_left, auto) done lemma add_integ_of_diff2 [simp]: - "[| v: bin; w: bin |] + "[| v \ bin; w \ bin |] ==> integ_of(v) $+ (c $- integ_of(w)) = integ_of (bin_add (v, bin_minus(w))) $+ (c)" apply (subst diff_integ_of_eq [symmetric]) @@ -509,10 +509,10 @@ lemma zdiff_self [simp]: "x $- x = #0" by (simp add: zdiff_def) -lemma znegative_iff_zless_0: "k: int ==> znegative(k) \ k $< #0" +lemma znegative_iff_zless_0: "k \ int ==> znegative(k) \ k $< #0" by (simp add: zless_def) -lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k: int|] ==> znegative($-k)" +lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k \ int|] ==> znegative($-k)" by (simp add: zless_def) lemma zero_zle_int_of [simp]: "#0 $<= $# n" @@ -521,7 +521,7 @@ lemma nat_of_0 [simp]: "nat_of(#0) = 0" by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of) -lemma nat_le_int0_lemma: "[| z $<= $#0; z: int |] ==> nat_of(z) = 0" +lemma nat_le_int0_lemma: "[| z $<= $#0; z \ int |] ==> nat_of(z) = 0" by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of) lemma nat_le_int0: "z $<= $#0 ==> nat_of(z) = 0" @@ -545,7 +545,7 @@ lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $<= z then intify(z) else #0)" by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless) -lemma zless_nat_iff_int_zless: "[| m: nat; z: int |] ==> (m < nat_of(z)) \ ($#m $< z)" +lemma zless_nat_iff_int_zless: "[| m \ nat; z \ int |] ==> (m < nat_of(z)) \ ($#m $< z)" apply (case_tac "znegative (z) ") apply (erule_tac [2] not_zneg_nat_of [THEN subst]) apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans] diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Cardinal.thy Thu Mar 15 16:35:02 2012 +0000 @@ -14,11 +14,11 @@ definition eqpoll :: "[i,i] => o" (infixl "eqpoll" 50) where - "A eqpoll B == \f. f: bij(A,B)" + "A eqpoll B == \f. f \ bij(A,B)" definition lepoll :: "[i,i] => o" (infixl "lepoll" 50) where - "A lepoll B == \f. f: inj(A,B)" + "A lepoll B == \f. f \ inj(A,B)" definition lesspoll :: "[i,i] => o" (infixl "lesspoll" 50) where @@ -56,7 +56,7 @@ by (rule bnd_monoI, blast+) lemma Banach_last_equation: - "g: Y->X + "g \ Y->X ==> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = X - lfp(X, %W. X - g``(Y - f``W))" apply (rule_tac P = "%u. ?v = X-u" @@ -65,7 +65,7 @@ done lemma decomposition: - "[| f: X->Y; g: Y->X |] ==> + "[| f \ X->Y; g \ Y->X |] ==> \XA XB YA YB. (XA \ XB = 0) & (XA \ XB = X) & (YA \ YB = 0) & (YA \ YB = Y) & f``XA=YA & g``YB=XB" @@ -77,7 +77,7 @@ done lemma schroeder_bernstein: - "[| f: inj(X,Y); g: inj(Y,X) |] ==> \h. h: bij(X,Y)" + "[| f \ inj(X,Y); g \ inj(Y,X) |] ==> \h. h \ bij(X,Y)" apply (insert decomposition [of f X Y g]) apply (simp add: inj_is_fun) apply (blast intro!: restrict_bij bij_disjoint_Un intro: bij_converse_bij) @@ -88,7 +88,7 @@ (** Equipollence is an equivalence relation **) -lemma bij_imp_eqpoll: "f: bij(A,B) ==> A \ B" +lemma bij_imp_eqpoll: "f \ bij(A,B) ==> A \ B" apply (unfold eqpoll_def) apply (erule exI) done @@ -128,10 +128,10 @@ done lemma eq_lepoll_trans [trans]: "[| X \ Y; Y \ Z |] ==> X \ Z" - by (blast intro: eqpoll_imp_lepoll lepoll_trans) + by (blast intro: eqpoll_imp_lepoll lepoll_trans) lemma lepoll_eq_trans [trans]: "[| X \ Y; Y \ Z |] ==> X \ Z" - by (blast intro: eqpoll_imp_lepoll lepoll_trans) + by (blast intro: eqpoll_imp_lepoll lepoll_trans) (*Asymmetry law*) lemma eqpollI: "[| X \ Y; Y \ X |] ==> X \ Y" @@ -234,11 +234,11 @@ lemma eq_lesspoll_trans [trans]: "[| X \ Y; Y \ Z |] ==> X \ Z" - by (blast intro: eqpoll_imp_lepoll lesspoll_trans1) + by (blast intro: eqpoll_imp_lepoll lesspoll_trans1) lemma lesspoll_eq_trans [trans]: "[| X \ Y; Y \ Z |] ==> X \ Z" - by (blast intro: eqpoll_imp_lepoll lesspoll_trans2) + by (blast intro: eqpoll_imp_lepoll lesspoll_trans2) (** LEAST -- the least number operator [from HOL/Univ.ML] **) @@ -328,12 +328,12 @@ by (rule Ord_cardinal_eqpoll [THEN cardinal_cong]) lemma well_ord_cardinal_eqE: - assumes woX: "well_ord(X,r)" and woY: "well_ord(Y,s)" and eq: "|X| = |Y|" + assumes woX: "well_ord(X,r)" and woY: "well_ord(Y,s)" and eq: "|X| = |Y|" shows "X \ Y" proof - - have "X \ |X|" by (blast intro: well_ord_cardinal_eqpoll [OF woX] eqpoll_sym) + have "X \ |X|" by (blast intro: well_ord_cardinal_eqpoll [OF woX] eqpoll_sym) also have "... = |Y|" by (rule eq) - also have "... \ Y" by (rule well_ord_cardinal_eqpoll [OF woY]) + also have "... \ Y" by (rule well_ord_cardinal_eqpoll [OF woY]) finally show ?thesis . qed @@ -413,45 +413,45 @@ next case True --{*real case: @{term A} is isomorphic to some ordinal*} then obtain i where i: "Ord(i)" "i \ A" by blast - show ?thesis + show ?thesis proof (rule CardI [OF Ord_Least], rule notI) fix j - assume j: "j < (\ i. i \ A)" + assume j: "j < (\ i. i \ A)" assume "j \ (\ i. i \ A)" also have "... \ A" using i by (auto intro: LeastI) finally have "j \ A" . - thus False + thus False by (rule less_LeastE [OF _ j]) qed qed qed (*Kunen's Lemma 10.5*) -lemma cardinal_eq_lemma: +lemma cardinal_eq_lemma: assumes i:"|i| \ j" and j: "j \ i" shows "|j| = |i|" proof (rule eqpollI [THEN cardinal_cong]) show "j \ i" by (rule le_imp_lepoll [OF j]) next have Oi: "Ord(i)" using j by (rule le_Ord2) - hence "i \ |i|" - by (blast intro: Ord_cardinal_eqpoll eqpoll_sym) - also have "... \ j" - by (blast intro: le_imp_lepoll i) + hence "i \ |i|" + by (blast intro: Ord_cardinal_eqpoll eqpoll_sym) + also have "... \ j" + by (blast intro: le_imp_lepoll i) finally show "i \ j" . qed -lemma cardinal_mono: +lemma cardinal_mono: assumes ij: "i \ j" shows "|i| \ |j|" proof (cases rule: Ord_linear_le [OF Ord_cardinal Ord_cardinal]) assume "|i| \ |j|" thus ?thesis . next assume cj: "|j| \ |i|" have i: "Ord(i)" using ij - by (simp add: lt_Ord) - have ci: "|i| \ j" - by (blast intro: Ord_cardinal_le ij le_trans i) - have "|i| = ||i||" - by (auto simp add: Ord_cardinal_idem i) + by (simp add: lt_Ord) + have ci: "|i| \ j" + by (blast intro: Ord_cardinal_le ij le_trans i) + have "|i| = ||i||" + by (auto simp add: Ord_cardinal_idem i) also have "... = |j|" by (rule cardinal_eq_lemma [OF cj ci]) finally have "|i| = |j|" . @@ -482,11 +482,11 @@ assume BA: "|B| \ |A|" from lepoll_well_ord [OF AB wB] obtain s where s: "well_ord(A, s)" by blast - have "B \ |B|" by (blast intro: wB eqpoll_sym well_ord_cardinal_eqpoll) + have "B \ |B|" by (blast intro: wB eqpoll_sym well_ord_cardinal_eqpoll) also have "... \ |A|" by (rule le_imp_lepoll [OF BA]) also have "... \ A" by (rule well_ord_cardinal_eqpoll [OF s]) finally have "B \ A" . - hence "A \ B" by (blast intro: eqpollI AB) + hence "A \ B" by (blast intro: eqpollI AB) hence "|A| = |B|" by (rule cardinal_cong) thus ?thesis by simp qed @@ -556,7 +556,7 @@ qed qed -lemma nat_eqpoll_iff: "[| m \ nat; n: nat |] ==> m \ n \ m = n" +lemma nat_eqpoll_iff: "[| m \ nat; n \ nat |] ==> m \ n \ m = n" apply (rule iffI) apply (blast intro: nat_lepoll_imp_le le_anti_sym elim!: eqpollE) apply (simp add: eqpoll_refl) @@ -564,7 +564,7 @@ (*The object of all this work: every natural number is a (finite) cardinal*) lemma nat_into_Card: - "n: nat ==> Card(n)" + "n \ nat ==> Card(n)" apply (unfold Card_def cardinal_def) apply (subst Least_equality) apply (rule eqpoll_refl) @@ -601,13 +601,13 @@ assumes A: "A \ m" and m: "m \ nat" shows "A \ succ(m)" proof - - { assume "A \ succ(m)" + { assume "A \ succ(m)" hence "succ(m) \ A" by (rule eqpoll_sym) also have "... \ m" by (rule A) finally have "succ(m) \ m" . hence False by (rule succ_lepoll_natE) (rule m) } moreover have "A \ succ(m)" by (blast intro: lepoll_trans A lepoll_succ) - ultimately show ?thesis by (auto simp add: lesspoll_def) + ultimately show ?thesis by (auto simp add: lesspoll_def) qed lemma lesspoll_succ_imp_lepoll: @@ -642,7 +642,7 @@ proof - { assume i: "i \ n" have "succ(n) \ i" using n - by (elim ltE, blast intro: Ord_succ_subsetI [THEN subset_imp_lepoll]) + by (elim ltE, blast intro: Ord_succ_subsetI [THEN subset_imp_lepoll]) also have "... \ n" by (rule i) finally have "succ(n) \ n" . hence False by (rule succ_lepoll_natE) (rule n) } @@ -657,13 +657,13 @@ next assume "i < n" hence "i \ nat" by (rule lt_nat_in_nat) (rule n) - thus ?thesis by (simp add: nat_eqpoll_iff n) + thus ?thesis by (simp add: nat_eqpoll_iff n) next assume "i = n" - thus ?thesis by (simp add: eqpoll_refl) + thus ?thesis by (simp add: eqpoll_refl) next assume "n < i" - hence "~ i \ n" using n by (rule lt_not_lepoll) + hence "~ i \ n" using n by (rule lt_not_lepoll) hence "~ i \ n" using n by (blast intro: eqpoll_imp_lepoll) moreover have "i \ n" using `n nat" - hence "~ nat \ i" - by (simp add: lt_def lt_not_lepoll) - hence False using i + assume i: "i < nat" "i \ nat" + hence "~ nat \ i" + by (simp add: lt_def lt_not_lepoll) + hence False using i by (simp add: eqpoll_iff) } - hence "(\ i. i \ nat) = nat" by (blast intro: Least_equality eqpoll_refl) + hence "(\ i. i \ nat) = nat" by (blast intro: Least_equality eqpoll_refl) thus ?thesis - by (auto simp add: Card_def cardinal_def) + by (auto simp add: Card_def cardinal_def) qed (*Allows showing that |i| is a limit cardinal*) @@ -701,7 +701,7 @@ "[| A \ B; b \ B |] ==> cons(a,A) \ cons(b,B)" apply (unfold lepoll_def, safe) apply (rule_tac x = "\y\cons (a,A) . if y=a then b else f`y" in exI) -apply (rule_tac d = "%z. if z:B then converse (f) `z else a" in lam_injective) +apply (rule_tac d = "%z. if z \ B then converse (f) `z else a" in lam_injective) apply (safe elim!: consE') apply simp_all apply (blast intro: inj_is_fun [THEN apply_type])+ @@ -756,11 +756,11 @@ done lemma inj_disjoint_eqpoll: - "[| f: inj(A,B); A \ B = 0 |] ==> A \ (B - range(f)) \ B" + "[| f \ inj(A,B); A \ B = 0 |] ==> A \ (B - range(f)) \ B" apply (unfold eqpoll_def) apply (rule exI) -apply (rule_tac c = "%x. if x:A then f`x else x" - and d = "%y. if y: range (f) then converse (f) `y else y" +apply (rule_tac c = "%x. if x \ A then f`x else x" + and d = "%y. if y \ range (f) then converse (f) `y else y" in lam_bijective) apply (blast intro!: if_type inj_is_fun [THEN apply_type]) apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype]) @@ -774,7 +774,7 @@ (*New proofs using cons_lepoll_cons. Could generalise from succ to cons.*) -text{*If @{term A} has at most @{term"n+1"} elements and @{term"a \ A"} +text{*If @{term A} has at most @{term"n+1"} elements and @{term"a \ A"} then @{term"A-{a}"} has at most @{term n}.*} lemma Diff_sing_lepoll: "[| a \ A; A \ succ(n) |] ==> A - {a} \ n" @@ -790,11 +790,11 @@ proof - have "cons(n,n) \ A" using A by (unfold succ_def) - also have "... \ cons(a, A-{a})" + also have "... \ cons(a, A-{a})" by (blast intro: subset_imp_lepoll) finally have "cons(n,n) \ cons(a, A-{a})" . thus ?thesis - by (blast intro: cons_lepoll_consD mem_irrefl) + by (blast intro: cons_lepoll_consD mem_irrefl) qed lemma Diff_sing_eqpoll: "[| a \ A; A \ succ(n) |] ==> A - {a} \ n" @@ -855,8 +855,8 @@ lemma lepoll_Finite: assumes Y: "Y \ X" and X: "Finite(X)" shows "Finite(Y)" proof - - obtain n where n: "n \ nat" "X \ n" using X - by (auto simp add: Finite_def) + obtain n where n: "n \ nat" "X \ n" using X + by (auto simp add: Finite_def) have "Y \ X" by (rule Y) also have "... \ n" by (rule n) finally have "Y \ n" . @@ -872,7 +872,7 @@ lemma Finite_cons: "Finite(x) ==> Finite(cons(y,x))" apply (unfold Finite_def) -apply (case_tac "y:x") +apply (case_tac "y \ x") apply (simp add: cons_absorb) apply (erule bexE) apply (rule bexI) @@ -936,10 +936,10 @@ apply (blast intro: eqpoll_trans eqpoll_sym) done -lemma Fin_lemma [rule_format]: "n: nat ==> \A. A \ n \ A \ Fin(A)" +lemma Fin_lemma [rule_format]: "n \ nat ==> \A. A \ n \ A \ Fin(A)" apply (induct_tac n) apply (simp add: eqpoll_0_iff, clarify) -apply (subgoal_tac "\u. u:A") +apply (subgoal_tac "\u. u \ A") apply (erule exE) apply (rule Diff_sing_eqpoll [elim_format]) prefer 2 apply assumption @@ -1004,7 +1004,7 @@ [| ~Finite(A); Finite(B) |] ==> ~Finite(A-B) *) lemma Diff_Finite [rule_format]: "Finite(B) ==> Finite(A-B) \ Finite(A)" apply (erule Finite_induct, auto) -apply (case_tac "x:A") +apply (case_tac "x \ A") apply (subgoal_tac [2] "A-cons (x, B) = A - B") apply (subgoal_tac "A - cons (x, B) = (A - B) - {x}", simp) apply (drule Diff_sing_Finite, auto) @@ -1060,8 +1060,8 @@ apply (blast intro: wf_onI) apply (rule wf_onI) apply (simp add: wf_on_def wf_def) -apply (case_tac "x:Z") - txt{*x:Z case*} +apply (case_tac "x \ Z") + txt{*true case*} apply (drule_tac x = x in bspec, assumption) apply (blast elim: mem_irrefl mem_asym) txt{*other case*} @@ -1085,15 +1085,15 @@ done lemma ordertype_eq_n: - assumes r: "well_ord(A,r)" and A: "A \ n" and n: "n \ nat" + assumes r: "well_ord(A,r)" and A: "A \ n" and n: "n \ nat" shows "ordertype(A,r) = n" proof - - have "ordertype(A,r) \ A" - by (blast intro: bij_imp_eqpoll bij_converse_bij ordermap_bij r) + have "ordertype(A,r) \ A" + by (blast intro: bij_imp_eqpoll bij_converse_bij ordermap_bij r) also have "... \ n" by (rule A) finally have "ordertype(A,r) \ n" . thus ?thesis - by (simp add: Ord_nat_eqpoll_iff Ord_ordertype n r) + by (simp add: Ord_nat_eqpoll_iff Ord_ordertype n r) qed lemma Finite_well_ord_converse: @@ -1112,15 +1112,15 @@ proof - { fix n assume n: "n \ nat" "nat \ n" - have "n \ nat" by (rule n) + have "n \ nat" by (rule n) also have "... = n" using n - by (simp add: Ord_nat_eqpoll_iff Ord_nat) + by (simp add: Ord_nat_eqpoll_iff Ord_nat) finally have "n \ n" . - hence False - by (blast elim: mem_irrefl) + hence False + by (blast elim: mem_irrefl) } thus ?thesis - by (auto simp add: Finite_def) + by (auto simp add: Finite_def) qed end diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/CardinalArith.thy Thu Mar 15 16:35:02 2012 +0000 @@ -31,7 +31,7 @@ --{*This def is more complex than Kunen's but it more easily proved to be a cardinal*} "jump_cardinal(K) == - \X\Pow(K). {z. r: Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}" + \X\Pow(K). {z. r \ Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}" definition csucc :: "i=>i" where @@ -48,10 +48,10 @@ cmult (infixl "\" 70) -lemma Card_Union [simp,intro,TC]: +lemma Card_Union [simp,intro,TC]: assumes A: "\x. x\A \ Card(x)" shows "Card(\(A))" proof (rule CardI) - show "Ord(\A)" using A + show "Ord(\A)" using A by (simp add: Card_is_Ord) next fix j @@ -60,24 +60,24 @@ by (auto simp add: lt_def intro: Card_is_Ord) then obtain c where c: "c\A" "j < c" "Card(c)" by blast - hence jls: "j \ c" - by (simp add: lt_Card_imp_lesspoll) + hence jls: "j \ c" + by (simp add: lt_Card_imp_lesspoll) { assume eqp: "j \ \A" have "c \ \A" using c by (blast intro: subset_imp_lepoll) also have "... \ j" by (rule eqpoll_sym [OF eqp]) also have "... \ c" by (rule jls) finally have "c \ c" . - hence False + hence False by auto } thus "\ j \ \A" by blast qed -lemma Card_UN: "(!!x. x:A ==> Card(K(x))) ==> Card(\x\A. K(x))" +lemma Card_UN: "(!!x. x \ A ==> Card(K(x))) ==> Card(\x\A. K(x))" by blast lemma Card_OUN [simp,intro,TC]: - "(!!x. x:A ==> Card(K(x))) ==> Card(\x A ==> Card(K(x))) ==> Card(\x K |] ==> b \ K" @@ -99,7 +99,7 @@ lemma sum_commute_eqpoll: "A+B \ B+A" proof (unfold eqpoll_def, rule exI) show "(\z\A+B. case(Inr,Inl,z)) \ bij(A+B, B+A)" - by (auto intro: lam_bijective [where d = "case(Inr,Inl)"]) + by (auto intro: lam_bijective [where d = "case(Inr,Inl)"]) qed lemma cadd_commute: "i \ j = j \ i" @@ -121,11 +121,11 @@ shows "(i \ j) \ k = i \ (j \ k)" proof (unfold cadd_def, rule cardinal_cong) have "|i + j| + k \ (i + j) + k" - by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) + by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) also have "... \ i + (j + k)" - by (rule sum_assoc_eqpoll) + by (rule sum_assoc_eqpoll) also have "... \ i + |j + k|" - by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd j k eqpoll_sym) + by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd j k eqpoll_sym) finally show "|i + j| + k \ i + |j + k|" . qed @@ -148,7 +148,7 @@ lemma sum_lepoll_self: "A \ A+B" proof (unfold lepoll_def, rule exI) show "(\x\A. Inl (x)) \ inj(A, A + B)" - by (simp add: inj_def) + by (simp add: inj_def) qed (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) @@ -157,12 +157,12 @@ assumes K: "Card(K)" and L: "Ord(L)" shows "K \ (K \ L)" proof (unfold cadd_def) have "K \ |K|" - by (rule Card_cardinal_le [OF K]) + by (rule Card_cardinal_le [OF K]) moreover have "|K| \ |K + L|" using K L by (blast intro: well_ord_lepoll_imp_Card_le sum_lepoll_self - well_ord_radd well_ord_Memrel Card_is_Ord) - ultimately show "K \ |K + L|" - by (blast intro: le_trans) + well_ord_radd well_ord_Memrel Card_is_Ord) + ultimately show "K \ |K + L|" + by (blast intro: le_trans) qed subsubsection{*Monotonicity of addition*} @@ -197,7 +197,7 @@ apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+ done -(*Pulling the succ(...) outside the |...| requires m, n: nat *) +(*Pulling the succ(...) outside the |...| requires m, n \ nat *) (*Unconditional version requires AC*) lemma cadd_succ_lemma: assumes "Ord(m)" "Ord(n)" shows "succ(m) \ n = |succ(m \ n)|" @@ -206,14 +206,14 @@ by (blast intro: eqpoll_sym well_ord_cardinal_eqpoll well_ord_radd well_ord_Memrel) have "|succ(m) + n| = |succ(m + n)|" - by (rule sum_succ_eqpoll [THEN cardinal_cong]) - also have "... = |succ(|m + n|)|" + by (rule sum_succ_eqpoll [THEN cardinal_cong]) + also have "... = |succ(|m + n|)|" by (blast intro: succ_eqpoll_cong cardinal_cong) finally show "|succ(m) + n| = |succ(|m + n|)|" . qed lemma nat_cadd_eq_add: - assumes m: "m: nat" and [simp]: "n: nat" shows"m \ n = m #+ n" + assumes m: "m \ nat" and [simp]: "n \ nat" shows"m \ n = m #+ n" using m proof (induct m) case 0 thus ?case by (simp add: nat_into_Card cadd_0) @@ -252,11 +252,11 @@ shows "(i \ j) \ k = i \ (j \ k)" proof (unfold cmult_def, rule cardinal_cong) have "|i * j| * k \ (i * j) * k" - by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult i j) + by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult i j) also have "... \ i * (j * k)" - by (rule prod_assoc_eqpoll) + by (rule prod_assoc_eqpoll) also have "... \ i * |j * k|" - by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult j k eqpoll_sym) + by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult j k eqpoll_sym) finally show "|i * j| * k \ i * |j * k|" . qed @@ -273,11 +273,11 @@ shows "(i \ j) \ k = (i \ k) \ (j \ k)" proof (unfold cadd_def cmult_def, rule cardinal_cong) have "|i + j| * k \ (i + j) * k" - by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) + by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) also have "... \ i * k + j * k" - by (rule sum_prod_distrib_eqpoll) + by (rule sum_prod_distrib_eqpoll) also have "... \ |i * k| + |j * k|" - by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll well_ord_rmult i j k eqpoll_sym) + by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll well_ord_rmult i j k eqpoll_sym) finally show "|i + j| * k \ |i * k| + |j * k|" . qed @@ -324,7 +324,7 @@ subsubsection{*Multiplication by a non-zero cardinal*} -lemma prod_lepoll_self: "b: B ==> A \ A*B" +lemma prod_lepoll_self: "b \ B ==> A \ A*B" apply (unfold lepoll_def inj_def) apply (rule_tac x = "\x\A. " in exI, simp) done @@ -381,7 +381,7 @@ apply (blast intro: well_ord_rmult well_ord_Memrel) done -lemma nat_cmult_eq_mult: "[| m: nat; n: nat |] ==> m \ n = m#*n" +lemma nat_cmult_eq_mult: "[| m \ nat; n \ nat |] ==> m \ n = m#*n" apply (induct_tac m) apply (simp_all add: cmult_succ_lemma nat_cadd_eq_add) done @@ -389,13 +389,13 @@ lemma cmult_2: "Card(n) ==> 2 \ n = n \ n" by (simp add: cmult_succ_lemma Card_is_Ord cadd_commute [of _ 0]) -lemma sum_lepoll_prod: +lemma sum_lepoll_prod: assumes C: "2 \ C" shows "B+B \ C*B" proof - have "B+B \ 2*B" - by (simp add: sum_eq_2_times) + by (simp add: sum_eq_2_times) also have "... \ C*B" - by (blast intro: prod_lepoll_mono lepoll_refl C) + by (blast intro: prod_lepoll_mono lepoll_refl C) finally show "B+B \ C*B" . qed @@ -407,18 +407,18 @@ (*This proof is modelled upon one assuming nat<=A, with injection \z\cons(u,A). if z=u then 0 else if z \ nat then succ(z) else z - and inverse %y. if y:nat then nat_case(u, %z. z, y) else y. \ - If f: inj(nat,A) then range(f) behaves like the natural numbers.*) + and inverse %y. if y \ nat then nat_case(u, %z. z, y) else y. \ + If f \ inj(nat,A) then range(f) behaves like the natural numbers.*) lemma nat_cons_lepoll: "nat \ A ==> cons(u,A) \ A" apply (unfold lepoll_def) apply (erule exE) apply (rule_tac x = "\z\cons (u,A). if z=u then f`0 - else if z: range (f) then f`succ (converse (f) `z) else z" + else if z \ range (f) then f`succ (converse (f) `z) else z" in exI) apply (rule_tac d = - "%y. if y: range(f) then nat_case (u, %z. f`z, converse(f) `y) + "%y. if y \ range(f) then nat_case (u, %z. f`z, converse(f) `y) else y" in lam_injective) apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun) @@ -475,7 +475,7 @@ (*A general fact about ordermap*) lemma ordermap_eqpoll_pred: - "[| well_ord(A,r); x:A |] ==> ordermap(A,r)`x \ Order.pred(A,x,r)" + "[| well_ord(A,r); x \ A |] ==> ordermap(A,r)`x \ Order.pred(A,x,r)" apply (unfold eqpoll_def) apply (rule exI) apply (simp add: ordermap_eq_image well_ord_is_wf) @@ -486,7 +486,7 @@ subsubsection{*Establishing the well-ordering*} -lemma well_ord_csquare: +lemma well_ord_csquare: assumes K: "Ord(K)" shows "well_ord(K*K, csquare_rel(K))" proof (unfold csquare_rel_def, rule well_ord_rvimage) show "(\\x,y\\K \ K. \x \ y, x, y\) \ inj(K \ K, K \ K \ K)" using K @@ -553,23 +553,23 @@ text{*Kunen: "each @{term"\x,y\ \ K \ K"} has no more than @{term"z \ z"} predecessors..." (page 29) *} lemma ordermap_csquare_le: - assumes K: "Limit(K)" and x: "x succ(x \ y)" shows "|ordermap(K \ K, csquare_rel(K)) ` \x,y\| \ |succ(z)| \ |succ(z)|" proof (unfold cmult_def, rule well_ord_lepoll_imp_Card_le) - show "well_ord(|succ(z)| \ |succ(z)|, + show "well_ord(|succ(z)| \ |succ(z)|, rmult(|succ(z)|, Memrel(|succ(z)|), |succ(z)|, Memrel(|succ(z)|)))" - by (blast intro: Ord_cardinal well_ord_Memrel well_ord_rmult) + by (blast intro: Ord_cardinal well_ord_Memrel well_ord_rmult) next have zK: "z K, csquare_rel(K)) ` \x,y\ \ ordermap(K \ K, csquare_rel(K)) ` \z,z\" using z_def - by (blast intro: ordermap_z_lt leI le_imp_lepoll K x y) + by (blast intro: ordermap_z_lt leI le_imp_lepoll K x y) also have "... \ Order.pred(K \ K, \z,z\, csquare_rel(K))" proof (rule ordermap_eqpoll_pred) - show "well_ord(K \ K, csquare_rel(K))" using K + show "well_ord(K \ K, csquare_rel(K))" using K by (rule Limit_is_Ord [THEN well_ord_csquare]) next show "\z, z\ \ K \ K" using zK @@ -578,7 +578,7 @@ also have "... \ succ(z) \ succ(z)" using zK by (rule pred_csquare_subset [THEN subset_imp_lepoll]) also have "... \ |succ(z)| \ |succ(z)|" using oz - by (blast intro: prod_eqpoll_cong Ord_succ Ord_cardinal_eqpoll eqpoll_sym) + by (blast intro: prod_eqpoll_cong Ord_succ Ord_cardinal_eqpoll eqpoll_sym) finally show "ordermap(K \ K, csquare_rel(K)) ` \x,y\ \ |succ(z)| \ |succ(z)|" . qed @@ -587,8 +587,8 @@ assumes IK: "InfCard(K)" and eq: "\y. y\K \ InfCard(y) \ y \ y = y" shows "ordertype(K*K, csquare_rel(K)) \ K" proof - - have CK: "Card(K)" using IK by (rule InfCard_is_Card) - hence OK: "Ord(K)" by (rule Card_is_Ord) + have CK: "Card(K)" using IK by (rule InfCard_is_Card) + hence OK: "Ord(K)" by (rule Card_is_Ord) moreover have "Ord(ordertype(K \ K, csquare_rel(K)))" using OK by (rule well_ord_csquare [THEN Ord_ordertype]) ultimately show ?thesis @@ -596,18 +596,18 @@ fix i assume i: "i < ordertype(K \ K, csquare_rel(K))" hence Oi: "Ord(i)" by (elim ltE) - obtain x y where x: "x \ K" and y: "y \ K" + obtain x y where x: "x \ K" and y: "y \ K" and ieq: "i = ordermap(K \ K, csquare_rel(K)) ` \x,y\" using i by (auto simp add: ordertype_unfold elim: ltE) - hence xy: "Ord(x)" "Ord(y)" "x < K" "y < K" using OK + hence xy: "Ord(x)" "Ord(y)" "x < K" "y < K" using OK by (blast intro: Ord_in_Ord ltI)+ hence ou: "Ord(x \ y)" - by (simp add: Ord_Un) + by (simp add: Ord_Un) show "i < K" proof (rule Card_lt_imp_lt [OF _ Oi CK]) have "|i| \ |succ(succ(x \ y))| \ |succ(succ(x \ y))|" using IK xy by (auto simp add: ieq intro: InfCard_is_Limit [THEN ordermap_csquare_le]) - moreover have "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" + moreover have "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" proof (cases rule: Ord_linear2 [OF ou Ord_nat]) assume "x \ y < nat" hence "|succ(succ(x \ y))| \ |succ(succ(x \ y))| \ nat" @@ -615,46 +615,46 @@ nat_into_Card [THEN Card_cardinal_eq] Ord_nat) also have "... \ K" using IK by (simp add: InfCard_def le_imp_subset) - finally show "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" - by (simp add: ltI OK) + finally show "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" + by (simp add: ltI OK) next assume natxy: "nat \ x \ y" - hence seq: "|succ(succ(x \ y))| = |x \ y|" using xy + hence seq: "|succ(succ(x \ y))| = |x \ y|" using xy by (simp add: le_imp_subset nat_succ_eqpoll [THEN cardinal_cong] le_succ_iff) - also have "... < K" using xy + also have "... < K" using xy by (simp add: Un_least_lt Ord_cardinal_le [THEN lt_trans1]) finally have "|succ(succ(x \ y))| < K" . moreover have "InfCard(|succ(succ(x \ y))|)" using xy natxy by (simp add: seq InfCard_def Card_cardinal nat_le_cardinal) - ultimately show ?thesis by (simp add: eq ltD) + ultimately show ?thesis by (simp add: eq ltD) qed - ultimately show "|i| < K" by (blast intro: lt_trans1) + ultimately show "|i| < K" by (blast intro: lt_trans1) qed qed qed (*Main result: Kunen's Theorem 10.12*) -lemma InfCard_csquare_eq: +lemma InfCard_csquare_eq: assumes IK: "InfCard(K)" shows "InfCard(K) ==> K \ K = K" proof - - have OK: "Ord(K)" using IK by (simp add: Card_is_Ord InfCard_is_Card) + have OK: "Ord(K)" using IK by (simp add: Card_is_Ord InfCard_is_Card) show "InfCard(K) ==> K \ K = K" using OK proof (induct rule: trans_induct) case (step i) show "i \ i = i" proof (rule le_anti_sym) - have "|i \ i| = |ordertype(i \ i, csquare_rel(i))|" - by (rule cardinal_cong, + have "|i \ i| = |ordertype(i \ i, csquare_rel(i))|" + by (rule cardinal_cong, simp add: step.hyps well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll]) - hence "i \ i \ ordertype(i \ i, csquare_rel(i))" + hence "i \ i \ ordertype(i \ i, csquare_rel(i))" by (simp add: step.hyps cmult_def Ord_cardinal_le well_ord_csquare [THEN Ord_ordertype]) moreover have "ordertype(i \ i, csquare_rel(i)) \ i" using step - by (simp add: ordertype_csquare_le) + by (simp add: ordertype_csquare_le) ultimately show "i \ i \ i" by (rule le_trans) next show "i \ i \ i" using step - by (blast intro: cmult_square_le InfCard_is_Card) + by (blast intro: cmult_square_le InfCard_is_Card) qed qed qed @@ -664,7 +664,7 @@ assumes r: "well_ord(A,r)" and I: "InfCard(|A|)" shows "A \ A \ A" proof - have "A \ A \ |A| \ |A|" - by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_sym r) + by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_sym r) also have "... \ A" proof (rule well_ord_cardinal_eqE [OF _ r]) show "well_ord(|A| \ |A|, rmult(|A|, Memrel(|A|), |A|, Memrel(|A|)))" @@ -672,7 +672,7 @@ next show "||A| \ |A|| = |A|" using InfCard_csquare_eq I by (simp add: cmult_def) - qed + qed finally show ?thesis . qed @@ -842,21 +842,21 @@ { fix X have "Finite(X) ==> a \ X \ cons(a,X) \ X \ False" proof (induct X rule: Finite_induct) - case 0 thus False by (simp add: lepoll_0_iff) + case 0 thus False by (simp add: lepoll_0_iff) next - case (cons x Y) - hence "cons(x, cons(a, Y)) \ cons(x, Y)" by (simp add: cons_commute) + case (cons x Y) + hence "cons(x, cons(a, Y)) \ cons(x, Y)" by (simp add: cons_commute) hence "cons(a, Y) \ Y" using cons by (blast dest: cons_lepoll_consD) thus False using cons by auto qed - } + } hence [simp]: "~ cons(a,A) \ A" using a FA by auto have [simp]: "|A| \ A" using Finite_imp_well_ord [OF FA] by (blast intro: well_ord_cardinal_eqpoll) - have "(\ i. i \ cons(a, A)) = succ(|A|)" + have "(\ i. i \ cons(a, A)) = succ(|A|)" proof (rule Least_equality [OF _ _ notI]) - show "succ(|A|) \ cons(a, A)" - by (simp add: succ_def cons_eqpoll_cong mem_not_refl a) + show "succ(|A|) \ cons(a, A)" + by (simp add: succ_def cons_eqpoll_cong mem_not_refl a) next show "Ord(succ(|A|))" by simp next @@ -868,17 +868,17 @@ finally have "cons(a, A) \ A" . thus False by simp qed - thus ?thesis by (simp add: cardinal_def) + thus ?thesis by (simp add: cardinal_def) qed lemma Finite_imp_succ_cardinal_Diff: - "[| Finite(A); a:A |] ==> succ(|A-{a}|) = |A|" + "[| Finite(A); a \ A |] ==> succ(|A-{a}|) = |A|" apply (rule_tac b = A in cons_Diff [THEN subst], assumption) apply (simp add: Finite_imp_cardinal_cons Diff_subset [THEN subset_Finite]) apply (simp add: cons_Diff) done -lemma Finite_imp_cardinal_Diff: "[| Finite(A); a:A |] ==> |A-{a}| < |A|" +lemma Finite_imp_cardinal_Diff: "[| Finite(A); a \ A |] ==> |A-{a}| < |A|" apply (rule succ_leE) apply (simp add: Finite_imp_succ_cardinal_Diff) done @@ -922,11 +922,11 @@ lemmas nat_implies_well_ord = nat_into_Ord [THEN well_ord_Memrel] -lemma nat_sum_eqpoll_sum: +lemma nat_sum_eqpoll_sum: assumes m: "m \ nat" and n: "n \ nat" shows "m + n \ m #+ n" proof - have "m + n \ |m+n|" using m n - by (blast intro: nat_implies_well_ord well_ord_radd well_ord_cardinal_eqpoll eqpoll_sym) + by (blast intro: nat_implies_well_ord well_ord_radd well_ord_cardinal_eqpoll eqpoll_sym) also have "... = m #+ n" using m n by (simp add: nat_cadd_eq_add [symmetric] cadd_def) finally show ?thesis . diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Constructible/Formula.thy --- a/src/ZF/Constructible/Formula.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Constructible/Formula.thy Thu Mar 15 16:35:02 2012 +0000 @@ -13,10 +13,10 @@ consts formula :: i datatype - "formula" = Member ("x: nat", "y: nat") - | Equal ("x: nat", "y: nat") - | Nand ("p: formula", "q: formula") - | Forall ("p: formula") + "formula" = Member ("x \ nat", "y \ nat") + | Equal ("x \ nat", "y \ nat") + | Nand ("p \ formula", "q \ formula") + | Forall ("p \ formula") declare formula.intros [TC] @@ -488,7 +488,7 @@ DPow(B)"}.*} (*This may be true but the proof looks difficult, requiring relativization -lemma DPow_insert: "DPow (cons(a,A)) = DPow(A) \ {cons(a,X) . X: DPow(A)}" +lemma DPow_insert: "DPow (cons(a,A)) = DPow(A) \ {cons(a,X) . X \ DPow(A)}" apply (rule equalityI, safe) oops *) @@ -656,7 +656,7 @@ text{*This version lets us remove the premise @{term "Ord(i)"} sometimes.*} lemma Lset_mono_mem [rule_format]: - "\j. i:j \ Lset(i) \ Lset(j)" + "\j. i \ j \ Lset(i) \ Lset(j)" proof (induct i rule: eps_induct, intro allI impI) fix x j assume "\y\x. \j. y \ j \ Lset(y) \ Lset(j)" @@ -712,12 +712,12 @@ "Limit(i) ==> Lset(i) = (\y\i. Lset(y))" by (simp add: Lset_Union [symmetric] Limit_Union_eq) -lemma lt_LsetI: "[| a: Lset(j); j a \ Lset(i)" +lemma lt_LsetI: "[| a \ Lset(j); j a \ Lset(i)" by (blast dest: Lset_mono [OF le_imp_subset [OF leI]]) lemma Limit_LsetE: - "[| a: Lset(i); ~R ==> Limit(i); - !!x. [| x R + "[| a \ Lset(i); ~R ==> Limit(i); + !!x. [| x Lset(x) |] ==> R |] ==> R" apply (rule classical) apply (rule Limit_Lset_eq [THEN equalityD1, THEN subsetD, THEN UN_E]) @@ -728,7 +728,7 @@ subsubsection{* Basic closure properties *} -lemma zero_in_Lset: "y:x ==> 0 \ Lset(x)" +lemma zero_in_Lset: "y \ x ==> 0 \ Lset(x)" by (subst Lset, blast intro: empty_in_DPow) lemma notin_Lset: "x \ Lset(x)" @@ -792,15 +792,15 @@ subsubsection{* Finite sets and ordered pairs *} -lemma singleton_in_Lset: "a: Lset(i) ==> {a} \ Lset(succ(i))" +lemma singleton_in_Lset: "a \ Lset(i) ==> {a} \ Lset(succ(i))" by (simp add: Lset_succ singleton_in_DPow) lemma doubleton_in_Lset: - "[| a: Lset(i); b: Lset(i) |] ==> {a,b} \ Lset(succ(i))" + "[| a \ Lset(i); b \ Lset(i) |] ==> {a,b} \ Lset(succ(i))" by (simp add: Lset_succ empty_in_DPow cons_in_DPow) lemma Pair_in_Lset: - "[| a: Lset(i); b: Lset(i); Ord(i) |] ==> \ Lset(succ(succ(i)))" + "[| a \ Lset(i); b \ Lset(i); Ord(i) |] ==> \ Lset(succ(succ(i)))" apply (unfold Pair_def) apply (blast intro: doubleton_in_Lset) done @@ -808,9 +808,9 @@ lemmas Lset_UnI1 = Un_upper1 [THEN Lset_mono [THEN subsetD]] lemmas Lset_UnI2 = Un_upper2 [THEN Lset_mono [THEN subsetD]] -text{*Hard work is finding a single j:i such that {a,b}<=Lset(j)*} +text{*Hard work is finding a single @{term"j \ i"} such that @{term"{a,b} \ Lset(j)"}*} lemma doubleton_in_LLimit: - "[| a: Lset(i); b: Lset(i); Limit(i) |] ==> {a,b} \ Lset(i)" + "[| a \ Lset(i); b \ Lset(i); Limit(i) |] ==> {a,b} \ Lset(i)" apply (erule Limit_LsetE, assumption) apply (erule Limit_LsetE, assumption) apply (blast intro: lt_LsetI [OF doubleton_in_Lset] @@ -824,7 +824,7 @@ done lemma Pair_in_LLimit: - "[| a: Lset(i); b: Lset(i); Limit(i) |] ==> \ Lset(i)" + "[| a \ Lset(i); b \ Lset(i); Limit(i) |] ==> \ Lset(i)" txt{*Infer that a, b occur at ordinals x,xa < i.*} apply (erule Limit_LsetE, assumption) apply (erule Limit_LsetE, assumption) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Constructible/Relative.thy Thu Mar 15 16:35:02 2012 +0000 @@ -745,7 +745,7 @@ (*The first premise can't simply be assumed as a schema. It is essential to take care when asserting instances of Replacement. Let K be a nonconstructible subset of nat and define - f(x) = x if x:K and f(x)=0 otherwise. Then RepFun(nat,f) = cons(0,K), a + f(x) = x if x \ K and f(x)=0 otherwise. Then RepFun(nat,f) = cons(0,K), a nonconstructible set. So we cannot assume that M(X) implies M(RepFun(X,f)) even for f \ M -> M. *) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Epsilon.thy Thu Mar 15 16:35:02 2012 +0000 @@ -67,7 +67,7 @@ lemmas arg_into_eclose_sing = arg_in_eclose_sing [THEN ecloseD] (* This is epsilon-induction for eclose(A); see also eclose_induct_down... - [| a: eclose(A); !!x. [| x: eclose(A); \y\x. P(y) |] ==> P(x) + [| a \ eclose(A); !!x. [| x \ eclose(A); \y\x. P(y) |] ==> P(x) |] ==> P(a) *) lemmas eclose_induct = @@ -85,7 +85,7 @@ (** eclose(A) is the least transitive set including A as a subset. **) lemma eclose_least_lemma: - "[| Transset(X); A<=X; n: nat |] ==> nat_rec(n, A, %m r. \(r)) \ X" + "[| Transset(X); A<=X; n \ nat |] ==> nat_rec(n, A, %m r. \(r)) \ X" apply (unfold Transset_def) apply (erule nat_induct) apply (simp add: nat_rec_0) @@ -100,9 +100,9 @@ (*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) lemma eclose_induct_down [consumes 1]: - "[| a: eclose(b); - !!y. [| y: b |] ==> P(y); - !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) + "[| a \ eclose(b); + !!y. [| y \ b |] ==> P(y); + !!y z. [| y \ eclose(b); P(y); z \ y |] ==> P(z) |] ==> P(a)" apply (rule eclose_least [THEN subsetD, THEN CollectD2, of "eclose(b)"]) prefer 3 apply assumption @@ -131,17 +131,17 @@ subsection{*Epsilon Recursion*} (*Unused...*) -lemma mem_eclose_trans: "[| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)" +lemma mem_eclose_trans: "[| A \ eclose(B); B \ eclose(C) |] ==> A \ eclose(C)" by (rule eclose_least [OF Transset_eclose eclose_subset, THEN subsetD], assumption+) (*Variant of the previous lemma in a useable form for the sequel*) lemma mem_eclose_sing_trans: - "[| A: eclose({B}); B: eclose({C}) |] ==> A: eclose({C})" + "[| A \ eclose({B}); B \ eclose({C}) |] ==> A \ eclose({C})" by (rule eclose_least [OF Transset_eclose singleton_subsetI, THEN subsetD], assumption+) -lemma under_Memrel: "[| Transset(i); j:i |] ==> Memrel(i)-``{j} = j" +lemma under_Memrel: "[| Transset(i); j \ i |] ==> Memrel(i)-``{j} = j" by (unfold Transset_def, blast) lemma lt_Memrel: "j < i ==> Memrel(i) -`` {j} = j" @@ -153,7 +153,7 @@ lemmas wfrec_ssubst = wf_Memrel [THEN wfrec, THEN ssubst] lemma wfrec_eclose_eq: - "[| k:eclose({j}); j:eclose({i}) |] ==> + "[| k \ eclose({j}); j \ eclose({i}) |] ==> wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)" apply (erule eclose_induct) apply (rule wfrec_ssubst) @@ -162,7 +162,7 @@ done lemma wfrec_eclose_eq2: - "k: i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)" + "k \ i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)" apply (rule arg_in_eclose_sing [THEN wfrec_eclose_eq]) apply (erule arg_into_eclose_sing) done @@ -181,7 +181,7 @@ done lemma transrec_type: - "[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) \ B(x) |] + "[| !!x u. [| x \ eclose({a}); u \ Pi(x,B) |] ==> H(x,u) \ B(x) |] ==> transrec(a,H) \ B(a)" apply (rule_tac i = a in arg_in_eclose_sing [THEN eclose_induct]) apply (subst transrec) @@ -205,9 +205,9 @@ done lemma Ord_transrec_type: - assumes jini: "j: i" + assumes jini: "j \ i" and ordi: "Ord(i)" - and minor: " !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) \ B(x)" + and minor: " !!x u. [| x \ i; u \ Pi(x,B) |] ==> H(x,u) \ B(x)" shows "transrec(j,H) \ B(j)" apply (rule transrec_type) apply (insert jini ordi) @@ -235,13 +235,13 @@ apply (simp add: Ord_equality) done -lemma rank_lt: "a:b ==> rank(a) < rank(b)" +lemma rank_lt: "a \ b ==> rank(a) < rank(b)" apply (rule_tac a1 = b in rank [THEN ssubst]) apply (erule UN_I [THEN ltI]) apply (rule_tac [2] Ord_UN, auto) done -lemma eclose_rank_lt: "a: eclose(b) ==> rank(a) < rank(b)" +lemma eclose_rank_lt: "a \ eclose(b) ==> rank(a) < rank(b)" apply (erule eclose_induct_down) apply (erule rank_lt) apply (erule rank_lt [THEN lt_trans], assumption) @@ -321,7 +321,7 @@ subsection{*Corollaries of Leastness*} -lemma mem_eclose_subset: "A:B ==> eclose(A)<=eclose(B)" +lemma mem_eclose_subset: "A \ B ==> eclose(A)<=eclose(B)" apply (rule Transset_eclose [THEN eclose_least]) apply (erule arg_into_eclose [THEN eclose_subset]) done @@ -390,9 +390,9 @@ done lemma rec_type: - "[| n: nat; - a: C(0); - !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) |] + "[| n \ nat; + a \ C(0); + !!m z. [| m \ nat; z \ C(m) |] ==> b(m,z): C(succ(m)) |] ==> rec(n,a,b) \ C(n)" by (erule nat_induct, auto) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/EquivClass.thy --- a/src/ZF/EquivClass.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/EquivClass.thy Thu Mar 15 16:35:02 2012 +0000 @@ -9,7 +9,7 @@ definition quotient :: "[i,i]=>i" (infixl "'/'/" 90) (*set of equiv classes*) where - "A//r == {r``{x} . x:A}" + "A//r == {r``{x} . x \ A}" definition congruent :: "[i,i=>i]=>o" where @@ -72,15 +72,15 @@ done lemma equiv_class_self: - "[| equiv(A,r); a: A |] ==> a: r``{a}" + "[| equiv(A,r); a \ A |] ==> a \ r``{a}" by (unfold equiv_def refl_def, blast) (*Lemma for the next result*) lemma subset_equiv_class: - "[| equiv(A,r); r``{b} \ r``{a}; b: A |] ==> : r" + "[| equiv(A,r); r``{b} \ r``{a}; b \ A |] ==> : r" by (unfold equiv_def refl_def, blast) -lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> : r" +lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b \ A |] ==> : r" by (assumption | rule equalityD2 subset_equiv_class)+ (*thus r``{a} = r``{b} as well*) @@ -92,24 +92,24 @@ by (unfold equiv_def, blast) lemma equiv_class_eq_iff: - "equiv(A,r) ==> : r \ r``{x} = r``{y} & x:A & y:A" + "equiv(A,r) ==> : r \ r``{x} = r``{y} & x \ A & y \ A" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) lemma eq_equiv_class_iff: - "[| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} \ : r" + "[| equiv(A,r); x \ A; y \ A |] ==> r``{x} = r``{y} \ : r" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) (*** Quotients ***) (** Introduction/elimination rules -- needed? **) -lemma quotientI [TC]: "x:A ==> r``{x}: A//r" +lemma quotientI [TC]: "x \ A ==> r``{x}: A//r" apply (unfold quotient_def) apply (erule RepFunI) done lemma quotientE: - "[| X: A//r; !!x. [| X = r``{x}; x:A |] ==> P |] ==> P" + "[| X \ A//r; !!x. [| X = r``{x}; x \ A |] ==> P |] ==> P" by (unfold quotient_def, blast) lemma Union_quotient: @@ -117,7 +117,7 @@ by (unfold equiv_def refl_def quotient_def, blast) lemma quotient_disj: - "[| equiv(A,r); X: A//r; Y: A//r |] ==> X=Y | (X \ Y \ 0)" + "[| equiv(A,r); X \ A//r; Y \ A//r |] ==> X=Y | (X \ Y \ 0)" apply (unfold quotient_def) apply (safe intro!: equiv_class_eq, assumption) apply (unfold equiv_def trans_def sym_def, blast) @@ -130,7 +130,7 @@ (*Conversion rule*) lemma UN_equiv_class: - "[| equiv(A,r); b respects r; a: A |] ==> (\x\r``{a}. b(x)) = b(a)" + "[| equiv(A,r); b respects r; a \ A |] ==> (\x\r``{a}. b(x)) = b(a)" apply (subgoal_tac "\x \ r``{a}. b(x) = b(a)") apply simp apply (blast intro: equiv_class_self) @@ -139,19 +139,19 @@ (*type checking of @{term"\x\r``{a}. b(x)"} *) lemma UN_equiv_class_type: - "[| equiv(A,r); b respects r; X: A//r; !!x. x \ A ==> b(x) \ B |] + "[| equiv(A,r); b respects r; X \ A//r; !!x. x \ A ==> b(x) \ B |] ==> (\x\X. b(x)) \ B" apply (unfold quotient_def, safe) apply (simp (no_asm_simp) add: UN_equiv_class) done (*Sufficient conditions for injectiveness. Could weaken premises! - major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B + major premise could be an inclusion; bcong could be !!y. y \ A ==> b(y):B *) lemma UN_equiv_class_inject: "[| equiv(A,r); b respects r; - (\x\X. b(x))=(\y\Y. b(y)); X: A//r; Y: A//r; - !!x y. [| x:A; y:A; b(x)=b(y) |] ==> :r |] + (\x\X. b(x))=(\y\Y. b(y)); X \ A//r; Y \ A//r; + !!x y. [| x \ A; y \ A; b(x)=b(y) |] ==> :r |] ==> X=Y" apply (unfold quotient_def, safe) apply (rule equiv_class_eq, assumption) @@ -162,11 +162,11 @@ subsection{*Defining Binary Operations upon Equivalence Classes*} lemma congruent2_implies_congruent: - "[| equiv(A,r1); congruent2(r1,r2,b); a: A |] ==> congruent(r2,b(a))" + "[| equiv(A,r1); congruent2(r1,r2,b); a \ A |] ==> congruent(r2,b(a))" by (unfold congruent_def congruent2_def equiv_def refl_def, blast) lemma congruent2_implies_congruent_UN: - "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a: A2 |] ==> + "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \ A2 |] ==> congruent(r1, %x1. \x2 \ r2``{a}. b(x1,x2))" apply (unfold congruent_def, safe) apply (frule equiv_type [THEN subsetD], assumption) @@ -206,8 +206,8 @@ lemma congruent2_commuteI: assumes equivA: "equiv(A,r)" - and commute: "!! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y)" - and congt: "!! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z)" + and commute: "!! y z. [| y \ A; z \ A |] ==> b(y,z) = b(z,y)" + and congt: "!! y z w. [| w \ A; : r |] ==> b(w,y) = b(w,z)" shows "b respects2 r" apply (insert equivA [THEN equiv_type, THEN subsetD]) apply (rule congruent2I [OF equivA equivA]) @@ -219,9 +219,9 @@ (*Obsolete?*) lemma congruent_commuteI: - "[| equiv(A,r); Z: A//r; - !!w. [| w: A |] ==> congruent(r, %z. b(w,z)); - !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) + "[| equiv(A,r); Z \ A//r; + !!w. [| w \ A |] ==> congruent(r, %z. b(w,z)); + !!x y. [| x \ A; y \ A |] ==> b(y,x) = b(x,y) |] ==> congruent(r, %w. \z\Z. b(w,z))" apply (simp (no_asm) add: congruent_def) apply (safe elim!: quotientE) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Finite.thy --- a/src/ZF/Finite.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Finite.thy Thu Mar 15 16:35:02 2012 +0000 @@ -2,7 +2,7 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge -prove: b: Fin(A) ==> inj(b,b) \ surj(b,b) +prove: b \ Fin(A) ==> inj(b,b) \ surj(b,b) *) header{*Finite Powerset Operator and Finite Function Space*} @@ -25,7 +25,7 @@ domains "Fin(A)" \ "Pow(A)" intros emptyI: "0 \ Fin(A)" - consI: "[| a: A; b: Fin(A) |] ==> cons(a,b) \ Fin(A)" + consI: "[| a \ A; b \ Fin(A) |] ==> cons(a,b) \ Fin(A)" type_intros empty_subsetI cons_subsetI PowI type_elims PowD [elim_format] @@ -33,7 +33,7 @@ domains "FiniteFun(A,B)" \ "Fin(A*B)" intros emptyI: "0 \ A -||> B" - consI: "[| a: A; b: B; h: A -||> B; a \ domain(h) |] + consI: "[| a \ A; b \ B; h \ A -||> B; a \ domain(h) |] ==> cons(,h) \ A -||> B" type_intros Fin.intros @@ -54,12 +54,12 @@ (*Discharging @{term"x\y"} entails extra work*) lemma Fin_induct [case_names 0 cons, induct set: Fin]: - "[| b: Fin(A); + "[| b \ Fin(A); P(0); - !!x y. [| x: A; y: Fin(A); x\y; P(y) |] ==> P(cons(x,y)) + !!x y. [| x \ A; y \ Fin(A); x\y; P(y) |] ==> P(cons(x,y)) |] ==> P(b)" apply (erule Fin.induct, simp) -apply (case_tac "a:b") +apply (case_tac "a \ b") apply (erule cons_absorb [THEN ssubst], assumption) (*backtracking!*) apply simp done @@ -72,7 +72,7 @@ by (blast intro: Fin.emptyI dest: FinD) (*The union of two finite sets is finite.*) -lemma Fin_UnI [simp]: "[| b: Fin(A); c: Fin(A) |] ==> b \ c \ Fin(A)" +lemma Fin_UnI [simp]: "[| b \ Fin(A); c \ Fin(A) |] ==> b \ c \ Fin(A)" apply (erule Fin_induct) apply (simp_all add: Un_cons) done @@ -83,25 +83,25 @@ by (erule Fin_induct, simp_all) (*Every subset of a finite set is finite.*) -lemma Fin_subset_lemma [rule_format]: "b: Fin(A) ==> \z. z<=b \ z: Fin(A)" +lemma Fin_subset_lemma [rule_format]: "b \ Fin(A) ==> \z. z<=b \ z \ Fin(A)" apply (erule Fin_induct) apply (simp add: subset_empty_iff) apply (simp add: subset_cons_iff distrib_simps, safe) apply (erule_tac b = z in cons_Diff [THEN subst], simp) done -lemma Fin_subset: "[| c<=b; b: Fin(A) |] ==> c: Fin(A)" +lemma Fin_subset: "[| c<=b; b \ Fin(A) |] ==> c \ Fin(A)" by (blast intro: Fin_subset_lemma) -lemma Fin_IntI1 [intro,simp]: "b: Fin(A) ==> b \ c \ Fin(A)" +lemma Fin_IntI1 [intro,simp]: "b \ Fin(A) ==> b \ c \ Fin(A)" by (blast intro: Fin_subset) -lemma Fin_IntI2 [intro,simp]: "c: Fin(A) ==> b \ c \ Fin(A)" +lemma Fin_IntI2 [intro,simp]: "c \ Fin(A) ==> b \ c \ Fin(A)" by (blast intro: Fin_subset) lemma Fin_0_induct_lemma [rule_format]: - "[| c: Fin(A); b: Fin(A); P(b); - !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) + "[| c \ Fin(A); b \ Fin(A); P(b); + !!x y. [| x \ A; y \ Fin(A); x \ y; P(y) |] ==> P(y-{x}) |] ==> c<=b \ P(b-c)" apply (erule Fin_induct, simp) apply (subst Diff_cons) @@ -109,16 +109,16 @@ done lemma Fin_0_induct: - "[| b: Fin(A); + "[| b \ Fin(A); P(b); - !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) + !!x y. [| x \ A; y \ Fin(A); x \ y; P(y) |] ==> P(y-{x}) |] ==> P(0)" apply (rule Diff_cancel [THEN subst]) apply (blast intro: Fin_0_induct_lemma) done (*Functions from a finite ordinal*) -lemma nat_fun_subset_Fin: "n: nat ==> n->A \ Fin(nat*A)" +lemma nat_fun_subset_Fin: "n \ nat ==> n->A \ Fin(nat*A)" apply (induct_tac "n") apply (simp add: subset_iff) apply (simp add: succ_def mem_not_refl [THEN cons_fun_eq]) @@ -139,19 +139,19 @@ lemma FiniteFun_mono1: "A<=B ==> A -||> A \ B -||> B" by (blast dest: FiniteFun_mono) -lemma FiniteFun_is_fun: "h: A -||>B ==> h: domain(h) -> B" +lemma FiniteFun_is_fun: "h \ A -||>B ==> h \ domain(h) -> B" apply (erule FiniteFun.induct, simp) apply (simp add: fun_extend3) done -lemma FiniteFun_domain_Fin: "h: A -||>B ==> domain(h) \ Fin(A)" +lemma FiniteFun_domain_Fin: "h \ A -||>B ==> domain(h) \ Fin(A)" by (erule FiniteFun.induct, simp, simp) lemmas FiniteFun_apply_type = FiniteFun_is_fun [THEN apply_type] (*Every subset of a finite function is a finite function.*) lemma FiniteFun_subset_lemma [rule_format]: - "b: A-||>B ==> \z. z<=b \ z: A-||>B" + "b \ A-||>B ==> \z. z<=b \ z \ A-||>B" apply (erule FiniteFun.induct) apply (simp add: subset_empty_iff FiniteFun.intros) apply (simp add: subset_cons_iff distrib_simps, safe) @@ -160,15 +160,15 @@ apply (fast intro!: FiniteFun.intros) done -lemma FiniteFun_subset: "[| c<=b; b: A-||>B |] ==> c: A-||>B" +lemma FiniteFun_subset: "[| c<=b; b \ A-||>B |] ==> c \ A-||>B" by (blast intro: FiniteFun_subset_lemma) (** Some further results by Sidi O. Ehmety **) -lemma fun_FiniteFunI [rule_format]: "A:Fin(X) ==> \f. f:A->B \ f:A-||>B" +lemma fun_FiniteFunI [rule_format]: "A \ Fin(X) ==> \f. f \ A->B \ f \ A-||>B" apply (erule Fin.induct) apply (simp add: FiniteFun.intros, clarify) -apply (case_tac "a:b") +apply (case_tac "a \ b") apply (simp add: cons_absorb) apply (subgoal_tac "restrict (f,b) \ b -||> B") prefer 2 apply (blast intro: restrict_type2) @@ -178,11 +178,11 @@ FiniteFun_mono [THEN [2] rev_subsetD]) done -lemma lam_FiniteFun: "A: Fin(X) ==> (\x\A. b(x)) \ A -||> {b(x). x:A}" +lemma lam_FiniteFun: "A \ Fin(X) ==> (\x\A. b(x)) \ A -||> {b(x). x \ A}" by (blast intro: fun_FiniteFunI lam_funtype) lemma FiniteFun_Collect_iff: - "f \ FiniteFun(A, {y:B. P(y)}) + "f \ FiniteFun(A, {y \ B. P(y)}) \ f \ FiniteFun(A,B) & (\x\domain(f). P(f`x))" apply auto apply (blast intro: FiniteFun_mono [THEN [2] rev_subsetD]) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Induct/Multiset.thy Thu Mar 15 16:35:02 2012 +0000 @@ -93,7 +93,7 @@ definition multirel :: "[i, i] => i" where - "multirel(A, r) == multirel1(A, r)^+" + "multirel(A, r) == multirel1(A, r)^+" (* ordinal multiset orderings *) @@ -446,7 +446,7 @@ by (induct_tac n) auto lemma munion_is_single: - "[|multiset(M); multiset(N)|] + "[|multiset(M); multiset(N)|] ==> (M +# N = {#a#}) \ (M={#a#} & N=0) | (M = 0 & N = {#a#})" apply (simp (no_asm_simp) add: multiset_equality) apply safe @@ -747,7 +747,7 @@ done lemma multirel1_mono2: "r\s ==> multirel1(A,r)\multirel1(A, s)" -apply (simp add: multirel1_def, auto) +apply (simp add: multirel1_def, auto) apply (rule_tac x = a in bexI) apply (rule_tac x = M0 in bexI) apply (simp_all add: Mult_iff_multiset) @@ -807,9 +807,9 @@ apply (erule_tac P = "mset_of (K) \A" in rev_mp) apply (erule_tac M = K in multiset_induct) (* three subgoals *) -(* subgoal 1: the induction base case *) +(* subgoal 1 \ the induction base case *) apply (simp (no_asm_simp)) -(* subgoal 2: the induction general case *) +(* subgoal 2 \ the induction general case *) apply (simp add: Ball_def Un_subset_iff, clarify) apply (drule_tac x = aa in spec, simp) apply (subgoal_tac "aa \ A") @@ -817,7 +817,7 @@ apply (drule_tac x = "M0 +# M" and P = "%x. x \ acc(multirel1(A, r)) \ ?Q(x)" in spec) apply (simp add: munion_assoc [symmetric]) -(* subgoal 3: additional conditions *) +(* subgoal 3 \ additional conditions *) apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset) done diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/IntArith.thy --- a/src/ZF/IntArith.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/IntArith.thy Thu Mar 15 16:35:02 2012 +0000 @@ -34,7 +34,7 @@ lemma zless_iff_zdiff_zless_0: "(x $< y) \ (x$-y $< #0)" by (simp add: zcompare_rls) -lemma eq_iff_zdiff_eq_0: "[| x: int; y: int |] ==> (x = y) \ (x$-y = #0)" +lemma eq_iff_zdiff_eq_0: "[| x \ int; y \ int |] ==> (x = y) \ (x$-y = #0)" by (simp add: zcompare_rls) lemma zle_iff_zdiff_zle_0: "(x $<= y) \ (x$-y $<= #0)" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Int_ZF.thy --- a/src/ZF/Int_ZF.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Int_ZF.thy Thu Mar 15 16:35:02 2012 +0000 @@ -9,12 +9,12 @@ definition intrel :: i where - "intrel == {p \ (nat*nat)*(nat*nat). + "intrel == {p \ (nat*nat)*(nat*nat). \x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" definition int :: i where - "int == (nat*nat)//intrel" + "int == (nat*nat)//intrel" definition int_of :: "i=>i" --{*coercion from nat to int*} ("$# _" [80] 80) where @@ -39,7 +39,7 @@ definition iszero :: "i=>o" where "iszero(z) == z = $# 0" - + definition raw_nat_of :: "i=>i" where "raw_nat_of(z) == natify (\\z. x#-y)" @@ -60,8 +60,8 @@ (*Cannot use UN here or in zadd because of the form of congruent2. Perhaps a "curried" or even polymorphic congruent predicate would be better.*) - "raw_zmult(z1,z2) == - \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. + "raw_zmult(z1,z2) == + \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. intrel``{}, p2), p1)" definition @@ -70,8 +70,8 @@ definition raw_zadd :: "[i,i]=>i" where - "raw_zadd (z1, z2) == - \z1\z1. \z2\z2. let =z1; =z2 + "raw_zadd (z1, z2) == + \z1\z1. \z2\z2. let =z1; =z2 in intrel``{}" definition @@ -85,11 +85,11 @@ definition zless :: "[i,i]=>o" (infixl "$<" 50) where "z1 $< z2 == znegative(z1 $- z2)" - + definition zle :: "[i,i]=>o" (infixl "$<=" 50) where "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)" - + notation (xsymbols) zmult (infixl "$\" 70) and @@ -106,22 +106,22 @@ (** Natural deduction for intrel **) -lemma intrel_iff [simp]: - "<,>: intrel \ +lemma intrel_iff [simp]: + "<,>: intrel \ x1\nat & y1\nat & x2\nat & y2\nat & x1#+y2 = x2#+y1" by (simp add: intrel_def) -lemma intrelI [intro!]: - "[| x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat |] +lemma intrelI [intro!]: + "[| x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat |] ==> <,>: intrel" by (simp add: intrel_def) lemma intrelE [elim!]: - "[| p: intrel; - !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; - x1\nat; y1\nat; x2\nat; y2\nat |] ==> Q |] + "[| p \ intrel; + !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; + x1\nat; y1\nat; x2\nat; y2\nat |] ==> Q |] ==> Q" -by (simp add: intrel_def, blast) +by (simp add: intrel_def, blast) lemma int_trans_lemma: "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" @@ -228,8 +228,8 @@ lemma zminus_type [TC,iff]: "$-z \ int" by (simp add: zminus_def raw_zminus_type) -lemma raw_zminus_inject: - "[| raw_zminus(z) = raw_zminus(w); z: int; w: int |] ==> z=w" +lemma raw_zminus_inject: + "[| raw_zminus(z) = raw_zminus(w); z \ int; w \ int |] ==> z=w" apply (simp add: int_def raw_zminus_def) apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) apply (auto dest: eq_intrelD simp add: add_ac) @@ -240,16 +240,16 @@ apply (blast dest!: raw_zminus_inject) done -lemma zminus_inject: "[| $-z = $-w; z: int; w: int |] ==> z=w" +lemma zminus_inject: "[| $-z = $-w; z \ int; w \ int |] ==> z=w" by auto -lemma raw_zminus: +lemma raw_zminus: "[| x\nat; y\nat |] ==> raw_zminus(intrel``{}) = intrel `` {}" apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) done -lemma zminus: - "[| x\nat; y\nat |] +lemma zminus: + "[| x\nat; y\nat |] ==> $- (intrel``{}) = intrel `` {}" by (simp add: zminus_def raw_zminus image_intrel_int) @@ -269,15 +269,15 @@ subsection{*@{term znegative}: the test for negative integers*} lemma znegative: "[| x\nat; y\nat |] ==> znegative(intrel``{}) \ xx. (\\x,y\. x #- y)(x)) respects intrel" by (auto simp add: congruent_def split add: nat_diff_split) -lemma raw_nat_of: +lemma raw_nat_of: "[| x\nat; y\nat |] ==> raw_nat_of(intrel``{}) = x#-y" by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) @@ -332,24 +332,24 @@ apply (rule theI2, auto) done -lemma not_zneg_int_of: - "[| z: int; ~ znegative(z) |] ==> \n\nat. z = $# n" +lemma not_zneg_int_of: + "[| z \ int; ~ znegative(z) |] ==> \n\nat. z = $# n" apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) -apply (rename_tac x y) -apply (rule_tac x="x#-y" in bexI) -apply (auto simp add: add_diff_inverse2) +apply (rename_tac x y) +apply (rule_tac x="x#-y" in bexI) +apply (auto simp add: add_diff_inverse2) done lemma not_zneg_mag [simp]: - "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" + "[| z \ int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" by (drule not_zneg_int_of, auto) -lemma zneg_int_of: - "[| znegative(z); z: int |] ==> \n\nat. z = $- ($# succ(n))" +lemma zneg_int_of: + "[| znegative(z); z \ int |] ==> \n\nat. z = $- ($# succ(n))" by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) lemma zneg_mag [simp]: - "[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z" + "[| znegative(z); z \ int |] ==> $# (zmagnitude(z)) = $- z" by (drule zneg_int_of, auto) lemma int_cases: "z \ int ==> \n\nat. z = $# n | z = $- ($# succ(n))" @@ -359,7 +359,7 @@ done lemma not_zneg_raw_nat_of: - "[| ~ znegative(z); z: int |] ==> $# (raw_nat_of(z)) = z" + "[| ~ znegative(z); z \ int |] ==> $# (raw_nat_of(z)) = z" apply (drule not_zneg_int_of) apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) done @@ -368,23 +368,23 @@ "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) -lemma not_zneg_nat_of: "[| ~ znegative(z); z: int |] ==> $# (nat_of(z)) = z" +lemma not_zneg_nat_of: "[| ~ znegative(z); z \ int |] ==> $# (nat_of(z)) = z" apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) done lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" apply (subgoal_tac "intify(z) \ int") -apply (simp add: int_def) -apply (auto simp add: znegative nat_of_def raw_nat_of - split add: nat_diff_split) +apply (simp add: int_def) +apply (auto simp add: znegative nat_of_def raw_nat_of + split add: nat_diff_split) done subsection{*@{term zadd}: addition on int*} text{*Congruence Property for Addition*} -lemma zadd_congruent2: - "(%z1 z2. let =z1; =z2 +lemma zadd_congruent2: + "(%z1 z2. let =z1; =z2 in intrel``{}) respects2 intrel" apply (simp add: congruent2_def) @@ -398,7 +398,7 @@ apply (simp (no_asm_simp) add: add_assoc [symmetric]) done -lemma raw_zadd_type: "[| z: int; w: int |] ==> raw_zadd(z,w) \ int" +lemma raw_zadd_type: "[| z \ int; w \ int |] ==> raw_zadd(z,w) \ int" apply (simp add: int_def raw_zadd_def) apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) apply (simp add: Let_def) @@ -407,18 +407,18 @@ lemma zadd_type [iff,TC]: "z $+ w \ int" by (simp add: zadd_def raw_zadd_type) -lemma raw_zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zadd (intrel``{}, intrel``{}) = +lemma raw_zadd: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> raw_zadd (intrel``{}, intrel``{}) = intrel `` {}" -apply (simp add: raw_zadd_def +apply (simp add: raw_zadd_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) apply (simp add: Let_def) done -lemma zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $+ (intrel``{}) = +lemma zadd: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> (intrel``{}) $+ (intrel``{}) = intrel `` {}" by (simp add: zadd_def raw_zadd image_intrel_int) @@ -428,25 +428,25 @@ lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" by (simp add: zadd_def raw_zadd_int0) -lemma zadd_int0: "z: int ==> $#0 $+ z = z" +lemma zadd_int0: "z \ int ==> $#0 $+ z = z" by simp -lemma raw_zminus_zadd_distrib: - "[| z: int; w: int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" +lemma raw_zminus_zadd_distrib: + "[| z \ int; w \ int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" by (auto simp add: zminus raw_zadd int_def) lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" by (simp add: zadd_def raw_zminus_zadd_distrib) lemma raw_zadd_commute: - "[| z: int; w: int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" + "[| z \ int; w \ int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" by (auto simp add: raw_zadd add_ac int_def) lemma zadd_commute: "z $+ w = w $+ z" by (simp add: zadd_def raw_zadd_commute) -lemma raw_zadd_assoc: - "[| z1: int; z2: int; z3: int |] +lemma raw_zadd_assoc: + "[| z1: int; z2: int; z3: int |] ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" by (auto simp add: int_def raw_zadd add_assoc) @@ -468,7 +468,7 @@ lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)" by (simp add: int_of_add [symmetric] natify_succ) -lemma int_of_diff: +lemma int_of_diff: "[| m\nat; n \ m |] ==> $# (m #- n) = ($#m) $- ($#n)" apply (simp add: int_of_def zdiff_def) apply (frule lt_nat_in_nat) @@ -490,7 +490,7 @@ lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" by (rule trans [OF zadd_commute zadd_int0_intify]) -lemma zadd_int0_right: "z:int ==> z $+ $#0 = z" +lemma zadd_int0_right: "z \ int ==> z $+ $#0 = z" by simp @@ -498,7 +498,7 @@ text{*Congruence property for multiplication*} lemma zmult_congruent2: - "(%p1 p2. split(%x1 y1. split(%x2 y2. + "(%p1 p2. split(%x1 y1. split(%x2 y2. intrel``{}, p2), p1)) respects2 intrel" apply (rule equiv_intrel [THEN congruent2_commuteI], auto) @@ -511,7 +511,7 @@ done -lemma raw_zmult_type: "[| z: int; w: int |] ==> raw_zmult(z,w) \ int" +lemma raw_zmult_type: "[| z \ int; w \ int |] ==> raw_zmult(z,w) \ int" apply (simp add: int_def raw_zmult_def) apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) apply (simp add: Let_def) @@ -520,16 +520,16 @@ lemma zmult_type [iff,TC]: "z $* w \ int" by (simp add: zmult_def raw_zmult_type) -lemma raw_zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zmult(intrel``{}, intrel``{}) = +lemma raw_zmult: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> raw_zmult(intrel``{}, intrel``{}) = intrel `` {}" -by (simp add: raw_zmult_def +by (simp add: raw_zmult_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) -lemma zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $* (intrel``{}) = +lemma zmult: + "[| x1\nat; y1\nat; x2\nat; y2\nat |] + ==> (intrel``{}) $* (intrel``{}) = intrel `` {}" by (simp add: zmult_def raw_zmult image_intrel_int) @@ -549,14 +549,14 @@ by simp lemma raw_zmult_commute: - "[| z: int; w: int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" + "[| z \ int; w \ int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" by (auto simp add: int_def raw_zmult add_ac mult_ac) lemma zmult_commute: "z $* w = w $* z" by (simp add: zmult_def raw_zmult_commute) -lemma raw_zmult_zminus: - "[| z: int; w: int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" +lemma raw_zmult_zminus: + "[| z \ int; w \ int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" by (auto simp add: int_def zminus raw_zmult add_ac) lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" @@ -567,8 +567,8 @@ lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)" by (simp add: zmult_commute [of w]) -lemma raw_zmult_assoc: - "[| z1: int; z2: int; z3: int |] +lemma raw_zmult_assoc: + "[| z1: int; z2: int; z3: int |] ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) @@ -584,20 +584,20 @@ (*Integer multiplication is an AC operator*) lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute -lemma raw_zadd_zmult_distrib: - "[| z1: int; z2: int; w: int |] - ==> raw_zmult(raw_zadd(z1,z2), w) = +lemma raw_zadd_zmult_distrib: + "[| z1: int; z2: int; w \ int |] + ==> raw_zmult(raw_zadd(z1,z2), w) = raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)" -by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type +by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type raw_zadd_zmult_distrib) lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)" by (simp add: zmult_commute [of w] zadd_zmult_distrib) -lemmas int_typechecks = +lemmas int_typechecks = int_of_type zminus_type zmagnitude_type zadd_type zmult_type @@ -628,8 +628,8 @@ subsection{*The "Less Than" Relation*} (*"Less than" is a linear ordering*) -lemma zless_linear_lemma: - "[| z: int; w: int |] ==> z$ int; w \ int |] ==> z$ (x \ y) \ (x $< y | y $< x)" +lemma neq_iff_zless: "[| x \ int; y \ int |] ==> (x \ y) \ (x $< y | y $< x)" by (cut_tac z = x and w = y in zless_linear, auto) lemma zless_imp_intify_neq: "w $< z ==> intify(w) \ intify(z)" @@ -656,8 +656,8 @@ done (*This lemma allows direct proofs of other <-properties*) -lemma zless_imp_succ_zadd_lemma: - "[| w $< z; w: int; z: int |] ==> (\n\nat. z = w $+ $#(succ(n)))" +lemma zless_imp_succ_zadd_lemma: + "[| w $< z; w \ int; z \ int |] ==> (\n\nat. z = w $+ $#(succ(n)))" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) apply (rule_tac x = k in bexI) @@ -671,7 +671,7 @@ apply auto done -lemma zless_succ_zadd_lemma: +lemma zless_succ_zadd_lemma: "w \ int ==> w $< w $+ $# succ(n)" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto simp add: zadd zminus int_of_def image_iff) @@ -694,8 +694,8 @@ apply (blast intro: sym) done -lemma zless_trans_lemma: - "[| x $< y; y $< z; x: int; y \ int; z: int |] ==> x $< z" +lemma zless_trans_lemma: + "[| x $< y; y $< z; x \ int; y \ int; z \ int |] ==> x $< z" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto simp add: zadd zminus image_iff) apply (rename_tac x1 x2 y1 y2) @@ -741,11 +741,11 @@ apply (blast dest: zless_trans) done -lemma zle_anti_sym: "[| x $<= y; y $<= x; x: int; y: int |] ==> x=y" +lemma zle_anti_sym: "[| x $<= y; y $<= x; x \ int; y \ int |] ==> x=y" by (drule zle_anti_sym_intify, auto) lemma zle_trans_lemma: - "[| x: int; y: int; z: int; x $<= y; y $<= z |] ==> x $<= z" + "[| x \ int; y \ int; z \ int; x $<= y; y $<= z |] ==> x $<= z" apply (simp add: zle_def, auto) apply (blast intro: zless_trans) done @@ -792,21 +792,21 @@ lemma zless_zdiff_iff: "(x $< z$-y) \ (x $+ y $< z)" by (simp add: zless_def zdiff_def zadd_ac) -lemma zdiff_eq_iff: "[| x: int; z: int |] ==> (x$-y = z) \ (x = z $+ y)" +lemma zdiff_eq_iff: "[| x \ int; z \ int |] ==> (x$-y = z) \ (x = z $+ y)" by (auto simp add: zdiff_def zadd_assoc) -lemma eq_zdiff_iff: "[| x: int; z: int |] ==> (x = z$-y) \ (x $+ y = z)" +lemma eq_zdiff_iff: "[| x \ int; z \ int |] ==> (x = z$-y) \ (x $+ y = z)" by (auto simp add: zdiff_def zadd_assoc) lemma zdiff_zle_iff_lemma: - "[| x: int; z: int |] ==> (x$-y $<= z) \ (x $<= z $+ y)" + "[| x \ int; z \ int |] ==> (x$-y $<= z) \ (x $<= z $+ y)" by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) lemma zdiff_zle_iff: "(x$-y $<= z) \ (x $<= z $+ y)" by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) lemma zle_zdiff_iff_lemma: - "[| x: int; z: int |] ==>(x $<= z$-y) \ (x $+ y $<= z)" + "[| x \ int; z \ int |] ==>(x $<= z$-y) \ (x $+ y $<= z)" apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) apply (auto simp add: zdiff_def zadd_assoc) done @@ -815,12 +815,12 @@ by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp) text{*This list of rewrites simplifies (in)equalities by bringing subtractions - to the top and then moving negative terms to the other side. + to the top and then moving negative terms to the other side. Use with @{text zadd_ac}*} lemmas zcompare_rls = zdiff_def [symmetric] - zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 - zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff + zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 + zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff zdiff_eq_iff eq_zdiff_iff @@ -828,7 +828,7 @@ of the CancelNumerals Simprocs*} lemma zadd_left_cancel: - "[| w: int; w': int |] ==> (z $+ w' = z $+ w) \ (w' = w)" + "[| w \ int; w': int |] ==> (z $+ w' = z $+ w) \ (w' = w)" apply safe apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) @@ -841,7 +841,7 @@ done lemma zadd_right_cancel: - "[| w: int; w': int |] ==> (w' $+ z = w $+ z) \ (w' = w)" + "[| w \ int; w': int |] ==> (w' $+ z = w $+ z) \ (w' = w)" apply safe apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) @@ -895,10 +895,10 @@ subsubsection{*More inequality lemmas*} -lemma equation_zminus: "[| x: int; y: int |] ==> (x = $- y) \ (y = $- x)" +lemma equation_zminus: "[| x \ int; y \ int |] ==> (x = $- y) \ (y = $- x)" by auto -lemma zminus_equation: "[| x: int; y: int |] ==> ($- x = y) \ ($- y = x)" +lemma zminus_equation: "[| x \ int; y \ int |] ==> ($- x = y) \ ($- y = x)" by auto lemma equation_zminus_intify: "(intify(x) = $- y) \ (intify(y) = $- x)" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/List_ZF.thy --- a/src/ZF/List_ZF.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/List_ZF.thy Thu Mar 15 16:35:02 2012 +0000 @@ -11,7 +11,7 @@ list :: "i=>i" datatype - "list(A)" = Nil | Cons ("a:A", "l: list(A)") + "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") syntax @@ -171,13 +171,13 @@ (*These two theorems justify datatypes involving list(nat), list(A), ...*) lemmas list_subset_univ = subset_trans [OF list_mono list_univ] -lemma list_into_univ: "[| l: list(A); A \ univ(B) |] ==> l: univ(B)" +lemma list_into_univ: "[| l \ list(A); A \ univ(B) |] ==> l \ univ(B)" by (blast intro: list_subset_univ [THEN subsetD]) lemma list_case_type: - "[| l: list(A); - c: C(Nil); - !!x y. [| x: A; y: list(A) |] ==> h(x,y): C(Cons(x,y)) + "[| l \ list(A); + c \ C(Nil); + !!x y. [| x \ A; y \ list(A) |] ==> h(x,y): C(Cons(x,y)) |] ==> list_case(c,h,l) \ C(l)" by (erule list.induct, auto) @@ -189,26 +189,26 @@ (*** List functions ***) -lemma tl_type: "l: list(A) ==> tl(l) \ list(A)" +lemma tl_type: "l \ list(A) ==> tl(l) \ list(A)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list.intros) done (** drop **) -lemma drop_Nil [simp]: "i:nat ==> drop(i, Nil) = Nil" +lemma drop_Nil [simp]: "i \ nat ==> drop(i, Nil) = Nil" apply (induct_tac "i") apply (simp_all (no_asm_simp)) done -lemma drop_succ_Cons [simp]: "i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)" +lemma drop_succ_Cons [simp]: "i \ nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)" apply (rule sym) apply (induct_tac "i") apply (simp (no_asm)) apply (simp (no_asm_simp)) done -lemma drop_type [simp,TC]: "[| i:nat; l: list(A) |] ==> drop(i,l) \ list(A)" +lemma drop_type [simp,TC]: "[| i \ nat; l \ list(A) |] ==> drop(i,l) \ list(A)" apply (induct_tac "i") apply (simp_all (no_asm_simp) add: tl_type) done @@ -219,28 +219,28 @@ (** Type checking -- proved by induction, as usual **) lemma list_rec_type [TC]: - "[| l: list(A); - c: C(Nil); - !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) + "[| l \ list(A); + c \ C(Nil); + !!x y r. [| x \ A; y \ list(A); r \ C(y) |] ==> h(x,y,r): C(Cons(x,y)) |] ==> list_rec(c,h,l) \ C(l)" by (induct_tac "l", auto) (** map **) lemma map_type [TC]: - "[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) \ list(B)" + "[| l \ list(A); !!x. x \ A ==> h(x): B |] ==> map(h,l) \ list(B)" apply (simp add: map_list_def) apply (typecheck add: list.intros list_rec_type, blast) done -lemma map_type2 [TC]: "l: list(A) ==> map(h,l) \ list({h(u). u:A})" +lemma map_type2 [TC]: "l \ list(A) ==> map(h,l) \ list({h(u). u \ A})" apply (erule map_type) apply (erule RepFunI) done (** length **) -lemma length_type [TC]: "l: list(A) ==> length(l) \ nat" +lemma length_type [TC]: "l \ list(A) ==> length(l) \ nat" by (simp add: length_list_def) lemma lt_length_in_nat: @@ -266,7 +266,7 @@ (** set_of_list **) -lemma set_of_list_type [TC]: "l: list(A) ==> set_of_list(l) \ Pow(A)" +lemma set_of_list_type [TC]: "l \ list(A) ==> set_of_list(l) \ Pow(A)" apply (unfold set_of_list_list_def) apply (erule list_rec_type, auto) done @@ -286,12 +286,12 @@ (*** theorems about map ***) -lemma map_ident [simp]: "l: list(A) ==> map(%u. u, l) = l" +lemma map_ident [simp]: "l \ list(A) ==> map(%u. u, l) = l" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done -lemma map_compose: "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)" +lemma map_compose: "l \ list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done @@ -307,7 +307,7 @@ done lemma list_rec_map: - "l: list(A) ==> + "l \ list(A) ==> list_rec(c, d, map(h,l)) = list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)" apply (induct_tac "l") @@ -319,7 +319,7 @@ (* @{term"c \ list(Collect(B,P)) ==> c \ list"} *) lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD] -lemma map_list_Collect: "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)" +lemma map_list_Collect: "l \ list({x \ A. h(x)=j(x)}) ==> map(h,l) = map(j,l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done @@ -354,7 +354,7 @@ by (erule list.induct, simp_all) lemma drop_length [rule_format]: - "l: list(A) ==> \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" + "l \ list(A) ==> \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" apply (erule list.induct, simp_all, safe) apply (erule drop_length_Cons) apply (rule natE) @@ -378,7 +378,7 @@ (*** theorems about rev ***) -lemma rev_map_distrib: "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))" +lemma rev_map_distrib: "l \ list(A) ==> rev(map(h,l)) = map(h,rev(l))" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: map_app_distrib) done @@ -393,7 +393,7 @@ apply (simp_all add: app_assoc) done -lemma rev_rev_ident [simp]: "l: list(A) ==> rev(rev(l))=l" +lemma rev_rev_ident [simp]: "l \ list(A) ==> rev(rev(l))=l" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: rev_app_distrib) done @@ -412,7 +412,7 @@ apply (induct_tac "xs", simp_all) done -lemma list_add_rev: "l: list(nat) ==> list_add(rev(l)) = list_add(l)" +lemma list_add_rev: "l \ list(nat) ==> list_add(rev(l)) = list_add(l)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list_add_app) done @@ -426,9 +426,9 @@ (** New induction rules **) lemma list_append_induct [case_names Nil snoc, consumes 1]: - "[| l: list(A); + "[| l \ list(A); P(Nil); - !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) + !!x y. [| x \ A; y \ list(A); P(y) |] ==> P(y @ [x]) |] ==> P(l)" apply (subgoal_tac "P(rev(rev(l)))", simp) apply (erule rev_type [THEN list.induct], simp_all) @@ -462,31 +462,31 @@ (** min FIXME: replace by Int! **) (* Min theorems are also true for i, j ordinals *) -lemma min_sym: "[| i:nat; j:nat |] ==> min(i,j)=min(j,i)" +lemma min_sym: "[| i \ nat; j \ nat |] ==> min(i,j)=min(j,i)" apply (unfold min_def) apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym) done -lemma min_type [simp,TC]: "[| i:nat; j:nat |] ==> min(i,j):nat" +lemma min_type [simp,TC]: "[| i \ nat; j \ nat |] ==> min(i,j):nat" by (unfold min_def, auto) -lemma min_0 [simp]: "i:nat ==> min(0,i) = 0" +lemma min_0 [simp]: "i \ nat ==> min(0,i) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done -lemma min_02 [simp]: "i:nat ==> min(i, 0) = 0" +lemma min_02 [simp]: "i \ nat ==> min(i, 0) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done -lemma lt_min_iff: "[| i:nat; j:nat; k:nat |] ==> i i nat; j \ nat; k \ nat |] ==> i i min(succ(i), succ(j))= succ(min(i, j))" + "[| i \ nat; j \ nat |] ==> min(succ(i), succ(j))= succ(min(i, j))" apply (unfold min_def, auto) done @@ -603,7 +603,7 @@ by simp (* Can also be proved from append_eq_append_iff2, -but the proof requires two more hypotheses: x:A and y:A *) +but the proof requires two more hypotheses: x \ A and y \ A *) lemma append1_eq_iff [rule_format,simp]: "xs:list(A) ==> \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys & x=y)" apply (erule list.induct) @@ -656,26 +656,26 @@ (** more theorems about drop **) lemma length_drop [rule_format,simp]: - "n:nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" + "n \ nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" apply (erule nat_induct) apply (auto elim: list.cases) done lemma drop_all [rule_format,simp]: - "n:nat ==> \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" + "n \ nat ==> \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" apply (erule nat_induct) apply (auto elim: list.cases) done lemma drop_append [rule_format]: - "n:nat ==> + "n \ nat ==> \xs \ list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)" apply (induct_tac "n") apply (auto elim: list.cases) done lemma drop_drop: - "m:nat ==> \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" + "m \ nat ==> \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" apply (induct_tac "m") apply (auto elim: list.cases) done @@ -688,15 +688,15 @@ done lemma take_succ_Cons [simp]: - "n:nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" + "n \ nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" by (simp add: take_def) (* Needed for proving take_all *) -lemma take_Nil [simp]: "n:nat ==> take(n, Nil) = Nil" +lemma take_Nil [simp]: "n \ nat ==> take(n, Nil) = Nil" by (unfold take_def, auto) lemma take_all [rule_format,simp]: - "n:nat ==> \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" + "n \ nat ==> \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" apply (erule nat_induct) apply (auto elim: list.cases) done @@ -730,7 +730,7 @@ lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a" by (simp add: nth_def) -lemma nth_Cons [simp]: "n:nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)" +lemma nth_Cons [simp]: "n \ nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)" by (simp add: nth_def) lemma nth_empty [simp]: "nth(n, Nil) = 0" @@ -759,7 +759,7 @@ lemma set_of_list_conv_nth: "xs:list(A) - ==> set_of_list(xs) = {x:A. \i\nat. i set_of_list(xs) = {x \ A. \i\nat. i + "k \ nat ==> \xs \ list(A). (\ys \ list(A). k \ length(xs) \ k \ length(ys) \ (\i \ nat. i nth(i,xs) = nth(i,ys))\ take(k,xs) = take(k,ys))" apply (induct_tac "k") @@ -811,7 +811,7 @@ done lemma nth_drop [rule_format]: - "n:nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" + "n \ nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" apply (induct_tac "n", simp_all, clarify) apply (erule list.cases, auto) done @@ -886,7 +886,7 @@ done lemma zip_Cons_Cons [simp]: - "[| xs:list(A); ys:list(B); x:A; y:B |] ==> + "[| xs:list(A); ys:list(B); x \ A; y \ B |] ==> zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" apply (simp add: zip_def, auto) apply (rule zip_aux_unique, auto) @@ -951,7 +951,7 @@ done lemma set_of_list_zip [rule_format]: - "[| xs:list(A); ys:list(B); i:nat |] + "[| xs:list(A); ys:list(B); i \ nat |] ==> set_of_list(zip(xs, ys)) = {:A*B. \i\nat. i < min(length(xs), length(ys)) & x = nth(i, xs) & y = nth(i, ys)}" @@ -959,20 +959,20 @@ (** list_update **) -lemma list_update_Nil [simp]: "i:nat ==>list_update(Nil, i, v) = Nil" +lemma list_update_Nil [simp]: "i \ nat ==>list_update(Nil, i, v) = Nil" by (unfold list_update_def, auto) lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)" by (unfold list_update_def, auto) lemma list_update_Cons_succ [simp]: - "n:nat ==> + "n \ nat ==> list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))" apply (unfold list_update_def, auto) done lemma list_update_type [rule_format,simp,TC]: - "[| xs:list(A); v:A |] ==> \n \ nat. list_update(xs, n, v):list(A)" + "[| xs:list(A); v \ A |] ==> \n \ nat. list_update(xs, n, v):list(A)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify @@ -1056,7 +1056,7 @@ done lemma set_of_list_update_subsetI: - "[| set_of_list(xs) \ A; xs:list(A); x:A; i:nat|] + "[| set_of_list(xs) \ A; xs:list(A); x \ A; i \ nat|] ==> set_of_list(list_update(xs, i,x)) \ A" apply (rule subset_trans) apply (rule set_update_subset_cons, auto) @@ -1065,13 +1065,13 @@ (** upt **) lemma upt_rec: - "j:nat ==> upt(i,j) = (if i nat ==> upt(i,j) = (if i i; j:nat |] ==> upt(i,j) = Nil" +lemma upt_conv_Nil [simp]: "[| j \ i; j \ nat |] ==> upt(i,j) = Nil" apply (subst upt_rec, auto) apply (auto simp add: le_iff) apply (drule lt_asym [THEN notE], auto) @@ -1079,34 +1079,34 @@ (*Only needed if upt_Suc is deleted from the simpset*) lemma upt_succ_append: - "[| i \ j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" + "[| i \ j; j \ nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" by simp lemma upt_conv_Cons: - "[| i upt(i,j) = Cons(i,upt(succ(i),j))" + "[| i nat |] ==> upt(i,j) = Cons(i,upt(succ(i),j))" apply (rule trans) apply (rule upt_rec, auto) done -lemma upt_type [simp,TC]: "j:nat ==> upt(i,j):list(nat)" +lemma upt_type [simp,TC]: "j \ nat ==> upt(i,j):list(nat)" by (induct_tac "j", auto) (*LOOPS as a simprule, since j<=j*) lemma upt_add_eq_append: - "[| i \ j; j:nat; k:nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" + "[| i \ j; j \ nat; k \ nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" apply (induct_tac "k") apply (auto simp add: app_assoc app_type) apply (rule_tac j = j in le_trans, auto) done -lemma length_upt [simp]: "[| i:nat; j:nat |] ==>length(upt(i,j)) = j #- i" +lemma length_upt [simp]: "[| i \ nat; j \ nat |] ==>length(upt(i,j)) = j #- i" apply (induct_tac "j") apply (rule_tac [2] sym) apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff) done lemma nth_upt [rule_format,simp]: - "[| i:nat; j:nat; k:nat |] ==> i #+ k < j \ nth(k, upt(i,j)) = i #+ k" + "[| i \ nat; j \ nat; k \ nat |] ==> i #+ k < j \ nth(k, upt(i,j)) = i #+ k" apply (induct_tac "j", simp) apply (simp add: nth_append le_iff) apply (auto dest!: not_lt_imp_le @@ -1114,7 +1114,7 @@ done lemma take_upt [rule_format,simp]: - "[| m:nat; n:nat |] ==> + "[| m \ nat; n \ nat |] ==> \i \ nat. i #+ m \ n \ take(m, upt(i,n)) = upt(i,i#+m)" apply (induct_tac "m") apply (simp (no_asm_simp) add: take_0) @@ -1128,7 +1128,7 @@ done lemma map_succ_upt: - "[| m:nat; n:nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" + "[| m \ nat; n \ nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" apply (induct_tac "n") apply (auto simp add: map_app_distrib) done @@ -1142,7 +1142,7 @@ done lemma nth_map_upt [rule_format]: - "[| m:nat; n:nat |] ==> + "[| m \ nat; n \ nat |] ==> \i \ nat. i < n #- m \ nth(i, map(f, upt(m,n))) = f(m #+ i)" apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) apply (subst map_succ_upt [symmetric], simp_all, clarify) @@ -1170,9 +1170,9 @@ by (unfold sublist_def, auto) lemma sublist_shift_lemma: - "[| xs:list(B); i:nat |] ==> + "[| xs:list(B); i \ nat |] ==> map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = - map(fst, filter(%p. snd(p):nat & snd(p) #+ i:A, zip(xs,upt(0,length(xs)))))" + map(fst, filter(%p. snd(p):nat & snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" apply (erule list_append_induct) apply (simp (no_asm_simp)) apply (auto simp add: add_commute length_app filter_append map_app_distrib) @@ -1186,12 +1186,12 @@ done lemma upt_add_eq_append2: - "[| i:nat; j:nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" + "[| i \ nat; j \ nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" by (simp add: upt_add_eq_append [of 0] nat_0_le) lemma sublist_append: "[| xs:list(B); ys:list(B) |] ==> - sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j:nat. j #+ length(xs): A})" + sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" apply (unfold sublist_def) apply (erule_tac l = ys in list_append_induct, simp) apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric]) @@ -1201,9 +1201,9 @@ lemma sublist_Cons: - "[| xs:list(B); x:B |] ==> + "[| xs:list(B); x \ B |] ==> sublist(Cons(x, xs), A) = - (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) \ A})" + (if 0 \ A then [x] else []) @ sublist(xs, {j \ nat. succ(j) \ A})" apply (erule_tac l = xs in list_append_induct) apply (simp (no_asm_simp) add: sublist_def) apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Main_ZF.thy --- a/src/ZF/Main_ZF.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Main_ZF.thy Thu Mar 15 16:35:02 2012 +0000 @@ -23,21 +23,21 @@ iterates_omega ("(_^\ '(_'))" [60,1000] 60) lemma iterates_triv: - "[| n\nat; F(x) = x |] ==> F^n (x) = x" + "[| n\nat; F(x) = x |] ==> F^n (x) = x" by (induct n rule: nat_induct, simp_all) lemma iterates_type [TC]: - "[| n:nat; a: A; !!x. x:A ==> F(x) \ A |] - ==> F^n (a) \ A" + "[| n \ nat; a \ A; !!x. x \ A ==> F(x) \ A |] + ==> F^n (a) \ A" by (induct n rule: nat_induct, simp_all) lemma iterates_omega_triv: - "F(x) = x ==> F^\ (x) = x" -by (simp add: iterates_omega_def iterates_triv) + "F(x) = x ==> F^\ (x) = x" +by (simp add: iterates_omega_def iterates_triv) lemma Ord_iterates [simp]: - "[| n\nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |] - ==> Ord(F^n (x))" + "[| n\nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |] + ==> Ord(F^n (x))" by (induct n rule: nat_induct, simp_all) lemma iterates_commute: "n \ nat ==> F(F^n (x)) = F^n (F(x))" @@ -46,12 +46,12 @@ subsection{* Transfinite Recursion *} -text{*Transfinite recursion for definitions based on the +text{*Transfinite recursion for definitions based on the three cases of ordinals*} definition transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where - "transrec3(k, a, b, c) == + "transrec3(k, a, b, c) == transrec(k, \x r. if x=0 then a else if Limit(x) then c(x, \y\x. r`y) @@ -65,7 +65,7 @@ by (rule transrec3_def [THEN def_transrec, THEN trans], simp) lemma transrec3_Limit: - "Limit(i) ==> + "Limit(i) ==> transrec3(i,a,b,c) = c(i, \j\i. transrec3(j,a,b,c))" by (rule transrec3_def [THEN def_transrec, THEN trans], force) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Nat_ZF.thy --- a/src/ZF/Nat_ZF.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Nat_ZF.thy Thu Mar 15 16:35:02 2012 +0000 @@ -9,7 +9,7 @@ definition nat :: i where - "nat == lfp(Inf, %X. {0} \ {succ(i). i:X})" + "nat == lfp(Inf, %X. {0} \ {succ(i). i \ X})" definition quasinat :: "i => o" where @@ -45,18 +45,18 @@ definition greater_than :: "i=>i" where - "greater_than(n) == {i:nat. n < i}" + "greater_than(n) == {i \ nat. n < i}" text{*No need for a less-than operator: a natural number is its list of predecessors!*} -lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i:X})" +lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i \ X})" apply (rule bnd_monoI) apply (cut_tac infinity, blast, blast) done -(* @{term"nat = {0} \ {succ(x). x:nat}"} *) +(* @{term"nat = {0} \ {succ(x). x \ nat}"} *) lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold]] (** Type checking of 0 and successor **) @@ -87,22 +87,22 @@ (*Mathematical induction*) lemma nat_induct [case_names 0 succ, induct set: nat]: - "[| n \ nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" + "[| n \ nat; P(0); !!x. [| x \ nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" by (erule def_induct [OF nat_def nat_bnd_mono], blast) lemma natE: assumes "n \ nat" - obtains (0) "n=0" | (succ) x where "x \ nat" "n=succ(x)" + obtains (0) "n=0" | (succ) x where "x \ nat" "n=succ(x)" using assms by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto lemma nat_into_Ord [simp]: "n \ nat ==> Ord(n)" by (erule nat_induct, auto) -(* @{term"i: nat ==> 0 \ i"}; same thing as @{term"0 nat ==> 0 \ i"}; same thing as @{term"0 i \ i"}; same thing as @{term"i nat ==> i \ i"}; same thing as @{term"i nat ==> ~ Limit(a)" by (induct a rule: nat_induct, auto) -lemma succ_natD: "succ(i): nat ==> i: nat" +lemma succ_natD: "succ(i): nat ==> i \ nat" by (rule Ord_trans [OF succI1], auto) lemma nat_succ_iff [iff]: "succ(n): nat \ n \ nat" @@ -137,15 +137,15 @@ apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) done -(* [| succ(i): k; k: nat |] ==> i: k *) +(* [| succ(i): k; k \ nat |] ==> i \ k *) lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] -lemma lt_nat_in_nat: "[| m nat |] ==> m: nat" +lemma lt_nat_in_nat: "[| m nat |] ==> m \ nat" apply (erule ltE) apply (erule Ord_trans, assumption, simp) done -lemma le_in_nat: "[| m \ n; n:nat |] ==> m:nat" +lemma le_in_nat: "[| m \ n; n \ nat |] ==> m \ nat" by (blast dest!: lt_nat_in_nat) @@ -160,8 +160,8 @@ lemma nat_induct_from_lemma [rule_format]: - "[| n \ nat; m: nat; - !!x. [| x: nat; m \ x; P(x) |] ==> P(succ(x)) |] + "[| n \ nat; m \ nat; + !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] ==> m \ n \ P(m) \ P(n)" apply (erule nat_induct) apply (simp_all add: distrib_simps le0_iff le_succ_iff) @@ -169,19 +169,19 @@ (*Induction starting from m rather than 0*) lemma nat_induct_from: - "[| m \ n; m: nat; n \ nat; + "[| m \ n; m \ nat; n \ nat; P(m); - !!x. [| x: nat; m \ x; P(x) |] ==> P(succ(x)) |] + !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] ==> P(n)" apply (blast intro: nat_induct_from_lemma) done (*Induction suitable for subtraction and less-than*) lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: - "[| m: nat; n \ nat; - !!x. x: nat ==> P(x,0); - !!y. y: nat ==> P(0,succ(y)); - !!x y. [| x: nat; y: nat; P(x,y) |] ==> P(succ(x),succ(y)) |] + "[| m \ nat; n \ nat; + !!x. x \ nat ==> P(x,0); + !!y. y \ nat ==> P(0,succ(y)); + !!x y. [| x \ nat; y \ nat; P(x,y) |] ==> P(succ(x),succ(y)) |] ==> P(m,n)" apply (erule_tac x = m in rev_bspec) apply (erule nat_induct, simp) @@ -194,7 +194,7 @@ (** Induction principle analogous to trancl_induct **) lemma succ_lt_induct_lemma [rule_format]: - "m: nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ + "m \ nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ (\n\nat. m P(m,n))" apply (erule nat_induct) apply (intro impI, rule nat_induct [THEN ballI]) @@ -205,7 +205,7 @@ lemma succ_lt_induct: "[| m nat; P(m,succ(m)); - !!x. [| x: nat; P(m,x) |] ==> P(m,succ(x)) |] + !!x. [| x \ nat; P(m,x) |] ==> P(m,succ(x)) |] ==> P(m,n)" by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) @@ -243,7 +243,7 @@ by (simp add: nat_case_def) lemma nat_case_type [TC]: - "[| n \ nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) |] + "[| n \ nat; a \ C(0); !!m. m \ nat ==> b(m): C(succ(m)) |] ==> nat_case(a,b,n) \ C(n)"; by (erule nat_induct, auto) @@ -266,7 +266,7 @@ apply (rule nat_case_0) done -lemma nat_rec_succ: "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" +lemma nat_rec_succ: "m \ nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) apply (rule wf_Memrel) apply (simp add: vimage_singleton_iff) @@ -274,12 +274,12 @@ (** The union of two natural numbers is a natural number -- their maximum **) -lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i \ j: nat" +lemma Un_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" apply (rule Un_least_lt [THEN ltD]) apply (simp_all add: lt_def) done -lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i \ j: nat" +lemma Int_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" apply (rule Int_greatest_lt [THEN ltD]) apply (simp_all add: lt_def) done diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/OrdQuant.thy Thu Mar 15 16:35:02 2012 +0000 @@ -98,12 +98,12 @@ by (blast intro: OUN_least_le OUN_upper_le le_Ord2 Ord_OUN) lemma OUN_UN_eq: - "(!!x. x:A ==> Ord(B(x))) + "(!!x. x \ A ==> Ord(B(x))) ==> (\z < (\x\A. B(x)). C(z)) = (\x\A. \z < B(x). C(z))" by (simp add: OUnion_def) lemma OUN_Union_eq: - "(!!x. x:X ==> Ord(x)) + "(!!x. x \ X ==> Ord(x)) ==> (\z < \(X). C(z)) = (\x\X. \z < x. C(z))" by (simp add: OUnion_def) @@ -168,11 +168,11 @@ subsubsection {*Rules for Ordinal-Indexed Unions*} -lemma OUN_I [intro]: "[| a b: (\z B(a) |] ==> b: (\z (\z R |] ==> R" + "[| b \ (\z B(a); a R |] ==> R" apply (unfold OUnion_def lt_def, blast) done diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Order.thy --- a/src/ZF/Order.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Order.thy Thu Mar 15 16:35:02 2012 +0000 @@ -46,16 +46,16 @@ definition mono_map :: "[i,i,i,i]=>i" (*Order-preserving maps*) where "mono_map(A,r,B,s) == - {f: A->B. \x\A. \y\A. :r \ :s}" + {f \ A->B. \x\A. \y\A. :r \ :s}" definition ord_iso :: "[i,i,i,i]=>i" (*Order isomorphisms*) where "ord_iso(A,r,B,s) == - {f: bij(A,B). \x\A. \y\A. :r \ :s}" + {f \ bij(A,B). \x\A. \y\A. :r \ :s}" definition pred :: "[i,i,i]=>i" (*Set of predecessors*) where - "pred(A,x,r) == {y:A. :r}" + "pred(A,x,r) == {y \ A. :r}" definition ord_iso_map :: "[i,i,i,i]=>i" (*Construction for linearity theorem*) where @@ -64,7 +64,7 @@ definition first :: "[i, i, i] => o" where - "first(u, X, R) == u:X & (\v\X. v\u \ \ R)" + "first(u, X, R) == u \ X & (\v\X. v\u \ \ R)" notation (xsymbols) @@ -78,7 +78,7 @@ by (unfold part_ord_def irrefl_def trans_on_def asym_def, blast) lemma linearE: - "[| linear(A,r); x:A; y:A; + "[| linear(A,r); x \ A; y \ A; :r ==> P; x=y ==> P; :r ==> P |] ==> P" by (simp add: linear_def, blast) @@ -107,12 +107,12 @@ (** Derived rules for pred(A,x,r) **) -lemma pred_iff: "y \ pred(A,x,r) \ :r & y:A" +lemma pred_iff: "y \ pred(A,x,r) \ :r & y \ A" by (unfold pred_def, blast) lemmas predI = conjI [THEN pred_iff [THEN iffD2]] -lemma predE: "[| y: pred(A,x,r); [| y:A; :r |] ==> P |] ==> P" +lemma predE: "[| y \ pred(A,x,r); [| y \ A; :r |] ==> P |] ==> P" by (simp add: pred_def) lemma pred_subset_under: "pred(A,x,r) \ r -`` {x}" @@ -126,7 +126,7 @@ by (simp add: pred_def, blast) lemma trans_pred_pred_eq: - "[| trans[A](r); :r; x:A; y:A |] + "[| trans[A](r); :r; x \ A; y \ A |] ==> pred(pred(A,x,r), y, r) = pred(A,y,r)" by (unfold trans_on_def pred_def, blast) @@ -244,39 +244,39 @@ (** Order-preserving (monotone) maps **) -lemma mono_map_is_fun: "f: mono_map(A,r,B,s) ==> f: A->B" +lemma mono_map_is_fun: "f \ mono_map(A,r,B,s) ==> f \ A->B" by (simp add: mono_map_def) lemma mono_map_is_inj: - "[| linear(A,r); wf[B](s); f: mono_map(A,r,B,s) |] ==> f: inj(A,B)" + "[| linear(A,r); wf[B](s); f \ mono_map(A,r,B,s) |] ==> f \ inj(A,B)" apply (unfold mono_map_def inj_def, clarify) apply (erule_tac x=w and y=x in linearE, assumption+) apply (force intro: apply_type dest: wf_on_not_refl)+ done lemma ord_isoI: - "[| f: bij(A, B); - !!x y. [| x:A; y:A |] ==> \ r \ \ s |] - ==> f: ord_iso(A,r,B,s)" + "[| f \ bij(A, B); + !!x y. [| x \ A; y \ A |] ==> \ r \ \ s |] + ==> f \ ord_iso(A,r,B,s)" by (simp add: ord_iso_def) lemma ord_iso_is_mono_map: - "f: ord_iso(A,r,B,s) ==> f: mono_map(A,r,B,s)" + "f \ ord_iso(A,r,B,s) ==> f \ mono_map(A,r,B,s)" apply (simp add: ord_iso_def mono_map_def) apply (blast dest!: bij_is_fun) done lemma ord_iso_is_bij: - "f: ord_iso(A,r,B,s) ==> f: bij(A,B)" + "f \ ord_iso(A,r,B,s) ==> f \ bij(A,B)" by (simp add: ord_iso_def) (*Needed? But ord_iso_converse is!*) lemma ord_iso_apply: - "[| f: ord_iso(A,r,B,s); : r; x:A; y:A |] ==> \ s" + "[| f \ ord_iso(A,r,B,s); : r; x \ A; y \ A |] ==> \ s" by (simp add: ord_iso_def) lemma ord_iso_converse: - "[| f: ord_iso(A,r,B,s); : s; x:B; y:B |] + "[| f \ ord_iso(A,r,B,s); : s; x \ B; y \ B |] ==> \ r" apply (simp add: ord_iso_def, clarify) apply (erule bspec [THEN bspec, THEN iffD2]) @@ -292,7 +292,7 @@ by (rule id_bij [THEN ord_isoI], simp) (*Symmetry of similarity*) -lemma ord_iso_sym: "f: ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)" +lemma ord_iso_sym: "f \ ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)" apply (simp add: ord_iso_def) apply (auto simp add: right_inverse_bij bij_converse_bij bij_is_fun [THEN apply_funtype]) @@ -300,7 +300,7 @@ (*Transitivity of similarity*) lemma mono_map_trans: - "[| g: mono_map(A,r,B,s); f: mono_map(B,s,C,t) |] + "[| g \ mono_map(A,r,B,s); f \ mono_map(B,s,C,t) |] ==> (f O g): mono_map(A,r,C,t)" apply (unfold mono_map_def) apply (auto simp add: comp_fun) @@ -308,7 +308,7 @@ (*Transitivity of similarity: the order-isomorphism relation*) lemma ord_iso_trans: - "[| g: ord_iso(A,r,B,s); f: ord_iso(B,s,C,t) |] + "[| g \ ord_iso(A,r,B,s); f \ ord_iso(B,s,C,t) |] ==> (f O g): ord_iso(A,r,C,t)" apply (unfold ord_iso_def, clarify) apply (frule bij_is_fun [of f]) @@ -319,8 +319,8 @@ (** Two monotone maps can make an order-isomorphism **) lemma mono_ord_isoI: - "[| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); - f O g = id(B); g O f = id(A) |] ==> f: ord_iso(A,r,B,s)" + "[| f \ mono_map(A,r,B,s); g \ mono_map(B,s,A,r); + f O g = id(B); g O f = id(A) |] ==> f \ ord_iso(A,r,B,s)" apply (simp add: ord_iso_def mono_map_def, safe) apply (intro fg_imp_bijective, auto) apply (subgoal_tac " \ r") @@ -330,8 +330,8 @@ lemma well_ord_mono_ord_isoI: "[| well_ord(A,r); well_ord(B,s); - f: mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) |] - ==> f: ord_iso(A,r,B,s)" + f \ mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) |] + ==> f \ ord_iso(A,r,B,s)" apply (intro mono_ord_isoI, auto) apply (frule mono_map_is_fun [THEN fun_is_rel]) apply (erule converse_converse [THEN subst], rule left_comp_inverse) @@ -343,13 +343,13 @@ (** Order-isomorphisms preserve the ordering's properties **) lemma part_ord_ord_iso: - "[| part_ord(B,s); f: ord_iso(A,r,B,s) |] ==> part_ord(A,r)" + "[| part_ord(B,s); f \ ord_iso(A,r,B,s) |] ==> part_ord(A,r)" apply (simp add: part_ord_def irrefl_def trans_on_def ord_iso_def) apply (fast intro: bij_is_fun [THEN apply_type]) done lemma linear_ord_iso: - "[| linear(B,s); f: ord_iso(A,r,B,s) |] ==> linear(A,r)" + "[| linear(B,s); f \ ord_iso(A,r,B,s) |] ==> linear(A,r)" apply (simp add: linear_def ord_iso_def, safe) apply (drule_tac x1 = "f`x" and x = "f`y" in bspec [THEN bspec]) apply (safe elim!: bij_is_fun [THEN apply_type]) @@ -358,15 +358,15 @@ done lemma wf_on_ord_iso: - "[| wf[B](s); f: ord_iso(A,r,B,s) |] ==> wf[A](r)" + "[| wf[B](s); f \ ord_iso(A,r,B,s) |] ==> wf[A](r)" apply (simp add: wf_on_def wf_def ord_iso_def, safe) -apply (drule_tac x = "{f`z. z:Z \ A}" in spec) +apply (drule_tac x = "{f`z. z \ Z \ A}" in spec) apply (safe intro!: equalityI) apply (blast dest!: equalityD1 intro: bij_is_fun [THEN apply_type])+ done lemma well_ord_ord_iso: - "[| well_ord(B,s); f: ord_iso(A,r,B,s) |] ==> well_ord(A,r)" + "[| well_ord(B,s); f \ ord_iso(A,r,B,s) |] ==> well_ord(A,r)" apply (unfold well_ord_def tot_ord_def) apply (fast elim!: part_ord_ord_iso linear_ord_iso wf_on_ord_iso) done @@ -377,7 +377,7 @@ (*Inductive argument for Kunen's Lemma 6.1, etc. Simple proof from Halmos, page 72*) lemma well_ord_iso_subset_lemma: - "[| well_ord(A,r); f: ord_iso(A,r, A',r); A'<= A; y: A |] + "[| well_ord(A,r); f \ ord_iso(A,r, A',r); A'<= A; y \ A |] ==> ~ : r" apply (simp add: well_ord_def ord_iso_def) apply (elim conjE CollectE) @@ -385,10 +385,10 @@ apply (blast dest: bij_is_fun [THEN apply_type]) done -(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment +(*Kunen's Lemma 6.1 \ there's no order-isomorphism to an initial segment of a well-ordering*) lemma well_ord_iso_predE: - "[| well_ord(A,r); f \ ord_iso(A, r, pred(A,x,r), r); x:A |] ==> P" + "[| well_ord(A,r); f \ ord_iso(A, r, pred(A,x,r), r); x \ A |] ==> P" apply (insert well_ord_iso_subset_lemma [of A r f "pred(A,x,r)" x]) apply (simp add: pred_subset) (*Now we know f`x < x *) @@ -400,7 +400,7 @@ (*Simple consequence of Lemma 6.1*) lemma well_ord_iso_pred_eq: "[| well_ord(A,r); f \ ord_iso(pred(A,a,r), r, pred(A,c,r), r); - a:A; c:A |] ==> a=c" + a \ A; c \ A |] ==> a=c" apply (frule well_ord_is_trans_on) apply (frule well_ord_is_linear) apply (erule_tac x=a and y=c in linearE, assumption+) @@ -413,7 +413,7 @@ (*Does not assume r is a wellordering!*) lemma ord_iso_image_pred: - "[|f \ ord_iso(A,r,B,s); a:A|] ==> f `` pred(A,a,r) = pred(B, f`a, s)" + "[|f \ ord_iso(A,r,B,s); a \ A|] ==> f `` pred(A,a,r) = pred(B, f`a, s)" apply (unfold ord_iso_def pred_def) apply (erule CollectE) apply (simp (no_asm_simp) add: image_fun [OF bij_is_fun Collect_subset]) @@ -434,7 +434,7 @@ (*But in use, A and B may themselves be initial segments. Then use trans_pred_pred_eq to simplify the pred(pred...) terms. See just below.*) lemma ord_iso_restrict_pred: - "[| f \ ord_iso(A,r,B,s); a:A |] + "[| f \ ord_iso(A,r,B,s); a \ A |] ==> restrict(f, pred(A,a,r)) \ ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)" apply (simp add: ord_iso_image_pred [symmetric]) apply (blast intro: ord_iso_restrict_image elim: predE) @@ -445,7 +445,7 @@ "[| well_ord(A,r); well_ord(B,s); : r; f \ ord_iso(pred(A,a,r), r, pred(B,b,s), s); g \ ord_iso(pred(A,c,r), r, pred(B,d,s), s); - a:A; c:A; b:B; d:B |] ==> : s" + a \ A; c \ A; b \ B; d \ B |] ==> : s" apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+) apply (subgoal_tac "b = g`a") apply (simp (no_asm_simp)) @@ -458,7 +458,7 @@ (*See Halmos, page 72*) lemma well_ord_iso_unique_lemma: "[| well_ord(A,r); - f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s); y: A |] + f \ ord_iso(A,r, B,s); g \ ord_iso(A,r, B,s); y \ A |] ==> ~ \ s" apply (frule well_ord_iso_subset_lemma) apply (rule_tac f = "converse (f) " and g = g in ord_iso_trans) @@ -476,7 +476,7 @@ (*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*) lemma well_ord_iso_unique: "[| well_ord(A,r); - f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s) |] ==> f = g" + f \ ord_iso(A,r, B,s); g \ ord_iso(A,r, B,s) |] ==> f = g" apply (rule fun_extension) apply (erule ord_iso_is_bij [THEN bij_is_fun])+ apply (subgoal_tac "f`x \ B & g`x \ B & linear(B,s)") @@ -522,7 +522,7 @@ apply (unfold mono_map_def) apply (simp (no_asm_simp) add: ord_iso_map_fun) apply safe -apply (subgoal_tac "x:A & ya:A & y:B & yb:B") +apply (subgoal_tac "x \ A & ya:A & y \ B & yb:B") apply (simp add: apply_equality [OF _ ord_iso_map_fun]) apply (unfold ord_iso_map_def) apply (blast intro: well_ord_iso_preserving, blast) @@ -545,7 +545,7 @@ (*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*) lemma domain_ord_iso_map_subset: "[| well_ord(A,r); well_ord(B,s); - a: A; a \ domain(ord_iso_map(A,r,B,s)) |] + a \ A; a \ domain(ord_iso_map(A,r,B,s)) |] ==> domain(ord_iso_map(A,r,B,s)) \ pred(A, a, r)" apply (unfold ord_iso_map_def) apply (safe intro!: predI) @@ -642,7 +642,7 @@ (** By Krzysztof Grabczewski. Lemmas involving the first element of a well ordered set **) -lemma first_is_elem: "first(b,B,r) ==> b:B" +lemma first_is_elem: "first(b,B,r) ==> b \ B" by (unfold first_def, blast) lemma well_ord_imp_ex1_first: diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/OrderArith.thy Thu Mar 15 16:35:02 2012 +0000 @@ -10,24 +10,24 @@ definition (*disjoint sum of two relations; underlies ordinal addition*) radd :: "[i,i,i,i]=>i" where - "radd(A,r,B,s) == - {z: (A+B) * (A+B). - (\x y. z = ) | - (\x' x. z = & :r) | + "radd(A,r,B,s) == + {z: (A+B) * (A+B). + (\x y. z = ) | + (\x' x. z = & :r) | (\y' y. z = & :s)}" definition (*lexicographic product of two relations; underlies ordinal multiplication*) rmult :: "[i,i,i,i]=>i" where - "rmult(A,r,B,s) == - {z: (A*B) * (A*B). - \x' y' x y. z = <, > & + "rmult(A,r,B,s) == + {z: (A*B) * (A*B). + \x' y' x y. z = <, > & (: r | (x'=x & : s))}" definition (*inverse image of a relation*) rvimage :: "[i,i,i]=>i" where - "rvimage(A,f,r) == {z: A*A. \x y. z = & : r}" + "rvimage(A,f,r) == {z \ A*A. \x y. z = & : r}" definition measure :: "[i, i\i] \ i" where @@ -38,33 +38,33 @@ subsubsection{*Rewrite rules. Can be used to obtain introduction rules*} -lemma radd_Inl_Inr_iff [iff]: - " \ radd(A,r,B,s) \ a:A & b:B" +lemma radd_Inl_Inr_iff [iff]: + " \ radd(A,r,B,s) \ a \ A & b \ B" by (unfold radd_def, blast) -lemma radd_Inl_iff [iff]: - " \ radd(A,r,B,s) \ a':A & a:A & :r" +lemma radd_Inl_iff [iff]: + " \ radd(A,r,B,s) \ a':A & a \ A & :r" by (unfold radd_def, blast) -lemma radd_Inr_iff [iff]: - " \ radd(A,r,B,s) \ b':B & b:B & :s" +lemma radd_Inr_iff [iff]: + " \ radd(A,r,B,s) \ b':B & b \ B & :s" by (unfold radd_def, blast) -lemma radd_Inr_Inl_iff [simp]: +lemma radd_Inr_Inl_iff [simp]: " \ radd(A,r,B,s) \ False" by (unfold radd_def, blast) -declare radd_Inr_Inl_iff [THEN iffD1, dest!] +declare radd_Inr_Inl_iff [THEN iffD1, dest!] subsubsection{*Elimination Rule*} lemma raddE: - "[| \ radd(A,r,B,s); - !!x y. [| p'=Inl(x); x:A; p=Inr(y); y:B |] ==> Q; - !!x' x. [| p'=Inl(x'); p=Inl(x); : r; x':A; x:A |] ==> Q; - !!y' y. [| p'=Inr(y'); p=Inr(y); : s; y':B; y:B |] ==> Q + "[| \ radd(A,r,B,s); + !!x y. [| p'=Inl(x); x \ A; p=Inr(y); y \ B |] ==> Q; + !!x' x. [| p'=Inl(x'); p=Inl(x); : r; x':A; x \ A |] ==> Q; + !!y' y. [| p'=Inr(y'); p=Inr(y); : s; y':B; y \ B |] ==> Q |] ==> Q" -by (unfold radd_def, blast) +by (unfold radd_def, blast) subsubsection{*Type checking*} @@ -77,9 +77,9 @@ subsubsection{*Linearity*} -lemma linear_radd: +lemma linear_radd: "[| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))" -by (unfold linear_def, blast) +by (unfold linear_def, blast) subsubsection{*Well-foundedness*} @@ -92,17 +92,17 @@ apply (erule_tac V = "y \ A + B" in thin_rl) apply (rule_tac ballI) apply (erule_tac r = r and a = x in wf_on_induct, assumption) - apply blast + apply blast txt{*Returning to main part of proof*} apply safe apply blast -apply (erule_tac r = s and a = ya in wf_on_induct, assumption, blast) +apply (erule_tac r = s and a = ya in wf_on_induct, assumption, blast) done lemma wf_radd: "[| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A [OF _ field_radd]) -apply (blast intro: wf_on_radd) +apply (blast intro: wf_on_radd) done lemma well_ord_radd: @@ -115,17 +115,17 @@ subsubsection{*An @{term ord_iso} congruence law*} lemma sum_bij: - "[| f: bij(A,C); g: bij(B,D) |] + "[| f \ bij(A,C); g \ bij(B,D) |] ==> (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ bij(A+B, C+D)" -apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))" +apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))" in lam_bijective) -apply (typecheck add: bij_is_inj inj_is_fun) -apply (auto simp add: left_inverse_bij right_inverse_bij) +apply (typecheck add: bij_is_inj inj_is_fun) +apply (auto simp add: left_inverse_bij right_inverse_bij) done -lemma sum_ord_iso_cong: - "[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> - (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) +lemma sum_ord_iso_cong: + "[| f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s') |] ==> + (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))" apply (unfold ord_iso_def) apply (safe intro!: sum_bij) @@ -133,27 +133,27 @@ apply (auto cong add: conj_cong simp add: bij_is_fun [THEN apply_type]) done -(*Could we prove an ord_iso result? Perhaps +(*Could we prove an ord_iso result? Perhaps ord_iso(A+B, radd(A,r,B,s), A \ B, r \ s) *) -lemma sum_disjoint_bij: "A \ B = 0 ==> +lemma sum_disjoint_bij: "A \ B = 0 ==> (\z\A+B. case(%x. x, %y. y, z)) \ bij(A+B, A \ B)" -apply (rule_tac d = "%z. if z:A then Inl (z) else Inr (z) " in lam_bijective) +apply (rule_tac d = "%z. if z \ A then Inl (z) else Inr (z) " in lam_bijective) apply auto done subsubsection{*Associativity*} lemma sum_assoc_bij: - "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) + "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) \ bij((A+B)+C, A+(B+C))" -apply (rule_tac d = "case (%x. Inl (Inl (x)), case (%x. Inl (Inr (x)), Inr))" +apply (rule_tac d = "case (%x. Inl (Inl (x)), case (%x. Inl (Inr (x)), Inr))" in lam_bijective) apply auto done lemma sum_assoc_ord_iso: - "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) - \ ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), + "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) + \ ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))" by (rule sum_assoc_bij [THEN ord_isoI], auto) @@ -162,19 +162,19 @@ subsubsection{*Rewrite rule. Can be used to obtain introduction rules*} -lemma rmult_iff [iff]: - "<, > \ rmult(A,r,B,s) \ - (: r & a':A & a:A & b': B & b: B) | - (: s & a'=a & a:A & b': B & b: B)" +lemma rmult_iff [iff]: + "<, > \ rmult(A,r,B,s) \ + (: r & a':A & a \ A & b': B & b \ B) | + (: s & a'=a & a \ A & b': B & b \ B)" by (unfold rmult_def, blast) -lemma rmultE: - "[| <, > \ rmult(A,r,B,s); - [| : r; a':A; a:A; b':B; b:B |] ==> Q; - [| : s; a:A; a'=a; b':B; b:B |] ==> Q +lemma rmultE: + "[| <, > \ rmult(A,r,B,s); + [| : r; a':A; a \ A; b':B; b \ B |] ==> Q; + [| : s; a \ A; a'=a; b':B; b \ B |] ==> Q |] ==> Q" -by blast +by blast subsubsection{*Type checking*} @@ -187,7 +187,7 @@ lemma linear_rmult: "[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))" -by (simp add: linear_def, blast) +by (simp add: linear_def, blast) subsubsection{*Well-foundedness*} @@ -206,7 +206,7 @@ lemma wf_rmult: "[| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A [OF _ field_rmult]) -apply (blast intro: wf_on_rmult) +apply (blast intro: wf_on_rmult) done lemma well_ord_rmult: @@ -220,17 +220,17 @@ subsubsection{*An @{term ord_iso} congruence law*} lemma prod_bij: - "[| f: bij(A,C); g: bij(B,D) |] + "[| f \ bij(A,C); g \ bij(B,D) |] ==> (lam :A*B. ) \ bij(A*B, C*D)" -apply (rule_tac d = "%. " +apply (rule_tac d = "%. " in lam_bijective) -apply (typecheck add: bij_is_inj inj_is_fun) -apply (auto simp add: left_inverse_bij right_inverse_bij) +apply (typecheck add: bij_is_inj inj_is_fun) +apply (auto simp add: left_inverse_bij right_inverse_bij) done -lemma prod_ord_iso_cong: - "[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] - ==> (lam :A*B. ) +lemma prod_ord_iso_cong: + "[| f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s') |] + ==> (lam :A*B. ) \ ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))" apply (unfold ord_iso_def) apply (safe intro!: prod_bij) @@ -243,7 +243,7 @@ (*Used??*) lemma singleton_prod_ord_iso: - "well_ord({x},xr) ==> + "well_ord({x},xr) ==> (\z\A. ) \ ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))" apply (rule singleton_prod_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) @@ -253,8 +253,8 @@ (*Here we build a complicated function term, then simplify it using case_cong, id_conv, comp_lam, case_case.*) lemma prod_sum_singleton_bij: - "a\C ==> - (\x\C*B + D. case(%x. x, %y., x)) + "a\C ==> + (\x\C*B + D. case(%x. x, %y., x)) \ bij(C*B + D, C*B \ {a}*D)" apply (rule subst_elem) apply (rule id_bij [THEN sum_bij, THEN comp_bij]) @@ -267,10 +267,10 @@ done lemma prod_sum_singleton_ord_iso: - "[| a:A; well_ord(A,r) |] ==> - (\x\pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y., x)) - \ ord_iso(pred(A,a,r)*B + pred(B,b,s), - radd(A*B, rmult(A,r,B,s), B, s), + "[| a \ A; well_ord(A,r) |] ==> + (\x\pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y., x)) + \ ord_iso(pred(A,a,r)*B + pred(B,b,s), + radd(A*B, rmult(A,r,B,s), B, s), pred(A,a,r)*B \ {a}*pred(B,b,s), rmult(A,r,B,s))" apply (rule prod_sum_singleton_bij [THEN ord_isoI]) apply (simp (no_asm_simp) add: pred_iff well_ord_is_wf [THEN wf_on_not_refl]) @@ -280,14 +280,14 @@ subsubsection{*Distributive law*} lemma sum_prod_distrib_bij: - "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) + "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) \ bij((A+B)*C, (A*C)+(B*C))" -by (rule_tac d = "case (%., %.) " +by (rule_tac d = "case (%., %.) " in lam_bijective, auto) lemma sum_prod_distrib_ord_iso: - "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) - \ ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t), + "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) + \ ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t), (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))" by (rule sum_prod_distrib_bij [THEN ord_isoI], auto) @@ -298,8 +298,8 @@ by (rule_tac d = "%>. <, z>" in lam_bijective, auto) lemma prod_assoc_ord_iso: - "(lam <, z>:(A*B)*C. >) - \ ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), + "(lam <, z>:(A*B)*C. >) + \ ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))" by (rule prod_assoc_bij [THEN ord_isoI], auto) @@ -307,7 +307,7 @@ subsubsection{*Rewrite rule*} -lemma rvimage_iff: " \ rvimage(A,f,r) \ : r & a:A & b:A" +lemma rvimage_iff: " \ rvimage(A,f,r) \ : r & a \ A & b \ A" by (unfold rvimage_def, blast) subsubsection{*Type checking*} @@ -323,20 +323,20 @@ subsubsection{*Partial Ordering Properties*} -lemma irrefl_rvimage: - "[| f: inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))" +lemma irrefl_rvimage: + "[| f \ inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))" apply (unfold irrefl_def rvimage_def) apply (blast intro: inj_is_fun [THEN apply_type]) done -lemma trans_on_rvimage: - "[| f: inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))" +lemma trans_on_rvimage: + "[| f \ inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))" apply (unfold trans_on_def rvimage_def) apply (blast intro: inj_is_fun [THEN apply_type]) done -lemma part_ord_rvimage: - "[| f: inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))" +lemma part_ord_rvimage: + "[| f \ inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))" apply (unfold part_ord_def) apply (blast intro!: irrefl_rvimage trans_on_rvimage) done @@ -344,13 +344,13 @@ subsubsection{*Linearity*} lemma linear_rvimage: - "[| f: inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))" -apply (simp add: inj_def linear_def rvimage_iff) -apply (blast intro: apply_funtype) + "[| f \ inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))" +apply (simp add: inj_def linear_def rvimage_iff) +apply (blast intro: apply_funtype) done -lemma tot_ord_rvimage: - "[| f: inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))" +lemma tot_ord_rvimage: + "[| f \ inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))" apply (unfold tot_ord_def) apply (blast intro!: part_ord_rvimage linear_rvimage) done @@ -361,19 +361,19 @@ lemma wf_rvimage [intro!]: "wf(r) ==> wf(rvimage(A,f,r))" apply (simp (no_asm_use) add: rvimage_def wf_eq_minimal) apply clarify -apply (subgoal_tac "\w. w \ {w: {f`x. x:Q}. \x. x: Q & (f`x = w) }") +apply (subgoal_tac "\w. w \ {w: {f`x. x \ Q}. \x. x \ Q & (f`x = w) }") apply (erule allE) apply (erule impE) apply assumption apply blast -apply blast +apply blast done text{*But note that the combination of @{text wf_imp_wf_on} and @{text wf_rvimage} gives @{prop "wf(r) ==> wf[C](rvimage(A,f,r))"}*} -lemma wf_on_rvimage: "[| f: A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))" +lemma wf_on_rvimage: "[| f \ A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))" apply (rule wf_onI2) -apply (subgoal_tac "\z\A. f`z=f`y \ z: Ba") +apply (subgoal_tac "\z\A. f`z=f`y \ z \ Ba") apply blast apply (erule_tac a = "f`y" in wf_on_induct) apply (blast intro!: apply_funtype) @@ -382,21 +382,21 @@ (*Note that we need only wf[A](...) and linear(A,...) to get the result!*) lemma well_ord_rvimage: - "[| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))" + "[| f \ inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))" apply (rule well_ordI) apply (unfold well_ord_def tot_ord_def) apply (blast intro!: wf_on_rvimage inj_is_fun) apply (blast intro!: linear_rvimage) done -lemma ord_iso_rvimage: - "f: bij(A,B) ==> f: ord_iso(A, rvimage(A,f,s), B, s)" +lemma ord_iso_rvimage: + "f \ bij(A,B) ==> f \ ord_iso(A, rvimage(A,f,s), B, s)" apply (unfold ord_iso_def) apply (simp add: rvimage_iff) done -lemma ord_iso_rvimage_eq: - "f: ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r \ A*A" +lemma ord_iso_rvimage_eq: + "f \ ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r \ A*A" by (unfold ord_iso_def rvimage_def, blast) @@ -463,14 +463,14 @@ text{*Could also be used to prove @{text wf_radd}*} lemma wf_Un: "[| range(r) \ domain(s) = 0; wf(r); wf(s) |] ==> wf(r \ s)" -apply (simp add: wf_def, clarify) -apply (rule equalityI) - prefer 2 apply blast -apply clarify +apply (simp add: wf_def, clarify) +apply (rule equalityI) + prefer 2 apply blast +apply clarify apply (drule_tac x=Z in spec) apply (drule_tac x="Z \ domain(s)" in spec) -apply simp -apply (blast intro: elim: equalityE) +apply simp +apply (blast intro: elim: equalityE) done subsubsection{*The Empty Relation*} @@ -496,29 +496,29 @@ lemma wf_measure [iff]: "wf(measure(A,f))" by (simp (no_asm) add: measure_eq_rvimage_Memrel wf_Memrel wf_rvimage) -lemma measure_iff [iff]: " \ measure(A,f) \ x:A & y:A & f(x) \ measure(A,f) \ x \ A & y \ A & f(x) A ==> Ord(f(x))" and inj: "!!x y. [|x \ A; y \ A; f(x) = f(y) |] ==> x=y" shows "linear(A, measure(A,f))" -apply (auto simp add: linear_def) -apply (rule_tac i="f(x)" and j="f(y)" in Ord_linear_lt) - apply (simp_all add: Ordf) -apply (blast intro: inj) +apply (auto simp add: linear_def) +apply (rule_tac i="f(x)" and j="f(y)" in Ord_linear_lt) + apply (simp_all add: Ordf) +apply (blast intro: inj) done lemma wf_on_measure: "wf[B](measure(A,f))" by (rule wf_imp_wf_on [OF wf_measure]) -lemma well_ord_measure: +lemma well_ord_measure: assumes Ordf: "!!x. x \ A ==> Ord(f(x))" and inj: "!!x y. [|x \ A; y \ A; f(x) = f(y) |] ==> x=y" shows "well_ord(A, measure(A,f))" apply (rule well_ordI) -apply (rule wf_on_measure) -apply (blast intro: linear_measure Ordf inj) +apply (rule wf_on_measure) +apply (blast intro: linear_measure Ordf inj) done lemma measure_type: "measure(A,f) \ A*A" @@ -529,7 +529,7 @@ lemma wf_on_Union: assumes wfA: "wf[A](r)" and wfB: "!!a. a\A ==> wf[B(a)](s)" - and ok: "!!a u v. [| \ s; v \ B(a); a \ A|] + and ok: "!!a u v. [| \ s; v \ B(a); a \ A|] ==> (\a'\A. \ r & u \ B(a')) | u \ B(a)" shows "wf[\a\A. B(a)](s)" apply (rule wf_onI2) @@ -538,25 +538,25 @@ apply (rule_tac a = a in wf_on_induct [OF wfA], assumption) apply (rule ballI) apply (rule_tac a = z in wf_on_induct [OF wfB], assumption, assumption) -apply (rename_tac u) -apply (drule_tac x=u in bspec, blast) +apply (rename_tac u) +apply (drule_tac x=u in bspec, blast) apply (erule mp, clarify) -apply (frule ok, assumption+, blast) +apply (frule ok, assumption+, blast) done subsubsection{*Bijections involving Powersets*} lemma Pow_sum_bij: - "(\Z \ Pow(A+B). <{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}>) + "(\Z \ Pow(A+B). <{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}>) \ bij(Pow(A+B), Pow(A)*Pow(B))" -apply (rule_tac d = "%. {Inl (x). x \ X} \ {Inr (y). y \ Y}" +apply (rule_tac d = "%. {Inl (x). x \ X} \ {Inr (y). y \ Y}" in lam_bijective) apply force+ done text{*As a special case, we have @{term "bij(Pow(A*B), A -> Pow(B))"} *} lemma Pow_Sigma_bij: - "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) + "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) \ bij(Pow(Sigma(A,B)), \ x \ A. Pow(B(x)))" apply (rule_tac d = "%f. \x \ A. \y \ f`x. {}" in lam_bijective) apply (blast intro: lam_type) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/OrderType.thy Thu Mar 15 16:35:02 2012 +0000 @@ -79,11 +79,11 @@ done lemma pred_Memrel: - "x:A ==> pred(A, x, Memrel(A)) = A \ x" + "x \ A ==> pred(A, x, Memrel(A)) = A \ x" by (unfold pred_def Memrel_def, blast) lemma Ord_iso_implies_eq_lemma: - "[| j R" + "[| j ord_iso(i,Memrel(i),j,Memrel(j)) |] ==> R" apply (frule lt_pred_Memrel) apply (erule ltE) apply (rule well_ord_Memrel [THEN well_ord_iso_predE, of i f j], auto) @@ -95,7 +95,7 @@ (*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*) lemma Ord_iso_implies_eq: - "[| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) |] + "[| Ord(i); Ord(j); f \ ord_iso(i,Memrel(i),j,Memrel(j)) |] ==> i=j" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: ord_iso_sym Ord_iso_implies_eq_lemma)+ @@ -115,7 +115,7 @@ (*Useful for cardinality reasoning; see CardinalArith.ML*) lemma ordermap_eq_image: - "[| wf[A](r); x:A |] + "[| wf[A](r); x \ A |] ==> ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)" apply (unfold ordermap_def pred_def) apply (simp (no_asm_simp)) @@ -125,7 +125,7 @@ (*Useful for rewriting PROVIDED pred is not unfolded until later!*) lemma ordermap_pred_unfold: - "[| wf[A](r); x:A |] + "[| wf[A](r); x \ A |] ==> ordermap(A,r) ` x = {ordermap(A,r)`y . y \ pred(A,x,r)}" by (simp add: ordermap_eq_image pred_subset ordermap_type [THEN image_fun]) @@ -135,14 +135,14 @@ (*The theorem above is [| wf[A](r); x \ A |] -==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y: A . \ r}} +==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \ r}} NOTE: the definition of ordermap used here delivers ordinals only if r is transitive. If r is the predecessor relation on the naturals then ordermap(nat,predr) ` n equals {n-1} and not n. A more complicated definition, like - ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y: A . \ r}}, + ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \ A . \ r}}, might eliminate the need for r to be transitive. *) @@ -151,7 +151,7 @@ subsubsection{*Showing that ordermap, ordertype yield ordinals *} lemma Ord_ordermap: - "[| well_ord(A,r); x:A |] ==> Ord(ordermap(A,r) ` x)" + "[| well_ord(A,r); x \ A |] ==> Ord(ordermap(A,r) ` x)" apply (unfold well_ord_def tot_ord_def part_ord_def, safe) apply (rule_tac a=x in wf_on_induct, assumption+) apply (simp (no_asm_simp) add: ordermap_pred_unfold) @@ -176,14 +176,14 @@ subsubsection{*ordermap preserves the orderings in both directions *} lemma ordermap_mono: - "[| : r; wf[A](r); w: A; x: A |] + "[| : r; wf[A](r); w \ A; x \ A |] ==> ordermap(A,r)`w \ ordermap(A,r)`x" apply (erule_tac x1 = x in ordermap_unfold [THEN ssubst], assumption, blast) done (*linearity of r is crucial here*) lemma converse_ordermap_mono: - "[| ordermap(A,r)`w \ ordermap(A,r)`x; well_ord(A,r); w: A; x: A |] + "[| ordermap(A,r)`w \ ordermap(A,r)`x; well_ord(A,r); w \ A; x \ A |] ==> : r" apply (unfold well_ord_def tot_ord_def, safe) apply (erule_tac x=w and y=x in linearE, assumption+) @@ -214,7 +214,7 @@ done lemma ordertype_eq: - "[| f: ord_iso(A,r,B,s); well_ord(B,s) |] + "[| f \ ord_iso(A,r,B,s); well_ord(B,s) |] ==> ordertype(A,r) = ordertype(B,s)" apply (frule well_ord_ord_iso, assumption) apply (rule Ord_iso_implies_eq, (erule Ord_ordertype)+) @@ -223,7 +223,7 @@ lemma ordertype_eq_imp_ord_iso: "[| ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s) |] - ==> \f. f: ord_iso(A,r,B,s)" + ==> \f. f \ ord_iso(A,r,B,s)" apply (rule exI) apply (rule ordertype_ord_iso [THEN ord_iso_trans], assumption) apply (erule ssubst) @@ -254,7 +254,7 @@ apply (rule Ord_0 [THEN ordertype_Memrel]) done -(*Ordertype of rvimage: [| f: bij(A,B); well_ord(B,s) |] ==> +(*Ordertype of rvimage: [| f \ bij(A,B); well_ord(B,s) |] ==> ordertype(A, rvimage(A,f,s)) = ordertype(B,s) *) lemmas bij_ordertype_vimage = ord_iso_rvimage [THEN ordertype_eq] @@ -262,7 +262,7 @@ (*Ordermap returns the same result if applied to an initial segment*) lemma ordermap_pred_eq_ordermap: - "[| well_ord(A,r); y:A; z: pred(A,y,r) |] + "[| well_ord(A,r); y \ A; z \ pred(A,y,r) |] ==> ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z" apply (frule wf_on_subset_A [OF well_ord_is_wf pred_subset]) apply (rule_tac a=z in wf_on_induct, assumption+) @@ -284,14 +284,14 @@ text{*Theorems by Krzysztof Grabczewski; proofs simplified by lcp *} -lemma ordertype_pred_subset: "[| well_ord(A,r); x:A |] ==> +lemma ordertype_pred_subset: "[| well_ord(A,r); x \ A |] ==> ordertype(pred(A,x,r),r) \ ordertype(A,r)" apply (simp add: ordertype_unfold well_ord_subset [OF _ pred_subset]) apply (fast intro: ordermap_pred_eq_ordermap elim: predE) done lemma ordertype_pred_lt: - "[| well_ord(A,r); x:A |] + "[| well_ord(A,r); x \ A |] ==> ordertype(pred(A,x,r),r) < ordertype(A,r)" apply (rule ordertype_pred_subset [THEN subset_imp_le, THEN leE]) apply (simp_all add: Ord_ordertype well_ord_subset [OF _ pred_subset]) @@ -304,7 +304,7 @@ well_ord(pred(A,x,r), r) *) lemma ordertype_pred_unfold: "well_ord(A,r) - ==> ordertype(A,r) = {ordertype(pred(A,x,r),r). x:A}" + ==> ordertype(A,r) = {ordertype(pred(A,x,r),r). x \ A}" apply (rule equalityI) apply (safe intro!: ordertype_pred_lt [THEN ltD]) apply (auto simp add: ordertype_def well_ord_is_wf [THEN ordermap_eq_image] @@ -367,7 +367,7 @@ (*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *) lemma pred_Inl_bij: - "a:A ==> (\x\pred(A,a,r). Inl(x)) + "a \ A ==> (\x\pred(A,a,r). Inl(x)) \ bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))" apply (unfold pred_def) apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective) @@ -375,7 +375,7 @@ done lemma ordertype_pred_Inl_eq: - "[| a:A; well_ord(A,r) |] + "[| a \ A; well_ord(A,r) |] ==> ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) = ordertype(pred(A,a,r), r)" apply (rule pred_Inl_bij [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq]) @@ -384,7 +384,7 @@ done lemma pred_Inr_bij: - "b:B ==> + "b \ B ==> id(A+pred(B,b,s)) \ bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))" apply (unfold pred_def id_def) @@ -392,7 +392,7 @@ done lemma ordertype_pred_Inr_eq: - "[| b:B; well_ord(A,r); well_ord(B,s) |] + "[| b \ B; well_ord(A,r); well_ord(B,s) |] ==> ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))" apply (rule pred_Inr_bij [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq]) @@ -476,7 +476,7 @@ done lemma subset_ord_iso_Memrel: - "[| f: ord_iso(A,Memrel(B),C,r); A<=B |] ==> f: ord_iso(A,Memrel(A),C,r)" + "[| 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) @@ -594,7 +594,7 @@ text{*Ordinal addition with limit ordinals *} lemma oadd_UN: - "[| !!x. x:A ==> Ord(j(x)); a:A |] + "[| !!x. x \ A ==> Ord(j(x)); a \ A |] ==> i ++ (\x\A. j(x)) = (\x\A. i++j(x))" by (blast intro: ltI Ord_UN Ord_oadd lt_oadd1 [THEN ltD] oadd_lt_mono2 [THEN ltD] @@ -632,15 +632,15 @@ lemma oadd_le_self2: "Ord(i) ==> i \ j++i" proof (induct i rule: trans_induct3) - case 0 thus ?case by (simp add: Ord_0_le) + case 0 thus ?case by (simp add: Ord_0_le) next - case (succ i) thus ?case by (simp add: oadd_succ succ_leI) + case (succ i) thus ?case by (simp add: oadd_succ succ_leI) next case (limit l) - hence "l = (\x\l. x)" + hence "l = (\x\l. x)" by (simp add: Union_eq_UN [symmetric] Limit_Union_eq) - also have "... \ (\x\l. j++x)" - by (rule le_implies_UN_le_UN) (rule limit.hyps) + also have "... \ (\x\l. j++x)" + by (rule le_implies_UN_le_UN) (rule limit.hyps) finally have "l \ (\x\l. j++x)" . thus ?case using limit.hyps by (simp add: oadd_Limit) qed @@ -691,7 +691,7 @@ It's probably simpler to define the difference recursively!*} lemma bij_sum_Diff: - "A<=B ==> (\y\B. if(y:A, Inl(y), Inr(y))) \ bij(B, A+(B-A))" + "A<=B ==> (\y\B. if(y \ A, Inl(y), Inr(y))) \ bij(B, A+(B-A))" apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective) apply (blast intro!: if_type) apply (fast intro!: case_type) @@ -763,13 +763,13 @@ subsubsection{*A useful unfolding law *} lemma pred_Pair_eq: - "[| a:A; b:B |] ==> pred(A*B, , rmult(A,r,B,s)) = + "[| a \ A; b \ B |] ==> pred(A*B, , rmult(A,r,B,s)) = pred(A,a,r)*B \ ({a} * pred(B,b,s))" apply (unfold pred_def, blast) done lemma ordertype_pred_Pair_eq: - "[| a:A; b:B; well_ord(A,r); well_ord(B,s) |] ==> + "[| a \ A; b \ B; well_ord(A,r); well_ord(B,s) |] ==> ordertype(pred(A*B, , rmult(A,r,B,s)), rmult(A,r,B,s)) = ordertype(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s))" @@ -908,7 +908,7 @@ text{*Ordinal multiplication with limit ordinals *} lemma omult_UN: - "[| Ord(i); !!x. x:A ==> Ord(j(x)) |] + "[| Ord(i); !!x. x \ A ==> Ord(j(x)) |] ==> i ** (\x\A. j(x)) = (\x\A. i**j(x))" by (simp (no_asm_simp) add: Ord_UN omult_unfold, blast) @@ -934,17 +934,17 @@ have o: "Ord(k)" "Ord(j)" by (rule lt_Ord [OF kj] le_Ord2 [OF kj])+ show ?thesis using i proof (induct i rule: trans_induct3) - case 0 thus ?case + case 0 thus ?case by simp next - case (succ i) thus ?case - by (simp add: o kj omult_succ oadd_le_mono) + case (succ i) thus ?case + by (simp add: o kj omult_succ oadd_le_mono) next case (limit l) - thus ?case - by (auto simp add: o kj omult_Limit le_implies_UN_le_UN) + thus ?case + by (auto simp add: o kj omult_Limit le_implies_UN_le_UN) qed -qed +qed lemma omult_lt_mono2: "[| k i**k < i**j" apply (rule ltI) @@ -966,30 +966,30 @@ lemma omult_lt_mono: "[| i' \ i; j' i'**j' < i**j" by (blast intro: lt_trans1 omult_le_mono1 omult_lt_mono2 Ord_succD elim: ltE) -lemma omult_le_self2: +lemma omult_le_self2: assumes i: "Ord(i)" and j: "0 j**i" proof - have oj: "Ord(j)" by (rule lt_Ord2 [OF j]) show ?thesis using i proof (induct i rule: trans_induct3) - case 0 thus ?case + case 0 thus ?case by simp next - case (succ i) - have "j \\ i ++ 0 < j \\ i ++ j" - by (rule oadd_lt_mono2 [OF j]) - with succ.hyps show ?case + case (succ i) + have "j \\ i ++ 0 < j \\ i ++ j" + by (rule oadd_lt_mono2 [OF j]) + with succ.hyps show ?case by (simp add: oj j omult_succ ) (rule lt_trans1) next case (limit l) - hence "l = (\x\l. x)" + hence "l = (\x\l. x)" by (simp add: Union_eq_UN [symmetric] Limit_Union_eq) - also have "... \ (\x\l. j**x)" - by (rule le_implies_UN_le_UN) (rule limit.hyps) + also have "... \ (\x\l. j**x)" + by (rule le_implies_UN_le_UN) (rule limit.hyps) finally have "l \ (\x\l. j**x)" . thus ?case using limit.hyps by (simp add: oj omult_Limit) qed -qed +qed text{*Further properties of ordinal multiplication *} diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Ordinal.thy --- a/src/ZF/Ordinal.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Ordinal.thy Thu Mar 15 16:35:02 2012 +0000 @@ -66,11 +66,11 @@ done lemma Transset_includes_domain: - "[| Transset(C); A*B \ C; b: B |] ==> A \ C" + "[| Transset(C); A*B \ C; b \ B |] ==> A \ C" by (blast dest: Transset_Pair_D) lemma Transset_includes_range: - "[| Transset(C); A*B \ C; a: A |] ==> B \ C" + "[| Transset(C); A*B \ C; a \ A |] ==> B \ C" by (blast dest: Transset_Pair_D) subsubsection{*Closure Properties*} @@ -276,12 +276,12 @@ lemma Memrel_iff [simp]: " \ Memrel(A) <-> a\b & a\A & b\A" by (unfold Memrel_def, blast) -lemma MemrelI [intro!]: "[| a: b; a: A; b: A |] ==> \ Memrel(A)" +lemma MemrelI [intro!]: "[| a \ b; a \ A; b \ A |] ==> \ Memrel(A)" by auto lemma MemrelE [elim!]: "[| \ Memrel(A); - [| a: A; b: A; a\b |] ==> P |] + [| a \ A; b \ A; a\b |] ==> P |] ==> P" by auto @@ -327,8 +327,8 @@ (*Epsilon induction over a transitive set*) lemma Transset_induct: - "[| i: k; Transset(k); - !!x.[| x: k; \y\x. P(y) |] ==> P(x) |] + "[| i \ k; Transset(k); + !!x.[| x \ k; \y\x. P(y) |] ==> P(x) |] ==> P(i)" apply (simp add: Transset_def) apply (erule wf_Memrel [THEN wf_induct2], blast+) @@ -364,7 +364,7 @@ text{*The trichotomy law for ordinals*} lemma Ord_linear_lt: assumes o: "Ord(i)" "Ord(j)" - obtains (lt) "i i" + obtains (lt) "i i" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: leI le_eqI sym o) + done lemma Ord_linear_le: assumes o: "Ord(i)" "Ord(j)" - obtains (le) "i \ j" | (ge) "j \ i" + obtains (le) "i \ j" | (ge) "j \ i" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: leI le_eqI o) + done @@ -598,7 +598,7 @@ by (unfold lt_def, blast) lemma UN_upper_le: - "[| a: A; i \ b(a); Ord(\x\A. b(x)) |] ==> i \ (\x\A. b(x))" + "[| a \ A; i \ b(a); Ord(\x\A. b(x)) |] ==> i \ (\x\A. b(x))" apply (frule ltD) apply (rule le_imp_subset [THEN subset_trans, THEN subset_imp_le]) apply (blast intro: lt_Ord UN_upper)+ @@ -608,7 +608,7 @@ by (auto simp: lt_def Ord_Union) lemma Union_upper_le: - "[| j: J; i\j; Ord(\(J)) |] ==> i \ \J" + "[| j \ J; i\j; Ord(\(J)) |] ==> i \ \J" apply (subst Union_eq_UN) apply (rule UN_upper_le, auto) done @@ -677,10 +677,10 @@ { fix y assume yi: "y y" using yi by (blast dest: le_imp_not_lt) - hence "succ(y) < i" using nsucc [of y] + have "~ i \ y" using yi by (blast dest: le_imp_not_lt) + hence "succ(y) < i" using nsucc [of y] by (blast intro: Ord_linear_lt [OF Osy Oi]) } - thus ?thesis using i Oi by (auto simp add: Limit_def) + thus ?thesis using i Oi by (auto simp add: Limit_def) qed lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P" @@ -703,7 +703,7 @@ lemma Ord_cases: assumes i: "Ord(i)" - obtains (0) "i=0" | (succ) j where "Ord(j)" "i=succ(j)" | (limit) "Limit(i)" + obtains (0) "i=0" | (succ) j where "Ord(j)" "i=succ(j)" | (limit) "Limit(i)" by (insert Ord_cases_disj [OF i], auto) lemma trans_induct3_raw: @@ -722,7 +722,7 @@ union is a limit ordinal.*} lemma Union_le: "[| !!x. x\I ==> x\j; Ord(j) |] ==> \(I) \ j" - by (auto simp add: le_subset_iff Union_least) + by (auto simp add: le_subset_iff Union_least) lemma Ord_set_cases: assumes I: "\i\I. Ord(i)" @@ -734,20 +734,20 @@ next fix j assume j: "Ord(j)" and UIj:"\(I) = succ(j)" - { assume "\i\I. i\j" - hence "\(I) \ j" - by (simp add: Union_le j) - hence False + { assume "\i\I. i\j" + hence "\(I) \ j" + by (simp add: Union_le j) + hence False by (simp add: UIj lt_not_refl) } then obtain i where i: "i \ I" "succ(j) \ i" using I j - by (atomize, auto simp add: not_le_iff_lt) + by (atomize, auto simp add: not_le_iff_lt) have "\(I) \ succ(j)" using UIj j by auto hence "i \ succ(j)" using i - by (simp add: le_subset_iff Union_subset_iff) - hence "succ(j) = i" using i - by (blast intro: le_anti_sym) + by (simp add: le_subset_iff Union_subset_iff) + hence "succ(j) = i" using i + by (blast intro: le_anti_sym) hence "succ(j) \ I" by (simp add: i) - thus ?thesis by (simp add: UIj) + thus ?thesis by (simp add: UIj) next assume "Limit(\I)" thus ?thesis by auto qed diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Perm.thy --- a/src/ZF/Perm.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Perm.thy Thu Mar 15 16:35:02 2012 +0000 @@ -26,12 +26,12 @@ definition (*one-to-one functions from A to B*) inj :: "[i,i]=>i" where - "inj(A,B) == { f: A->B. \w\A. \x\A. f`w=f`x \ w=x}" + "inj(A,B) == { f \ A->B. \w\A. \x\A. f`w=f`x \ w=x}" definition (*onto functions from A to B*) surj :: "[i,i]=>i" where - "surj(A,B) == { f: A->B . \y\B. \x\A. f`x=y}" + "surj(A,B) == { f \ A->B . \y\B. \x\A. f`x=y}" definition (*one-to-one and onto functions*) @@ -41,17 +41,17 @@ subsection{*Surjective Function Space*} -lemma surj_is_fun: "f: surj(A,B) ==> f: A->B" +lemma surj_is_fun: "f \ surj(A,B) ==> f \ A->B" apply (unfold surj_def) apply (erule CollectD1) done -lemma fun_is_surj: "f \ Pi(A,B) ==> f: surj(A,range(f))" +lemma fun_is_surj: "f \ Pi(A,B) ==> f \ surj(A,range(f))" apply (unfold surj_def) apply (blast intro: apply_equality range_of_fun domain_type) done -lemma surj_range: "f: surj(A,B) ==> range(f)=B" +lemma surj_range: "f \ surj(A,B) ==> range(f)=B" apply (unfold surj_def) apply (best intro: apply_Pair elim: range_type) done @@ -59,14 +59,14 @@ text{* A function with a right inverse is a surjection *} lemma f_imp_surjective: - "[| f: A->B; !!y. y:B ==> d(y): A; !!y. y:B ==> f`d(y) = y |] - ==> f: surj(A,B)" + "[| f \ A->B; !!y. y \ B ==> d(y): A; !!y. y \ B ==> f`d(y) = y |] + ==> f \ surj(A,B)" by (simp add: surj_def, blast) lemma lam_surjective: - "[| !!x. x:A ==> c(x): B; - !!y. y:B ==> d(y): A; - !!y. y:B ==> c(d(y)) = y + "[| !!x. x \ A ==> c(x): B; + !!y. y \ B ==> d(y): A; + !!y. y \ B ==> c(d(y)) = y |] ==> (\x\A. c(x)) \ surj(A,B)" apply (rule_tac d = d in f_imp_surjective) apply (simp_all add: lam_type) @@ -82,31 +82,31 @@ subsection{*Injective Function Space*} -lemma inj_is_fun: "f: inj(A,B) ==> f: A->B" +lemma inj_is_fun: "f \ inj(A,B) ==> f \ A->B" apply (unfold inj_def) apply (erule CollectD1) done text{*Good for dealing with sets of pairs, but a bit ugly in use [used in AC]*} lemma inj_equality: - "[| :f; :f; f: inj(A,B) |] ==> a=c" + "[| :f; :f; f \ inj(A,B) |] ==> a=c" apply (unfold inj_def) apply (blast dest: Pair_mem_PiD) done -lemma inj_apply_equality: "[| f:inj(A,B); f`a=f`b; a:A; b:A |] ==> a=b" +lemma inj_apply_equality: "[| f \ inj(A,B); f`a=f`b; a \ A; b \ A |] ==> a=b" by (unfold inj_def, blast) text{* A function with a left inverse is an injection *} -lemma f_imp_injective: "[| f: A->B; \x\A. d(f`x)=x |] ==> f: inj(A,B)" +lemma f_imp_injective: "[| f \ A->B; \x\A. d(f`x)=x |] ==> f \ inj(A,B)" apply (simp (no_asm_simp) add: inj_def) apply (blast intro: subst_context [THEN box_equals]) done lemma lam_injective: - "[| !!x. x:A ==> c(x): B; - !!x. x:A ==> d(c(x)) = x |] + "[| !!x. x \ A ==> c(x): B; + !!x. x \ A ==> d(c(x)) = x |] ==> (\x\A. c(x)) \ inj(A,B)" apply (rule_tac d = d in f_imp_injective) apply (simp_all add: lam_type) @@ -114,31 +114,31 @@ subsection{*Bijections*} -lemma bij_is_inj: "f: bij(A,B) ==> f: inj(A,B)" +lemma bij_is_inj: "f \ bij(A,B) ==> f \ inj(A,B)" apply (unfold bij_def) apply (erule IntD1) done -lemma bij_is_surj: "f: bij(A,B) ==> f: surj(A,B)" +lemma bij_is_surj: "f \ bij(A,B) ==> f \ surj(A,B)" apply (unfold bij_def) apply (erule IntD2) done -text{* f: bij(A,B) ==> f: A->B *} -lemmas bij_is_fun = bij_is_inj [THEN inj_is_fun] +lemma bij_is_fun: "f \ bij(A,B) ==> f \ A->B" + by (rule bij_is_inj [THEN inj_is_fun]) lemma lam_bijective: - "[| !!x. x:A ==> c(x): B; - !!y. y:B ==> d(y): A; - !!x. x:A ==> d(c(x)) = x; - !!y. y:B ==> c(d(y)) = y + "[| !!x. x \ A ==> c(x): B; + !!y. y \ B ==> d(y): A; + !!x. x \ A ==> d(c(x)) = x; + !!y. y \ B ==> c(d(y)) = y |] ==> (\x\A. c(x)) \ bij(A,B)" apply (unfold bij_def) apply (blast intro!: lam_injective lam_surjective) done lemma RepFun_bijective: "(\y\x. EX! y'. f(y') = f(y)) - ==> (\z\{f(y). y:x}. THE y. f(y) = z) \ bij({f(y). y:x}, x)" + ==> (\z\{f(y). y \ x}. THE y. f(y) = z) \ bij({f(y). y \ x}, x)" apply (rule_tac d = f in lam_bijective) apply (auto simp add: the_equality2) done @@ -146,12 +146,12 @@ subsection{*Identity Function*} -lemma idI [intro!]: "a:A ==> \ id(A)" +lemma idI [intro!]: "a \ A ==> \ id(A)" apply (unfold id_def) apply (erule lamI) done -lemma idE [elim!]: "[| p: id(A); !!x.[| x:A; p= |] ==> P |] ==> P" +lemma idE [elim!]: "[| p \ id(A); !!x.[| x \ A; p= |] ==> P |] ==> P" by (simp add: id_def lam_def, blast) lemma id_type: "id(A) \ A->A" @@ -159,7 +159,7 @@ apply (rule lam_type, assumption) done -lemma id_conv [simp]: "x:A ==> id(A)`x = x" +lemma id_conv [simp]: "x \ A ==> id(A)`x = x" apply (unfold id_def) apply (simp (no_asm_simp)) done @@ -198,7 +198,7 @@ subsection{*Converse of a Function*} -lemma inj_converse_fun: "f: inj(A,B) ==> converse(f) \ range(f)->A" +lemma inj_converse_fun: "f \ inj(A,B) ==> converse(f) \ range(f)->A" apply (unfold inj_def) apply (simp (no_asm_simp) add: Pi_iff function_def) apply (erule CollectE) @@ -210,10 +210,10 @@ text{*The premises are equivalent to saying that f is injective...*} lemma left_inverse_lemma: - "[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a" + "[| f \ A->B; converse(f): C->A; a \ A |] ==> converse(f)`(f`a) = a" by (blast intro: apply_Pair apply_equality converseI) -lemma left_inverse [simp]: "[| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a" +lemma left_inverse [simp]: "[| f \ inj(A,B); a \ A |] ==> converse(f)`(f`a) = a" by (blast intro: left_inverse_lemma inj_converse_fun inj_is_fun) lemma left_inverse_eq: @@ -223,21 +223,21 @@ lemmas left_inverse_bij = bij_is_inj [THEN left_inverse] lemma right_inverse_lemma: - "[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b" + "[| f \ A->B; converse(f): C->A; b \ C |] ==> f`(converse(f)`b) = b" by (rule apply_Pair [THEN converseD [THEN apply_equality]], auto) -(*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse? +(*Should the premises be f \ surj(A,B), b \ B for symmetry with left_inverse? No: they would not imply that converse(f) was a function! *) lemma right_inverse [simp]: - "[| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b" + "[| f \ inj(A,B); b \ range(f) |] ==> f`(converse(f)`b) = b" by (blast intro: right_inverse_lemma inj_converse_fun inj_is_fun) -lemma right_inverse_bij: "[| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b" +lemma right_inverse_bij: "[| f \ bij(A,B); b \ B |] ==> f`(converse(f)`b) = b" by (force simp add: bij_def surj_range) subsection{*Converses of Injections, Surjections, Bijections*} -lemma inj_converse_inj: "f: inj(A,B) ==> converse(f): inj(range(f), A)" +lemma inj_converse_inj: "f \ inj(A,B) ==> converse(f): inj(range(f), A)" apply (rule f_imp_injective) apply (erule inj_converse_fun, clarify) apply (rule right_inverse) @@ -245,12 +245,12 @@ apply blast done -lemma inj_converse_surj: "f: inj(A,B) ==> converse(f): surj(range(f), A)" +lemma inj_converse_surj: "f \ inj(A,B) ==> converse(f): surj(range(f), A)" by (blast intro: f_imp_surjective inj_converse_fun left_inverse inj_is_fun range_of_fun [THEN apply_type]) text{*Adding this as an intro! rule seems to cause looping*} -lemma bij_converse_bij [TC]: "f: bij(A,B) ==> converse(f): bij(B,A)" +lemma bij_converse_bij [TC]: "f \ bij(A,B) ==> converse(f): bij(B,A)" apply (unfold bij_def) apply (fast elim: surj_range [THEN subst] inj_converse_inj inj_converse_surj) done @@ -298,10 +298,10 @@ lemma image_comp: "(r O s)``A = r``(s``A)" by blast -lemma inj_inj_range: "f: inj(A,B) ==> f \ inj(A,range(f))" +lemma inj_inj_range: "f \ inj(A,B) ==> f \ inj(A,range(f))" by (auto simp add: inj_def Pi_iff function_def) -lemma inj_bij_range: "f: inj(A,B) ==> f \ bij(A,range(f))" +lemma inj_bij_range: "f \ inj(A,B) ==> f \ bij(A,range(f))" by (auto simp add: bij_def intro: inj_inj_range inj_is_fun fun_is_surj) @@ -337,14 +337,14 @@ by (unfold function_def, blast) text{*Don't think the premises can be weakened much*} -lemma comp_fun: "[| g: A->B; f: B->C |] ==> (f O g) \ A->C" +lemma comp_fun: "[| g \ A->B; f \ B->C |] ==> (f O g) \ A->C" apply (auto simp add: Pi_def comp_function Pow_iff comp_rel) apply (subst range_rel_subset [THEN domain_comp_eq], auto) done -(*Thanks to the new definition of "apply", the premise f: B->C is gone!*) +(*Thanks to the new definition of "apply", the premise f \ B->C is gone!*) lemma comp_fun_apply [simp]: - "[| g: A->B; a:A |] ==> (f O g)`a = f`(g`a)" + "[| g \ A->B; a \ A |] ==> (f O g)`a = f`(g`a)" apply (frule apply_Pair, assumption) apply (simp add: apply_def image_comp) apply (blast dest: apply_equality) @@ -352,7 +352,7 @@ text{*Simplifies compositions of lambda-abstractions*} lemma comp_lam: - "[| !!x. x:A ==> b(x): B |] + "[| !!x. x \ A ==> b(x): B |] ==> (\y\B. c(y)) O (\x\A. b(x)) = (\x\A. c(b(x)))" apply (subgoal_tac "(\x\A. b(x)) \ A -> B") apply (rule fun_extension) @@ -363,7 +363,7 @@ done lemma comp_inj: - "[| g: inj(A,B); f: inj(B,C) |] ==> (f O g) \ inj(A,C)" + "[| g \ inj(A,B); f \ inj(B,C) |] ==> (f O g) \ inj(A,C)" apply (frule inj_is_fun [of g]) apply (frule inj_is_fun [of f]) apply (rule_tac d = "%y. converse (g) ` (converse (f) ` y)" in f_imp_injective) @@ -371,13 +371,13 @@ done lemma comp_surj: - "[| g: surj(A,B); f: surj(B,C) |] ==> (f O g) \ surj(A,C)" + "[| g \ surj(A,B); f \ surj(B,C) |] ==> (f O g) \ surj(A,C)" apply (unfold surj_def) apply (blast intro!: comp_fun comp_fun_apply) done lemma comp_bij: - "[| g: bij(A,B); f: bij(B,C) |] ==> (f O g) \ bij(A,C)" + "[| g \ bij(A,B); f \ bij(B,C) |] ==> (f O g) \ bij(A,C)" apply (unfold bij_def) apply (blast intro: comp_inj comp_surj) done @@ -390,11 +390,11 @@ Artificial Intelligence, 10:1--27, 1978.*} lemma comp_mem_injD1: - "[| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)" + "[| (f O g): inj(A,C); g \ A->B; f \ B->C |] ==> g \ inj(A,B)" by (unfold inj_def, force) lemma comp_mem_injD2: - "[| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)" + "[| (f O g): inj(A,C); g \ surj(A,B); f \ B->C |] ==> f \ inj(B,C)" apply (unfold inj_def surj_def, safe) apply (rule_tac x1 = x in bspec [THEN bexE]) apply (erule_tac [3] x1 = w in bspec [THEN bexE], assumption+, safe) @@ -404,14 +404,14 @@ done lemma comp_mem_surjD1: - "[| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)" + "[| (f O g): surj(A,C); g \ A->B; f \ B->C |] ==> f \ surj(B,C)" apply (unfold surj_def) apply (blast intro!: comp_fun_apply [symmetric] apply_funtype) done lemma comp_mem_surjD2: - "[| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)" + "[| (f O g): surj(A,C); g \ A->B; f \ inj(B,C) |] ==> g \ surj(A,B)" apply (unfold inj_def surj_def, safe) apply (drule_tac x = "f`y" in bspec, auto) apply (blast intro: apply_funtype) @@ -420,17 +420,17 @@ subsubsection{*Inverses of Composition*} text{*left inverse of composition; one inclusion is - @{term "f: A->B ==> id(A) \ converse(f) O f"} *} -lemma left_comp_inverse: "f: inj(A,B) ==> converse(f) O f = id(A)" + @{term "f \ A->B ==> id(A) \ converse(f) O f"} *} +lemma left_comp_inverse: "f \ inj(A,B) ==> converse(f) O f = id(A)" apply (unfold inj_def, clarify) apply (rule equalityI) apply (auto simp add: apply_iff, blast) done text{*right inverse of composition; one inclusion is - @{term "f: A->B ==> f O converse(f) \ id(B)"} *} + @{term "f \ A->B ==> f O converse(f) \ id(B)"} *} lemma right_comp_inverse: - "f: surj(A,B) ==> f O converse(f) = id(B)" + "f \ surj(A,B) ==> f O converse(f) = id(B)" apply (simp add: surj_def, clarify) apply (rule equalityI) apply (best elim: domain_type range_type dest: apply_equality2) @@ -441,7 +441,7 @@ subsubsection{*Proving that a Function is a Bijection*} lemma comp_eq_id_iff: - "[| f: A->B; g: B->A |] ==> f O g = id(B) \ (\y\B. f`(g`y)=y)" + "[| f \ A->B; g \ B->A |] ==> f O g = id(B) \ (\y\B. f`(g`y)=y)" apply (unfold id_def, safe) apply (drule_tac t = "%h. h`y " in subst_context) apply simp @@ -451,17 +451,17 @@ done lemma fg_imp_bijective: - "[| f: A->B; g: B->A; f O g = id(B); g O f = id(A) |] ==> f \ bij(A,B)" + "[| f \ A->B; g \ B->A; f O g = id(B); g O f = id(A) |] ==> f \ bij(A,B)" apply (unfold bij_def) apply (simp add: comp_eq_id_iff) apply (blast intro: f_imp_injective f_imp_surjective apply_funtype) done -lemma nilpotent_imp_bijective: "[| f: A->A; f O f = id(A) |] ==> f \ bij(A,A)" +lemma nilpotent_imp_bijective: "[| f \ A->A; f O f = id(A) |] ==> f \ bij(A,A)" by (blast intro: fg_imp_bijective) lemma invertible_imp_bijective: - "[| converse(f): B->A; f: A->B |] ==> f \ bij(A,B)" + "[| converse(f): B->A; f \ A->B |] ==> f \ bij(A,B)" by (simp add: fg_imp_bijective comp_eq_id_iff left_inverse_lemma right_inverse_lemma) @@ -471,15 +471,15 @@ text{*Theorem by KG, proof by LCP*} lemma inj_disjoint_Un: - "[| f: inj(A,B); g: inj(C,D); B \ D = 0 |] - ==> (\a\A \ C. if a:A then f`a else g`a) \ inj(A \ C, B \ D)" -apply (rule_tac d = "%z. if z:B then converse (f) `z else converse (g) `z" + "[| f \ inj(A,B); g \ inj(C,D); B \ D = 0 |] + ==> (\a\A \ C. if a \ A then f`a else g`a) \ inj(A \ C, B \ D)" +apply (rule_tac d = "%z. if z \ B then converse (f) `z else converse (g) `z" in lam_injective) apply (auto simp add: inj_is_fun [THEN apply_type]) done lemma surj_disjoint_Un: - "[| f: surj(A,B); g: surj(C,D); A \ C = 0 |] + "[| f \ surj(A,B); g \ surj(C,D); A \ C = 0 |] ==> (f \ g) \ surj(A \ C, B \ D)" apply (simp add: surj_def fun_disjoint_Un) apply (blast dest!: domain_of_fun @@ -487,9 +487,9 @@ done text{*A simple, high-level proof; the version for injections follows from it, - using @{term "f:inj(A,B) \ f:bij(A,range(f))"} *} + using @{term "f \ inj(A,B) \ f \ bij(A,range(f))"} *} lemma bij_disjoint_Un: - "[| f: bij(A,B); g: bij(C,D); A \ C = 0; B \ D = 0 |] + "[| f \ bij(A,B); g \ bij(C,D); A \ C = 0; B \ D = 0 |] ==> (f \ g) \ bij(A \ C, B \ D)" apply (rule invertible_imp_bijective) apply (subst converse_Un) @@ -500,7 +500,7 @@ subsubsection{*Restrictions as Surjections and Bijections*} lemma surj_image: - "f: Pi(A,B) ==> f: surj(A, f``A)" + "f \ Pi(A,B) ==> f \ surj(A, f``A)" apply (simp add: surj_def) apply (blast intro: apply_equality apply_Pair Pi_type) done @@ -509,18 +509,18 @@ by (auto simp add: restrict_def) lemma restrict_inj: - "[| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)" + "[| f \ inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)" apply (unfold inj_def) apply (safe elim!: restrict_type2, auto) done -lemma restrict_surj: "[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)" +lemma restrict_surj: "[| f \ Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)" apply (insert restrict_type2 [THEN surj_image]) apply (simp add: restrict_image) done lemma restrict_bij: - "[| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)" + "[| f \ inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)" apply (simp add: inj_def bij_def) apply (blast intro: restrict_surj surj_is_fun) done @@ -528,13 +528,13 @@ subsubsection{*Lemmas for Ramsey's Theorem*} -lemma inj_weaken_type: "[| f: inj(A,B); B<=D |] ==> f: inj(A,D)" +lemma inj_weaken_type: "[| f \ inj(A,B); B<=D |] ==> f \ inj(A,D)" apply (unfold inj_def) apply (blast intro: fun_weaken_type) done lemma inj_succ_restrict: - "[| f: inj(succ(m), A) |] ==> restrict(f,m) \ inj(m, A-{f`m})" + "[| f \ inj(succ(m), A) |] ==> restrict(f,m) \ inj(m, A-{f`m})" apply (rule restrict_bij [THEN bij_is_inj, THEN inj_weaken_type], assumption, blast) apply (unfold inj_def) apply (fast elim: range_type mem_irrefl dest: apply_equality) @@ -542,7 +542,7 @@ lemma inj_extend: - "[| f: inj(A,B); a\A; b\B |] + "[| f \ inj(A,B); a\A; b\B |] ==> cons(,f) \ inj(cons(a,A), cons(b,B))" apply (unfold inj_def) apply (force intro: apply_type simp add: fun_extend) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/QPair.thy --- a/src/ZF/QPair.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/QPair.thy Thu Mar 15 16:35:02 2012 +0000 @@ -5,7 +5,7 @@ Many proofs are borrowed from pair.thy and sum.thy Do we EVER have rank(a) < rank() ? Perhaps if the latter rank -is not a limit ordinal? +is not a limit ordinal? *) header{*Quine-Inspired Ordered Pairs and Disjoint Sums*} @@ -38,16 +38,16 @@ definition qconverse :: "i => i" where - "qconverse(r) == {z. w:r, \x y. w= & z=}" + "qconverse(r) == {z. w \ r, \x y. w= & z=}" definition QSigma :: "[i, i => i] => i" where "QSigma(A,B) == \x\A. \y\B(x). {}" syntax - "_QSUM" :: "[idt, i, i] => i" ("(3QSUM _:_./ _)" 10) + "_QSUM" :: "[idt, i, i] => i" ("(3QSUM _ \ _./ _)" 10) translations - "QSUM x:A. B" => "CONST QSigma(A, %x. B)" + "QSUM x \ A. B" => "CONST QSigma(A, %x. B)" abbreviation qprod (infixr "<*>" 80) where @@ -94,21 +94,21 @@ subsubsection{*QSigma: Disjoint union of a family of sets Generalizes Cartesian product*} -lemma QSigmaI [intro!]: "[| a:A; b:B(a) |] ==> \ QSigma(A,B)" +lemma QSigmaI [intro!]: "[| a \ A; b \ B(a) |] ==> \ QSigma(A,B)" by (simp add: QSigma_def) (** Elimination rules for :A*B -- introducing no eigenvariables **) lemma QSigmaE [elim!]: - "[| c: QSigma(A,B); - !!x y.[| x:A; y:B(x); c= |] ==> P + "[| c \ QSigma(A,B); + !!x y.[| x \ A; y \ B(x); c= |] ==> P |] ==> P" -by (simp add: QSigma_def, blast) +by (simp add: QSigma_def, blast) lemma QSigmaE2 [elim!]: - "[| : QSigma(A,B); [| a:A; b:B(a) |] ==> P |] ==> P" -by (simp add: QSigma_def) + "[| : QSigma(A,B); [| a \ A; b \ B(a) |] ==> P |] ==> P" +by (simp add: QSigma_def) lemma QSigmaD1: " \ QSigma(A,B) ==> a \ A" by blast @@ -117,9 +117,9 @@ by blast lemma QSigma_cong: - "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> + "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> QSigma(A,B) = QSigma(A',B')" -by (simp add: QSigma_def) +by (simp add: QSigma_def) lemma QSigma_empty1 [simp]: "QSigma(0,B) = 0" by blast @@ -136,13 +136,13 @@ lemma qsnd_conv [simp]: "qsnd() = b" by (simp add: qsnd_def) -lemma qfst_type [TC]: "p:QSigma(A,B) ==> qfst(p) \ A" +lemma qfst_type [TC]: "p \ QSigma(A,B) ==> qfst(p) \ A" by auto -lemma qsnd_type [TC]: "p:QSigma(A,B) ==> qsnd(p) \ B(qfst(p))" +lemma qsnd_type [TC]: "p \ QSigma(A,B) ==> qsnd(p) \ B(qfst(p))" by auto -lemma QPair_qfst_qsnd_eq: "a: QSigma(A,B) ==> = a" +lemma QPair_qfst_qsnd_eq: "a \ QSigma(A,B) ==> = a" by auto @@ -154,13 +154,13 @@ lemma qsplit_type [elim!]: - "[| p:QSigma(A,B); - !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() + "[| p \ QSigma(A,B); + !!x y.[| x \ A; y \ B(x) |] ==> c(x,y):C() |] ==> qsplit(%x y. c(x,y), p) \ C(p)" -by auto +by auto -lemma expand_qsplit: - "u: A<*>B ==> R(qsplit(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" +lemma expand_qsplit: + "u \ A<*>B ==> R(qsplit(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" apply (simp add: qsplit_def, auto) done @@ -172,10 +172,10 @@ lemma qsplitE: - "[| qsplit(R,z); z:QSigma(A,B); - !!x y. [| z = ; R(x,y) |] ==> P + "[| qsplit(R,z); z \ QSigma(A,B); + !!x y. [| z = ; R(x,y) |] ==> P |] ==> P" -by (simp add: qsplit_def, auto) +by (simp add: qsplit_def, auto) lemma qsplitD: "qsplit(R,) ==> R(a,b)" by (simp add: qsplit_def) @@ -190,10 +190,10 @@ by (simp add: qconverse_def, blast) lemma qconverseE [elim!]: - "[| yx \ qconverse(r); - !!x y. [| yx=; :r |] ==> P + "[| yx \ qconverse(r); + !!x y. [| yx=; :r |] ==> P |] ==> P" -by (simp add: qconverse_def, blast) +by (simp add: qconverse_def, blast) lemma qconverse_qconverse: "r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" by blast @@ -223,11 +223,11 @@ (** Elimination rules **) lemma qsumE [elim!]: - "[| u: A <+> B; - !!x. [| x:A; u=QInl(x) |] ==> P; - !!y. [| y:B; u=QInr(y) |] ==> P + "[| u \ A <+> B; + !!x. [| x \ A; u=QInl(x) |] ==> P; + !!y. [| y \ B; u=QInr(y) |] ==> P |] ==> P" -by (simp add: qsum_defs, blast) +by (simp add: qsum_defs, blast) (** Injection and freeness equivalences, for rewriting **) @@ -254,16 +254,16 @@ lemmas QInl_neq_QInr = QInl_QInr_iff [THEN iffD1, THEN FalseE, elim!] lemmas QInr_neq_QInl = QInr_QInl_iff [THEN iffD1, THEN FalseE, elim!] -lemma QInlD: "QInl(a): A<+>B ==> a: A" +lemma QInlD: "QInl(a): A<+>B ==> a \ A" by blast -lemma QInrD: "QInr(b): A<+>B ==> b: B" +lemma QInrD: "QInr(b): A<+>B ==> b \ B" by blast (** <+> is itself injective... who cares?? **) lemma qsum_iff: - "u: A <+> B \ (\x. x:A & u=QInl(x)) | (\y. y:B & u=QInr(y))" + "u \ A <+> B \ (\x. x \ A & u=QInl(x)) | (\y. y \ B & u=QInr(y))" by blast lemma qsum_subset_iff: "A <+> B \ C <+> D \ A<=C & B<=D" @@ -284,21 +284,21 @@ by (simp add: qsum_defs ) lemma qcase_type: - "[| u: A <+> B; - !!x. x: A ==> c(x): C(QInl(x)); - !!y. y: B ==> d(y): C(QInr(y)) + "[| u \ A <+> B; + !!x. x \ A ==> c(x): C(QInl(x)); + !!y. y \ B ==> d(y): C(QInr(y)) |] ==> qcase(c,d,u) \ C(u)" -by (simp add: qsum_defs, auto) +by (simp add: qsum_defs, auto) (** Rules for the Part primitive **) -lemma Part_QInl: "Part(A <+> B,QInl) = {QInl(x). x: A}" +lemma Part_QInl: "Part(A <+> B,QInl) = {QInl(x). x \ A}" by blast -lemma Part_QInr: "Part(A <+> B,QInr) = {QInr(y). y: B}" +lemma Part_QInr: "Part(A <+> B,QInr) = {QInr(y). y \ B}" by blast -lemma Part_QInr2: "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y: Part(B,h)}" +lemma Part_QInr2: "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y \ Part(B,h)}" by blast lemma Part_qsum_equality: "C \ A <+> B ==> Part(C,QInl) \ Part(C,QInr) = C" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Sum.thy --- a/src/ZF/Sum.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Sum.thy Thu Mar 15 16:35:02 2012 +0000 @@ -23,24 +23,24 @@ (*operator for selecting out the various summands*) definition Part :: "[i,i=>i] => i" where - "Part(A,h) == {x: A. \z. x = h(z)}" + "Part(A,h) == {x \ A. \z. x = h(z)}" subsection{*Rules for the @{term Part} Primitive*} -lemma Part_iff: - "a \ Part(A,h) \ a:A & (\y. a=h(y))" +lemma Part_iff: + "a \ Part(A,h) \ a \ A & (\y. a=h(y))" apply (unfold Part_def) apply (rule separation) done -lemma Part_eqI [intro]: +lemma Part_eqI [intro]: "[| a \ A; a=h(b) |] ==> a \ Part(A,h)" by (unfold Part_def, blast) lemmas PartI = refl [THEN [2] Part_eqI] -lemma PartE [elim!]: - "[| a \ Part(A,h); !!z. [| a \ A; a=h(z) |] ==> P +lemma PartE [elim!]: + "[| a \ Part(A,h); !!z. [| a \ A; a=h(z) |] ==> P |] ==> P" apply (unfold Part_def, blast) done @@ -69,11 +69,11 @@ (** Elimination rules **) lemma sumE [elim!]: - "[| u: A+B; - !!x. [| x:A; u=Inl(x) |] ==> P; - !!y. [| y:B; u=Inr(y) |] ==> P + "[| u \ A+B; + !!x. [| x \ A; u=Inl(x) |] ==> P; + !!y. [| y \ B; u=Inr(y) |] ==> P |] ==> P" -by (unfold sum_defs, blast) +by (unfold sum_defs, blast) (** Injection and freeness equivalences, for rewriting **) @@ -100,13 +100,13 @@ lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE, elim!] -lemma InlD: "Inl(a): A+B ==> a: A" +lemma InlD: "Inl(a): A+B ==> a \ A" by blast -lemma InrD: "Inr(b): A+B ==> b: B" +lemma InrD: "Inr(b): A+B ==> b \ B" by blast -lemma sum_iff: "u: A+B \ (\x. x:A & u=Inl(x)) | (\y. y:B & u=Inr(y))" +lemma sum_iff: "u \ A+B \ (\x. x \ A & u=Inl(x)) | (\y. y \ B & u=Inr(y))" by blast lemma Inl_in_sum_iff [simp]: "(Inl(x) \ A+B) \ (x \ A)"; @@ -134,27 +134,27 @@ by (simp add: sum_defs) lemma case_type [TC]: - "[| u: A+B; - !!x. x: A ==> c(x): C(Inl(x)); - !!y. y: B ==> d(y): C(Inr(y)) + "[| u \ A+B; + !!x. x \ A ==> c(x): C(Inl(x)); + !!y. y \ B ==> d(y): C(Inr(y)) |] ==> case(c,d,u) \ C(u)" by auto -lemma expand_case: "u: A+B ==> - R(case(c,d,u)) \ - ((\x\A. u = Inl(x) \ R(c(x))) & +lemma expand_case: "u \ A+B ==> + R(case(c,d,u)) \ + ((\x\A. u = Inl(x) \ R(c(x))) & (\y\B. u = Inr(y) \ R(d(y))))" by auto lemma case_cong: - "[| z: A+B; - !!x. x:A ==> c(x)=c'(x); - !!y. y:B ==> d(y)=d'(y) + "[| z \ A+B; + !!x. x \ A ==> c(x)=c'(x); + !!y. y \ B ==> d(y)=d'(y) |] ==> case(c,d,z) = case(c',d',z)" -by auto +by auto -lemma case_case: "z: A+B ==> - case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = +lemma case_case: "z \ A+B ==> + case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = case(%x. c(c'(x)), %y. d(d'(y)), z)" by auto @@ -170,10 +170,10 @@ lemmas Part_CollectE = Part_Collect [THEN equalityD1, THEN subsetD, THEN CollectE] -lemma Part_Inl: "Part(A+B,Inl) = {Inl(x). x: A}" +lemma Part_Inl: "Part(A+B,Inl) = {Inl(x). x \ A}" by blast -lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y: B}" +lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y \ B}" by blast lemma PartD1: "a \ Part(A,h) ==> a \ A" @@ -182,7 +182,7 @@ lemma Part_id: "Part(A,%x. x) = A" by blast -lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}" +lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y \ Part(B,h)}" by blast lemma Part_sum_equality: "C \ A+B ==> Part(C,Inl) \ Part(C,Inr) = C" diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/Trancl.thy --- a/src/ZF/Trancl.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/Trancl.thy Thu Mar 15 16:35:02 2012 +0000 @@ -54,10 +54,10 @@ subsubsection{*irreflexivity*} lemma irreflI: - "[| !!x. x:A ==> \ r |] ==> irrefl(A,r)" + "[| !!x. x \ A ==> \ r |] ==> irrefl(A,r)" by (simp add: irrefl_def) -lemma irreflE: "[| irrefl(A,r); x:A |] ==> \ r" +lemma irreflE: "[| irrefl(A,r); x \ A |] ==> \ r" by (simp add: irrefl_def) subsubsection{*symmetry*} @@ -84,7 +84,7 @@ by (unfold trans_def, blast) lemma trans_onD: - "[| trans[A](r); :r; :r; a:A; b:A; c:A |] ==> :r" + "[| trans[A](r); :r; :r; a \ A; b \ A; c \ A |] ==> :r" by (unfold trans_on_def, blast) lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)" @@ -122,7 +122,7 @@ done (*Reflexivity of rtrancl*) -lemma rtrancl_refl: "[| a: field(r) |] ==> \ r^*" +lemma rtrancl_refl: "[| a \ field(r) |] ==> \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (erule idI [THEN UnI1]) done @@ -149,13 +149,13 @@ lemma rtrancl_full_induct [case_names initial step, consumes 1]: "[| \ r^*; - !!x. x: field(r) ==> P(); + !!x. x \ field(r) ==> P(); !!x y z.[| P(); : r^*; : r |] ==> P() |] ==> P()" by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast) (*nice induction rule. - Tried adding the typing hypotheses y,z:field(r), but these + Tried adding the typing hypotheses y,z \ field(r), but these caused expensive case splits!*) lemma rtrancl_induct [case_names initial step, induct set: rtrancl]: "[| \ r^*; diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/UNITY/ClientImpl.thy --- a/src/ZF/UNITY/ClientImpl.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/UNITY/ClientImpl.thy Thu Mar 15 16:35:02 2012 +0000 @@ -14,10 +14,10 @@ axiomatization where type_assumes: - "type_of(ask) = list(tokbag) & type_of(giv) = list(tokbag) & + "type_of(ask) = list(tokbag) & type_of(giv) = list(tokbag) & type_of(rel) = list(tokbag) & type_of(tok) = nat" and default_val_assumes: - "default_val(ask) = Nil & default_val(giv) = Nil & + "default_val(ask) = Nil & default_val(giv) = Nil & default_val(rel) = Nil & default_val(tok) = 0" @@ -31,7 +31,7 @@ t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) & nrel < length(s`giv) & nth(nrel, s`ask) \ nth(nrel, s`giv)}" - + (** Choose a new token requirement **) (** Including t=s suppresses fairness, allowing the non-trivial part of the action to be ignored **) @@ -41,7 +41,7 @@ definition "client_ask_act == { \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" - + definition "client_prog == mk_program({s \ state. s`tok \ NbT & s`giv = Nil & @@ -91,8 +91,8 @@ declare client_ask_act_def [THEN def_act_simp, simp] lemma client_prog_ok_iff: - "\G \ program. (client_prog ok G) \ - (G \ preserves(lift(rel)) & G \ preserves(lift(ask)) & + "\G \ program. (client_prog ok G) \ + (G \ preserves(lift(rel)) & G \ preserves(lift(ask)) & G \ preserves(lift(tok)) & client_prog \ Allowed(G))" by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed]) @@ -107,19 +107,19 @@ lemma preserves_lift_imp_stable: "G \ preserves(lift(ff)) ==> G \ stable({s \ state. P(s`ff)})"; apply (drule preserves_imp_stable) -apply (simp add: lift_def) +apply (simp add: lift_def) done lemma preserves_imp_prefix: - "G \ preserves(lift(ff)) + "G \ preserves(lift(ff)) ==> G \ stable({s \ state. \k, s`ff\ \ prefix(nat)})"; -by (erule preserves_lift_imp_stable) +by (erule preserves_lift_imp_stable) -(*Safety property 1: ask, rel are increasing: (24) *) -lemma client_prog_Increasing_ask_rel: +(*Safety property 1 \ ask, rel are increasing: (24) *) +lemma client_prog_Increasing_ask_rel: "client_prog: program guarantees Incr(lift(ask)) \ Incr(lift(rel))" apply (unfold guar_def) -apply (auto intro!: increasing_imp_Increasing +apply (auto intro!: increasing_imp_Increasing simp add: client_prog_ok_iff Increasing.increasing_def preserves_imp_prefix) apply (safety, force, force)+ done @@ -131,17 +131,17 @@ apply (rule Ord_0_lt, auto) done -(*Safety property 2: the client never requests too many tokens. +(*Safety property 2 \ the client never requests too many tokens. With no Substitution Axiom, we must prove the two invariants simultaneously. *) -lemma ask_Bounded_lemma: -"[| client_prog ok G; G \ program |] - ==> client_prog \ G \ - Always({s \ state. s`tok \ NbT} \ +lemma ask_Bounded_lemma: +"[| client_prog ok G; G \ program |] + ==> client_prog \ G \ + Always({s \ state. s`tok \ NbT} \ {s \ state. \elt \ set_of_list(s`ask). elt \ NbT})" apply (rotate_tac -1) apply (auto simp add: client_prog_ok_iff) -apply (rule invariantI [THEN stable_Join_Always2], force) +apply (rule invariantI [THEN stable_Join_Always2], force) prefer 2 apply (fast intro: stable_Int preserves_lift_imp_stable, safety) apply (auto dest: ActsD) @@ -152,8 +152,8 @@ (* Export version, with no mention of tok in the postcondition, but unfortunately tok must be declared local.*) -lemma client_prog_ask_Bounded: - "client_prog \ program guarantees +lemma client_prog_ask_Bounded: + "client_prog \ program guarantees Always({s \ state. \elt \ set_of_list(s`ask). elt \ NbT})" apply (rule guaranteesI) apply (erule ask_Bounded_lemma [THEN Always_weaken], auto) @@ -161,19 +161,19 @@ (*** Towards proving the liveness property ***) -lemma client_prog_stable_rel_le_giv: +lemma client_prog_stable_rel_le_giv: "client_prog \ stable({s \ state. \ prefix(nat)})" by (safety, auto) -lemma client_prog_Join_Stable_rel_le_giv: -"[| client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel)) |] +lemma client_prog_Join_Stable_rel_le_giv: +"[| client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel)) |] ==> client_prog \ G \ Stable({s \ state. \ prefix(nat)})" apply (rule client_prog_stable_rel_le_giv [THEN Increasing_preserves_Stable]) apply (auto simp add: lift_def) done lemma client_prog_Join_Always_rel_le_giv: - "[| client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel)) |] + "[| client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel)) |] ==> client_prog \ G \ Always({s \ state. \ prefix(nat)})" by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv) @@ -184,9 +184,9 @@ lemma act_subset: "A={ \ state*state. P(s, t)} ==> A<=state*state" by auto -lemma transient_lemma: -"client_prog \ - transient({s \ state. s`rel = k & \ strict_prefix(nat) +lemma transient_lemma: +"client_prog \ + transient({s \ state. s`rel = k & \ strict_prefix(nat) & \ prefix(nat) & h pfixGe s`ask})" apply (rule_tac act = client_rel_act in transientI) apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts]) @@ -208,20 +208,20 @@ apply (auto simp add: id_def lam_def) done -lemma strict_prefix_is_prefix: +lemma strict_prefix_is_prefix: " \ strict_prefix(A) \ \ prefix(A) & xs\ys" apply (unfold strict_prefix_def id_def lam_def) apply (auto dest: prefix_type [THEN subsetD]) done -lemma induct_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] - ==> client_prog \ G \ - {s \ state. s`rel = k & \ strict_prefix(nat) - & \ prefix(nat) & h pfixGe s`ask} - LeadsTo {s \ state. \ strict_prefix(nat) - & \ prefix(nat) & - \ prefix(nat) & +lemma induct_lemma: +"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] + ==> client_prog \ G \ + {s \ state. s`rel = k & \ strict_prefix(nat) + & \ prefix(nat) & h pfixGe s`ask} + LeadsTo {s \ state. \ strict_prefix(nat) + & \ prefix(nat) & + \ prefix(nat) & h pfixGe s`ask}" apply (rule single_LeadsTo_I) prefer 2 apply simp @@ -239,68 +239,68 @@ apply (erule client_prog_Join_Stable_rel_le_giv, blast, simp_all) prefer 2 apply (blast intro: sym strict_prefix_is_prefix [THEN iffD2] prefix_trans prefix_imp_pfixGe pfixGe_trans) -apply (auto intro: strict_prefix_is_prefix [THEN iffD1, THEN conjunct1] +apply (auto intro: strict_prefix_is_prefix [THEN iffD1, THEN conjunct1] prefix_trans) done -lemma rel_progress_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] - ==> client_prog \ G \ - {s \ state. \ strict_prefix(nat) - & \ prefix(nat) & h pfixGe s`ask} +lemma rel_progress_lemma: +"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] + ==> client_prog \ G \ + {s \ state. \ strict_prefix(nat) + & \ prefix(nat) & h pfixGe s`ask} LeadsTo {s \ state. \ prefix(nat)}" -apply (rule_tac f = "\x \ state. length(h) #- length(x`rel)" +apply (rule_tac f = "\x \ state. length(h) #- length(x`rel)" in LessThan_induct) apply (auto simp add: vimage_def) - prefer 2 apply (force simp add: lam_def) + prefer 2 apply (force simp add: lam_def) apply (rule single_LeadsTo_I) - prefer 2 apply simp + prefer 2 apply simp apply (subgoal_tac "h \ list(nat)") - prefer 2 apply (blast dest: prefix_type [THEN subsetD]) + prefer 2 apply (blast dest: prefix_type [THEN subsetD]) apply (rule induct_lemma [THEN LeadsTo_weaken]) apply (simp add: length_type lam_def) apply (auto intro: strict_prefix_is_prefix [THEN iffD2] dest: common_prefix_linear prefix_type [THEN subsetD]) apply (erule swap) apply (rule imageI) - apply (force dest!: simp add: lam_def) -apply (simp add: length_type lam_def, clarify) + apply (force dest!: simp add: lam_def) +apply (simp add: length_type lam_def, clarify) apply (drule strict_prefix_length_lt)+ apply (drule less_imp_succ_add, simp)+ -apply clarify -apply simp +apply clarify +apply simp apply (erule diff_le_self [THEN ltD]) done -lemma progress_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] +lemma progress_lemma: +"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] ==> client_prog \ G - \ {s \ state. \ prefix(nat) & h pfixGe s`ask} + \ {s \ state. \ prefix(nat) & h pfixGe s`ask} LeadsTo {s \ state. \ prefix(nat)}" -apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI], +apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI], assumption) apply (force simp add: client_prog_ok_iff) -apply (rule LeadsTo_weaken_L) -apply (rule LeadsTo_Un [OF rel_progress_lemma +apply (rule LeadsTo_weaken_L) +apply (rule LeadsTo_Un [OF rel_progress_lemma subset_refl [THEN subset_imp_LeadsTo]]) apply (auto intro: strict_prefix_is_prefix [THEN iffD2] dest: common_prefix_linear prefix_type [THEN subsetD]) done (*Progress property: all tokens that are given will be released*) -lemma client_prog_progress: -"client_prog \ Incr(lift(giv)) guarantees - (\h \ list(nat). {s \ state. \ prefix(nat) & +lemma client_prog_progress: +"client_prog \ Incr(lift(giv)) guarantees + (\h \ list(nat). {s \ state. \ prefix(nat) & h pfixGe s`ask} LeadsTo {s \ state. \ prefix(nat)})" apply (rule guaranteesI) apply (blast intro: progress_lemma, auto) done lemma client_prog_Allowed: - "Allowed(client_prog) = + "Allowed(client_prog) = preserves(lift(rel)) \ preserves(lift(ask)) \ preserves(lift(tok))" apply (cut_tac v = "lift (ask)" in preserves_type) -apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed] +apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed] cons_Int_distrib safety_prop_Acts_iff) done diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/UNITY/GenPrefix.thy --- a/src/ZF/UNITY/GenPrefix.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/UNITY/GenPrefix.thy Thu Mar 15 16:35:02 2012 +0000 @@ -30,7 +30,7 @@ intros Nil: "<[],[]>:gen_prefix(A, r)" - prepend: "[| :gen_prefix(A, r); :r; x:A; y:A |] + prepend: "[| :gen_prefix(A, r); :r; x \ A; y \ A |] ==> : gen_prefix(A, r)" append: "[| :gen_prefix(A, r); zs:list(A) |] diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/UNITY/SubstAx.thy Thu Mar 15 16:35:02 2012 +0000 @@ -8,17 +8,17 @@ header{*Weak LeadsTo relation (restricted to the set of reachable states)*} theory SubstAx -imports WFair Constrains +imports WFair Constrains begin definition (* The definitions below are not `conventional', but yield simpler rules *) Ensures :: "[i,i] => i" (infixl "Ensures" 60) where - "A Ensures B == {F:program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" + "A Ensures B == {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" definition LeadsTo :: "[i, i] => i" (infixl "LeadsTo" 60) where - "A LeadsTo B == {F:program. F:(reachable(F) \ A) leadsTo (reachable(F) \ B)}" + "A LeadsTo B == {F \ program. F:(reachable(F) \ A) leadsTo (reachable(F) \ B)}" notation (xsymbols) LeadsTo (infixl " \w " 60) @@ -28,7 +28,7 @@ (*Resembles the previous definition of LeadsTo*) (* Equivalence with the HOL-like definition *) -lemma LeadsTo_eq: +lemma LeadsTo_eq: "st_set(B)==> A LeadsTo B = {F \ program. F:(reachable(F) \ A) leadsTo B}" apply (unfold LeadsTo_def) apply (blast dest: psp_stable2 leadsToD2 constrainsD2 intro: leadsTo_weaken) @@ -107,7 +107,7 @@ done (*Lets us look at the starting state*) -lemma single_LeadsTo_I: +lemma single_LeadsTo_I: "[|(!!s. s \ A ==> F:{s} LeadsTo B); F \ program|]==>F \ A LeadsTo B" apply (subst UN_singleton [symmetric], rule LeadsTo_UN, auto) done @@ -117,7 +117,7 @@ apply (blast intro: subset_imp_leadsTo) done -lemma empty_LeadsTo: "F:0 LeadsTo A \ F \ program" +lemma empty_LeadsTo: "F \ 0 LeadsTo A \ F \ program" by (auto dest: LeadsTo_type [THEN subsetD] intro: empty_subsetI [THEN subset_imp_LeadsTo]) declare empty_LeadsTo [iff] @@ -139,8 +139,8 @@ lemma LeadsTo_weaken: "[| F \ A LeadsTo A'; B<=A; A'<=B' |] ==> F \ B LeadsTo B'" by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans) -lemma Always_LeadsTo_weaken: -"[| F \ Always(C); F \ A LeadsTo A'; C \ B \ A; C \ A' \ B' |] +lemma Always_LeadsTo_weaken: +"[| F \ Always(C); F \ A LeadsTo A'; C \ B \ A; C \ A' \ B' |] ==> F \ B LeadsTo B'" apply (blast dest: Always_LeadsToI intro: LeadsTo_weaken Always_LeadsToD) done @@ -151,7 +151,7 @@ by (blast dest: LeadsTo_type [THEN subsetD] intro: LeadsTo_Un subset_imp_LeadsTo) -lemma LeadsTo_Trans_Un: "[| F \ A LeadsTo B; F \ B LeadsTo C |] +lemma LeadsTo_Trans_Un: "[| F \ A LeadsTo B; F \ B LeadsTo C |] ==> F \ (A \ B) LeadsTo C" apply (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans dest: LeadsTo_type [THEN subsetD]) done @@ -175,8 +175,8 @@ apply (blast intro: ensuresI constrains_weaken transient_strengthen dest: constrainsD2) done -lemma Always_LeadsTo_Basis: "[| F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); - F \ transient (I \ (A-A')) |] +lemma Always_LeadsTo_Basis: "[| F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); + F \ transient (I \ (A-A')) |] ==> F \ A LeadsTo A'" apply (rule Always_LeadsToI, assumption) apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) @@ -188,10 +188,10 @@ "[| F \ (A-B) LeadsTo C; F \ (A \ B) LeadsTo C |] ==> F \ A LeadsTo C" by (blast intro: LeadsTo_Un LeadsTo_weaken) -lemma LeadsTo_UN_UN: - "[|(!!i. i \ I ==> F \ A(i) LeadsTo A'(i)); F \ program |] +lemma LeadsTo_UN_UN: + "[|(!!i. i \ I ==> F \ A(i) LeadsTo A'(i)); F \ program |] ==> F \ (\i \ I. A(i)) LeadsTo (\i \ I. A'(i))" -apply (rule LeadsTo_Union, auto) +apply (rule LeadsTo_Union, auto) apply (blast intro: LeadsTo_weaken_R) done @@ -258,7 +258,7 @@ lemma PSP2: "[| F \ A LeadsTo A'; F \ B Co B' |]==> F:(B' \ A) LeadsTo ((B \ A') \ (B' - B))" by (simp (no_asm_simp) add: PSP Int_ac) -lemma PSP_Unless: +lemma PSP_Unless: "[| F \ A LeadsTo A'; F \ B Unless B'|]==> F:(A \ B) LeadsTo ((A' \ B) \ B')" apply (unfold op_Unless_def) apply (drule PSP, assumption) @@ -268,21 +268,21 @@ (*** Induction rules ***) (** Meta or object quantifier ????? **) -lemma LeadsTo_wf_induct: "[| wf(r); - \m \ I. F \ (A \ f-``{m}) LeadsTo - ((A \ f-``(converse(r) `` {m})) \ B); - field(r)<=I; A<=f-``I; F \ program |] +lemma LeadsTo_wf_induct: "[| wf(r); + \m \ I. F \ (A \ f-``{m}) LeadsTo + ((A \ f-``(converse(r) `` {m})) \ B); + field(r)<=I; A<=f-``I; F \ program |] ==> F \ A LeadsTo B" apply (simp (no_asm_use) add: LeadsTo_def) apply auto apply (erule_tac I = I and f = f in leadsTo_wf_induct, safe) apply (drule_tac [2] x = m in bspec, safe) apply (rule_tac [2] A' = "reachable (F) \ (A \ f -`` (converse (r) ``{m}) \ B) " in leadsTo_weaken_R) -apply (auto simp add: Int_assoc) +apply (auto simp add: Int_assoc) done -lemma LessThan_induct: "[| \m \ nat. F:(A \ f-``{m}) LeadsTo ((A \ f-``m) \ B); +lemma LessThan_induct: "[| \m \ nat. F:(A \ f-``{m}) LeadsTo ((A \ f-``m) \ B); A<=f-``nat; F \ program |] ==> F \ A LeadsTo B" apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct]) apply (simp_all add: nat_measure_field) @@ -290,7 +290,7 @@ done -(****** +(****** To be ported ??? I am not sure. integ_0_le_induct @@ -301,51 +301,51 @@ (*** Completion \ Binary and General Finite versions ***) -lemma Completion: "[| F \ A LeadsTo (A' \ C); F \ A' Co (A' \ C); - F \ B LeadsTo (B' \ C); F \ B' Co (B' \ C) |] +lemma Completion: "[| F \ A LeadsTo (A' \ C); F \ A' Co (A' \ C); + F \ B LeadsTo (B' \ C); F \ B' Co (B' \ C) |] ==> F \ (A \ B) LeadsTo ((A' \ B') \ C)" apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains Int_Un_distrib) apply (blast intro: completion leadsTo_weaken) done lemma Finite_completion_aux: - "[| I \ Fin(X);F \ program |] - ==> (\i \ I. F \ (A(i)) LeadsTo (A'(i) \ C)) \ - (\i \ I. F \ (A'(i)) Co (A'(i) \ C)) \ + "[| I \ Fin(X);F \ program |] + ==> (\i \ I. F \ (A(i)) LeadsTo (A'(i) \ C)) \ + (\i \ I. F \ (A'(i)) Co (A'(i) \ C)) \ F \ (\i \ I. A(i)) LeadsTo ((\i \ I. A'(i)) \ C)" apply (erule Fin_induct) apply (auto simp del: INT_simps simp add: Inter_0) -apply (rule Completion, auto) +apply (rule Completion, auto) apply (simp del: INT_simps add: INT_extend_simps) apply (blast intro: Constrains_INT) done -lemma Finite_completion: - "[| I \ Fin(X); !!i. i \ I ==> F \ A(i) LeadsTo (A'(i) \ C); - !!i. i \ I ==> F \ A'(i) Co (A'(i) \ C); - F \ program |] +lemma Finite_completion: + "[| I \ Fin(X); !!i. i \ I ==> F \ A(i) LeadsTo (A'(i) \ C); + !!i. i \ I ==> F \ A'(i) Co (A'(i) \ C); + F \ program |] ==> F \ (\i \ I. A(i)) LeadsTo ((\i \ I. A'(i)) \ C)" by (blast intro: Finite_completion_aux [THEN mp, THEN mp]) -lemma Stable_completion: - "[| F \ A LeadsTo A'; F \ Stable(A'); - F \ B LeadsTo B'; F \ Stable(B') |] +lemma Stable_completion: + "[| F \ A LeadsTo A'; F \ Stable(A'); + F \ B LeadsTo B'; F \ Stable(B') |] ==> F \ (A \ B) LeadsTo (A' \ B')" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Completion [THEN LeadsTo_weaken_R]) prefer 5 - apply blast -apply auto + apply blast +apply auto done -lemma Finite_stable_completion: - "[| I \ Fin(X); - (!!i. i \ I ==> F \ A(i) LeadsTo A'(i)); - (!!i. i \ I ==>F \ Stable(A'(i))); F \ program |] +lemma Finite_stable_completion: + "[| I \ Fin(X); + (!!i. i \ I ==> F \ A(i) LeadsTo A'(i)); + (!!i. i \ I ==>F \ Stable(A'(i))); F \ program |] ==> F \ (\i \ I. A(i)) LeadsTo (\i \ I. A'(i))" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Finite_completion [THEN LeadsTo_weaken_R], simp_all) -apply (rule_tac [3] subset_refl, auto) +apply (rule_tac [3] subset_refl, auto) done ML {* @@ -354,7 +354,7 @@ let val ss = simpset_of ctxt in SELECT_GOAL (EVERY [REPEAT (Always_Int_tac 1), - etac @{thm Always_LeadsTo_Basis} 1 + etac @{thm Always_LeadsTo_Basis} 1 ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, @{thm EnsuresI}, @{thm ensuresI}] 1), @@ -362,7 +362,7 @@ simp_tac (ss addsimps (Program_Defs.get ctxt)) 2, res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2, (*simplify the command's domain*) - simp_tac (ss addsimps [@{thm domain_def}]) 3, + simp_tac (ss addsimps [@{thm domain_def}]) 3, (* proving the domain part *) clarify_tac ctxt 3, dtac @{thm swap} 3, force_tac ctxt 4, rtac @{thm ReplaceI} 3, force_tac ctxt 3, force_tac ctxt 4, diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/UNITY/Union.thy --- a/src/ZF/UNITY/Union.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/UNITY/Union.thy Thu Mar 15 16:35:02 2012 +0000 @@ -4,7 +4,7 @@ Unions of programs -Partly from Misra's Chapter 5: Asynchronous Compositions of Programs +Partly from Misra's Chapter 5 \ Asynchronous Compositions of Programs Theory ported form HOL.. @@ -14,13 +14,13 @@ begin definition - (*FIXME: conjoin Init(F) \ Init(G) \ 0 *) + (*FIXME: conjoin Init(F) \ Init(G) \ 0 *) ok :: "[i, i] => o" (infixl "ok" 65) where "F ok G == Acts(F) \ AllowedActs(G) & Acts(G) \ AllowedActs(F)" definition - (*FIXME: conjoin (\i \ I. Init(F(i))) \ 0 *) + (*FIXME: conjoin (\i \ I. Init(F(i))) \ 0 *) OK :: "[i, i=>i] => o" where "OK(I,F) == (\i \ I. \j \ I-{i}. Acts(F(i)) \ AllowedActs(F(j)))" @@ -39,20 +39,20 @@ safety_prop :: "i => o" where "safety_prop(X) == X\program & SKIP \ X & (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" - + notation (xsymbols) SKIP ("\") and Join (infixl "\" 65) syntax "_JOIN1" :: "[pttrns, i] => i" ("(3JN _./ _)" 10) - "_JOIN" :: "[pttrn, i, i] => i" ("(3JN _:_./ _)" 10) + "_JOIN" :: "[pttrn, i, i] => i" ("(3JN _ \ _./ _)" 10) syntax (xsymbols) "_JOIN1" :: "[pttrns, i] => i" ("(3\ _./ _)" 10) "_JOIN" :: "[pttrn, i, i] => i" ("(3\ _ \ _./ _)" 10) translations - "JN x:A. B" == "CONST JOIN(A, (%x. B))" + "JN x \ A. B" == "CONST JOIN(A, (%x. B))" "JN x y. B" == "JN x. JN y. B" "JN x. B" == "CONST JOIN(CONST state,(%x. B))" @@ -105,7 +105,7 @@ lemma Acts_Join [simp]: "Acts(F Join G) = Acts(F) \ Acts(G)" by (simp add: Int_Un_distrib2 cons_absorb Join_def) -lemma AllowedActs_Join [simp]: "AllowedActs(F Join G) = +lemma AllowedActs_Join [simp]: "AllowedActs(F Join G) = AllowedActs(F) \ AllowedActs(G)" apply (simp add: Int_assoc cons_absorb Join_def) done @@ -164,7 +164,7 @@ "Init(\i \ I. F(i)) = (if I=0 then state else (\i \ I. Init(F(i))))" by (simp add: JOIN_def INT_extend_simps del: INT_simps) -lemma Acts_JN [simp]: +lemma Acts_JN [simp]: "Acts(JOIN(I,F)) = cons(id(state), \i \ I. Acts(F(i)))" apply (unfold JOIN_def) apply (auto simp del: INT_simps UN_simps) @@ -172,8 +172,8 @@ apply (auto dest: Acts_type [THEN subsetD]) done -lemma AllowedActs_JN [simp]: - "AllowedActs(\i \ I. F(i)) = +lemma AllowedActs_JN [simp]: + "AllowedActs(\i \ I. F(i)) = (if I=0 then Pow(state*state) else (\i \ I. AllowedActs(F(i))))" apply (unfold JOIN_def, auto) apply (rule equalityI) @@ -184,7 +184,7 @@ by (rule program_equalityI, auto) lemma JN_cong [cong]: - "[| I=J; !!i. i \ J ==> F(i) = G(i) |] ==> + "[| I=J; !!i. i \ J ==> F(i) = G(i) |] ==> (\i \ I. F(i)) = (\i \ J. G(i))" by (simp add: JOIN_def) @@ -208,7 +208,7 @@ lemma JN_Join_distrib: "(\i \ I. F(i) Join G(i)) = (\i \ I. F(i)) Join (\i \ I. G(i))" apply (rule program_equalityI) -apply (simp_all add: INT_Int_distrib, blast) +apply (simp_all add: INT_Int_distrib, blast) done lemma JN_Join_miniscope: "(\i \ I. F(i) Join G) = ((\i \ I. F(i) Join G))" @@ -227,7 +227,7 @@ alternative precondition is A\B, but most proofs using this rule require I to be nonempty for other reasons anyway.*) -lemma JN_constrains: +lemma JN_constrains: "i \ I==>(\i \ I. F(i)) \ A co B \ (\i \ I. programify(F(i)) \ A co B)" apply (unfold constrains_def JOIN_def st_set_def, auto) @@ -242,7 +242,7 @@ by (auto simp add: constrains_def) lemma Join_unless [iff]: - "(F Join G \ A unless B) \ + "(F Join G \ A unless B) \ (programify(F) \ A unless B & programify(G) \ A unless B)" by (simp add: Join_constrains unless_def) @@ -252,7 +252,7 @@ *) lemma Join_constrains_weaken: - "[| F \ A co A'; G \ B co B' |] + "[| F \ A co A'; G \ B co B' |] ==> F Join G \ (A \ B) co (A' \ B')" apply (subgoal_tac "st_set (A) & st_set (B) & F \ program & G \ program") prefer 2 apply (blast dest: constrainsD2, simp) @@ -280,17 +280,17 @@ apply (drule_tac x = act in bspec, auto) done -lemma initially_JN_I: +lemma initially_JN_I: assumes major: "(!!i. i \ I ==>F(i) \ initially(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ initially(A)" apply (cut_tac minor) -apply (auto elim!: not_emptyE simp add: Inter_iff initially_def) +apply (auto elim!: not_emptyE simp add: Inter_iff initially_def) apply (frule_tac i = x in major) -apply (auto simp add: initially_def) +apply (auto simp add: initially_def) done -lemma invariant_JN_I: +lemma invariant_JN_I: assumes major: "(!!i. i \ I ==> F(i) \ invariant(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ invariant(A)" @@ -304,7 +304,7 @@ done lemma Join_stable [iff]: - " (F Join G \ stable(A)) \ + " (F Join G \ stable(A)) \ (programify(F) \ stable(A) & programify(G) \ stable(A))" by (simp add: stable_def) @@ -313,7 +313,7 @@ by (unfold initially_def, auto) lemma invariant_JoinI: - "[| F \ invariant(A); G \ invariant(A) |] + "[| F \ invariant(A); G \ invariant(A) |] ==> F Join G \ invariant(A)" apply (subgoal_tac "F \ program&G \ program") prefer 2 apply (blast dest: invariantD2) @@ -329,7 +329,7 @@ subsection{*Progress: transient, ensures*} lemma JN_transient: - "i \ I ==> + "i \ I ==> (\i \ I. F(i)) \ transient(A) \ (\i \ I. programify(F(i)) \ transient(A))" apply (auto simp add: transient_def JOIN_def) apply (unfold st_set_def) @@ -338,7 +338,7 @@ done lemma Join_transient [iff]: - "F Join G \ transient(A) \ + "F Join G \ transient(A) \ (programify(F) \ transient(A) | programify(G) \ transient(A))" apply (auto simp add: transient_def Join_def Int_Un_distrib2) done @@ -352,28 +352,28 @@ (*If I=0 it degenerates to (SKIP \ A ensures B) = False, i.e. to ~(A\B) *) lemma JN_ensures: - "i \ I ==> - (\i \ I. F(i)) \ A ensures B \ - ((\i \ I. programify(F(i)) \ (A-B) co (A \ B)) & + "i \ I ==> + (\i \ I. F(i)) \ A ensures B \ + ((\i \ I. programify(F(i)) \ (A-B) co (A \ B)) & (\i \ I. programify(F(i)) \ A ensures B))" by (auto simp add: ensures_def JN_constrains JN_transient) -lemma Join_ensures: - "F Join G \ A ensures B \ - (programify(F) \ (A-B) co (A \ B) & programify(G) \ (A-B) co (A \ B) & +lemma Join_ensures: + "F Join G \ A ensures B \ + (programify(F) \ (A-B) co (A \ B) & programify(G) \ (A-B) co (A \ B) & (programify(F) \ transient (A-B) | programify(G) \ transient (A-B)))" apply (unfold ensures_def) apply (auto simp add: Join_transient) done -lemma stable_Join_constrains: - "[| F \ stable(A); G \ A co A' |] +lemma stable_Join_constrains: + "[| F \ stable(A); G \ A co A' |] ==> F Join G \ A co A'" apply (unfold stable_def constrains_def Join_def st_set_def) apply (cut_tac F = F in Acts_type) -apply (cut_tac F = G in Acts_type, force) +apply (cut_tac F = G in Acts_type, force) done (*Premise for G cannot use Always because F \ Stable A is @@ -462,10 +462,10 @@ by (simp add: OK_def) lemma OK_cons_iff: - "OK(cons(i, I), F) \ + "OK(cons(i, I), F) \ (i \ I & OK(I, F)) | (i\I & OK(I, F) & F(i) ok JOIN(I,F))" apply (simp add: OK_iff_ok) -apply (blast intro: ok_sym) +apply (blast intro: ok_sym) done @@ -475,25 +475,25 @@ by (auto dest: Acts_type [THEN subsetD] simp add: Allowed_def) lemma Allowed_Join [simp]: - "Allowed(F Join G) = + "Allowed(F Join G) = Allowed(programify(F)) \ Allowed(programify(G))" apply (auto simp add: Allowed_def) done lemma Allowed_JN [simp]: - "i \ I ==> + "i \ I ==> Allowed(JOIN(I,F)) = (\i \ I. Allowed(programify(F(i))))" apply (auto simp add: Allowed_def, blast) done lemma ok_iff_Allowed: - "F ok G \ (programify(F) \ Allowed(programify(G)) & + "F ok G \ (programify(F) \ Allowed(programify(G)) & programify(G) \ Allowed(programify(F)))" by (simp add: ok_def Allowed_def) lemma OK_iff_Allowed: - "OK(I,F) \ + "OK(I,F) \ (\i \ I. \j \ I-{i}. programify(F(i)) \ Allowed(programify(F(j))))" apply (auto simp add: OK_iff_ok ok_iff_Allowed) done @@ -510,10 +510,10 @@ done lemma safety_prop_AllowedActs_iff_Allowed: - "safety_prop(X) ==> + "safety_prop(X) ==> (\G \ X. Acts(G)) \ AllowedActs(F) \ (X \ Allowed(programify(F)))" -apply (simp add: Allowed_def safety_prop_Acts_iff [THEN iff_sym] - safety_prop_def, blast) +apply (simp add: Allowed_def safety_prop_Acts_iff [THEN iff_sym] + safety_prop_def, blast) done @@ -526,7 +526,7 @@ done lemma def_prg_Allowed: - "[| F == mk_program (init, acts, \F \ X. Acts(F)); safety_prop(X) |] + "[| F == mk_program (init, acts, \F \ X. Acts(F)); safety_prop(X) |] ==> Allowed(F) = X" by (simp add: Allowed_eq) @@ -571,8 +571,8 @@ apply blast+ done -lemma def_UNION_ok_iff: -"[| F == mk_program(init,acts, \G \ X. Acts(G)); safety_prop(X) |] +lemma def_UNION_ok_iff: +"[| F == mk_program(init,acts, \G \ X. Acts(G)); safety_prop(X) |] ==> F ok G \ (programify(G) \ X & acts \ Pow(state*state) \ AllowedActs(G))" apply (unfold ok_def) apply (drule_tac G = G in safety_prop_Acts_iff) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/UNITY/WFair.thy --- a/src/ZF/UNITY/WFair.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/UNITY/WFair.thy Thu Mar 15 16:35:02 2012 +0000 @@ -15,41 +15,41 @@ definition (* This definition specifies weak fairness. The rest of the theory - is generic to all forms of fairness.*) + is generic to all forms of fairness.*) transient :: "i=>i" where - "transient(A) =={F:program. (\act\Acts(F). A<=domain(act) & + "transient(A) =={F \ program. (\act\Acts(F). A<=domain(act) & act``A \ state-A) & st_set(A)}" definition ensures :: "[i,i] => i" (infixl "ensures" 60) where "A ensures B == ((A-B) co (A \ B)) \ transient(A-B)" - + consts (*LEADS-TO constant for the inductive definition*) leads :: "[i, i]=>i" -inductive +inductive domains "leads(D, F)" \ "Pow(D)*Pow(D)" - intros - Basis: "[| F:A ensures B; A:Pow(D); B:Pow(D) |] ==> :leads(D, F)" + intros + Basis: "[| F \ A ensures B; A \ Pow(D); B \ Pow(D) |] ==> :leads(D, F)" Trans: "[| \ leads(D, F); \ leads(D, F) |] ==> :leads(D, F)" - Union: "[| S:Pow({A:S. :leads(D, F)}); B:Pow(D); S:Pow(Pow(D)) |] ==> + Union: "[| S \ Pow({A \ S. :leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D)) |] ==> <\(S),B>:leads(D, F)" monos Pow_mono type_intros Union_Pow_iff [THEN iffD2] UnionI PowI - + definition (* The Visible version of the LEADS-TO relation*) leadsTo :: "[i, i] => i" (infixl "leadsTo" 60) where - "A leadsTo B == {F:program. :leads(state, F)}" - + "A leadsTo B == {F \ program. :leads(state, F)}" + definition (* wlt(F, B) is the largest set that leads to B*) wlt :: "[i, i] => i" where - "wlt(F, B) == \({A:Pow(state). F: A leadsTo B})" + "wlt(F, B) == \({A \ Pow(state). F \ A leadsTo B})" notation (xsymbols) leadsTo (infixl "\" 60) @@ -67,7 +67,7 @@ lemma transient_type: "transient(A)<=program" by (unfold transient_def, auto) -lemma transientD2: +lemma transientD2: "F \ transient(A) ==> F \ program & st_set(A)" apply (unfold transient_def, auto) done @@ -80,20 +80,20 @@ apply (blast intro!: rev_bexI) done -lemma transientI: -"[|act \ Acts(F); A \ domain(act); act``A \ state-A; +lemma transientI: +"[|act \ Acts(F); A \ domain(act); act``A \ state-A; F \ program; st_set(A)|] ==> F \ transient(A)" by (simp add: transient_def, blast) -lemma transientE: - "[| F \ transient(A); +lemma transientE: + "[| F \ transient(A); !!act. [| act \ Acts(F); A \ domain(act); act``A \ state-A|]==>P|] ==>P" by (simp add: transient_def, blast) lemma transient_state: "transient(state) = 0" apply (simp add: transient_def) -apply (rule equalityI, auto) +apply (rule equalityI, auto) apply (cut_tac F = x in Acts_type) apply (simp add: Diff_cancel) apply (auto intro: st0_in_state) @@ -117,7 +117,7 @@ lemma ensures_type: "A ensures B <=program" by (simp add: ensures_def constrains_def, auto) -lemma ensuresI: +lemma ensuresI: "[|F:(A-B) co (A \ B); F \ transient(A-B)|]==>F \ A ensures B" apply (unfold ensures_def) apply (auto simp add: transient_type [THEN subsetD]) @@ -138,10 +138,10 @@ apply (blast intro: transient_strengthen constrains_weaken) done -(*The L-version (precondition strengthening) fails, but we have this*) -lemma stable_ensures_Int: +(*The L-version (precondition strengthening) fails, but we have this*) +lemma stable_ensures_Int: "[| F \ stable(C); F \ A ensures B |] ==> F:(C \ A) ensures (C \ B)" - + apply (unfold ensures_def) apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric]) apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken) @@ -166,13 +166,13 @@ lemma leadsTo_type: "A leadsTo B \ program" by (unfold leadsTo_def, auto) -lemma leadsToD2: +lemma leadsToD2: "F \ A leadsTo B ==> F \ program & st_set(A) & st_set(B)" apply (unfold leadsTo_def st_set_def) apply (blast dest: leads_left leads_right) done -lemma leadsTo_Basis: +lemma leadsTo_Basis: "[|F \ A ensures B; st_set(A); st_set(B)|] ==> F \ A leadsTo B" apply (unfold leadsTo_def st_set_def) apply (cut_tac ensures_type) @@ -204,22 +204,22 @@ by (simp add: Un_ac) (*The Union introduction rule as we should have liked to state it*) -lemma leadsTo_Union: +lemma leadsTo_Union: "[|!!A. A \ S ==> F \ A leadsTo B; F \ program; st_set(B)|] ==> F \ \(S) leadsTo B" apply (unfold leadsTo_def st_set_def) apply (blast intro: leads.Union dest: leads_left) done -lemma leadsTo_Union_Int: - "[|!!A. A \ S ==>F \ (A \ C) leadsTo B; F \ program; st_set(B)|] +lemma leadsTo_Union_Int: + "[|!!A. A \ S ==>F \ (A \ C) leadsTo B; F \ program; st_set(B)|] ==> F \ (\(S)Int C)leadsTo B" apply (unfold leadsTo_def st_set_def) apply (simp only: Int_Union_Union) apply (blast dest: leads_left intro: leads.Union) done -lemma leadsTo_UN: +lemma leadsTo_UN: "[| !!i. i \ I ==> F \ A(i) leadsTo B; F \ program; st_set(B)|] ==> F:(\i \ I. A(i)) leadsTo B" apply (simp add: Int_Union_Union leadsTo_def st_set_def) @@ -234,10 +234,10 @@ done lemma single_leadsTo_I: - "[|!!x. x \ A==> F:{x} leadsTo B; F \ program; st_set(B) |] + "[|!!x. x \ A==> F:{x} leadsTo B; F \ program; st_set(B) |] ==> F \ A leadsTo B" apply (rule_tac b = A in UN_singleton [THEN subst]) -apply (rule leadsTo_UN, auto) +apply (rule leadsTo_UN, auto) done lemma leadsTo_refl: "[| F \ program; st_set(A) |] ==> F \ A leadsTo A" @@ -278,7 +278,7 @@ lemma leadsTo_Un_distrib: "F:(A \ B) leadsTo C \ (F \ A leadsTo C & F \ B leadsTo C)" by (blast intro: leadsTo_Un leadsTo_weaken_L) -lemma leadsTo_UN_distrib: +lemma leadsTo_UN_distrib: "(F:(\i \ I. A(i)) leadsTo B)\ ((\i \ I. F \ A(i) leadsTo B) & F \ program & st_set(B))" apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L) done @@ -293,10 +293,10 @@ by (blast intro: leadsTo_Un leadsTo_weaken dest: leadsToD2) lemma leadsTo_UN_UN: - "[|!!i. i \ I ==> F \ A(i) leadsTo A'(i); F \ program |] + "[|!!i. i \ I ==> F \ A(i) leadsTo A'(i); F \ program |] ==> F: (\i \ I. A(i)) leadsTo (\i \ I. A'(i))" apply (rule leadsTo_Union) -apply (auto intro: leadsTo_weaken_R dest: leadsToD2) +apply (auto intro: leadsTo_weaken_R dest: leadsToD2) done (*Binary union version*) @@ -336,17 +336,17 @@ lemma leadsTo_induct: assumes major: "F \ za leadsTo zb" and basis: "!!A B. [|F \ A ensures B; st_set(A); st_set(B)|] ==> P(A,B)" - and trans: "!!A B C. [| F \ A leadsTo B; P(A, B); + and trans: "!!A B C. [| F \ A leadsTo B; P(A, B); F \ B leadsTo C; P(B, C) |] ==> P(A,C)" - and union: "!!B S. [| \A \ S. F \ A leadsTo B; \A \ S. P(A,B); + and union: "!!B S. [| \A \ S. F \ A leadsTo B; \A \ S. P(A,B); st_set(B); \A \ S. st_set(A)|] ==> P(\(S), B)" shows "P(za, zb)" apply (cut_tac major) -apply (unfold leadsTo_def, clarify) -apply (erule leads.induct) +apply (unfold leadsTo_def, clarify) +apply (erule leads.induct) apply (blast intro: basis [unfolded st_set_def]) - apply (blast intro: trans [unfolded leadsTo_def]) -apply (force intro: union [unfolded st_set_def leadsTo_def]) + apply (blast intro: trans [unfolded leadsTo_def]) +apply (force intro: union [unfolded st_set_def leadsTo_def]) done @@ -354,11 +354,11 @@ lemma leadsTo_induct2: assumes major: "F \ za leadsTo zb" and basis1: "!!A B. [| A<=B; st_set(B) |] ==> P(A, B)" - and basis2: "!!A B. [| F \ A co A \ B; F \ transient(A); st_set(B) |] + and basis2: "!!A B. [| F \ A co A \ B; F \ transient(A); st_set(B) |] ==> P(A, B)" - and trans: "!!A B C. [| F \ A leadsTo B; P(A, B); + and trans: "!!A B C. [| F \ A leadsTo B; P(A, B); F \ B leadsTo C; P(B, C) |] ==> P(A,C)" - and union: "!!B S. [| \A \ S. F \ A leadsTo B; \A \ S. P(A,B); + and union: "!!B S. [| \A \ S. F \ A leadsTo B; \A \ S. P(A,B); st_set(B); \A \ S. st_set(A)|] ==> P(\(S), B)" shows "P(za, zb)" apply (cut_tac major) @@ -381,11 +381,11 @@ (** Variant induction rule: on the preconditions for B **) (*Lemma is the weak version: can't see how to do it in one step*) -lemma leadsTo_induct_pre_aux: - "[| F \ za leadsTo zb; - P(zb); - !!A B. [| F \ A ensures B; P(B); st_set(A); st_set(B) |] ==> P(A); - !!S. [| \A \ S. P(A); \A \ S. st_set(A) |] ==> P(\(S)) +lemma leadsTo_induct_pre_aux: + "[| F \ za leadsTo zb; + P(zb); + !!A B. [| F \ A ensures B; P(B); st_set(A); st_set(B) |] ==> P(A); + !!S. [| \A \ S. P(A); \A \ S. st_set(A) |] ==> P(\(S)) |] ==> P(za)" txt{*by induction on this formula*} apply (subgoal_tac "P (zb) \ P (za) ") @@ -396,15 +396,15 @@ done -lemma leadsTo_induct_pre: - "[| F \ za leadsTo zb; - P(zb); - !!A B. [| F \ A ensures B; F \ B leadsTo zb; P(B); st_set(A) |] ==> P(A); - !!S. \A \ S. F \ A leadsTo zb & P(A) & st_set(A) ==> P(\(S)) +lemma leadsTo_induct_pre: + "[| F \ za leadsTo zb; + P(zb); + !!A B. [| F \ A ensures B; F \ B leadsTo zb; P(B); st_set(A) |] ==> P(A); + !!S. \A \ S. F \ A leadsTo zb & P(A) & st_set(A) ==> P(\(S)) |] ==> P(za)" apply (subgoal_tac " (F \ za leadsTo zb) & P (za) ") apply (erule conjunct2) -apply (frule leadsToD2) +apply (frule leadsToD2) apply (erule leadsTo_induct_pre_aux) prefer 3 apply (blast dest: leadsToD2 intro: leadsTo_Union) prefer 2 apply (blast intro: leadsTo_Trans leadsTo_Basis) @@ -412,7 +412,7 @@ done (** The impossibility law **) -lemma leadsTo_empty: +lemma leadsTo_empty: "F \ A leadsTo 0 ==> A=0" apply (erule leadsTo_induct_pre) apply (auto simp add: ensures_def constrains_def transient_def st_set_def) @@ -425,10 +425,10 @@ text{*Special case of PSP: Misra's "stable conjunction"*} -lemma psp_stable: +lemma psp_stable: "[| F \ A leadsTo A'; F \ stable(B) |] ==> F:(A \ B) leadsTo (A' \ B)" apply (unfold stable_def) -apply (frule leadsToD2) +apply (frule leadsToD2) apply (erule leadsTo_induct) prefer 3 apply (blast intro: leadsTo_Union_Int) prefer 2 apply (blast intro: leadsTo_Trans) @@ -442,7 +442,7 @@ apply (simp (no_asm_simp) add: psp_stable Int_ac) done -lemma psp_ensures: +lemma psp_ensures: "[| F \ A ensures A'; F \ B co B' |]==> F: (A \ B') ensures ((A' \ B) \ (B' - B))" apply (unfold ensures_def constrains_def st_set_def) (*speeds up the proof*) @@ -450,7 +450,7 @@ apply (blast intro: transient_strengthen) done -lemma psp: +lemma psp: "[|F \ A leadsTo A'; F \ B co B'; st_set(B')|]==> F:(A \ B') leadsTo ((A' \ B) \ (B' - B))" apply (subgoal_tac "F \ program & st_set (A) & st_set (A') & st_set (B) ") prefer 2 apply (blast dest!: constrainsD2 leadsToD2) @@ -466,12 +466,12 @@ done -lemma psp2: "[| F \ A leadsTo A'; F \ B co B'; st_set(B') |] +lemma psp2: "[| F \ A leadsTo A'; F \ B co B'; st_set(B') |] ==> F \ (B' \ A) leadsTo ((B \ A') \ (B' - B))" by (simp (no_asm_simp) add: psp Int_ac) -lemma psp_unless: - "[| F \ A leadsTo A'; F \ B unless B'; st_set(B); st_set(B') |] +lemma psp_unless: + "[| F \ A leadsTo A'; F \ B unless B'; st_set(B); st_set(B') |] ==> F \ (A \ B) leadsTo ((A' \ B) \ B')" apply (unfold unless_def) apply (subgoal_tac "st_set (A) &st_set (A') ") @@ -484,12 +484,12 @@ subsection{*Proving the induction rules*} (** The most general rule \ r is any wf relation; f is any variant function **) -lemma leadsTo_wf_induct_aux: "[| wf(r); - m \ I; - field(r)<=I; - F \ program; st_set(B); - \m \ I. F \ (A \ f-``{m}) leadsTo - ((A \ f-``(converse(r)``{m})) \ B) |] +lemma leadsTo_wf_induct_aux: "[| wf(r); + m \ I; + field(r)<=I; + F \ program; st_set(B); + \m \ I. F \ (A \ f-``{m}) leadsTo + ((A \ f-``(converse(r)``{m})) \ B) |] ==> F \ (A \ f-``{m}) leadsTo B" apply (erule_tac a = m in wf_induct2, simp_all) apply (subgoal_tac "F \ (A \ (f-`` (converse (r) ``{x}))) leadsTo B") @@ -500,17 +500,17 @@ done (** Meta or object quantifier ? **) -lemma leadsTo_wf_induct: "[| wf(r); - field(r)<=I; - A<=f-``I; - F \ program; st_set(A); st_set(B); - \m \ I. F \ (A \ f-``{m}) leadsTo - ((A \ f-``(converse(r)``{m})) \ B) |] +lemma leadsTo_wf_induct: "[| wf(r); + field(r)<=I; + A<=f-``I; + F \ program; st_set(A); st_set(B); + \m \ I. F \ (A \ f-``{m}) leadsTo + ((A \ f-``(converse(r)``{m})) \ B) |] ==> F \ A leadsTo B" apply (rule_tac b = A in subst) defer 1 apply (rule_tac I = I in leadsTo_UN) - apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) + apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) done lemma nat_measure_field: "field(measure(nat, %x. x)) = nat" @@ -536,12 +536,12 @@ done (*Alternative proof is via the lemma F \ (A \ f-`(lessThan m)) leadsTo B*) -lemma lessThan_induct: - "[| A<=f-``nat; - F \ program; st_set(A); st_set(B); - \m \ nat. F:(A \ f-``{m}) leadsTo ((A \ f -`` m) \ B) |] +lemma lessThan_induct: + "[| A<=f-``nat; + F \ program; st_set(A); st_set(B); + \m \ nat. F:(A \ f-``{m}) leadsTo ((A \ f -`` m) \ B) |] ==> F \ A leadsTo B" -apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) +apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) apply (simp_all add: nat_measure_field) apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric]) done @@ -586,17 +586,17 @@ done (*Used in the Trans case below*) -lemma leadsTo_123_aux: - "[| B \ A2; - F \ (A1 - B) co (A1 \ B); - F \ (A2 - C) co (A2 \ C) |] +lemma leadsTo_123_aux: + "[| B \ A2; + F \ (A1 - B) co (A1 \ B); + F \ (A2 - C) co (A2 \ C) |] ==> F \ (A1 \ A2 - C) co (A1 \ A2 \ C)" apply (unfold constrains_def st_set_def, blast) done (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*) (* slightly different from the HOL one \ B here is bounded *) -lemma leadsTo_123: "F \ A leadsTo A' +lemma leadsTo_123: "F \ A leadsTo A' ==> \B \ Pow(state). A<=B & F \ B leadsTo A' & F \ (B-A') co (B \ A')" apply (frule leadsToD2) apply (erule leadsTo_induct) @@ -612,7 +612,7 @@ defer 1 apply (rule AC_ball_Pi, safe) apply (rotate_tac 1) -apply (drule_tac x = x in bspec, blast, blast) +apply (drule_tac x = x in bspec, blast, blast) apply (rule_tac x = "\A \ S. y`A" in bexI, safe) apply (rule_tac [3] I1 = S in constrains_UN [THEN constrains_weaken]) apply (rule_tac [2] leadsTo_Union) @@ -633,13 +633,13 @@ subsection{*Completion: Binary and General Finite versions*} -lemma completion_aux: "[| W = wlt(F, (B' \ C)); - F \ A leadsTo (A' \ C); F \ A' co (A' \ C); - F \ B leadsTo (B' \ C); F \ B' co (B' \ C) |] +lemma completion_aux: "[| W = wlt(F, (B' \ C)); + F \ A leadsTo (A' \ C); F \ A' co (A' \ C); + F \ B leadsTo (B' \ C); F \ B' co (B' \ C) |] ==> F \ (A \ B) leadsTo ((A' \ B') \ C)" apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \ program") - prefer 2 - apply simp + prefer 2 + apply simp apply (blast dest!: leadsToD2) apply (subgoal_tac "F \ (W-C) co (W \ B' \ C) ") prefer 2 @@ -668,9 +668,9 @@ lemmas completion = refl [THEN completion_aux] lemma finite_completion_aux: - "[| I \ Fin(X); F \ program; st_set(C) |] ==> - (\i \ I. F \ (A(i)) leadsTo (A'(i) \ C)) \ - (\i \ I. F \ (A'(i)) co (A'(i) \ C)) \ + "[| I \ Fin(X); F \ program; st_set(C) |] ==> + (\i \ I. F \ (A(i)) leadsTo (A'(i) \ C)) \ + (\i \ I. F \ (A'(i)) co (A'(i) \ C)) \ F \ (\i \ I. A(i)) leadsTo ((\i \ I. A'(i)) \ C)" apply (erule Fin_induct) apply (auto simp add: Inter_0) @@ -679,16 +679,16 @@ apply (blast intro: constrains_INT) done -lemma finite_completion: - "[| I \ Fin(X); - !!i. i \ I ==> F \ A(i) leadsTo (A'(i) \ C); - !!i. i \ I ==> F \ A'(i) co (A'(i) \ C); F \ program; st_set(C)|] +lemma finite_completion: + "[| I \ Fin(X); + !!i. i \ I ==> F \ A(i) leadsTo (A'(i) \ C); + !!i. i \ I ==> F \ A'(i) co (A'(i) \ C); F \ program; st_set(C)|] ==> F \ (\i \ I. A(i)) leadsTo ((\i \ I. A'(i)) \ C)" by (blast intro: finite_completion_aux [THEN mp, THEN mp]) -lemma stable_completion: - "[| F \ A leadsTo A'; F \ stable(A'); - F \ B leadsTo B'; F \ stable(B') |] +lemma stable_completion: + "[| F \ A leadsTo A'; F \ stable(A'); + F \ B leadsTo B'; F \ stable(B') |] ==> F \ (A \ B) leadsTo (A' \ B')" apply (unfold stable_def) apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+) @@ -696,15 +696,15 @@ done -lemma finite_stable_completion: - "[| I \ Fin(X); - (!!i. i \ I ==> F \ A(i) leadsTo A'(i)); - (!!i. i \ I ==> F \ stable(A'(i))); F \ program |] +lemma finite_stable_completion: + "[| I \ Fin(X); + (!!i. i \ I ==> F \ A(i) leadsTo A'(i)); + (!!i. i \ I ==> F \ stable(A'(i))); F \ program |] ==> F \ (\i \ I. A(i)) leadsTo (\i \ I. A'(i))" apply (unfold stable_def) apply (subgoal_tac "st_set (\i \ I. A' (i))") prefer 2 apply (blast dest: leadsToD2) -apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) +apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) done end diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/WF.thy --- a/src/ZF/WF.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/WF.thy Thu Mar 15 16:35:02 2012 +0000 @@ -21,7 +21,7 @@ definition wf :: "i=>o" where (*r is a well-founded relation*) - "wf(r) == \Z. Z=0 | (\x\Z. \y. :r \ ~ y:Z)" + "wf(r) == \Z. Z=0 | (\x\Z. \y. :r \ ~ y \ Z)" definition wf_on :: "[i,i]=>o" ("wf[_]'(_')") where @@ -80,7 +80,7 @@ text{*If every non-empty subset of @{term A} has an @{term r}-minimal element then we have @{term "wf[A](r)"}.*} lemma wf_onI: - assumes prem: "!!Z u. [| Z<=A; u:Z; \x\Z. \y\Z. :r |] ==> False" + assumes prem: "!!Z u. [| Z<=A; u \ Z; \x\Z. \y\Z. :r |] ==> False" shows "wf[A](r)" apply (unfold wf_on_def wf_def) apply (rule equals0I [THEN disjCI, THEN allI]) @@ -89,10 +89,10 @@ text{*If @{term r} allows well-founded induction over @{term A} then we have @{term "wf[A](r)"}. Premise is equivalent to - @{prop "!!B. \x\A. (\y. : r \ y:B) \ x:B ==> A<=B"} *} + @{prop "!!B. \x\A. (\y. : r \ y \ B) \ x \ B ==> A<=B"} *} lemma wf_onI2: - assumes prem: "!!y B. [| \x\A. (\y\A. :r \ y:B) \ x:B; y:A |] - ==> y:B" + assumes prem: "!!y B. [| \x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A |] + ==> y \ B" shows "wf[A](r)" apply (rule wf_onI) apply (rule_tac c=u in prem [THEN DiffE]) @@ -118,10 +118,10 @@ text{*The form of this rule is designed to match @{text wfI}*} lemma wf_induct2: - "[| wf(r); a:A; field(r)<=A; - !!x.[| x: A; \y. : r \ P(y) |] ==> P(x) |] + "[| wf(r); a \ A; field(r)<=A; + !!x.[| x \ A; \y. : r \ P(y) |] ==> P(x) |] ==> P(a)" -apply (erule_tac P="a:A" in rev_mp) +apply (erule_tac P="a \ A" in rev_mp) apply (erule_tac a=a in wf_induct, blast) done @@ -129,8 +129,8 @@ by blast lemma wf_on_induct [consumes 2, induct set: wf_on]: - "[| wf[A](r); a:A; - !!x.[| x: A; \y\A. : r \ P(y) |] ==> P(x) + "[| wf[A](r); a \ A; + !!x.[| x \ A; \y\A. : r \ P(y) |] ==> P(x) |] ==> P(a)" apply (unfold wf_on_def) apply (erule wf_induct2, assumption) @@ -145,8 +145,8 @@ then we have @{term "wf(r)"}.*} lemma wfI: "[| field(r)<=A; - !!y B. [| \x\A. (\y\A. :r \ y:B) \ x:B; y:A|] - ==> y:B |] + !!y B. [| \x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A|] + ==> y \ B |] ==> wf(r)" apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) apply (rule wf_onI2) @@ -166,11 +166,11 @@ (* @{term"[| wf(r); \ r; ~P ==> \ r |] ==> P"} *) lemmas wf_asym = wf_not_sym [THEN swap] -lemma wf_on_not_refl: "[| wf[A](r); a: A |] ==> \ r" +lemma wf_on_not_refl: "[| wf[A](r); a \ A |] ==> \ r" by (erule_tac a=a in wf_on_induct, assumption, blast) lemma wf_on_not_sym [rule_format]: - "[| wf[A](r); a:A |] ==> \b\A. :r \ \r" + "[| wf[A](r); a \ A |] ==> \b\A. :r \ \r" apply (erule_tac a=a in wf_on_induct, assumption, blast) done @@ -183,7 +183,7 @@ (*Needed to prove well_ordI. Could also reason that wf[A](r) means wf(r \ A*A); thus wf( (r \ A*A)^+ ) and use wf_not_refl *) lemma wf_on_chain3: - "[| wf[A](r); :r; :r; :r; a:A; b:A; c:A |] ==> P" + "[| wf[A](r); :r; :r; :r; a \ A; b \ A; c \ A |] ==> P" apply (subgoal_tac "\y\A. \z\A. :r \ :r \ :r \ P", blast) apply (erule_tac a=a in wf_on_induct, assumption, blast) @@ -218,7 +218,7 @@ subsection{*The Predicate @{term is_recfun}*} -lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f: r-``{a} -> range(f)" +lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f \ r-``{a} -> range(f)" apply (unfold is_recfun_def) apply (erule ssubst) apply (rule lamI [THEN rangeI, THEN lam_type], assumption) @@ -345,8 +345,8 @@ done lemma wfrec_type: - "[| wf(r); a:A; field(r)<=A; - !!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) \ B(x) + "[| wf(r); a \ A; field(r)<=A; + !!x u. [| x \ A; u \ Pi(r-``{x}, B) |] ==> H(x,u) \ B(x) |] ==> wfrec(r,a,H) \ B(a)" apply (rule_tac a = a in wf_induct2, assumption+) apply (subst wfrec, assumption) @@ -355,7 +355,7 @@ lemma wfrec_on: - "[| wf[A](r); a: A |] ==> + "[| wf[A](r); a \ A |] ==> wfrec[A](r,a,H) = H(a, \x\(r-``{a}) \ A. wfrec[A](r,x,H))" apply (unfold wf_on_def wfrec_on_def) apply (erule wfrec [THEN trans]) @@ -364,7 +364,7 @@ text{*Minimal-element characterization of well-foundedness*} lemma wf_eq_minimal: - "wf(r) \ (\Q x. x:Q \ (\z\Q. \y. :r \ y\Q))" + "wf(r) \ (\Q x. x \ Q \ (\z\Q. \y. :r \ y\Q))" by (unfold wf_def, blast) end diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/ex/Group.thy --- a/src/ZF/ex/Group.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/ex/Group.thy Thu Mar 15 16:35:02 2012 +0000 @@ -11,7 +11,7 @@ subsection {* Monoids *} (*First, we must simulate a record declaration: -record monoid = +record monoid = carrier :: i mult :: "[i,i] => i" (infixl "\\" 70) one :: i ("\\") @@ -41,7 +41,7 @@ assumes m_closed [intro, simp]: "\x \ carrier(G); y \ carrier(G)\ \ x \ y \ carrier(G)" and m_assoc: - "\x \ carrier(G); y \ carrier(G); z \ carrier(G)\ + "\x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ (x \ y) \ z = x \ (y \ z)" and one_closed [intro, simp]: "\ \ carrier(G)" and l_one [simp]: "x \ carrier(G) \ \ \ x = x" @@ -61,13 +61,13 @@ by (simp add: update_carrier_def) lemma carrier_update_carrier [simp]: "carrier(update_carrier(M,B)) = B" - by (simp add: update_carrier_def) + by (simp add: update_carrier_def) lemma mult_update_carrier [simp]: "mmult(update_carrier(M,B),x,y) = mmult(M,x,y)" - by (simp add: update_carrier_def mmult_def) + by (simp add: update_carrier_def mmult_def) lemma one_update_carrier [simp]: "one(update_carrier(M,B)) = one(M)" - by (simp add: update_carrier_def one_def) + by (simp add: update_carrier_def one_def) lemma (in monoid) inv_unique: @@ -109,7 +109,7 @@ proof fix x y z assume G: "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" - { + { assume eq: "x \ y = x \ z" with G l_inv_ex obtain x_inv where xG: "x_inv \ carrier(G)" and l_inv: "x_inv \ x = \" by fast @@ -147,15 +147,15 @@ by (fast intro: l_inv r_inv) qed show ?thesis - by (blast intro: group.intro monoid.intro group_axioms.intro + by (blast intro: group.intro monoid.intro group_axioms.intro assms r_one inv_ex) qed lemma (in group) inv [simp]: "x \ carrier(G) \ inv x \ carrier(G) & inv x \ x = \ & x \ inv x = \" - apply (frule inv_ex) + apply (frule inv_ex) apply (unfold Bex_def m_inv_def) - apply (erule exE) + apply (erule exE) apply (rule theI) apply (rule ex1I, assumption) apply (blast intro: inv_unique) @@ -221,10 +221,10 @@ lemma (in group) inv_one [simp]: "inv \ = \" - by (auto intro: inv_equality) + by (auto intro: inv_equality) lemma (in group) inv_inv [simp]: "x \ carrier(G) \ inv (inv x) = x" - by (auto intro: inv_equality) + by (auto intro: inv_equality) text{*This proof is by cancellation*} lemma (in group) inv_mult_group: @@ -360,7 +360,7 @@ lemma (in group) hom_compose: "\h \ hom(G,H); i \ hom(H,I)\ \ i O h \ hom(G,I)" -by (force simp add: hom_def comp_fun) +by (force simp add: hom_def comp_fun) lemma hom_is_fun: "h \ hom(G,H) \ h \ carrier(G) -> carrier(H)" @@ -374,19 +374,19 @@ "G \ H == hom(G,H) \ bij(carrier(G), carrier(H))" lemma (in group) iso_refl: "id(carrier(G)) \ G \ G" - by (simp add: iso_def hom_def id_type id_bij) + by (simp add: iso_def hom_def id_type id_bij) lemma (in group) iso_sym: "h \ G \ H \ converse(h) \ H \ G" -apply (simp add: iso_def bij_converse_bij, clarify) -apply (subgoal_tac "converse(h) \ carrier(H) \ carrier(G)") - prefer 2 apply (simp add: bij_converse_bij bij_is_fun) -apply (auto intro: left_inverse_eq [of _ "carrier(G)" "carrier(H)"] - simp add: hom_def bij_is_inj right_inverse_bij); +apply (simp add: iso_def bij_converse_bij, clarify) +apply (subgoal_tac "converse(h) \ carrier(H) \ carrier(G)") + prefer 2 apply (simp add: bij_converse_bij bij_is_fun) +apply (auto intro: left_inverse_eq [of _ "carrier(G)" "carrier(H)"] + simp add: hom_def bij_is_inj right_inverse_bij); done -lemma (in group) iso_trans: +lemma (in group) iso_trans: "\h \ G \ H; i \ H \ I\ \ i O h \ G \ I" by (auto simp add: iso_def hom_compose comp_bij) @@ -408,7 +408,7 @@ interpret group H by fact interpret group I by fact show ?thesis - by (auto intro: lam_type simp add: iso_def hom_def inj_def surj_def bij_def) + by (auto intro: lam_type simp add: iso_def hom_def inj_def surj_def bij_def) qed text{*Basis for homomorphism proofs: we assume two groups @{term G} and @@ -482,7 +482,7 @@ lemma (in group) subgroup_self: "subgroup (carrier(G),G)" -by (simp add: subgroup_def) +by (simp add: subgroup_def) lemma (in group) subgroup_imp_group: "subgroup(H,G) \ group (update_carrier(G,H))" @@ -512,8 +512,8 @@ theorem group_BijGroup: "group(BijGroup(S))" apply (simp add: BijGroup_def) -apply (rule groupI) - apply (simp_all add: id_bij comp_bij comp_assoc) +apply (rule groupI) + apply (simp_all add: id_bij comp_bij comp_assoc) apply (simp add: id_bij bij_is_fun left_comp_id [of _ S S] fun_is_rel) apply (blast intro: left_comp_inverse bij_is_inj bij_converse_bij) done @@ -521,14 +521,14 @@ subsection{*Automorphisms Form a Group*} -lemma Bij_Inv_mem: "\f \ bij(S,S); x \ S\ \ converse(f) ` x \ S" +lemma Bij_Inv_mem: "\f \ bij(S,S); x \ S\ \ converse(f) ` x \ S" by (blast intro: apply_funtype bij_is_fun bij_converse_bij) lemma inv_BijGroup: "f \ bij(S,S) \ m_inv (BijGroup(S), f) = converse(f)" apply (rule group.inv_equality) apply (rule group_BijGroup) -apply (simp_all add: BijGroup_def bij_converse_bij - left_comp_inverse [OF bij_is_inj]) +apply (simp_all add: BijGroup_def bij_converse_bij + left_comp_inverse [OF bij_is_inj]) done lemma iso_is_bij: "h \ G \ H ==> h \ bij(carrier(G), carrier(H))" @@ -554,17 +554,17 @@ by (auto simp add: auto_def BijGroup_def iso_def) next fix x y - assume "x \ auto(G)" "y \ auto(G)" + assume "x \ auto(G)" "y \ auto(G)" thus "x \\<^bsub>BijGroup (carrier(G))\<^esub> y \ auto(G)" - by (auto simp add: BijGroup_def auto_def iso_def bij_is_fun + by (auto simp add: BijGroup_def auto_def iso_def bij_is_fun group.hom_compose comp_bij) next show "\\<^bsub>BijGroup (carrier(G))\<^esub> \ auto(G)" by (simp add: BijGroup_def id_in_auto) next - fix x - assume "x \ auto(G)" + fix x + assume "x \ auto(G)" thus "inv\<^bsub>BijGroup (carrier(G))\<^esub> x \ auto(G)" - by (simp add: auto_def inv_BijGroup iso_is_bij iso_sym) + by (simp add: auto_def inv_BijGroup iso_is_bij iso_sym) qed theorem (in group) AutoGroup: "group (AutoGroup(G))" @@ -656,13 +656,13 @@ lemma normal_imp_subgroup: "H \ G ==> subgroup(H,G)" by (simp add: normal_def subgroup_def) -lemma (in group) normalI: +lemma (in group) normalI: "subgroup(H,G) \ (\x \ carrier(G). H #> x = x <# H) \ H \ G"; by (simp add: normal_def normal_axioms_def) lemma (in normal) inv_op_closed1: "\x \ carrier(G); h \ H\ \ (inv x) \ h \ x \ H" -apply (insert coset_eq) +apply (insert coset_eq) apply (auto simp add: l_coset_def r_coset_def) apply (drule bspec, assumption) apply (drule equalityD1 [THEN subsetD], blast, clarify) @@ -672,9 +672,9 @@ lemma (in normal) inv_op_closed2: "\x \ carrier(G); h \ H\ \ x \ h \ (inv x) \ H" -apply (subgoal_tac "inv (inv x) \ h \ (inv x) \ H") -apply simp -apply (blast intro: inv_op_closed1) +apply (subgoal_tac "inv (inv x) \ h \ (inv x) \ H") +apply simp +apply (blast intro: inv_op_closed1) done text{*Alternative characterization of normal subgroups*} @@ -685,12 +685,12 @@ proof assume N: "N \ G" show ?rhs - by (blast intro: N normal.inv_op_closed2 normal_imp_subgroup) + by (blast intro: N normal.inv_op_closed2 normal_imp_subgroup) next assume ?rhs - hence sg: "subgroup(N,G)" + hence sg: "subgroup(N,G)" and closed: "\x. x\carrier(G) \ \h\N. x \ h \ inv x \ N" by auto - hence sb: "N \ carrier(G)" by (simp add: subgroup.subset) + hence sb: "N \ carrier(G)" by (simp add: subgroup.subset) show "N \ G" proof (intro normalI [OF sg], simp add: l_coset_def r_coset_def, clarify) fix x @@ -700,9 +700,9 @@ show "(\h\N. {h \ x}) \ (\h\N. {x \ h})" proof clarify fix n - assume n: "n \ N" + assume n: "n \ N" show "n \ x \ (\h\N. {x \ h})" - proof (rule UN_I) + proof (rule UN_I) from closed [of "inv x"] show "inv x \ n \ x \ N" by (simp add: x n) show "n \ x \ {x \ (inv x \ n \ x)}" @@ -713,9 +713,9 @@ show "(\h\N. {x \ h}) \ (\h\N. {h \ x})" proof clarify fix n - assume n: "n \ N" + assume n: "n \ N" show "x \ n \ (\h\N. {h \ x})" - proof (rule UN_I) + proof (rule UN_I) show "x \ n \ inv x \ N" by (simp add: x n closed) show "x \ n \ {x \ n \ inv x \ x}" by (simp add: x n m_assoc sb [THEN subsetD]) @@ -779,7 +779,7 @@ by (auto simp add: set_mult_def subsetD) lemma (in group) subgroup_mult_id: "subgroup(H,G) \ H <#> H = H" -apply (rule equalityI) +apply (rule equalityI) apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def image_def) apply (rule_tac x = x in bexI) apply (rule bexI [of _ "\"]) @@ -870,15 +870,15 @@ interpret group G by fact show ?thesis proof (simp add: equiv_def, intro conjI) show "rcong H \ carrier(G) \ carrier(G)" - by (auto simp add: r_congruent_def) + by (auto simp add: r_congruent_def) next show "refl (carrier(G), rcong H)" - by (auto simp add: r_congruent_def refl_def) + by (auto simp add: r_congruent_def refl_def) next show "sym (rcong H)" proof (simp add: r_congruent_def sym_def, clarify) fix x y - assume [simp]: "x \ carrier(G)" "y \ carrier(G)" + assume [simp]: "x \ carrier(G)" "y \ carrier(G)" and "inv x \ y \ H" hence "inv (inv x \ y) \ H" by simp thus "inv y \ x \ H" by (simp add: inv_mult_group) @@ -890,7 +890,7 @@ assume [simp]: "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" and "inv x \ y \ H" and "inv y \ z \ H" hence "(inv x \ y) \ (inv y \ z) \ H" by simp - hence "inv x \ (y \ inv y) \ z \ H" by (simp add: m_assoc del: inv) + hence "inv x \ (y \ inv y) \ z \ H" by (simp add: m_assoc del: inv) thus "inv x \ z \ H" by simp qed qed @@ -902,18 +902,18 @@ lemma (in subgroup) l_coset_eq_rcong: assumes "group(G)" assumes a: "a \ carrier(G)" - shows "a <# H = (rcong H) `` {a}" + shows "a <# H = (rcong H) `` {a}" proof - interpret group G by fact show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a - Collect_image_eq) + Collect_image_eq) qed lemma (in group) rcos_equation: assumes "subgroup(H, G)" shows - "\ha \ a = h \ b; a \ carrier(G); b \ carrier(G); + "\ha \ a = h \ b; a \ carrier(G); b \ carrier(G); h \ H; ha \ H; hb \ H\ \ hb \ a \ (\h\H. {h \ b})" (is "PROP ?P") proof - @@ -982,15 +982,15 @@ show "|H #> a| = |H|" proof (rule eqpollI [THEN cardinal_cong]) show "H #> a \ H" - proof (simp add: lepoll_def, intro exI) + proof (simp add: lepoll_def, intro exI) show "(\y \ H#>a. y \ inv a) \ inj(H #> a, H)" - by (auto intro: lam_type + by (auto intro: lam_type simp add: inj_def r_coset_def m_assoc subsetD [OF H] a) qed show "H \ H #> a" - proof (simp add: lepoll_def, intro exI) + proof (simp add: lepoll_def, intro exI) show "(\y\ H. y \ a) \ inj(H, H #> a)" - by (auto intro: lam_type + by (auto intro: lam_type simp add: inj_def r_coset_def subsetD [OF H] a) qed qed @@ -1021,7 +1021,7 @@ definition FactGroup :: "[i,i] => i" (infixl "Mod" 65) where --{*Actually defined for groups rather than monoids*} - "G Mod H == + "G Mod H == G\<^esub> H, \ \ (rcosets\<^bsub>G\<^esub> H) \ (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>" lemma (in normal) setmult_closed: @@ -1066,7 +1066,7 @@ lemma (in normal) inv_FactGroup: "X \ carrier (G Mod H) \ inv\<^bsub>G Mod H\<^esub> X = set_inv X" -apply (rule group.inv_equality [OF factorgroup_is_group]) +apply (rule group.inv_equality [OF factorgroup_is_group]) apply (simp_all add: FactGroup_def setinv_closed rcosets_inv_mult_group_eq) done @@ -1074,12 +1074,12 @@ @{term "G Mod H"}*} lemma (in normal) r_coset_hom_Mod: "(\a \ carrier(G). H #> a) \ hom(G, G Mod H)" -by (auto simp add: FactGroup_def RCOSETS_def hom_def rcos_sum intro: lam_type) +by (auto simp add: FactGroup_def RCOSETS_def hom_def rcos_sum intro: lam_type) subsection{*The First Isomorphism Theorem*} -text{*The quotient by the kernel of a homomorphism is isomorphic to the +text{*The quotient by the kernel of a homomorphism is isomorphic to the range of that homomorphism.*} definition @@ -1088,14 +1088,14 @@ "kernel(G,H,h) == {x \ carrier(G). h ` x = \\<^bsub>H\<^esub>}"; lemma (in group_hom) subgroup_kernel: "subgroup (kernel(G,H,h), G)" -apply (rule subgroup.intro) +apply (rule subgroup.intro) apply (auto simp add: kernel_def group.intro) done text{*The kernel of a homomorphism is a normal subgroup*} lemma (in group_hom) normal_kernel: "(kernel(G,H,h)) \ G" apply (simp add: group.normal_inv_iff subgroup_kernel group.intro) -apply (simp add: kernel_def) +apply (simp add: kernel_def) done lemma (in group_hom) FactGroup_nonempty: @@ -1103,10 +1103,10 @@ shows "X \ 0" proof - from X - obtain g where "g \ carrier(G)" + obtain g where "g \ carrier(G)" and "X = kernel(G,H,h) #> g" by (auto simp add: FactGroup_def RCOSETS_def) - thus ?thesis + thus ?thesis by (auto simp add: kernel_def r_coset_def image_def intro: hom_one) qed @@ -1116,46 +1116,46 @@ shows "contents (h``X) \ carrier(H)" proof - from X - obtain g where g: "g \ carrier(G)" + obtain g where g: "g \ carrier(G)" and "X = kernel(G,H,h) #> g" by (auto simp add: FactGroup_def RCOSETS_def) hence "h `` X = {h ` g}" - by (auto simp add: kernel_def r_coset_def image_UN + by (auto simp add: kernel_def r_coset_def image_UN image_eq_UN [OF hom_is_fun] g) thus ?thesis by (auto simp add: g) qed lemma mult_FactGroup: - "[|X \ carrier(G Mod H); X' \ carrier(G Mod H)|] + "[|X \ carrier(G Mod H); X' \ carrier(G Mod H)|] ==> X \\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'" -by (simp add: FactGroup_def) +by (simp add: FactGroup_def) lemma (in normal) FactGroup_m_closed: - "[|X \ carrier(G Mod H); X' \ carrier(G Mod H)|] + "[|X \ carrier(G Mod H); X' \ carrier(G Mod H)|] ==> X <#>\<^bsub>G\<^esub> X' \ carrier(G Mod H)" -by (simp add: FactGroup_def setmult_closed) +by (simp add: FactGroup_def setmult_closed) lemma (in group_hom) FactGroup_hom: "(\X \ carrier(G Mod (kernel(G,H,h))). contents (h``X)) - \ hom (G Mod (kernel(G,H,h)), H)" -proof (simp add: hom_def FactGroup_contents_mem lam_type mult_FactGroup normal.FactGroup_m_closed [OF normal_kernel], intro ballI) + \ hom (G Mod (kernel(G,H,h)), H)" +proof (simp add: hom_def FactGroup_contents_mem lam_type mult_FactGroup normal.FactGroup_m_closed [OF normal_kernel], intro ballI) fix X and X' assume X: "X \ carrier (G Mod kernel(G,H,h))" and X': "X' \ carrier (G Mod kernel(G,H,h))" then obtain g and g' - where "g \ carrier(G)" and "g' \ carrier(G)" + where "g \ carrier(G)" and "g' \ carrier(G)" and "X = kernel(G,H,h) #> g" and "X' = kernel(G,H,h) #> g'" by (auto simp add: FactGroup_def RCOSETS_def) - hence all: "\x\X. h ` x = h ` g" "\x\X'. h ` x = h ` g'" + hence all: "\x\X. h ` x = h ` g" "\x\X'. h ` x = h ` g'" and Xsub: "X \ carrier(G)" and X'sub: "X' \ carrier(G)" by (force simp add: kernel_def r_coset_def image_def)+ hence "h `` (X <#> X') = {h ` g \\<^bsub>H\<^esub> h ` g'}" using X X' by (auto dest!: FactGroup_nonempty simp add: set_mult_def image_eq_UN [OF hom_is_fun] image_UN - subsetD [OF Xsub] subsetD [OF X'sub]) + subsetD [OF Xsub] subsetD [OF X'sub]) thus "contents (h `` (X <#> X')) = contents (h `` X) \\<^bsub>H\<^esub> contents (h `` X')" - by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty + by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty X X' Xsub X'sub) qed @@ -1165,21 +1165,21 @@ "\g \ carrier(G); g' \ carrier(G); h ` g = h ` g'\ \ kernel(G,H,h) #> g \ kernel(G,H,h) #> g'" apply (clarsimp simp add: kernel_def r_coset_def image_def) -apply (rename_tac y) -apply (rule_tac x="y \ g \ inv g'" in bexI) -apply (simp_all add: G.m_assoc) +apply (rename_tac y) +apply (rule_tac x="y \ g \ inv g'" in bexI) +apply (simp_all add: G.m_assoc) done lemma (in group_hom) FactGroup_inj: "(\X\carrier (G Mod kernel(G,H,h)). contents (h `` X)) \ inj(carrier (G Mod kernel(G,H,h)), carrier(H))" -proof (simp add: inj_def FactGroup_contents_mem lam_type, clarify) +proof (simp add: inj_def FactGroup_contents_mem lam_type, clarify) fix X and X' assume X: "X \ carrier (G Mod kernel(G,H,h))" and X': "X' \ carrier (G Mod kernel(G,H,h))" then obtain g and g' - where gX: "g \ carrier(G)" "g' \ carrier(G)" + where gX: "g \ carrier(G)" "g' \ carrier(G)" "X = kernel(G,H,h) #> g" "X' = kernel(G,H,h) #> g'" by (auto simp add: FactGroup_def RCOSETS_def) hence all: "\x\X. h ` x = h ` g" "\x\X'. h ` x = h ` g'" @@ -1187,16 +1187,16 @@ by (force simp add: kernel_def r_coset_def image_def)+ assume "contents (h `` X) = contents (h `` X')" hence h: "h ` g = h ` g'" - by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty + by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty X X' Xsub X'sub) - show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX) + show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX) qed lemma (in group_hom) kernel_rcoset_subset: assumes g: "g \ carrier(G)" shows "kernel(G,H,h) #> g \ carrier (G)" - by (auto simp add: g kernel_def r_coset_def) + by (auto simp add: g kernel_def r_coset_def) @@ -1210,12 +1210,12 @@ fix y assume y: "y \ carrier(H)" with h obtain g where g: "g \ carrier(G)" "h ` g = y" - by (auto simp add: surj_def) - hence "(\x\kernel(G,H,h) #> g. {h ` x}) = {y}" - by (auto simp add: y kernel_def r_coset_def) + by (auto simp add: surj_def) + hence "(\x\kernel(G,H,h) #> g. {h ` x}) = {y}" + by (auto simp add: y kernel_def r_coset_def) with g show "\x\carrier(G Mod kernel(G, H, h)). contents(h `` x) = y" --{*The witness is @{term "kernel(G,H,h) #> g"}*} - by (force simp add: FactGroup_def RCOSETS_def + by (force simp add: FactGroup_def RCOSETS_def image_eq_UN [OF hom_is_fun] kernel_rcoset_subset) qed @@ -1226,5 +1226,5 @@ "h \ surj(carrier(G), carrier(H)) \ (\X\carrier (G Mod kernel(G,H,h)). contents (h``X)) \ (G Mod (kernel(G,H,h))) \ H" by (simp add: iso_def FactGroup_hom FactGroup_inj bij_def FactGroup_surj) - + end diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/func.thy --- a/src/ZF/func.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/func.thy Thu Mar 15 16:35:02 2012 +0000 @@ -14,37 +14,37 @@ lemma relation_converse_converse [simp]: "relation(r) ==> converse(converse(r)) = r" -by (simp add: relation_def, blast) +by (simp add: relation_def, blast) lemma relation_restrict [simp]: "relation(restrict(r,A))" -by (simp add: restrict_def relation_def, blast) +by (simp add: restrict_def relation_def, blast) lemma Pi_iff: - "f: Pi(A,B) \ function(f) & f<=Sigma(A,B) & A<=domain(f)" + "f \ Pi(A,B) \ function(f) & f<=Sigma(A,B) & A<=domain(f)" by (unfold Pi_def, blast) (*For upward compatibility with the former definition*) lemma Pi_iff_old: - "f: Pi(A,B) \ f<=Sigma(A,B) & (\x\A. EX! y. : f)" + "f \ Pi(A,B) \ f<=Sigma(A,B) & (\x\A. EX! y. : f)" by (unfold Pi_def function_def, blast) -lemma fun_is_function: "f: Pi(A,B) ==> function(f)" +lemma fun_is_function: "f \ Pi(A,B) ==> function(f)" by (simp only: Pi_iff) lemma function_imp_Pi: "[|function(f); relation(f)|] ==> f \ domain(f) -> range(f)" -by (simp add: Pi_iff relation_def, blast) +by (simp add: Pi_iff relation_def, blast) -lemma functionI: +lemma functionI: "[| !!x y y'. [| :r; :r |] ==> y=y' |] ==> function(r)" -by (simp add: function_def, blast) +by (simp add: function_def, blast) (*Functions are relations*) -lemma fun_is_rel: "f: Pi(A,B) ==> f \ Sigma(A,B)" +lemma fun_is_rel: "f \ Pi(A,B) ==> f \ Sigma(A,B)" by (unfold Pi_def, blast) lemma Pi_cong: - "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')" + "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')" by (simp add: Pi_def cong add: Sigma_cong) (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause @@ -52,18 +52,18 @@ Sigmas and Pis are abbreviated as * or -> *) (*Weakening one function type to another; see also Pi_type*) -lemma fun_weaken_type: "[| f: A->B; B<=D |] ==> f: A->D" +lemma fun_weaken_type: "[| f \ A->B; B<=D |] ==> f \ A->D" by (unfold Pi_def, best) subsection{*Function Application*} -lemma apply_equality2: "[| : f; : f; f: Pi(A,B) |] ==> b=c" +lemma apply_equality2: "[| : f; : f; f \ Pi(A,B) |] ==> b=c" by (unfold Pi_def function_def, blast) lemma function_apply_equality: "[| : f; function(f) |] ==> f`a = b" by (unfold apply_def function_def, blast) -lemma apply_equality: "[| : f; f: Pi(A,B) |] ==> f`a = b" +lemma apply_equality: "[| : f; f \ Pi(A,B) |] ==> f`a = b" apply (unfold Pi_def) apply (blast intro: function_apply_equality) done @@ -72,72 +72,72 @@ lemma apply_0: "a \ domain(f) ==> f`a = 0" by (unfold apply_def, blast) -lemma Pi_memberD: "[| f: Pi(A,B); c: f |] ==> \x\A. c = " +lemma Pi_memberD: "[| f \ Pi(A,B); c \ f |] ==> \x\A. c = " apply (frule fun_is_rel) apply (blast dest: apply_equality) done lemma function_apply_Pair: "[| function(f); a \ domain(f)|] ==> : f" -apply (simp add: function_def, clarify) -apply (subgoal_tac "f`a = y", blast) -apply (simp add: apply_def, blast) +apply (simp add: function_def, clarify) +apply (subgoal_tac "f`a = y", blast) +apply (simp add: apply_def, blast) done -lemma apply_Pair: "[| f: Pi(A,B); a:A |] ==> : f" +lemma apply_Pair: "[| f \ Pi(A,B); a \ A |] ==> : f" apply (simp add: Pi_iff) apply (blast intro: function_apply_Pair) done (*Conclusion is flexible -- use rule_tac or else apply_funtype below!*) -lemma apply_type [TC]: "[| f: Pi(A,B); a:A |] ==> f`a \ B(a)" +lemma apply_type [TC]: "[| f \ Pi(A,B); a \ A |] ==> f`a \ B(a)" by (blast intro: apply_Pair dest: fun_is_rel) (*This version is acceptable to the simplifier*) -lemma apply_funtype: "[| f: A->B; a:A |] ==> f`a \ B" +lemma apply_funtype: "[| f \ A->B; a \ A |] ==> f`a \ B" by (blast dest: apply_type) -lemma apply_iff: "f: Pi(A,B) ==> : f \ a:A & f`a = b" +lemma apply_iff: "f \ Pi(A,B) ==> : f \ a \ A & f`a = b" apply (frule fun_is_rel) apply (blast intro!: apply_Pair apply_equality) done (*Refining one Pi type to another*) -lemma Pi_type: "[| f: Pi(A,C); !!x. x:A ==> f`x \ B(x) |] ==> f \ Pi(A,B)" +lemma Pi_type: "[| f \ Pi(A,C); !!x. x \ A ==> f`x \ B(x) |] ==> f \ Pi(A,B)" apply (simp only: Pi_iff) apply (blast dest: function_apply_equality) done (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) lemma Pi_Collect_iff: - "(f \ Pi(A, %x. {y:B(x). P(x,y)})) + "(f \ Pi(A, %x. {y \ B(x). P(x,y)})) \ f \ Pi(A,B) & (\x\A. P(x, f`x))" by (blast intro: Pi_type dest: apply_type) lemma Pi_weaken_type: - "[| f \ Pi(A,B); !!x. x:A ==> B(x)<=C(x) |] ==> f \ Pi(A,C)" + "[| f \ Pi(A,B); !!x. x \ A ==> B(x)<=C(x) |] ==> f \ Pi(A,C)" by (blast intro: Pi_type dest: apply_type) (** Elimination of membership in a function **) -lemma domain_type: "[| \ f; f: Pi(A,B) |] ==> a \ A" +lemma domain_type: "[| \ f; f \ Pi(A,B) |] ==> a \ A" by (blast dest: fun_is_rel) -lemma range_type: "[| \ f; f: Pi(A,B) |] ==> b \ B(a)" +lemma range_type: "[| \ f; f \ Pi(A,B) |] ==> b \ B(a)" by (blast dest: fun_is_rel) -lemma Pair_mem_PiD: "[| : f; f: Pi(A,B) |] ==> a:A & b:B(a) & f`a = b" +lemma Pair_mem_PiD: "[| : f; f \ Pi(A,B) |] ==> a \ A & b \ B(a) & f`a = b" by (blast intro: domain_type range_type apply_equality) subsection{*Lambda Abstraction*} -lemma lamI: "a:A ==> \ (\x\A. b(x))" +lemma lamI: "a \ A ==> \ (\x\A. b(x))" apply (unfold lam_def) apply (erule RepFunI) done lemma lamE: - "[| p: (\x\A. b(x)); !!x.[| x:A; p= |] ==> P + "[| p: (\x\A. b(x)); !!x.[| x \ A; p= |] ==> P |] ==> P" by (simp add: lam_def, blast) @@ -145,17 +145,17 @@ by (simp add: lam_def) lemma lam_type [TC]: - "[| !!x. x:A ==> b(x): B(x) |] ==> (\x\A. b(x)) \ Pi(A,B)" + "[| !!x. x \ A ==> b(x): B(x) |] ==> (\x\A. b(x)) \ Pi(A,B)" by (simp add: lam_def Pi_def function_def, blast) -lemma lam_funtype: "(\x\A. b(x)) \ A -> {b(x). x:A}" +lemma lam_funtype: "(\x\A. b(x)) \ A -> {b(x). x \ A}" by (blast intro: lam_type) lemma function_lam: "function (\x\A. b(x))" -by (simp add: function_def lam_def) +by (simp add: function_def lam_def) -lemma relation_lam: "relation (\x\A. b(x))" -by (simp add: relation_def lam_def) +lemma relation_lam: "relation (\x\A. b(x))" +by (simp add: relation_def lam_def) lemma beta_if [simp]: "(\x\A. b(x)) ` a = (if a \ A then b(a) else 0)" by (simp add: apply_def lam_def, blast) @@ -171,17 +171,17 @@ (*congruence rule for lambda abstraction*) lemma lam_cong [cong]: - "[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')" + "[| A=A'; !!x. x \ A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')" by (simp only: lam_def cong add: RepFun_cong) lemma lam_theI: - "(!!x. x:A ==> EX! y. Q(x,y)) ==> \f. \x\A. Q(x, f`x)" + "(!!x. x \ A ==> EX! y. Q(x,y)) ==> \f. \x\A. Q(x, f`x)" apply (rule_tac x = "\x\A. THE y. Q (x,y)" in exI) -apply simp +apply simp apply (blast intro: theI) done -lemma lam_eqE: "[| (\x\A. f(x)) = (\x\A. g(x)); a:A |] ==> f(a)=g(a)" +lemma lam_eqE: "[| (\x\A. f(x)) = (\x\A. g(x)); a \ A |] ==> f(a)=g(a)" by (fast intro!: lamI elim: equalityE lamE) @@ -207,13 +207,13 @@ (*Semi-extensionality!*) lemma fun_subset: - "[| f \ Pi(A,B); g: Pi(C,D); A<=C; - !!x. x:A ==> f`x = g`x |] ==> f<=g" + "[| f \ Pi(A,B); g \ Pi(C,D); A<=C; + !!x. x \ A ==> f`x = g`x |] ==> f<=g" by (force dest: Pi_memberD intro: apply_Pair) lemma fun_extension: - "[| f \ Pi(A,B); g: Pi(A,D); - !!x. x:A ==> f`x = g`x |] ==> f=g" + "[| f \ Pi(A,B); g \ Pi(A,D); + !!x. x \ A ==> f`x = g`x |] ==> f=g" by (blast del: subsetI intro: subset_refl sym fun_subset) lemma eta [simp]: "f \ Pi(A,B) ==> (\x\A. f`x) = f" @@ -222,18 +222,18 @@ done lemma fun_extension_iff: - "[| f:Pi(A,B); g:Pi(A,C) |] ==> (\a\A. f`a = g`a) \ f=g" + "[| f \ Pi(A,B); g \ Pi(A,C) |] ==> (\a\A. f`a = g`a) \ f=g" by (blast intro: fun_extension) (*thm by Mark Staples, proof by lcp*) -lemma fun_subset_eq: "[| f:Pi(A,B); g:Pi(A,C) |] ==> f \ g \ (f = g)" +lemma fun_subset_eq: "[| f \ Pi(A,B); g \ Pi(A,C) |] ==> f \ g \ (f = g)" by (blast dest: apply_Pair intro: fun_extension apply_equality [symmetric]) (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) lemma Pi_lamE: - assumes major: "f: Pi(A,B)" + assumes major: "f \ Pi(A,B)" and minor: "!!b. [| \x\A. b(x):B(x); f = (\x\A. b(x)) |] ==> P" shows "P" apply (rule minor) @@ -244,37 +244,37 @@ subsection{*Images of Functions*} -lemma image_lam: "C \ A ==> (\x\A. b(x)) `` C = {b(x). x:C}" +lemma image_lam: "C \ A ==> (\x\A. b(x)) `` C = {b(x). x \ C}" by (unfold lam_def, blast) lemma Repfun_function_if: - "function(f) - ==> {f`x. x:C} = (if C \ domain(f) then f``C else cons(0,f``C))"; + "function(f) + ==> {f`x. x \ C} = (if C \ domain(f) then f``C else cons(0,f``C))"; apply simp -apply (intro conjI impI) - apply (blast dest: function_apply_equality intro: function_apply_Pair) +apply (intro conjI impI) + apply (blast dest: function_apply_equality intro: function_apply_Pair) apply (rule equalityI) - apply (blast intro!: function_apply_Pair apply_0) -apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) + apply (blast intro!: function_apply_Pair apply_0) +apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) done -(*For this lemma and the next, the right-hand side could equivalently +(*For this lemma and the next, the right-hand side could equivalently be written \x\C. {f`x} *) lemma image_function: - "[| function(f); C \ domain(f) |] ==> f``C = {f`x. x:C}"; -by (simp add: Repfun_function_if) + "[| function(f); C \ domain(f) |] ==> f``C = {f`x. x \ C}"; +by (simp add: Repfun_function_if) -lemma image_fun: "[| f \ Pi(A,B); C \ A |] ==> f``C = {f`x. x:C}" -apply (simp add: Pi_iff) -apply (blast intro: image_function) +lemma image_fun: "[| f \ Pi(A,B); C \ A |] ==> f``C = {f`x. x \ C}" +apply (simp add: Pi_iff) +apply (blast intro: image_function) done -lemma image_eq_UN: +lemma image_eq_UN: assumes f: "f \ Pi(A,B)" "C \ A" shows "f``C = (\x\C. {f ` x})" -by (auto simp add: image_fun [OF f]) +by (auto simp add: image_fun [OF f]) lemma Pi_image_cons: - "[| f: Pi(A,B); x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)" + "[| f \ Pi(A,B); x \ A |] ==> f `` cons(x,y) = cons(f`x, f``y)" by (blast dest: apply_equality apply_Pair) @@ -287,7 +287,7 @@ "function(f) ==> function(restrict(f,A))" by (unfold restrict_def function_def, blast) -lemma restrict_type2: "[| f: Pi(C,B); A<=C |] ==> restrict(f,A) \ Pi(A,B)" +lemma restrict_type2: "[| f \ Pi(C,B); A<=C |] ==> restrict(f,A) \ Pi(A,B)" by (simp add: Pi_iff function_def restrict_def, blast) lemma restrict: "restrict(f,A) ` a = (if a \ A then f`a else 0)" @@ -297,7 +297,7 @@ by (unfold restrict_def, simp) lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" -by (simp add: restrict_def) +by (simp add: restrict_def) lemma restrict_restrict [simp]: "restrict(restrict(r,A),B) = restrict(r, A \ B)" @@ -346,10 +346,10 @@ "[| \x\S. function(x); \x\S. \y\S. x<=y | y<=x |] ==> function(\(S))" -by (unfold function_def, blast) +by (unfold function_def, blast) lemma fun_Union: - "[| \f\S. \C D. f:C->D; + "[| \f\S. \C D. f \ C->D; \f\S. \y\S. f<=y | y<=f |] ==> \(S) \ domain(\(S)) -> range(\(S))" apply (unfold Pi_def) @@ -358,7 +358,7 @@ lemma gen_relation_Union [rule_format]: "\f\F. relation(f) \ relation(\(F))" -by (simp add: relation_def) +by (simp add: relation_def) (** The Union of 2 disjoint functions is a function **) @@ -368,7 +368,7 @@ subset_trans [OF _ Un_upper2] lemma fun_disjoint_Un: - "[| f: A->B; g: C->D; A \ C = 0 |] + "[| f \ A->B; g \ C->D; A \ C = 0 |] ==> (f \ g) \ (A \ C) -> (B \ D)" (*Prove the product and domain subgoals using distributive laws*) apply (simp add: Pi_iff extension Un_rls) @@ -376,17 +376,17 @@ done lemma fun_disjoint_apply1: "a \ domain(g) ==> (f \ g)`a = f`a" -by (simp add: apply_def, blast) +by (simp add: apply_def, blast) lemma fun_disjoint_apply2: "c \ domain(f) ==> (f \ g)`c = g`c" -by (simp add: apply_def, blast) +by (simp add: apply_def, blast) subsection{*Domain and Range of a Function or Relation*} lemma domain_of_fun: "f \ Pi(A,B) ==> domain(f)=A" by (unfold Pi_def, blast) -lemma apply_rangeI: "[| f \ Pi(A,B); a: A |] ==> f`a \ range(f)" +lemma apply_rangeI: "[| f \ Pi(A,B); a \ A |] ==> f`a \ range(f)" by (erule apply_Pair [THEN rangeI], assumption) lemma range_of_fun: "f \ Pi(A,B) ==> f \ A->range(f)" @@ -395,23 +395,23 @@ subsection{*Extensions of Functions*} lemma fun_extend: - "[| f: A->B; c\A |] ==> cons(,f) \ cons(c,A) -> cons(b,B)" + "[| f \ A->B; c\A |] ==> cons(,f) \ cons(c,A) -> cons(b,B)" apply (frule singleton_fun [THEN fun_disjoint_Un], blast) -apply (simp add: cons_eq) +apply (simp add: cons_eq) done lemma fun_extend3: - "[| f: A->B; c\A; b: B |] ==> cons(,f) \ cons(c,A) -> B" + "[| f \ A->B; c\A; b \ B |] ==> cons(,f) \ cons(c,A) -> B" by (blast intro: fun_extend [THEN fun_weaken_type]) lemma extend_apply: "c \ domain(f) ==> cons(,f)`a = (if a=c then b else f`a)" -by (auto simp add: apply_def) +by (auto simp add: apply_def) lemma fun_extend_apply [simp]: - "[| f: A->B; c\A |] ==> cons(,f)`a = (if a=c then b else f`a)" -apply (rule extend_apply) -apply (simp add: Pi_def, blast) + "[| f \ A->B; c\A |] ==> cons(,f)`a = (if a=c then b else f`a)" +apply (rule extend_apply) +apply (simp add: Pi_def, blast) done lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp] @@ -425,13 +425,13 @@ apply (subgoal_tac "restrict (x, A) \ A -> B") prefer 2 apply (blast intro: restrict_type2) apply (rule UN_I, assumption) -apply (rule apply_funtype [THEN UN_I]) +apply (rule apply_funtype [THEN UN_I]) apply assumption - apply (rule consI1) + apply (rule consI1) apply (simp (no_asm)) -apply (rule fun_extension) +apply (rule fun_extension) apply assumption - apply (blast intro: fun_extend) + apply (blast intro: fun_extend) apply (erule consE, simp_all) done @@ -463,11 +463,11 @@ lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)" apply (simp add: update_def) -apply (case_tac "z \ domain(f)") +apply (case_tac "z \ domain(f)") apply (simp_all add: apply_0) done -lemma update_idem: "[| f`x = y; f: Pi(A,B); x: A |] ==> f(x:=y) = f" +lemma update_idem: "[| f`x = y; f \ Pi(A,B); x \ A |] ==> f(x:=y) = f" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb) apply (rule fun_extension) @@ -475,13 +475,13 @@ done -(* [| f: Pi(A, B); x:A |] ==> f(x := f`x) = f *) +(* [| f \ Pi(A, B); x \ A |] ==> f(x := f`x) = f *) declare refl [THEN update_idem, simp] lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))" by (unfold update_def, simp) -lemma update_type: "[| f:Pi(A,B); x \ A; y: B(x) |] ==> f(x:=y) \ Pi(A, B)" +lemma update_type: "[| f \ Pi(A,B); x \ A; y \ B(x) |] ==> f(x:=y) \ Pi(A, B)" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type) done @@ -496,7 +496,7 @@ lemma Replace_mono: "A<=B ==> Replace(A,P) \ Replace(B,P)" by (blast elim!: ReplaceE) -lemma RepFun_mono: "A<=B ==> {f(x). x:A} \ {f(x). x:B}" +lemma RepFun_mono: "A<=B ==> {f(x). x \ A} \ {f(x). x \ B}" by blast lemma Pow_mono: "A<=B ==> Pow(A) \ Pow(B)" @@ -506,8 +506,8 @@ by blast lemma UN_mono: - "[| A<=C; !!x. x:A ==> B(x)<=D(x) |] ==> (\x\A. B(x)) \ (\x\C. D(x))" -by blast + "[| A<=C; !!x. x \ A ==> B(x)<=D(x) |] ==> (\x\A. B(x)) \ (\x\C. D(x))" +by blast (*Intersection is ANTI-monotonic. There are TWO premises! *) lemma Inter_anti_mono: "[| A<=B; A\0 |] ==> \(B) \ \(A)" @@ -528,7 +528,7 @@ subsubsection{*Standard Products, Sums and Function Spaces *} lemma Sigma_mono [rule_format]: - "[| A<=C; !!x. x:A \ B(x) \ D(x) |] ==> Sigma(A,B) \ Sigma(C,D)" + "[| A<=C; !!x. x \ A \ B(x) \ D(x) |] ==> Sigma(A,B) \ Sigma(C,D)" by blast lemma sum_mono: "[| A<=C; B<=D |] ==> A+B \ C+D" @@ -569,11 +569,11 @@ lemma image_pair_mono: "[| !! x y. :r ==> :s; A<=B |] ==> r``A \ s``B" -by blast +by blast lemma vimage_pair_mono: "[| !! x y. :r ==> :s; A<=B |] ==> r-``A \ s-``B" -by blast +by blast lemma image_mono: "[| r<=s; A<=B |] ==> r``A \ s``B" by blast @@ -582,11 +582,11 @@ by blast lemma Collect_mono: - "[| A<=B; !!x. x:A ==> P(x) \ Q(x) |] ==> Collect(A,P) \ Collect(B,Q)" + "[| A<=B; !!x. x \ A ==> P(x) \ Q(x) |] ==> Collect(A,P) \ Collect(B,Q)" by blast (*Used in intr_elim.ML and in individual datatype definitions*) -lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono +lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono Part_mono in_mono (* Useful with simp; contributed by Clemens Ballarin. *) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/pair.thy --- a/src/ZF/pair.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/pair.thy Thu Mar 15 16:35:02 2012 +0000 @@ -63,7 +63,7 @@ have "{a, a} \ {{a, a}, {a, b}}" by (rule consI1) hence "{a, a} \ a" by (simp add: eq) moreover have "a \ {a, a}" by (rule consI1) - ultimately show "P" by (rule mem_asym) + ultimately show "P" by (rule mem_asym) qed lemma Pair_neq_snd: "=b ==> P" @@ -72,7 +72,7 @@ have "{a, b} \ {{a, a}, {a, b}}" by blast hence "{a, b} \ b" by (simp add: eq) moreover have "b \ {a, b}" by blast - ultimately show "P" by (rule mem_asym) + ultimately show "P" by (rule mem_asym) qed @@ -80,10 +80,10 @@ text{*Generalizes Cartesian product*} -lemma Sigma_iff [simp]: ": Sigma(A,B) \ a:A & b:B(a)" +lemma Sigma_iff [simp]: ": Sigma(A,B) \ a \ A & b \ B(a)" by (simp add: Sigma_def) -lemma SigmaI [TC,intro!]: "[| a:A; b:B(a) |] ==> \ Sigma(A,B)" +lemma SigmaI [TC,intro!]: "[| a \ A; b \ B(a) |] ==> \ Sigma(A,B)" by simp lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1] @@ -91,19 +91,19 @@ (*The general elimination rule*) lemma SigmaE [elim!]: - "[| c: Sigma(A,B); - !!x y.[| x:A; y:B(x); c= |] ==> P + "[| c \ Sigma(A,B); + !!x y.[| x \ A; y \ B(x); c= |] ==> P |] ==> P" -by (unfold Sigma_def, blast) +by (unfold Sigma_def, blast) lemma SigmaE2 [elim!]: - "[| \ Sigma(A,B); - [| a:A; b:B(a) |] ==> P + "[| \ Sigma(A,B); + [| a \ A; b \ B(a) |] ==> P |] ==> P" -by (unfold Sigma_def, blast) +by (unfold Sigma_def, blast) lemma Sigma_cong: - "[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> + "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> Sigma(A,B) = Sigma(A',B')" by (simp add: Sigma_def) @@ -129,13 +129,13 @@ lemma snd_conv [simp]: "snd() = b" by (simp add: snd_def) -lemma fst_type [TC]: "p:Sigma(A,B) ==> fst(p) \ A" +lemma fst_type [TC]: "p \ Sigma(A,B) ==> fst(p) \ A" by auto -lemma snd_type [TC]: "p:Sigma(A,B) ==> snd(p) \ B(fst(p))" +lemma snd_type [TC]: "p \ Sigma(A,B) ==> snd(p) \ B(fst(p))" by auto -lemma Pair_fst_snd_eq: "a: Sigma(A,B) ==> = a" +lemma Pair_fst_snd_eq: "a \ Sigma(A,B) ==> = a" by auto @@ -146,13 +146,13 @@ by (simp add: split_def) lemma split_type [TC]: - "[| p:Sigma(A,B); - !!x y.[| x:A; y:B(x) |] ==> c(x,y):C() + "[| p \ Sigma(A,B); + !!x y.[| x \ A; y \ B(x) |] ==> c(x,y):C() |] ==> split(%x y. c(x,y), p) \ C(p)" -by (erule SigmaE, auto) +by (erule SigmaE, auto) -lemma expand_split: - "u: A*B ==> +lemma expand_split: + "u \ A*B ==> R(split(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" by (auto simp add: split_def) @@ -163,8 +163,8 @@ by (simp add: split_def) lemma splitE: - "[| split(R,z); z:Sigma(A,B); - !!x y. [| z = ; R(x,y) |] ==> P + "[| split(R,z); z \ Sigma(A,B); + !!x y. [| z = ; R(x,y) |] ==> P |] ==> P" by (auto simp add: split_def) diff -r 5e1bcfdcb175 -r 2b6e55924af3 src/ZF/upair.thy --- a/src/ZF/upair.thy Thu Mar 15 15:54:22 2012 +0000 +++ b/src/ZF/upair.thy Thu Mar 15 16:35:02 2012 +0000 @@ -17,7 +17,7 @@ setup TypeCheck.setup lemma atomize_ball [symmetric, rulify]: - "(!!x. x:A ==> P(x)) == Trueprop (\x\A. P(x))" + "(!!x. x \ A ==> P(x)) == Trueprop (\x\A. P(x))" by (simp add: Ball_def atomize_all atomize_imp) @@ -37,7 +37,7 @@ subsection{*Rules for Binary Union, Defined via @{term Upair}*} -lemma Un_iff [simp]: "c \ A \ B \ (c:A | c:B)" +lemma Un_iff [simp]: "c \ A \ B \ (c \ A | c \ B)" apply (simp add: Un_def) apply (blast intro: UpairI1 UpairI2 elim: UpairE) done @@ -50,11 +50,11 @@ declare UnI1 [elim?] UnI2 [elim?] -lemma UnE [elim!]: "[| c \ A \ B; c:A ==> P; c:B ==> P |] ==> P" +lemma UnE [elim!]: "[| c \ A \ B; c \ A ==> P; c \ B ==> P |] ==> P" by (simp, blast) (*Stronger version of the rule above*) -lemma UnE': "[| c \ A \ B; c:A ==> P; [| c:B; c\A |] ==> P |] ==> P" +lemma UnE': "[| c \ A \ B; c \ A ==> P; [| c \ B; c\A |] ==> P |] ==> P" by (simp, blast) (*Classical introduction rule: no commitment to A vs B*) @@ -63,7 +63,7 @@ subsection{*Rules for Binary Intersection, Defined via @{term Upair}*} -lemma Int_iff [simp]: "c \ A \ B \ (c:A & c:B)" +lemma Int_iff [simp]: "c \ A \ B \ (c \ A & c \ B)" apply (unfold Int_def) apply (blast intro: UpairI1 UpairI2 elim: UpairE) done @@ -77,13 +77,13 @@ lemma IntD2: "c \ A \ B ==> c \ B" by simp -lemma IntE [elim!]: "[| c \ A \ B; [| c:A; c:B |] ==> P |] ==> P" +lemma IntE [elim!]: "[| c \ A \ B; [| c \ A; c \ B |] ==> P |] ==> P" by simp subsection{*Rules for Set Difference, Defined via @{term Upair}*} -lemma Diff_iff [simp]: "c \ A-B \ (c:A & c\B)" +lemma Diff_iff [simp]: "c \ A-B \ (c \ A & c\B)" by (unfold Diff_def, blast) lemma DiffI [intro!]: "[| c \ A; c \ B |] ==> c \ A - B" @@ -95,13 +95,13 @@ lemma DiffD2: "c \ A - B ==> c \ B" by simp -lemma DiffE [elim!]: "[| c \ A - B; [| c:A; c\B |] ==> P |] ==> P" +lemma DiffE [elim!]: "[| c \ A - B; [| c \ A; c\B |] ==> P |] ==> P" by simp subsection{*Rules for @{term cons}*} -lemma cons_iff [simp]: "a \ cons(b,A) \ (a=b | a:A)" +lemma cons_iff [simp]: "a \ cons(b,A) \ (a=b | a \ A)" apply (unfold cons_def) apply (blast intro: UpairI1 UpairI2 elim: UpairE) done @@ -115,16 +115,16 @@ lemma consI2: "a \ B ==> a \ cons(b,B)" by simp -lemma consE [elim!]: "[| a \ cons(b,A); a=b ==> P; a:A ==> P |] ==> P" +lemma consE [elim!]: "[| a \ cons(b,A); a=b ==> P; a \ A ==> P |] ==> P" by (simp, blast) (*Stronger version of the rule above*) lemma consE': - "[| a \ cons(b,A); a=b ==> P; [| a:A; a\b |] ==> P |] ==> P" + "[| a \ cons(b,A); a=b ==> P; [| a \ A; a\b |] ==> P |] ==> P" by (simp, blast) (*Classical introduction rule*) -lemma consCI [intro!]: "(a\B ==> a=b) ==> a: cons(b,B)" +lemma consCI [intro!]: "(a\B ==> a=b) ==> a \ cons(b,B)" by (simp, blast) lemma cons_not_0 [simp]: "cons(a,B) \ 0" @@ -207,7 +207,7 @@ ==> (if P then a else b) = (if Q then c else d)" by (simp add: if_def cong add: conj_cong) -(*Prevents simplification of x and y: faster and allows the execution +(*Prevents simplification of x and y \ faster and allows the execution of functional programs. NOW THE DEFAULT.*) lemma if_weak_cong: "P\Q ==> (if P then x else y) = (if Q then x else y)" by simp @@ -236,11 +236,11 @@ lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 (*Logically equivalent to split_if_mem2*) -lemma if_iff: "a: (if P then x else y) \ P & a:x | ~P & a:y" +lemma if_iff: "a: (if P then x else y) \ P & a \ x | ~P & a \ y" by simp lemma if_type [TC]: - "[| P ==> a: A; ~P ==> b: A |] ==> (if P then a else b): A" + "[| P ==> a \ A; ~P ==> b \ A |] ==> (if P then a else b): A" by simp (** Splitting IFs in the assumptions **) @@ -254,14 +254,14 @@ subsection{*Consequences of Foundation*} (*was called mem_anti_sym*) -lemma mem_asym: "[| a:b; ~P ==> b:a |] ==> P" +lemma mem_asym: "[| a \ b; ~P ==> b \ a |] ==> P" apply (rule classical) apply (rule_tac A1 = "{a,b}" in foundation [THEN disjE]) apply (blast elim!: equalityE)+ done (*was called mem_anti_refl*) -lemma mem_irrefl: "a:a ==> P" +lemma mem_irrefl: "a \ a ==> P" by (blast intro: mem_asym) (*mem_irrefl should NOT be added to default databases: @@ -273,7 +273,7 @@ done (*Good for proving inequalities by rewriting*) -lemma mem_imp_not_eq: "a:A ==> a \ A" +lemma mem_imp_not_eq: "a \ A ==> a \ A" by (blast elim!: mem_irrefl) lemma eq_imp_not_mem: "a=A ==> a \ A" @@ -281,7 +281,7 @@ subsection{*Rules for Successor*} -lemma succ_iff: "i \ succ(j) \ i=j | i:j" +lemma succ_iff: "i \ succ(j) \ i=j | i \ j" by (unfold succ_def, blast) lemma succI1 [simp]: "i \ succ(i)" @@ -291,12 +291,12 @@ by (simp add: succ_iff) lemma succE [elim!]: - "[| i \ succ(j); i=j ==> P; i:j ==> P |] ==> P" + "[| i \ succ(j); i=j ==> P; i \ j ==> P |] ==> P" apply (simp add: succ_iff, blast) done (*Classical introduction rule*) -lemma succCI [intro!]: "(i\j ==> i=j) ==> i: succ(j)" +lemma succCI [intro!]: "(i\j ==> i=j) ==> i \ succ(j)" by (simp add: succ_iff, blast) lemma succ_not_0 [simp]: "succ(n) \ 0" @@ -383,22 +383,22 @@ (** One-point rule for bounded quantifiers: see HOL/Set.ML **) -lemma bex_triv_one_point1 [simp]: "(\x\A. x=a) \ (a:A)" +lemma bex_triv_one_point1 [simp]: "(\x\A. x=a) \ (a \ A)" by blast -lemma bex_triv_one_point2 [simp]: "(\x\A. a=x) \ (a:A)" +lemma bex_triv_one_point2 [simp]: "(\x\A. a=x) \ (a \ A)" by blast -lemma bex_one_point1 [simp]: "(\x\A. x=a & P(x)) \ (a:A & P(a))" +lemma bex_one_point1 [simp]: "(\x\A. x=a & P(x)) \ (a \ A & P(a))" by blast -lemma bex_one_point2 [simp]: "(\x\A. a=x & P(x)) \ (a:A & P(a))" +lemma bex_one_point2 [simp]: "(\x\A. a=x & P(x)) \ (a \ A & P(a))" by blast -lemma ball_one_point1 [simp]: "(\x\A. x=a \ P(x)) \ (a:A \ P(a))" +lemma ball_one_point1 [simp]: "(\x\A. x=a \ P(x)) \ (a \ A \ P(a))" by blast -lemma ball_one_point2 [simp]: "(\x\A. a=x \ P(x)) \ (a:A \ P(a))" +lemma ball_one_point2 [simp]: "(\x\A. a=x \ P(x)) \ (a \ A \ P(a))" by blast @@ -406,9 +406,9 @@ text{*These cover both @{term Replace} and @{term Collect}*} lemma Rep_simps [simp]: - "{x. y:0, R(x,y)} = 0" - "{x:0. P(x)} = 0" - "{x:A. Q} = (if Q then A else 0)" + "{x. y \ 0, R(x,y)} = 0" + "{x \ 0. P(x)} = 0" + "{x \ A. Q} = (if Q then A else 0)" "RepFun(0,f) = 0" "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))" "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))"