--- a/src/ZF/AC.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/AC.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,9 +9,9 @@
text{*This definition comes from Halmos (1960), page 59.*}
axiomatization where
- AC: "[| a: A; !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"
+ AC: "[| a: A; !!x. x:A ==> (\<exists>y. y:B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
-(*The same as AC, but no premise a \<in> A*)
+(*The same as AC, but no premise @{term"a \<in> A"}*)
lemma AC_Pi: "[| !!x. x \<in> A ==> (\<exists>y. y \<in> B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
apply (case_tac "A=0")
apply (simp add: Pi_empty1)
@@ -31,17 +31,17 @@
done
lemma AC_func:
- "[| !!x. x \<in> A ==> (\<exists>y. y \<in> x) |] ==> \<exists>f \<in> A->Union(A). \<forall>x \<in> A. f`x \<in> x"
+ "[| !!x. x \<in> A ==> (\<exists>y. y \<in> x) |] ==> \<exists>f \<in> A->\<Union>(A). \<forall>x \<in> A. f`x \<in> x"
apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE])
-prefer 2 apply (blast dest: apply_type intro: Pi_type, blast)
+prefer 2 apply (blast dest: apply_type intro: Pi_type, blast)
done
lemma non_empty_family: "[| 0 \<notin> A; x \<in> A |] ==> \<exists>y. y \<in> x"
by (subgoal_tac "x \<noteq> 0", blast+)
-lemma AC_func0: "0 \<notin> A ==> \<exists>f \<in> A->Union(A). \<forall>x \<in> A. f`x \<in> x"
+lemma AC_func0: "0 \<notin> A ==> \<exists>f \<in> A->\<Union>(A). \<forall>x \<in> A. f`x \<in> x"
apply (rule AC_func)
-apply (simp_all add: non_empty_family)
+apply (simp_all add: non_empty_family)
done
lemma AC_func_Pow: "\<exists>f \<in> (Pow(C)-{0}) -> C. \<forall>x \<in> Pow(C)-{0}. f`x \<in> x"
@@ -53,7 +53,7 @@
lemma AC_Pi0: "0 \<notin> A ==> \<exists>f. f \<in> (\<Pi> x \<in> A. x)"
apply (rule AC_Pi)
-apply (simp_all add: non_empty_family)
+apply (simp_all add: non_empty_family)
done
end
--- a/src/ZF/Arith.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Arith.thy Tue Mar 06 15:15:49 2012 +0000
@@ -6,10 +6,10 @@
(*"Difference" is subtraction of natural numbers.
There are no negative numbers; we have
m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n.
- Also, rec(m, 0, %z w.z) is pred(m).
+ Also, rec(m, 0, %z w.z) is pred(m).
*)
-header{*Arithmetic Operators and Their Definitions*}
+header{*Arithmetic Operators and Their Definitions*}
theory Arith imports Univ begin
@@ -87,7 +87,7 @@
apply (induct_tac "k", auto)
done
-(* [| 0 < k; k \<in> nat; !!j. [| j \<in> nat; k = succ(j) |] ==> Q |] ==> Q *)
+(* @{term"[| 0 < k; k \<in> nat; !!j. [| j \<in> nat; k = succ(j) |] ==> Q |] ==> Q"} *)
lemmas zero_lt_natE = zero_lt_lemma [THEN bexE]
@@ -102,7 +102,7 @@
lemma natify_0 [simp]: "natify(0) = 0"
by (rule natify_def [THEN def_Vrecursor, THEN trans], auto)
-lemma natify_non_succ: "\<forall>z. x ~= succ(z) ==> natify(x) = 0"
+lemma natify_non_succ: "\<forall>z. x \<noteq> succ(z) ==> natify(x) = 0"
by (rule natify_def [THEN def_Vrecursor, THEN trans], auto)
lemma natify_in_nat [iff,TC]: "natify(x) \<in> nat"
@@ -214,8 +214,8 @@
lemma diff_0 [simp]: "m #- 0 = natify(m)"
by (simp add: diff_def)
-lemma diff_le_self: "m\<in>nat ==> (m #- n) le m"
-apply (subgoal_tac " (m #- natify (n)) le m")
+lemma diff_le_self: "m\<in>nat ==> (m #- n) \<le> m"
+apply (subgoal_tac " (m #- natify (n)) \<le> m")
apply (rule_tac [2] m = m and n = "natify (n) " in diff_induct)
apply (erule_tac [6] leE)
apply (simp_all add: le_iff)
@@ -293,13 +293,13 @@
by (force dest!: add_left_cancel_natify)
(*Thanks to Sten Agerholm*)
-lemma add_le_elim1_natify: "k#+m le k#+n ==> natify(m) le natify(n)"
-apply (rule_tac P = "natify(k) #+m le natify(k) #+n" in rev_mp)
+lemma add_le_elim1_natify: "k#+m \<le> k#+n ==> natify(m) \<le> natify(n)"
+apply (rule_tac P = "natify(k) #+m \<le> natify(k) #+n" in rev_mp)
apply (rule_tac [2] n = "natify(k) " in nat_induct)
apply auto
done
-lemma add_le_elim1: "[| k#+m le k#+n; m \<in> nat; n \<in> nat |] ==> m le n"
+lemma add_le_elim1: "[| k#+m \<le> k#+n; m \<in> nat; n \<in> nat |] ==> m \<le> n"
by (drule add_le_elim1_natify, auto)
lemma add_lt_elim1_natify: "k#+m < k#+n ==> natify(m) < natify(n)"
@@ -334,21 +334,21 @@
lemma Ord_lt_mono_imp_le_mono:
assumes lt_mono: "!!i j. [| i<j; j:k |] ==> f(i) < f(j)"
and ford: "!!i. i:k ==> Ord(f(i))"
- and leij: "i le j"
+ and leij: "i \<le> j"
and jink: "j:k"
- shows "f(i) le f(j)"
-apply (insert leij jink)
+ shows "f(i) \<le> f(j)"
+apply (insert leij jink)
apply (blast intro!: leCI lt_mono ford elim!: leE)
done
text{*@{text "\<le>"} monotonicity, 1st argument*}
-lemma add_le_mono1: "[| i le j; j\<in>nat |] ==> i#+k le j#+k"
-apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck)
+lemma add_le_mono1: "[| i \<le> j; j\<in>nat |] ==> i#+k \<le> j#+k"
+apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck)
apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+
done
text{*@{text "\<le>"} monotonicity, both arguments*}
-lemma add_le_mono: "[| i le j; k le l; j\<in>nat; l\<in>nat |] ==> i#+k le j#+l"
+lemma add_le_mono: "[| i \<le> j; k \<le> l; j\<in>nat; l\<in>nat |] ==> i#+k \<le> j#+l"
apply (rule add_le_mono1 [THEN le_trans], assumption+)
apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+)
done
@@ -365,8 +365,8 @@
text{*Less-than: in other words, strict in both arguments*}
lemma add_lt_mono: "[| i<j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
-apply (rule add_lt_le_mono)
-apply (auto intro: leI)
+apply (rule add_lt_le_mono)
+apply (auto intro: leI)
done
(** Subtraction is the inverse of addition. **)
@@ -400,12 +400,12 @@
by (simp add: pred_def)
lemma eq_succ_imp_eq_m1: "[|i = succ(j); i\<in>nat|] ==> j = i #- 1 & j \<in>nat"
-by simp
+by simp
lemma pred_Un_distrib:
- "[|i\<in>nat; j\<in>nat|] ==> pred(i Un j) = pred(i) Un pred(j)"
-apply (erule_tac n=i in natE, simp)
-apply (erule_tac n=j in natE, simp)
+ "[|i\<in>nat; j\<in>nat|] ==> pred(i \<union> j) = pred(i) \<union> pred(j)"
+apply (erule_tac n=i in natE, simp)
+apply (erule_tac n=j in natE, simp)
apply (simp add: succ_Un_distrib [symmetric])
done
@@ -414,23 +414,23 @@
by (simp add: pred_def split: split_nat_case)
lemma nat_diff_pred: "[|i\<in>nat; j\<in>nat|] ==> i #- succ(j) = pred(i #- j)";
-apply (rule_tac m=i and n=j in diff_induct)
+apply (rule_tac m=i and n=j in diff_induct)
apply (auto simp add: pred_def nat_imp_quasinat split: split_nat_case)
done
lemma diff_succ_eq_pred: "i #- succ(j) = pred(i #- j)";
apply (insert nat_diff_pred [of "natify(i)" "natify(j)"])
-apply (simp add: natify_succ [symmetric])
+apply (simp add: natify_succ [symmetric])
done
lemma nat_diff_Un_distrib:
- "[|i\<in>nat; j\<in>nat; k\<in>nat|] ==> (i Un j) #- k = (i#-k) Un (j#-k)"
-apply (rule_tac n=k in nat_induct)
-apply (simp_all add: diff_succ_eq_pred pred_Un_distrib)
+ "[|i\<in>nat; j\<in>nat; k\<in>nat|] ==> (i \<union> j) #- k = (i#-k) \<union> (j#-k)"
+apply (rule_tac n=k in nat_induct)
+apply (simp_all add: diff_succ_eq_pred pred_Un_distrib)
done
lemma diff_Un_distrib:
- "[|i\<in>nat; j\<in>nat|] ==> (i Un j) #- k = (i#-k) Un (j#-k)"
+ "[|i\<in>nat; j\<in>nat|] ==> (i \<union> j) #- k = (i#-k) \<union> (j#-k)"
by (insert nat_diff_Un_distrib [of i j "natify(k)"], simp)
text{*We actually prove @{term "i #- j #- k = i #- (j #+ k)"}*}
--- a/src/ZF/ArithSimp.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/ArithSimp.thy Tue Mar 06 15:15:49 2012 +0000
@@ -5,7 +5,7 @@
header{*Arithmetic with simplification*}
-theory ArithSimp
+theory ArithSimp
imports Arith
uses "~~/src/Provers/Arith/cancel_numerals.ML"
"~~/src/Provers/Arith/combine_numerals.ML"
@@ -22,20 +22,20 @@
(**Addition is the inverse of subtraction**)
(*We need m:nat even if we replace the RHS by natify(m), for consider e.g.
- n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 ~= 0 = natify(m).*)
-lemma add_diff_inverse: "[| n le m; m:nat |] ==> n #+ (m#-n) = m"
+ n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 \<noteq> 0 = natify(m).*)
+lemma add_diff_inverse: "[| n \<le> m; m:nat |] ==> n #+ (m#-n) = m"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (erule rev_mp)
apply (rule_tac m = m and n = n in diff_induct, auto)
done
-lemma add_diff_inverse2: "[| n le m; m:nat |] ==> (m#-n) #+ n = m"
+lemma add_diff_inverse2: "[| n \<le> m; m:nat |] ==> (m#-n) #+ n = m"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (simp (no_asm_simp) add: add_commute add_diff_inverse)
done
(*Proof is IDENTICAL to that of add_diff_inverse*)
-lemma diff_succ: "[| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"
+lemma diff_succ: "[| n \<le> m; m:nat |] ==> succ(m) #- n = succ(m#-n)"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (erule rev_mp)
apply (rule_tac m = m and n = n in diff_induct)
@@ -65,7 +65,7 @@
subsection{*Remainder*}
(*We need m:nat even with natify*)
-lemma div_termination: "[| 0<n; n le m; m:nat |] ==> m #- n < m"
+lemma div_termination: "[| 0<n; n \<le> m; m:nat |] ==> m #- n < m"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (erule rev_mp)
apply (erule rev_mp)
@@ -74,25 +74,25 @@
done
(*for mod and div*)
-lemmas div_rls =
- nat_typechecks Ord_transrec_type apply_funtype
+lemmas div_rls =
+ nat_typechecks Ord_transrec_type apply_funtype
div_termination [THEN ltD]
nat_into_Ord not_lt_iff_le [THEN iffD1]
-lemma raw_mod_type: "[| m:nat; n:nat |] ==> raw_mod (m, n) : nat"
+lemma raw_mod_type: "[| m:nat; n:nat |] ==> raw_mod (m, n) \<in> nat"
apply (unfold raw_mod_def)
apply (rule Ord_transrec_type)
apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff])
-apply (blast intro: div_rls)
+apply (blast intro: div_rls)
done
-lemma mod_type [TC,iff]: "m mod n : nat"
+lemma mod_type [TC,iff]: "m mod n \<in> nat"
apply (unfold mod_def)
apply (simp (no_asm) add: mod_def raw_mod_type)
done
-(** Aribtrary definitions for division by zero. Useful to simplify
+(** Aribtrary definitions for division by zero. Useful to simplify
certain equations **)
lemma DIVISION_BY_ZERO_DIV: "a div 0 = 0"
@@ -112,20 +112,20 @@
apply (simp (no_asm_simp) add: div_termination [THEN ltD])
done
-lemma mod_less [simp]: "[| m<n; n : nat |] ==> m mod n = m"
+lemma mod_less [simp]: "[| m<n; n \<in> nat |] ==> m mod n = m"
apply (frule lt_nat_in_nat, assumption)
apply (simp (no_asm_simp) add: mod_def raw_mod_less)
done
lemma raw_mod_geq:
- "[| 0<n; n le m; m:nat |] ==> raw_mod (m, n) = raw_mod (m#-n, n)"
+ "[| 0<n; n \<le> m; m:nat |] ==> raw_mod (m, n) = raw_mod (m#-n, n)"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (rule raw_mod_def [THEN def_transrec, THEN trans])
apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2], blast)
done
-lemma mod_geq: "[| n le m; m:nat |] ==> m mod n = (m#-n) mod n"
+lemma mod_geq: "[| n \<le> m; m:nat |] ==> m mod n = (m#-n) mod n"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (case_tac "n=0")
apply (simp add: DIVISION_BY_ZERO_MOD)
@@ -135,14 +135,14 @@
subsection{*Division*}
-lemma raw_div_type: "[| m:nat; n:nat |] ==> raw_div (m, n) : nat"
+lemma raw_div_type: "[| m:nat; n:nat |] ==> raw_div (m, n) \<in> nat"
apply (unfold raw_div_def)
apply (rule Ord_transrec_type)
apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff])
-apply (blast intro: div_rls)
+apply (blast intro: div_rls)
done
-lemma div_type [TC,iff]: "m div n : nat"
+lemma div_type [TC,iff]: "m div n \<in> nat"
apply (unfold div_def)
apply (simp (no_asm) add: div_def raw_div_type)
done
@@ -152,21 +152,21 @@
apply (simp (no_asm_simp) add: div_termination [THEN ltD])
done
-lemma div_less [simp]: "[| m<n; n : nat |] ==> m div n = 0"
+lemma div_less [simp]: "[| m<n; n \<in> nat |] ==> m div n = 0"
apply (frule lt_nat_in_nat, assumption)
apply (simp (no_asm_simp) add: div_def raw_div_less)
done
-lemma raw_div_geq: "[| 0<n; n le m; m:nat |] ==> raw_div(m,n) = succ(raw_div(m#-n, n))"
-apply (subgoal_tac "n ~= 0")
+lemma raw_div_geq: "[| 0<n; n \<le> m; m:nat |] ==> raw_div(m,n) = succ(raw_div(m#-n, n))"
+apply (subgoal_tac "n \<noteq> 0")
prefer 2 apply blast
apply (frule lt_nat_in_nat, erule nat_succI)
apply (rule raw_div_def [THEN def_transrec, THEN trans])
-apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2] )
+apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2] )
done
lemma div_geq [simp]:
- "[| 0<n; n le m; m:nat |] ==> m div n = succ ((m#-n) div n)"
+ "[| 0<n; n \<le> m; m:nat |] ==> m div n = succ ((m#-n) div n)"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (simp (no_asm_simp) add: div_def raw_div_geq)
done
@@ -183,13 +183,13 @@
apply (case_tac "x<n")
txt{*case x<n*}
apply (simp (no_asm_simp))
-txt{*case n le x*}
+txt{*case @{term"n \<le> x"}*}
apply (simp add: not_lt_iff_le add_assoc mod_geq div_termination [THEN ltD] add_diff_inverse)
done
lemma mod_div_equality_natify: "(m div n)#*n #+ m mod n = natify(m)"
apply (subgoal_tac " (natify (m) div natify (n))#*natify (n) #+ natify (m) mod natify (n) = natify (m) ")
-apply force
+apply force
apply (subst mod_div_lemma, auto)
done
@@ -203,14 +203,14 @@
text{*(mainly for mutilated chess board)*}
lemma mod_succ_lemma:
- "[| 0<n; m:nat; n:nat |]
+ "[| 0<n; m:nat; n:nat |]
==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"
apply (erule complete_induct)
apply (case_tac "succ (x) <n")
txt{* case succ(x) < n *}
apply (simp (no_asm_simp) add: nat_le_refl [THEN lt_trans] succ_neq_self)
apply (simp add: ltD [THEN mem_imp_not_eq])
-txt{* case n le succ(x) *}
+txt{* case @{term"n \<le> succ(x)"} *}
apply (simp add: mod_geq not_lt_iff_le)
apply (erule leE)
apply (simp (no_asm_simp) add: mod_geq div_termination [THEN ltD] diff_succ)
@@ -232,8 +232,8 @@
lemma mod_less_divisor: "[| 0<n; n:nat |] ==> m mod n < n"
apply (subgoal_tac "natify (m) mod n < n")
apply (rule_tac [2] i = "natify (m) " in complete_induct)
-apply (case_tac [3] "x<n", auto)
-txt{* case n le x*}
+apply (case_tac [3] "x<n", auto)
+txt{* case @{term"n \<le> x"}*}
apply (simp add: mod_geq not_lt_iff_le div_termination [THEN ltD])
done
@@ -264,25 +264,25 @@
subsection{*Additional theorems about @{text "\<le>"}*}
-lemma add_le_self: "m:nat ==> m le (m #+ n)"
+lemma add_le_self: "m:nat ==> m \<le> (m #+ n)"
apply (simp (no_asm_simp))
done
-lemma add_le_self2: "m:nat ==> m le (n #+ m)"
+lemma add_le_self2: "m:nat ==> m \<le> (n #+ m)"
apply (simp (no_asm_simp))
done
(*** Monotonicity of Multiplication ***)
-lemma mult_le_mono1: "[| i le j; j:nat |] ==> (i#*k) le (j#*k)"
-apply (subgoal_tac "natify (i) #*natify (k) le j#*natify (k) ")
+lemma mult_le_mono1: "[| i \<le> j; j:nat |] ==> (i#*k) \<le> (j#*k)"
+apply (subgoal_tac "natify (i) #*natify (k) \<le> j#*natify (k) ")
apply (frule_tac [2] lt_nat_in_nat)
apply (rule_tac [3] n = "natify (k) " in nat_induct)
apply (simp_all add: add_le_mono)
done
-(* le monotonicity, BOTH arguments*)
-lemma mult_le_mono: "[| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l"
+(* @{text"\<le>"} monotonicity, BOTH arguments*)
+lemma mult_le_mono: "[| i \<le> j; k \<le> l; j:nat; l:nat |] ==> i#*k \<le> j#*l"
apply (rule mult_le_mono1 [THEN le_trans], assumption+)
apply (subst mult_commute, subst mult_commute, rule mult_le_mono1, assumption+)
done
@@ -359,20 +359,20 @@
apply (simp (no_asm) add: mult_less_cancel2 mult_commute [of k])
done
-lemma mult_le_cancel2 [simp]: "(m#*k le n#*k) <-> (0 < natify(k) --> natify(m) le natify(n))"
+lemma mult_le_cancel2 [simp]: "(m#*k \<le> n#*k) <-> (0 < natify(k) \<longrightarrow> natify(m) \<le> natify(n))"
apply (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym])
apply auto
done
-lemma mult_le_cancel1 [simp]: "(k#*m le k#*n) <-> (0 < natify(k) --> natify(m) le natify(n))"
+lemma mult_le_cancel1 [simp]: "(k#*m \<le> k#*n) <-> (0 < natify(k) \<longrightarrow> natify(m) \<le> natify(n))"
apply (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym])
apply auto
done
-lemma mult_le_cancel_le1: "k : nat ==> k #* m le k \<longleftrightarrow> (0 < k \<longrightarrow> natify(m) le 1)"
+lemma mult_le_cancel_le1: "k \<in> nat ==> k #* m \<le> k \<longleftrightarrow> (0 < k \<longrightarrow> natify(m) \<le> 1)"
by (cut_tac k = k and m = m and n = 1 in mult_le_cancel1, auto)
-lemma Ord_eq_iff_le: "[| Ord(m); Ord(n) |] ==> m=n <-> (m le n & n le m)"
+lemma Ord_eq_iff_le: "[| Ord(m); Ord(n) |] ==> m=n <-> (m \<le> n & n \<le> m)"
by (blast intro: le_anti_sym)
lemma mult_cancel2_lemma:
@@ -406,7 +406,7 @@
lemma div_cancel:
"[| 0 < natify(n); 0 < natify(k) |] ==> (k#*m) div (k#*n) = m div n"
-apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)"
+apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)"
in div_cancel_raw)
apply auto
done
@@ -424,12 +424,12 @@
apply (erule_tac i = m in complete_induct)
apply (case_tac "x<n")
apply (simp (no_asm_simp) add: mod_less zero_lt_mult_iff mult_lt_mono2)
-apply (simp add: not_lt_iff_le zero_lt_mult_iff le_refl [THEN mult_le_mono]
+apply (simp add: not_lt_iff_le zero_lt_mult_iff le_refl [THEN mult_le_mono]
mod_geq diff_mult_distrib2 [symmetric] div_termination [THEN ltD])
done
lemma mod_mult_distrib2: "k #* (m mod n) = (k#*m) mod (k#*n)"
-apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)"
+apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)"
in mult_mod_distrib_raw)
apply auto
done
@@ -440,8 +440,8 @@
lemma mod_add_self2_raw: "n \<in> nat ==> (m #+ n) mod n = m mod n"
apply (subgoal_tac " (n #+ m) mod n = (n #+ m #- n) mod n")
-apply (simp add: add_commute)
-apply (subst mod_geq [symmetric], auto)
+apply (simp add: add_commute)
+apply (subst mod_geq [symmetric], auto)
done
lemma mod_add_self2 [simp]: "(m #+ n) mod n = m mod n"
@@ -470,21 +470,21 @@
(*Lemma for gcd*)
lemma mult_eq_self_implies_10: "m = m#*n ==> natify(n)=1 | m=0"
apply (subgoal_tac "m: nat")
- prefer 2
+ prefer 2
apply (erule ssubst)
- apply simp
+ apply simp
apply (rule disjCI)
apply (drule sym)
apply (rule Ord_linear_lt [of "natify(n)" 1])
-apply simp_all
- apply (subgoal_tac "m #* n = 0", simp)
+apply simp_all
+ apply (subgoal_tac "m #* n = 0", simp)
apply (subst mult_natify2 [symmetric])
apply (simp del: mult_natify2)
apply (drule nat_into_Ord [THEN Ord_0_lt, THEN [2] mult_lt_mono2], auto)
done
lemma less_imp_succ_add [rule_format]:
- "[| m<n; n: nat |] ==> EX k: nat. n = succ(m#+k)"
+ "[| m<n; n: nat |] ==> \<exists>k\<in>nat. n = succ(m#+k)"
apply (frule lt_nat_in_nat, assumption)
apply (erule rev_mp)
apply (induct_tac "n")
@@ -493,45 +493,45 @@
done
lemma less_iff_succ_add:
- "[| m: nat; n: nat |] ==> (m<n) <-> (EX k: nat. n = succ(m#+k))"
+ "[| m: nat; n: nat |] ==> (m<n) <-> (\<exists>k\<in>nat. n = succ(m#+k))"
by (auto intro: less_imp_succ_add)
lemma add_lt_elim2:
"\<lbrakk>a #+ d = b #+ c; a < b; b \<in> nat; c \<in> nat; d \<in> nat\<rbrakk> \<Longrightarrow> c < d"
-by (drule less_imp_succ_add, auto)
+by (drule less_imp_succ_add, auto)
lemma add_le_elim2:
- "\<lbrakk>a #+ d = b #+ c; a le b; b \<in> nat; c \<in> nat; d \<in> nat\<rbrakk> \<Longrightarrow> c le d"
-by (drule less_imp_succ_add, auto)
+ "\<lbrakk>a #+ d = b #+ c; a \<le> b; b \<in> nat; c \<in> nat; d \<in> nat\<rbrakk> \<Longrightarrow> c \<le> d"
+by (drule less_imp_succ_add, auto)
subsubsection{*More Lemmas About Difference*}
lemma diff_is_0_lemma:
- "[| m: nat; n: nat |] ==> m #- n = 0 <-> m le n"
+ "[| m: nat; n: nat |] ==> m #- n = 0 <-> m \<le> n"
apply (rule_tac m = m and n = n in diff_induct, simp_all)
done
-lemma diff_is_0_iff: "m #- n = 0 <-> natify(m) le natify(n)"
+lemma diff_is_0_iff: "m #- n = 0 <-> natify(m) \<le> natify(n)"
by (simp add: diff_is_0_lemma [symmetric])
lemma nat_lt_imp_diff_eq_0:
"[| a:nat; b:nat; a<b |] ==> a #- b = 0"
-by (simp add: diff_is_0_iff le_iff)
+by (simp add: diff_is_0_iff le_iff)
lemma raw_nat_diff_split:
- "[| a:nat; b:nat |] ==>
- (P(a #- b)) <-> ((a < b -->P(0)) & (ALL d:nat. a = b #+ d --> P(d)))"
+ "[| a:nat; b:nat |] ==>
+ (P(a #- b)) <-> ((a < b \<longrightarrow>P(0)) & (\<forall>d\<in>nat. a = b #+ d \<longrightarrow> P(d)))"
apply (case_tac "a < b")
apply (force simp add: nat_lt_imp_diff_eq_0)
-apply (rule iffI, force, simp)
+apply (rule iffI, force, simp)
apply (drule_tac x="a#-b" in bspec)
-apply (simp_all add: Ordinal.not_lt_iff_le add_diff_inverse)
+apply (simp_all add: Ordinal.not_lt_iff_le add_diff_inverse)
done
lemma nat_diff_split:
- "(P(a #- b)) <->
- (natify(a) < natify(b) -->P(0)) & (ALL d:nat. natify(a) = b #+ d --> P(d))"
+ "(P(a #- b)) <->
+ (natify(a) < natify(b) \<longrightarrow>P(0)) & (\<forall>d\<in>nat. natify(a) = b #+ d \<longrightarrow> P(d))"
apply (cut_tac P=P and a="natify(a)" and b="natify(b)" in raw_nat_diff_split)
apply simp_all
done
@@ -544,10 +544,10 @@
apply (blast intro: add_le_self lt_trans1)
apply (rule not_le_iff_lt [THEN iffD1], auto)
apply (subgoal_tac "i #+ da < j #+ d", force)
-apply (blast intro: add_le_lt_mono)
+apply (blast intro: add_le_lt_mono)
done
-lemma lt_imp_diff_lt: "[|j<i; i\<le>k; k\<in>nat|] ==> (k#-i) < (k#-j)"
+lemma lt_imp_diff_lt: "[|j<i; i\<le>k; k\<in>nat|] ==> (k#-i) < (k#-j)"
apply (frule le_in_nat, assumption)
apply (frule lt_nat_in_nat, assumption)
apply (simp split add: nat_diff_split, auto)
@@ -555,13 +555,13 @@
apply (blast intro: lt_irrefl lt_trans2)
apply (rule not_le_iff_lt [THEN iffD1], auto)
apply (subgoal_tac "j #+ d < i #+ da", force)
-apply (blast intro: add_lt_le_mono)
+apply (blast intro: add_lt_le_mono)
done
lemma diff_lt_iff_lt: "[|i\<le>k; j\<in>nat; k\<in>nat|] ==> (k#-i) < (k#-j) <-> j<i"
apply (frule le_in_nat, assumption)
-apply (blast intro: lt_imp_diff_lt diff_lt_imp_lt)
+apply (blast intro: lt_imp_diff_lt diff_lt_imp_lt)
done
end
--- a/src/ZF/Bin.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Bin.thy Tue Mar 06 15:15:49 2012 +0000
@@ -68,14 +68,14 @@
primrec (*sum*)
bin_adder_Pls:
- "bin_adder (Pls) = (lam w:bin. w)"
+ "bin_adder (Pls) = (\<lambda>w\<in>bin. w)"
bin_adder_Min:
- "bin_adder (Min) = (lam w:bin. bin_pred(w))"
+ "bin_adder (Min) = (\<lambda>w\<in>bin. bin_pred(w))"
bin_adder_BIT:
- "bin_adder (v BIT x) =
- (lam w:bin.
- bin_case (v BIT x, bin_pred(v BIT x),
- %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w),
+ "bin_adder (v BIT x) =
+ (\<lambda>w\<in>bin.
+ bin_case (v BIT x, bin_pred(v BIT x),
+ %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w),
x xor y),
w))"
@@ -83,7 +83,7 @@
primrec
"adding (v,x,Pls) = v BIT x"
"adding (v,x,Min) = bin_pred(v BIT x)"
- "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)),
+ "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)),
x xor y)"
*)
@@ -125,33 +125,33 @@
lemma NCons_BIT: "NCons(w BIT x,b) = w BIT x BIT b"
by (simp add: bin.case_eqns)
-lemmas NCons_simps [simp] =
+lemmas NCons_simps [simp] =
NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT
(** Type checking **)
-lemma integ_of_type [TC]: "w: bin ==> integ_of(w) : int"
+lemma integ_of_type [TC]: "w: bin ==> integ_of(w) \<in> 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) \<in> 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) \<in> 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) \<in> 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) \<in> bin"
by (induct_tac "w", auto)
(*This proof is complicated by the mutual recursion*)
lemma bin_add_type [rule_format,TC]:
- "v: bin ==> ALL w: bin. bin_add(v,w) : bin"
+ "v: bin ==> \<forall>w\<in>bin. bin_add(v,w) \<in> bin"
apply (unfold bin_add_def)
apply (induct_tac "v")
apply (rule_tac [3] ballI)
@@ -160,30 +160,30 @@
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) \<in> bin"
by (induct_tac "v", auto)
-subsubsection{*The Carry and Borrow Functions,
+subsubsection{*The Carry and Borrow Functions,
@{term bin_succ} and @{term bin_pred}*}
(*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)"
apply (erule bin.cases)
-apply (auto elim!: boolE)
+apply (auto elim!: boolE)
done
lemma integ_of_succ [simp]:
"w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)"
apply (erule bin.induct)
-apply (auto simp add: zadd_ac elim!: boolE)
+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)"
apply (erule bin.induct)
-apply (auto simp add: zadd_ac elim!: boolE)
+apply (auto simp add: zadd_ac elim!: boolE)
done
@@ -191,7 +191,7 @@
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)
+apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE)
done
@@ -220,23 +220,23 @@
by (unfold bin_add_def, simp)
lemma bin_add_BIT_BIT [simp]:
- "[| w: bin; y: bool |]
- ==> bin_add(v BIT x, w BIT y) =
+ "[| 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 ==>
- ALL w: bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)"
+ "v: bin ==>
+ \<forall>w\<in>bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)"
apply (erule bin.induct, simp, simp)
apply (rule ballI)
apply (induct_tac "wa")
-apply (auto simp add: zadd_ac elim!: boolE)
+apply (auto simp add: zadd_ac elim!: boolE)
done
(*Subtraction*)
-lemma diff_integ_of_eq:
- "[| v: bin; w: bin |]
+lemma diff_integ_of_eq:
+ "[| 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,11 +246,11 @@
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)
-apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib elim!: boolE)
+apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib elim!: boolE)
done
@@ -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
@@ -344,9 +344,9 @@
(** Equals (=) **)
-lemma eq_integ_of_eq:
- "[| v: bin; w: bin |]
- ==> ((integ_of(v)) = integ_of(w)) <->
+lemma eq_integ_of_eq:
+ "[| v: bin; w: bin |]
+ ==> ((integ_of(v)) = integ_of(w)) <->
iszero (integ_of (bin_add (v, bin_minus(w))))"
apply (unfold iszero_def)
apply (simp add: zcompare_rls integ_of_add integ_of_minus)
@@ -361,11 +361,11 @@
apply (simp add: zminus_equation)
done
-lemma iszero_integ_of_BIT:
- "[| w: bin; x: bool |]
+lemma iszero_integ_of_BIT:
+ "[| 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")
+apply (subgoal_tac "integ_of (w) \<in> int")
apply typecheck
apply (drule int_cases)
apply (safe elim!: boolE)
@@ -375,7 +375,7 @@
lemma iszero_integ_of_0:
"w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))"
-by (simp only: iszero_integ_of_BIT, blast)
+by (simp only: iszero_integ_of_BIT, blast)
lemma iszero_integ_of_1: "w: bin ==> ~ iszero (integ_of (w BIT 1))"
by (simp only: iszero_integ_of_BIT, blast)
@@ -384,9 +384,9 @@
(** Less-than (<) **)
-lemma less_integ_of_eq_neg:
- "[| v: bin; w: bin |]
- ==> integ_of(v) $< integ_of(w)
+lemma less_integ_of_eq_neg:
+ "[| 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)
apply (simp add: integ_of_minus integ_of_add)
@@ -399,14 +399,14 @@
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")
+apply (subgoal_tac "integ_of (w) \<in> int")
apply typecheck
apply (drule int_cases)
apply (auto elim!: boolE simp add: int_of_add [symmetric] zcompare_rls)
-apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def
+apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def
int_of_add [symmetric])
apply (subgoal_tac "$#1 $- $# succ (succ (n #+ n)) = $- $# succ (n #+ n) ")
apply (simp add: zdiff_def)
@@ -421,8 +421,8 @@
(*Delete the original rewrites, with their clumsy conditional expressions*)
-declare bin_succ_BIT [simp del]
- bin_pred_BIT [simp del]
+declare bin_succ_BIT [simp del]
+ bin_pred_BIT [simp del]
bin_minus_BIT [simp del]
NCons_Pls [simp del]
NCons_Min [simp del]
@@ -434,12 +434,12 @@
lemmas bin_arith_extra_simps =
- integ_of_add [symmetric]
- integ_of_minus [symmetric]
- integ_of_mult [symmetric]
- bin_succ_1 bin_succ_0
- bin_pred_1 bin_pred_0
- bin_minus_1 bin_minus_0
+ integ_of_add [symmetric]
+ integ_of_minus [symmetric]
+ integ_of_mult [symmetric]
+ bin_succ_1 bin_succ_0
+ bin_pred_1 bin_pred_0
+ bin_minus_1 bin_minus_0
bin_add_Pls_right bin_add_Min_right
bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11
diff_integ_of_eq
@@ -453,7 +453,7 @@
bin_succ_Pls bin_succ_Min
bin_add_Pls bin_add_Min
bin_minus_Pls bin_minus_Min
- bin_mult_Pls bin_mult_Min
+ bin_mult_Pls bin_mult_Min
bin_arith_extra_simps
(*Simplification of relational operations*)
@@ -471,25 +471,25 @@
(** 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 |]
+lemma add_integ_of_diff1 [simp]:
+ "[| 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 |]
- ==> integ_of(v) $+ (c $- integ_of(w)) =
+ "[| 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])
apply (simp_all add: zdiff_def zadd_ac)
@@ -555,7 +555,7 @@
(** nat_of and zless **)
-(*An alternative condition is $#0 <= w *)
+(*An alternative condition is @{term"$#0 \<subseteq> w"} *)
lemma zless_nat_conj_lemma: "$#0 $< z ==> (nat_of(w) < nat_of(z)) <-> (w $< z)"
apply (rule iff_trans)
apply (rule zless_int_of [THEN iff_sym])
--- a/src/ZF/Bool.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Bool.thy Tue Mar 06 15:15:49 2012 +0000
@@ -43,13 +43,13 @@
(* Introduction rules *)
-lemma bool_1I [simp,TC]: "1 : bool"
+lemma bool_1I [simp,TC]: "1 \<in> bool"
by (simp add: bool_defs )
-lemma bool_0I [simp,TC]: "0 : bool"
+lemma bool_0I [simp,TC]: "0 \<in> bool"
by (simp add: bool_defs)
-lemma one_not_0: "1~=0"
+lemma one_not_0: "1\<noteq>0"
by (simp add: bool_defs )
(** 1=0 ==> R **)
@@ -94,16 +94,16 @@
lemmas xor_1 = xor_def [THEN def_cond_1, simp]
lemmas xor_0 = xor_def [THEN def_cond_0, simp]
-lemma not_type [TC]: "a:bool ==> not(a) : bool"
+lemma not_type [TC]: "a:bool ==> not(a) \<in> bool"
by (simp add: not_def)
-lemma and_type [TC]: "[| a:bool; b:bool |] ==> a and b : bool"
+lemma and_type [TC]: "[| a:bool; b:bool |] ==> a and b \<in> bool"
by (simp add: and_def)
-lemma or_type [TC]: "[| a:bool; b:bool |] ==> a or b : bool"
+lemma or_type [TC]: "[| a:bool; b:bool |] ==> a or b \<in> bool"
by (simp add: or_def)
-lemma xor_type [TC]: "[| a:bool; b:bool |] ==> a xor b : bool"
+lemma xor_type [TC]: "[| a:bool; b:bool |] ==> a xor b \<in> bool"
by (simp add: xor_def)
lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type
--- a/src/ZF/Cardinal.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Cardinal.thy Tue Mar 06 15:15:49 2012 +0000
@@ -10,15 +10,15 @@
definition
(*least ordinal operator*)
Least :: "(i=>o) => i" (binder "LEAST " 10) where
- "Least(P) == THE i. Ord(i) & P(i) & (ALL j. j<i --> ~P(j))"
+ "Least(P) == THE i. Ord(i) & P(i) & (\<forall>j. j<i \<longrightarrow> ~P(j))"
definition
eqpoll :: "[i,i] => o" (infixl "eqpoll" 50) where
- "A eqpoll B == EX f. f: bij(A,B)"
+ "A eqpoll B == \<exists>f. f: bij(A,B)"
definition
lepoll :: "[i,i] => o" (infixl "lepoll" 50) where
- "A lepoll B == EX f. f: inj(A,B)"
+ "A lepoll B == \<exists>f. f: inj(A,B)"
definition
lesspoll :: "[i,i] => o" (infixl "lesspoll" 50) where
@@ -30,7 +30,7 @@
definition
Finite :: "i=>o" where
- "Finite(A) == EX n:nat. A eqpoll n"
+ "Finite(A) == \<exists>n\<in>nat. A eqpoll n"
definition
Card :: "i=>o" where
@@ -57,31 +57,31 @@
lemma Banach_last_equation:
"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"
+ ==> 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"
in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst])
apply (simp add: double_complement fun_is_rel [THEN image_subset])
done
lemma decomposition:
- "[| f: X->Y; g: Y->X |] ==>
- EX XA XB YA YB. (XA Int XB = 0) & (XA Un XB = X) &
- (YA Int YB = 0) & (YA Un YB = Y) &
+ "[| f: X->Y; g: Y->X |] ==>
+ \<exists>XA XB YA YB. (XA \<inter> XB = 0) & (XA \<union> XB = X) &
+ (YA \<inter> YB = 0) & (YA \<union> YB = Y) &
f``XA=YA & g``YB=XB"
apply (intro exI conjI)
apply (rule_tac [6] Banach_last_equation)
apply (rule_tac [5] refl)
-apply (assumption |
+apply (assumption |
rule Diff_disjoint Diff_partition fun_is_rel image_subset lfp_subset)+
done
lemma schroeder_bernstein:
- "[| f: inj(X,Y); g: inj(Y,X) |] ==> EX h. h: bij(X,Y)"
-apply (insert decomposition [of f X Y g])
+ "[| f: inj(X,Y); g: inj(Y,X) |] ==> \<exists>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)
-(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))"
+(* The instantiation of exI to @{term"restrict(f,XA) \<union> converse(restrict(g,YB))"}
is forced by the context!! *)
done
@@ -101,7 +101,7 @@
apply (blast intro: bij_converse_bij)
done
-lemma eqpoll_trans:
+lemma eqpoll_trans:
"[| X \<approx> Y; Y \<approx> Z |] ==> X \<approx> Z"
apply (unfold eqpoll_def)
apply (blast intro: comp_bij)
@@ -136,7 +136,7 @@
lemma eqpollE:
"[| X \<approx> Y; [| X \<lesssim> Y; Y \<lesssim> X |] ==> P |] ==> P"
-by (blast intro: eqpoll_imp_lepoll eqpoll_sym)
+by (blast intro: eqpoll_imp_lepoll eqpoll_sym)
lemma eqpoll_iff: "X \<approx> Y <-> X \<lesssim> Y & Y \<lesssim> X"
by (blast intro: eqpollI elim!: eqpollE)
@@ -146,14 +146,14 @@
apply (blast dest: apply_type)
done
-(*0 \<lesssim> Y*)
+(*@{term"0 \<lesssim> Y"}*)
lemmas empty_lepollI = empty_subsetI [THEN subset_imp_lepoll]
lemma lepoll_0_iff: "A \<lesssim> 0 <-> A=0"
by (blast intro: lepoll_0_is_0 lepoll_refl)
-lemma Un_lepoll_Un:
- "[| A \<lesssim> B; C \<lesssim> D; B Int D = 0 |] ==> A Un C \<lesssim> B Un D"
+lemma Un_lepoll_Un:
+ "[| A \<lesssim> B; C \<lesssim> D; B \<inter> D = 0 |] ==> A \<union> C \<lesssim> B \<union> D"
apply (unfold lepoll_def)
apply (blast intro: inj_disjoint_Un)
done
@@ -164,9 +164,9 @@
lemma eqpoll_0_iff: "A \<approx> 0 <-> A=0"
by (blast intro: eqpoll_0_is_0 eqpoll_refl)
-lemma eqpoll_disjoint_Un:
- "[| A \<approx> B; C \<approx> D; A Int C = 0; B Int D = 0 |]
- ==> A Un C \<approx> B Un D"
+lemma eqpoll_disjoint_Un:
+ "[| A \<approx> B; C \<approx> D; A \<inter> C = 0; B \<inter> D = 0 |]
+ ==> A \<union> C \<approx> B \<union> D"
apply (unfold eqpoll_def)
apply (blast intro: bij_disjoint_Un)
done
@@ -175,15 +175,15 @@
subsection{*lesspoll: contributions by Krzysztof Grabczewski *}
lemma lesspoll_not_refl: "~ (i \<prec> i)"
-by (simp add: lesspoll_def)
+by (simp add: lesspoll_def)
lemma lesspoll_irrefl [elim!]: "i \<prec> i ==> P"
-by (simp add: lesspoll_def)
+by (simp add: lesspoll_def)
lemma lesspoll_imp_lepoll: "A \<prec> B ==> A \<lesssim> B"
by (unfold lesspoll_def, blast)
-lemma lepoll_well_ord: "[| A \<lesssim> B; well_ord(B,r) |] ==> EX s. well_ord(A,s)"
+lemma lepoll_well_ord: "[| A \<lesssim> B; well_ord(B,r) |] ==> \<exists>s. well_ord(A,s)"
apply (unfold lepoll_def)
apply (blast intro: well_ord_rvimage)
done
@@ -193,34 +193,34 @@
apply (blast intro!: eqpollI elim!: eqpollE)
done
-lemma inj_not_surj_succ:
- "[| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)"
-apply (unfold inj_def surj_def)
-apply (safe del: succE)
-apply (erule swap, rule exI)
-apply (rule_tac a = "lam z:A. if f`z=m then y else f`z" in CollectI)
+lemma inj_not_surj_succ:
+ "[| f \<in> inj(A, succ(m)); f \<notin> surj(A, succ(m)) |] ==> \<exists>f. f:inj(A,m)"
+apply (unfold inj_def surj_def)
+apply (safe del: succE)
+apply (erule swap, rule exI)
+apply (rule_tac a = "\<lambda>z\<in>A. if f`z=m then y else f`z" in CollectI)
txt{*the typing condition*}
apply (best intro!: if_type [THEN lam_type] elim: apply_funtype [THEN succE])
txt{*Proving it's injective*}
apply simp
-apply blast
+apply blast
done
(** Variations on transitivity **)
-lemma lesspoll_trans:
+lemma lesspoll_trans:
"[| X \<prec> Y; Y \<prec> Z |] ==> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
done
-lemma lesspoll_trans1:
+lemma lesspoll_trans1:
"[| X \<lesssim> Y; Y \<prec> Z |] ==> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
done
-lemma lesspoll_trans2:
+lemma lesspoll_trans2:
"[| X \<prec> Y; Y \<lesssim> Z |] ==> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
@@ -229,9 +229,9 @@
(** LEAST -- the least number operator [from HOL/Univ.ML] **)
-lemma Least_equality:
+lemma Least_equality:
"[| P(i); Ord(i); !!x. x<i ==> ~P(x) |] ==> (LEAST x. P(x)) = i"
-apply (unfold Least_def)
+apply (unfold Least_def)
apply (rule the_equality, blast)
apply (elim conjE)
apply (erule Ord_linear_lt, assumption, blast+)
@@ -239,16 +239,16 @@
lemma LeastI: "[| P(i); Ord(i) |] ==> P(LEAST x. P(x))"
apply (erule rev_mp)
-apply (erule_tac i=i in trans_induct)
+apply (erule_tac i=i in trans_induct)
apply (rule impI)
apply (rule classical)
apply (blast intro: Least_equality [THEN ssubst] elim!: ltE)
done
(*Proof is almost identical to the one above!*)
-lemma Least_le: "[| P(i); Ord(i) |] ==> (LEAST x. P(x)) le i"
+lemma Least_le: "[| P(i); Ord(i) |] ==> (LEAST x. P(x)) \<le> i"
apply (erule rev_mp)
-apply (erule_tac i=i in trans_induct)
+apply (erule_tac i=i in trans_induct)
apply (rule impI)
apply (rule classical)
apply (subst Least_equality, assumption+)
@@ -259,23 +259,23 @@
(*LEAST really is the smallest*)
lemma less_LeastE: "[| P(i); i < (LEAST x. P(x)) |] ==> Q"
apply (rule Least_le [THEN [2] lt_trans2, THEN lt_irrefl], assumption+)
-apply (simp add: lt_Ord)
+apply (simp add: lt_Ord)
done
(*Easier to apply than LeastI: conclusion has only one occurrence of P*)
lemma LeastI2:
"[| P(i); Ord(i); !!j. P(j) ==> Q(j) |] ==> Q(LEAST j. P(j))"
-by (blast intro: LeastI )
+by (blast intro: LeastI )
(*If there is no such P then LEAST is vacuously 0*)
-lemma Least_0:
- "[| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0"
+lemma Least_0:
+ "[| ~ (\<exists>i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0"
apply (unfold Least_def)
apply (rule the_0, blast)
done
lemma Ord_Least [intro,simp,TC]: "Ord(LEAST x. P(x))"
-apply (case_tac "\<exists>i. Ord(i) & P(i)")
+apply (case_tac "\<exists>i. Ord(i) & P(i)")
apply safe
apply (rule Least_le [THEN ltE])
prefer 3 apply assumption+
@@ -291,7 +291,7 @@
"(!!y. P(y) <-> Q(y)) ==> (LEAST x. P(x)) = (LEAST x. Q(x))"
by simp
-(*Need AC to get X \<lesssim> Y ==> |X| le |Y|; see well_ord_lepoll_imp_Card_le
+(*Need AC to get @{term"X \<lesssim> Y ==> |X| \<le> |Y|"}; see well_ord_lepoll_imp_Card_le
Converse also requires AC, but see well_ord_cardinal_eqE*)
lemma cardinal_cong: "X \<approx> Y ==> |X| = |Y|"
apply (unfold eqpoll_def cardinal_def)
@@ -300,7 +300,7 @@
done
(*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
-lemma well_ord_cardinal_eqpoll:
+lemma well_ord_cardinal_eqpoll:
"well_ord(A,r) ==> |A| \<approx> A"
apply (unfold cardinal_def)
apply (rule LeastI)
@@ -308,7 +308,7 @@
apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_imp_eqpoll])
done
-(* Ord(A) ==> |A| \<approx> A *)
+(* @{term"Ord(A) ==> |A| \<approx> A"} *)
lemmas Ord_cardinal_eqpoll = well_ord_Memrel [THEN well_ord_cardinal_eqpoll]
lemma well_ord_cardinal_eqE:
@@ -325,7 +325,7 @@
(** Observations from Kunen, page 28 **)
-lemma Ord_cardinal_le: "Ord(i) ==> |i| le i"
+lemma Ord_cardinal_le: "Ord(i) ==> |i| \<le> i"
apply (unfold cardinal_def)
apply (erule eqpoll_refl [THEN Least_le])
done
@@ -335,9 +335,9 @@
apply (erule sym)
done
-(* Could replace the ~(j \<approx> i) by ~(i \<lesssim> j) *)
+(* Could replace the @{term"~(j \<approx> i)"} by @{term"~(i \<lesssim> j)"}. *)
lemma CardI: "[| Ord(i); !!j. j<i ==> ~(j \<approx> i) |] ==> Card(i)"
-apply (unfold Card_def cardinal_def)
+apply (unfold Card_def cardinal_def)
apply (subst Least_equality)
apply (blast intro: eqpoll_refl )+
done
@@ -348,7 +348,7 @@
apply (rule Ord_Least)
done
-lemma Card_cardinal_le: "Card(K) ==> K le |K|"
+lemma Card_cardinal_le: "Card(K) ==> K \<le> |K|"
apply (simp (no_asm_simp) add: Card_is_Ord Card_cardinal_eq)
done
@@ -358,7 +358,7 @@
done
(*The cardinals are the initial ordinals*)
-lemma Card_iff_initial: "Card(K) <-> Ord(K) & (ALL j. j<K --> ~ j \<approx> K)"
+lemma Card_iff_initial: "Card(K) <-> Ord(K) & (\<forall>j. j<K \<longrightarrow> ~ j \<approx> K)"
apply (safe intro!: CardI Card_is_Ord)
prefer 2 apply blast
apply (unfold Card_def cardinal_def)
@@ -377,7 +377,7 @@
apply (blast elim!: ltE)
done
-lemma Card_Un: "[| Card(K); Card(L) |] ==> Card(K Un L)"
+lemma Card_Un: "[| Card(K); Card(L) |] ==> Card(K \<union> L)"
apply (rule Ord_linear_le [of K L])
apply (simp_all add: subset_Un_iff [THEN iffD1] Card_is_Ord le_imp_subset
subset_Un_iff2 [THEN iffD1])
@@ -387,7 +387,7 @@
lemma Card_cardinal: "Card(|A|)"
apply (unfold cardinal_def)
-apply (case_tac "EX i. Ord (i) & i \<approx> A")
+apply (case_tac "\<exists>i. Ord (i) & i \<approx> A")
txt{*degenerate case*}
prefer 2 apply (erule Least_0 [THEN ssubst], rule Card_0)
txt{*real case: A is isomorphic to some ordinal*}
@@ -395,11 +395,11 @@
apply (rule less_LeastE)
prefer 2 apply assumption
apply (erule eqpoll_trans)
-apply (best intro: LeastI )
+apply (best intro: LeastI )
done
(*Kunen's Lemma 10.5*)
-lemma cardinal_eq_lemma: "[| |i| le j; j le i |] ==> |j| = |i|"
+lemma cardinal_eq_lemma: "[| |i| \<le> j; j \<le> i |] ==> |j| = |i|"
apply (rule eqpollI [THEN cardinal_cong])
apply (erule le_imp_lepoll)
apply (rule lepoll_trans)
@@ -409,7 +409,7 @@
apply (elim ltE Ord_succD)
done
-lemma cardinal_mono: "i le j ==> |i| le |j|"
+lemma cardinal_mono: "i \<le> j ==> |i| \<le> |j|"
apply (rule_tac i = "|i|" and j = "|j|" in Ord_linear_le)
apply (safe intro!: Ord_cardinal le_eqI)
apply (rule cardinal_eq_lemma)
@@ -419,7 +419,7 @@
apply (erule Ord_cardinal_le)
done
-(*Since we have |succ(nat)| le |nat|, the converse of cardinal_mono fails!*)
+(*Since we have @{term"|succ(nat)| \<le> |nat|"}, the converse of cardinal_mono fails!*)
lemma cardinal_lt_imp_lt: "[| |i| < |j|; Ord(i); Ord(j) |] ==> i < j"
apply (rule Ord_linear2 [of i j], assumption+)
apply (erule lt_trans2 [THEN lt_irrefl])
@@ -433,12 +433,12 @@
lemma Card_lt_iff: "[| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)"
by (blast intro: Card_lt_imp_lt Ord_cardinal_le [THEN lt_trans1])
-lemma Card_le_iff: "[| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)"
+lemma Card_le_iff: "[| Ord(i); Card(K) |] ==> (K \<le> |i|) <-> (K \<le> i)"
by (simp add: Card_lt_iff Card_is_Ord Ord_cardinal not_lt_iff_le [THEN iff_sym])
(*Can use AC or finiteness to discharge first premise*)
lemma well_ord_lepoll_imp_Card_le:
- "[| well_ord(B,r); A \<lesssim> B |] ==> |A| le |B|"
+ "[| well_ord(B,r); A \<lesssim> B |] ==> |A| \<le> |B|"
apply (rule_tac i = "|A|" and j = "|B|" in Ord_linear_le)
apply (safe intro!: Ord_cardinal le_eqI)
apply (rule eqpollI [THEN cardinal_cong], assumption)
@@ -452,7 +452,7 @@
apply (erule well_ord_rvimage, assumption)
done
-lemma lepoll_cardinal_le: "[| A \<lesssim> i; Ord(i) |] ==> |A| le i"
+lemma lepoll_cardinal_le: "[| A \<lesssim> i; Ord(i) |] ==> |A| \<le> i"
apply (rule le_trans)
apply (erule well_ord_Memrel [THEN well_ord_lepoll_imp_Card_le], assumption)
apply (erule Ord_cardinal_le)
@@ -466,7 +466,7 @@
apply (blast intro: lepoll_Ord_imp_eqpoll)
done
-lemma cardinal_subset_Ord: "[|A<=i; Ord(i)|] ==> |A| <= i"
+lemma cardinal_subset_Ord: "[|A<=i; Ord(i)|] ==> |A| \<subseteq> i"
apply (drule subset_imp_lepoll [THEN lepoll_cardinal_le])
apply (auto simp add: lt_def)
apply (blast intro: Ord_trans)
@@ -474,10 +474,10 @@
subsection{*The finite cardinals *}
-lemma cons_lepoll_consD:
- "[| cons(u,A) \<lesssim> cons(v,B); u~:A; v~:B |] ==> A \<lesssim> B"
+lemma cons_lepoll_consD:
+ "[| cons(u,A) \<lesssim> cons(v,B); u\<notin>A; v\<notin>B |] ==> A \<lesssim> B"
apply (unfold lepoll_def inj_def, safe)
-apply (rule_tac x = "lam x:A. if f`x=v then f`u else f`x" in exI)
+apply (rule_tac x = "\<lambda>x\<in>A. if f`x=v then f`u else f`x" in exI)
apply (rule CollectI)
(*Proving it's in the function space A->B*)
apply (rule if_type [THEN lam_type])
@@ -488,7 +488,7 @@
apply blast
done
-lemma cons_eqpoll_consD: "[| cons(u,A) \<approx> cons(v,B); u~:A; v~:B |] ==> A \<approx> B"
+lemma cons_eqpoll_consD: "[| cons(u,A) \<approx> cons(v,B); u\<notin>A; v\<notin>B |] ==> A \<approx> B"
apply (simp add: eqpoll_iff)
apply (blast intro: cons_lepoll_consD)
done
@@ -501,7 +501,7 @@
done
lemma nat_lepoll_imp_le [rule_format]:
- "m:nat ==> ALL n: nat. m \<lesssim> n --> m le n"
+ "m:nat ==> \<forall>n\<in>nat. m \<lesssim> n \<longrightarrow> m \<le> n"
apply (induct_tac m)
apply (blast intro!: nat_0_le)
apply (rule ballI)
@@ -517,12 +517,12 @@
done
(*The object of all this work: every natural number is a (finite) cardinal*)
-lemma nat_into_Card:
+lemma nat_into_Card:
"n: nat ==> Card(n)"
apply (unfold Card_def cardinal_def)
apply (subst Least_equality)
apply (rule eqpoll_refl)
-apply (erule nat_into_Ord)
+apply (erule nat_into_Ord)
apply (simp (no_asm_simp) add: lt_nat_in_nat [THEN nat_eqpoll_iff])
apply (blast elim!: lt_irrefl)+
done
@@ -538,25 +538,25 @@
lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat"
apply (unfold lesspoll_def)
apply (fast elim!: Ord_nat [THEN [2] ltI [THEN leI, THEN le_imp_lepoll]]
- eqpoll_sym [THEN eqpoll_imp_lepoll]
- intro: Ord_nat [THEN [2] nat_succI [THEN ltI], THEN leI,
+ eqpoll_sym [THEN eqpoll_imp_lepoll]
+ intro: Ord_nat [THEN [2] nat_succI [THEN ltI], THEN leI,
THEN le_imp_lepoll, THEN lepoll_trans, THEN succ_lepoll_natE])
done
-lemma nat_lepoll_imp_ex_eqpoll_n:
+lemma nat_lepoll_imp_ex_eqpoll_n:
"[| n \<in> nat; nat \<lesssim> X |] ==> \<exists>Y. Y \<subseteq> X & n \<approx> Y"
apply (unfold lepoll_def eqpoll_def)
apply (fast del: subsetI subsetCE
intro!: subset_SIs
dest!: Ord_nat [THEN [2] OrdmemD, THEN [2] restrict_inj]
- elim!: restrict_bij
+ elim!: restrict_bij
inj_is_fun [THEN fun_is_rel, THEN image_subset])
done
(** lepoll, \<prec> and natural numbers **)
-lemma lepoll_imp_lesspoll_succ:
+lemma lepoll_imp_lesspoll_succ:
"[| A \<lesssim> m; m:nat |] ==> A \<prec> succ(m)"
apply (unfold lesspoll_def)
apply (rule conjI)
@@ -567,7 +567,7 @@
apply (erule succ_lepoll_natE, assumption)
done
-lemma lesspoll_succ_imp_lepoll:
+lemma lesspoll_succ_imp_lepoll:
"[| A \<prec> succ(m); m:nat |] ==> A \<lesssim> m"
apply (unfold lesspoll_def lepoll_def eqpoll_def bij_def, clarify)
apply (blast intro!: inj_not_surj_succ)
@@ -615,14 +615,14 @@
lemma Card_nat: "Card(nat)"
apply (unfold Card_def cardinal_def)
apply (subst Least_equality)
-apply (rule eqpoll_refl)
-apply (rule Ord_nat)
+apply (rule eqpoll_refl)
+apply (rule Ord_nat)
apply (erule ltE)
apply (simp_all add: eqpoll_iff lt_not_lepoll ltI)
done
(*Allows showing that |i| is a limit cardinal*)
-lemma nat_le_cardinal: "nat le i ==> nat le |i|"
+lemma nat_le_cardinal: "nat \<le> i ==> nat \<le> |i|"
apply (rule Card_nat [THEN Card_cardinal_eq, THEN subst])
apply (erule cardinal_mono)
done
@@ -632,26 +632,26 @@
(** Congruence laws for successor, cardinal addition and multiplication **)
(*Congruence law for cons under equipollence*)
-lemma cons_lepoll_cong:
- "[| A \<lesssim> B; b ~: B |] ==> cons(a,A) \<lesssim> cons(b,B)"
+lemma cons_lepoll_cong:
+ "[| A \<lesssim> B; b \<notin> B |] ==> cons(a,A) \<lesssim> cons(b,B)"
apply (unfold lepoll_def, safe)
-apply (rule_tac x = "lam y: cons (a,A) . if y=a then b else f`y" in exI)
+apply (rule_tac x = "\<lambda>y\<in>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 (safe elim!: consE')
+apply (safe elim!: consE')
apply simp_all
-apply (blast intro: inj_is_fun [THEN apply_type])+
+apply (blast intro: inj_is_fun [THEN apply_type])+
done
lemma cons_eqpoll_cong:
- "[| A \<approx> B; a ~: A; b ~: B |] ==> cons(a,A) \<approx> cons(b,B)"
+ "[| A \<approx> B; a \<notin> A; b \<notin> B |] ==> cons(a,A) \<approx> cons(b,B)"
by (simp add: eqpoll_iff cons_lepoll_cong)
lemma cons_lepoll_cons_iff:
- "[| a ~: A; b ~: B |] ==> cons(a,A) \<lesssim> cons(b,B) <-> A \<lesssim> B"
+ "[| a \<notin> A; b \<notin> B |] ==> cons(a,A) \<lesssim> cons(b,B) <-> A \<lesssim> B"
by (blast intro: cons_lepoll_cong cons_lepoll_consD)
lemma cons_eqpoll_cons_iff:
- "[| a ~: A; b ~: B |] ==> cons(a,A) \<approx> cons(b,B) <-> A \<approx> B"
+ "[| a \<notin> A; b \<notin> B |] ==> cons(a,A) \<approx> cons(b,B) <-> A \<approx> B"
by (blast intro: cons_eqpoll_cong cons_eqpoll_consD)
lemma singleton_eqpoll_1: "{a} \<approx> 1"
@@ -664,7 +664,7 @@
apply (simp (no_asm) add: nat_into_Card [THEN Card_cardinal_eq])
done
-lemma not_0_is_lepoll_1: "A ~= 0 ==> 1 \<lesssim> A"
+lemma not_0_is_lepoll_1: "A \<noteq> 0 ==> 1 \<lesssim> A"
apply (erule not_emptyE)
apply (rule_tac a = "cons (x, A-{x}) " in subst)
apply (rule_tac [2] a = "cons(0,0)" and P= "%y. y \<lesssim> cons (x, A-{x})" in subst)
@@ -684,24 +684,24 @@
done
(*Congruence law for * under equipollence*)
-lemma prod_eqpoll_cong:
+lemma prod_eqpoll_cong:
"[| A \<approx> C; B \<approx> D |] ==> A*B \<approx> C*D"
apply (unfold eqpoll_def)
apply (blast intro!: prod_bij)
done
-lemma inj_disjoint_eqpoll:
- "[| f: inj(A,B); A Int B = 0 |] ==> A Un (B - range(f)) \<approx> B"
+lemma inj_disjoint_eqpoll:
+ "[| f: inj(A,B); A \<inter> B = 0 |] ==> A \<union> (B - range(f)) \<approx> 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])
-apply (safe elim!: UnE')
+apply (safe elim!: UnE')
apply (simp_all add: inj_is_fun [THEN apply_rangeI])
-apply (blast intro: inj_converse_fun [THEN apply_type])+
+apply (blast intro: inj_converse_fun [THEN apply_type])+
done
@@ -710,7 +710,7 @@
(*New proofs using cons_lepoll_cons. Could generalise from succ to cons.*)
(*If A has at most n+1 elements and a:A then A-{a} has at most n.*)
-lemma Diff_sing_lepoll:
+lemma Diff_sing_lepoll:
"[| a:A; A \<lesssim> succ(n) |] ==> A - {a} \<lesssim> n"
apply (unfold succ_def)
apply (rule cons_lepoll_consD)
@@ -719,7 +719,7 @@
done
(*If A has at least n+1 elements then A-{a} has at least n.*)
-lemma lepoll_Diff_sing:
+lemma lepoll_Diff_sing:
"[| succ(n) \<lesssim> A |] ==> n \<lesssim> A - {a}"
apply (unfold succ_def)
apply (rule cons_lepoll_consD)
@@ -729,8 +729,8 @@
done
lemma Diff_sing_eqpoll: "[| a:A; A \<approx> succ(n) |] ==> A - {a} \<approx> n"
-by (blast intro!: eqpollI
- elim!: eqpollE
+by (blast intro!: eqpollI
+ elim!: eqpollE
intro: Diff_sing_lepoll lepoll_Diff_sing)
lemma lepoll_1_is_sing: "[| A \<lesssim> 1; a:A |] ==> A = {a}"
@@ -739,23 +739,23 @@
apply (blast elim: equalityE)
done
-lemma Un_lepoll_sum: "A Un B \<lesssim> A+B"
+lemma Un_lepoll_sum: "A \<union> B \<lesssim> A+B"
apply (unfold lepoll_def)
-apply (rule_tac x = "lam x: A Un B. if x:A then Inl (x) else Inr (x) " in exI)
+apply (rule_tac x = "\<lambda>x\<in>A \<union> B. if x:A then Inl (x) else Inr (x) " in exI)
apply (rule_tac d = "%z. snd (z) " in lam_injective)
-apply force
+apply force
apply (simp add: Inl_def Inr_def)
done
lemma well_ord_Un:
- "[| well_ord(X,R); well_ord(Y,S) |] ==> EX T. well_ord(X Un Y, T)"
-by (erule well_ord_radd [THEN Un_lepoll_sum [THEN lepoll_well_ord]],
+ "[| well_ord(X,R); well_ord(Y,S) |] ==> \<exists>T. well_ord(X \<union> Y, T)"
+by (erule well_ord_radd [THEN Un_lepoll_sum [THEN lepoll_well_ord]],
assumption)
(*Krzysztof Grabczewski*)
-lemma disj_Un_eqpoll_sum: "A Int B = 0 ==> A Un B \<approx> A + B"
+lemma disj_Un_eqpoll_sum: "A \<inter> B = 0 ==> A \<union> B \<approx> A + B"
apply (unfold eqpoll_def)
-apply (rule_tac x = "lam a:A Un B. if a:A then Inl (a) else Inr (a) " in exI)
+apply (rule_tac x = "\<lambda>a\<in>A \<union> B. if a:A then Inl (a) else Inr (a) " in exI)
apply (rule_tac d = "%z. case (%x. x, %x. x, z) " in lam_bijective)
apply auto
done
@@ -776,14 +776,14 @@
apply (blast dest!: lepoll_succ_disj)
done
-lemma lesspoll_nat_is_Finite:
+lemma lesspoll_nat_is_Finite:
"A \<prec> nat ==> Finite(A)"
apply (unfold Finite_def)
-apply (blast dest: ltD lesspoll_cardinal_lt
+apply (blast dest: ltD lesspoll_cardinal_lt
lesspoll_imp_eqpoll [THEN eqpoll_sym])
done
-lemma lepoll_Finite:
+lemma lepoll_Finite:
"[| Y \<lesssim> X; Finite(X) |] ==> Finite(Y)"
apply (unfold Finite_def)
apply (blast elim!: eqpollE
@@ -793,8 +793,8 @@
lemmas subset_Finite = subset_imp_lepoll [THEN lepoll_Finite]
-lemma Finite_Int: "Finite(A) | Finite(B) ==> Finite(A Int B)"
-by (blast intro: subset_Finite)
+lemma Finite_Int: "Finite(A) | Finite(B) ==> Finite(A \<inter> B)"
+by (blast intro: subset_Finite)
lemmas Finite_Diff = Diff_subset [THEN subset_Finite]
@@ -819,16 +819,16 @@
lemma Finite_succ_iff [iff]: "Finite(succ(x)) <-> Finite(x)"
by (simp add: succ_def)
-lemma nat_le_infinite_Ord:
- "[| Ord(i); ~ Finite(i) |] ==> nat le i"
+lemma nat_le_infinite_Ord:
+ "[| Ord(i); ~ Finite(i) |] ==> nat \<le> i"
apply (unfold Finite_def)
apply (erule Ord_nat [THEN [2] Ord_linear2])
prefer 2 apply assumption
apply (blast intro!: eqpoll_refl elim!: ltE)
done
-lemma Finite_imp_well_ord:
- "Finite(A) ==> EX r. well_ord(A,r)"
+lemma Finite_imp_well_ord:
+ "Finite(A) ==> \<exists>r. well_ord(A,r)"
apply (unfold Finite_def eqpoll_def)
apply (blast intro: well_ord_rvimage bij_is_inj well_ord_Memrel nat_into_Ord)
done
@@ -840,7 +840,7 @@
by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0])
lemma Finite_Fin_lemma [rule_format]:
- "n \<in> nat ==> \<forall>A. (A\<approx>n & A \<subseteq> X) --> A \<in> Fin(X)"
+ "n \<in> nat ==> \<forall>A. (A\<approx>n & A \<subseteq> X) \<longrightarrow> A \<in> Fin(X)"
apply (induct_tac n)
apply (rule allI)
apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
@@ -857,17 +857,17 @@
done
lemma Finite_Fin: "[| Finite(A); A \<subseteq> X |] ==> A \<in> Fin(X)"
-by (unfold Finite_def, blast intro: Finite_Fin_lemma)
+by (unfold Finite_def, blast intro: Finite_Fin_lemma)
lemma eqpoll_imp_Finite_iff: "A \<approx> B ==> Finite(A) <-> Finite(B)"
-apply (unfold Finite_def)
-apply (blast intro: eqpoll_trans eqpoll_sym)
+apply (unfold Finite_def)
+apply (blast intro: eqpoll_trans eqpoll_sym)
done
-lemma Fin_lemma [rule_format]: "n: nat ==> ALL A. A \<approx> n --> A : Fin(A)"
+lemma Fin_lemma [rule_format]: "n: nat ==> \<forall>A. A \<approx> n \<longrightarrow> A \<in> Fin(A)"
apply (induct_tac n)
apply (simp add: eqpoll_0_iff, clarify)
-apply (subgoal_tac "EX u. u:A")
+apply (subgoal_tac "\<exists>u. u:A")
apply (erule exE)
apply (rule Diff_sing_eqpoll [elim_format])
prefer 2 apply assumption
@@ -880,28 +880,28 @@
apply (blast intro: bij_converse_bij [THEN bij_is_fun, THEN apply_type])
done
-lemma Finite_into_Fin: "Finite(A) ==> A : Fin(A)"
+lemma Finite_into_Fin: "Finite(A) ==> A \<in> Fin(A)"
apply (unfold Finite_def)
apply (blast intro: Fin_lemma)
done
-lemma Fin_into_Finite: "A : Fin(U) ==> Finite(A)"
+lemma Fin_into_Finite: "A \<in> Fin(U) ==> Finite(A)"
by (fast intro!: Finite_0 Finite_cons elim: Fin_induct)
-lemma Finite_Fin_iff: "Finite(A) <-> A : Fin(A)"
+lemma Finite_Fin_iff: "Finite(A) <-> A \<in> Fin(A)"
by (blast intro: Finite_into_Fin Fin_into_Finite)
-lemma Finite_Un: "[| Finite(A); Finite(B) |] ==> Finite(A Un B)"
-by (blast intro!: Fin_into_Finite Fin_UnI
+lemma Finite_Un: "[| Finite(A); Finite(B) |] ==> Finite(A \<union> B)"
+by (blast intro!: Fin_into_Finite Fin_UnI
dest!: Finite_into_Fin
- intro: Un_upper1 [THEN Fin_mono, THEN subsetD]
+ intro: Un_upper1 [THEN Fin_mono, THEN subsetD]
Un_upper2 [THEN Fin_mono, THEN subsetD])
-lemma Finite_Un_iff [simp]: "Finite(A Un B) <-> (Finite(A) & Finite(B))"
-by (blast intro: subset_Finite Finite_Un)
+lemma Finite_Un_iff [simp]: "Finite(A \<union> B) <-> (Finite(A) & Finite(B))"
+by (blast intro: subset_Finite Finite_Un)
text{*The converse must hold too.*}
-lemma Finite_Union: "[| ALL y:X. Finite(y); Finite(X) |] ==> Finite(Union(X))"
+lemma Finite_Union: "[| \<forall>y\<in>X. Finite(y); Finite(X) |] ==> Finite(\<Union>(X))"
apply (simp add: Finite_Fin_iff)
apply (rule Fin_UnionI)
apply (erule Fin_induct, simp)
@@ -911,9 +911,9 @@
(* Induction principle for Finite(A), by Sidi Ehmety *)
lemma Finite_induct [case_names 0 cons, induct set: Finite]:
"[| Finite(A); P(0);
- !! x B. [| Finite(B); x ~: B; P(B) |] ==> P(cons(x, B)) |]
+ !! x B. [| Finite(B); x \<notin> B; P(B) |] ==> P(cons(x, B)) |]
==> P(A)"
-apply (erule Finite_into_Fin [THEN Fin_induct])
+apply (erule Finite_into_Fin [THEN Fin_induct])
apply (blast intro: Fin_into_Finite)+
done
@@ -930,7 +930,7 @@
(*Sidi Ehmety. And the contrapositive of this says
[| ~Finite(A); Finite(B) |] ==> ~Finite(A-B) *)
-lemma Diff_Finite [rule_format]: "Finite(B) ==> Finite(A-B) --> Finite(A)"
+lemma Diff_Finite [rule_format]: "Finite(B) ==> Finite(A-B) \<longrightarrow> Finite(A)"
apply (erule Finite_induct, auto)
apply (case_tac "x:A")
apply (subgoal_tac [2] "A-cons (x, B) = A - B")
@@ -942,37 +942,37 @@
by (erule Finite_induct, simp_all)
lemma Finite_RepFun_iff_lemma [rule_format]:
- "[|Finite(x); !!x y. f(x)=f(y) ==> x=y|]
- ==> \<forall>A. x = RepFun(A,f) --> Finite(A)"
+ "[|Finite(x); !!x y. f(x)=f(y) ==> x=y|]
+ ==> \<forall>A. x = RepFun(A,f) \<longrightarrow> Finite(A)"
apply (erule Finite_induct)
- apply clarify
+ apply clarify
apply (case_tac "A=0", simp)
- apply (blast del: allE, clarify)
-apply (subgoal_tac "\<exists>z\<in>A. x = f(z)")
- prefer 2 apply (blast del: allE elim: equalityE, clarify)
+ apply (blast del: allE, clarify)
+apply (subgoal_tac "\<exists>z\<in>A. x = f(z)")
+ prefer 2 apply (blast del: allE elim: equalityE, clarify)
apply (subgoal_tac "B = {f(u) . u \<in> A - {z}}")
- apply (blast intro: Diff_sing_Finite)
-apply (thin_tac "\<forall>A. ?P(A) --> Finite(A)")
-apply (rule equalityI)
- apply (blast intro: elim: equalityE)
-apply (blast intro: elim: equalityCE)
+ apply (blast intro: Diff_sing_Finite)
+apply (thin_tac "\<forall>A. ?P(A) \<longrightarrow> Finite(A)")
+apply (rule equalityI)
+ apply (blast intro: elim: equalityE)
+apply (blast intro: elim: equalityCE)
done
text{*I don't know why, but if the premise is expressed using meta-connectives
then the simplifier cannot prove it automatically in conditional rewriting.*}
lemma Finite_RepFun_iff:
- "(\<forall>x y. f(x)=f(y) --> x=y) ==> Finite(RepFun(A,f)) <-> Finite(A)"
-by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f])
+ "(\<forall>x y. f(x)=f(y) \<longrightarrow> x=y) ==> Finite(RepFun(A,f)) <-> Finite(A)"
+by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f])
lemma Finite_Pow: "Finite(A) ==> Finite(Pow(A))"
-apply (erule Finite_induct)
-apply (simp_all add: Pow_insert Finite_Un Finite_RepFun)
+apply (erule Finite_induct)
+apply (simp_all add: Pow_insert Finite_Un Finite_RepFun)
done
lemma Finite_Pow_imp_Finite: "Finite(Pow(A)) ==> Finite(A)"
apply (subgoal_tac "Finite({{x} . x \<in> A})")
- apply (simp add: Finite_RepFun_iff )
-apply (blast intro: subset_Finite)
+ apply (simp add: Finite_RepFun_iff )
+apply (blast intro: subset_Finite)
done
lemma Finite_Pow_iff [iff]: "Finite(Pow(A)) <-> Finite(A)"
@@ -992,8 +992,8 @@
txt{*x:Z case*}
apply (drule_tac x = x in bspec, assumption)
apply (blast elim: mem_irrefl mem_asym)
-txt{*other case*}
-apply (drule_tac x = Z in spec, blast)
+txt{*other case*}
+apply (drule_tac x = Z in spec, blast)
done
lemma nat_well_ord_converse_Memrel: "n:nat ==> well_ord(n,converse(Memrel(n)))"
@@ -1003,7 +1003,7 @@
done
lemma well_ord_converse:
- "[|well_ord(A,r);
+ "[|well_ord(A,r);
well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) |]
==> well_ord(A,converse(r))"
apply (rule well_ord_Int_iff [THEN iffD1])
@@ -1021,7 +1021,7 @@
apply (blast intro!: ordermap_bij [THEN bij_converse_bij])
done
-lemma Finite_well_ord_converse:
+lemma Finite_well_ord_converse:
"[| Finite(A); well_ord(A,r) |] ==> well_ord(A,converse(r))"
apply (unfold Finite_def)
apply (rule well_ord_converse, assumption)
@@ -1034,9 +1034,9 @@
done
lemma nat_not_Finite: "~Finite(nat)"
-apply (unfold Finite_def, clarify)
-apply (drule eqpoll_imp_lepoll [THEN lepoll_cardinal_le], simp)
-apply (insert Card_nat)
+apply (unfold Finite_def, clarify)
+apply (drule eqpoll_imp_lepoll [THEN lepoll_cardinal_le], simp)
+apply (insert Card_nat)
apply (simp add: Card_def)
apply (drule le_imp_subset)
apply (blast elim: mem_irrefl)
--- a/src/ZF/CardinalArith.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/CardinalArith.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,30 +9,30 @@
definition
InfCard :: "i=>o" where
- "InfCard(i) == Card(i) & nat le i"
+ "InfCard(i) == Card(i) & nat \<le> i"
definition
cmult :: "[i,i]=>i" (infixl "|*|" 70) where
"i |*| j == |i*j|"
-
+
definition
cadd :: "[i,i]=>i" (infixl "|+|" 65) where
"i |+| j == |i+j|"
definition
csquare_rel :: "i=>i" where
- "csquare_rel(K) ==
- rvimage(K*K,
- lam <x,y>:K*K. <x Un y, x, y>,
+ "csquare_rel(K) ==
+ rvimage(K*K,
+ lam <x,y>:K*K. <x \<union> y, x, y>,
rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))"
definition
jump_cardinal :: "i=>i" where
--{*This def is more complex than Kunen's but it more easily proved to
be a cardinal*}
- "jump_cardinal(K) ==
+ "jump_cardinal(K) ==
\<Union>X\<in>Pow(K). {z. r: Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
-
+
definition
csucc :: "i=>i" where
--{*needed because @{term "jump_cardinal(K)"} might not be the successor
@@ -48,25 +48,25 @@
cmult (infixl "\<otimes>" 70)
-lemma Card_Union [simp,intro,TC]: "(ALL x:A. Card(x)) ==> Card(Union(A))"
-apply (rule CardI)
- apply (simp add: Card_is_Ord)
+lemma Card_Union [simp,intro,TC]: "(\<forall>x\<in>A. Card(x)) ==> Card(\<Union>(A))"
+apply (rule CardI)
+ apply (simp add: Card_is_Ord)
apply (clarify dest!: ltD)
-apply (drule bspec, assumption)
-apply (frule lt_Card_imp_lesspoll, blast intro: ltI Card_is_Ord)
+apply (drule bspec, assumption)
+apply (frule lt_Card_imp_lesspoll, blast intro: ltI Card_is_Ord)
apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll])
-apply (drule lesspoll_trans1, assumption)
+apply (drule lesspoll_trans1, assumption)
apply (subgoal_tac "B \<lesssim> \<Union>A")
- apply (drule lesspoll_trans1, assumption, blast)
-apply (blast intro: subset_imp_lepoll)
+ apply (drule lesspoll_trans1, assumption, blast)
+apply (blast intro: subset_imp_lepoll)
done
-lemma Card_UN: "(!!x. x:A ==> Card(K(x))) ==> Card(\<Union>x\<in>A. K(x))"
-by (blast intro: Card_Union)
+lemma Card_UN: "(!!x. x:A ==> Card(K(x))) ==> Card(\<Union>x\<in>A. K(x))"
+by (blast intro: Card_Union)
lemma Card_OUN [simp,intro,TC]:
"(!!x. x:A ==> Card(K(x))) ==> Card(\<Union>x<A. K(x))"
-by (simp add: OUnion_def Card_0)
+by (simp add: OUnion_def Card_0)
lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat"
apply (unfold lesspoll_def)
@@ -75,8 +75,8 @@
apply (rule notI)
apply (erule eqpollE)
apply (rule succ_lepoll_natE)
-apply (blast intro: nat_succI [THEN OrdmemD, THEN subset_imp_lepoll]
- lepoll_trans, assumption)
+apply (blast intro: nat_succI [THEN OrdmemD, THEN subset_imp_lepoll]
+ lepoll_trans, assumption)
done
lemma in_Card_imp_lesspoll: "[| Card(K); b \<in> K |] ==> b \<prec> K"
@@ -88,7 +88,7 @@
lemma lesspoll_lemma: "[| ~ A \<prec> B; C \<prec> B |] ==> A - C \<noteq> 0"
apply (unfold lesspoll_def)
apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll]
- intro!: eqpollI elim: notE
+ intro!: eqpollI elim: notE
elim!: eqpollE lepoll_trans)
done
@@ -123,18 +123,18 @@
done
(*Unconditional version requires AC*)
-lemma well_ord_cadd_assoc:
+lemma well_ord_cadd_assoc:
"[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |]
==> (i |+| j) |+| k = i |+| (j |+| k)"
apply (unfold cadd_def)
apply (rule cardinal_cong)
apply (rule eqpoll_trans)
apply (rule sum_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl])
- apply (blast intro: well_ord_radd )
+ apply (blast intro: well_ord_radd )
apply (rule sum_assoc_eqpoll [THEN eqpoll_trans])
apply (rule eqpoll_sym)
apply (rule sum_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll])
-apply (blast intro: well_ord_radd )
+apply (blast intro: well_ord_radd )
done
subsubsection{*0 is the identity for addition*}
@@ -154,14 +154,14 @@
lemma sum_lepoll_self: "A \<lesssim> A+B"
apply (unfold lepoll_def inj_def)
-apply (rule_tac x = "lam x:A. Inl (x) " in exI)
+apply (rule_tac x = "\<lambda>x\<in>A. Inl (x) " in exI)
apply simp
done
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
-lemma cadd_le_self:
- "[| Card(K); Ord(L) |] ==> K le (K |+| L)"
+lemma cadd_le_self:
+ "[| Card(K); Ord(L) |] ==> K \<le> (K |+| L)"
apply (unfold cadd_def)
apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_Card_le],
assumption)
@@ -171,18 +171,18 @@
subsubsection{*Monotonicity of addition*}
-lemma sum_lepoll_mono:
+lemma sum_lepoll_mono:
"[| A \<lesssim> C; B \<lesssim> D |] ==> A + B \<lesssim> C + D"
apply (unfold lepoll_def)
apply (elim exE)
-apply (rule_tac x = "lam z:A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI)
+apply (rule_tac x = "\<lambda>z\<in>A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI)
apply (rule_tac d = "case (%w. Inl(converse(f) `w), %y. Inr(converse(fa) ` y))"
in lam_injective)
apply (typecheck add: inj_is_fun, auto)
done
lemma cadd_le_mono:
- "[| K' le K; L' le L |] ==> (K' |+| L') le (K |+| L)"
+ "[| K' \<le> K; L' \<le> L |] ==> (K' |+| L') \<le> (K |+| L)"
apply (unfold cadd_def)
apply (safe dest!: le_subset_iff [THEN iffD1])
apply (rule well_ord_lepoll_imp_Card_le)
@@ -195,7 +195,7 @@
lemma sum_succ_eqpoll: "succ(A)+B \<approx> succ(A+B)"
apply (unfold eqpoll_def)
apply (rule exI)
-apply (rule_tac c = "%z. if z=Inl (A) then A+B else z"
+apply (rule_tac c = "%z. if z=Inl (A) then A+B else z"
and d = "%z. if z=A+B then Inl (A) else z" in lam_bijective)
apply simp_all
apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+
@@ -227,8 +227,8 @@
lemma prod_commute_eqpoll: "A*B \<approx> B*A"
apply (unfold eqpoll_def)
apply (rule exI)
-apply (rule_tac c = "%<x,y>.<y,x>" and d = "%<x,y>.<y,x>" in lam_bijective,
- auto)
+apply (rule_tac c = "%<x,y>.<y,x>" and d = "%<x,y>.<y,x>" in lam_bijective,
+ auto)
done
lemma cmult_commute: "i |*| j = j |*| i"
@@ -250,11 +250,11 @@
==> (i |*| j) |*| k = i |*| (j |*| k)"
apply (unfold cmult_def)
apply (rule cardinal_cong)
-apply (rule eqpoll_trans)
+apply (rule eqpoll_trans)
apply (rule prod_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl])
apply (blast intro: well_ord_rmult)
apply (rule prod_assoc_eqpoll [THEN eqpoll_trans])
-apply (rule eqpoll_sym)
+apply (rule eqpoll_sym)
apply (rule prod_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll])
apply (blast intro: well_ord_rmult)
done
@@ -272,12 +272,12 @@
==> (i |+| j) |*| k = (i |*| k) |+| (j |*| k)"
apply (unfold cadd_def cmult_def)
apply (rule cardinal_cong)
-apply (rule eqpoll_trans)
+apply (rule eqpoll_trans)
apply (rule prod_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl])
apply (blast intro: well_ord_radd)
apply (rule sum_prod_distrib_eqpoll [THEN eqpoll_trans])
-apply (rule eqpoll_sym)
-apply (rule sum_eqpoll_cong [OF well_ord_cardinal_eqpoll
+apply (rule eqpoll_sym)
+apply (rule sum_eqpoll_cong [OF well_ord_cardinal_eqpoll
well_ord_cardinal_eqpoll])
apply (blast intro: well_ord_rmult)+
done
@@ -310,11 +310,11 @@
lemma prod_square_lepoll: "A \<lesssim> A*A"
apply (unfold lepoll_def inj_def)
-apply (rule_tac x = "lam x:A. <x,x>" in exI, simp)
+apply (rule_tac x = "\<lambda>x\<in>A. <x,x>" in exI, simp)
done
(*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
-lemma cmult_square_le: "Card(K) ==> K le K |*| K"
+lemma cmult_square_le: "Card(K) ==> K \<le> K |*| K"
apply (unfold cmult_def)
apply (rule le_trans)
apply (rule_tac [2] well_ord_lepoll_imp_Card_le)
@@ -327,12 +327,12 @@
lemma prod_lepoll_self: "b: B ==> A \<lesssim> A*B"
apply (unfold lepoll_def inj_def)
-apply (rule_tac x = "lam x:A. <x,b>" in exI, simp)
+apply (rule_tac x = "\<lambda>x\<in>A. <x,b>" in exI, simp)
done
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
lemma cmult_le_self:
- "[| Card(K); Ord(L); 0<L |] ==> K le (K |*| L)"
+ "[| Card(K); Ord(L); 0<L |] ==> K \<le> (K |*| L)"
apply (unfold cmult_def)
apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_Card_le])
apply assumption
@@ -347,13 +347,13 @@
apply (unfold lepoll_def)
apply (elim exE)
apply (rule_tac x = "lam <w,y>:A*B. <f`w, fa`y>" in exI)
-apply (rule_tac d = "%<w,y>. <converse (f) `w, converse (fa) `y>"
+apply (rule_tac d = "%<w,y>. <converse (f) `w, converse (fa) `y>"
in lam_injective)
apply (typecheck add: inj_is_fun, auto)
done
lemma cmult_le_mono:
- "[| K' le K; L' le L |] ==> (K' |*| L') le (K |*| L)"
+ "[| K' \<le> K; L' \<le> L |] ==> (K' |*| L') \<le> (K |*| L)"
apply (unfold cmult_def)
apply (safe dest!: le_subset_iff [THEN iffD1])
apply (rule well_ord_lepoll_imp_Card_le)
@@ -391,10 +391,10 @@
by (simp add: cmult_succ_lemma Card_is_Ord cadd_commute [of _ 0])
lemma sum_lepoll_prod: "2 \<lesssim> C ==> B+B \<lesssim> C*B"
-apply (rule lepoll_trans)
-apply (rule sum_eq_2_times [THEN equalityD1, THEN subset_imp_lepoll])
-apply (erule prod_lepoll_mono)
-apply (rule lepoll_refl)
+apply (rule lepoll_trans)
+apply (rule sum_eq_2_times [THEN equalityD1, THEN subset_imp_lepoll])
+apply (erule prod_lepoll_mono)
+apply (rule lepoll_refl)
done
lemma lepoll_imp_sum_lepoll_prod: "[| A \<lesssim> B; 2 \<lesssim> A |] ==> A+B \<lesssim> A*B"
@@ -404,20 +404,20 @@
subsection{*Infinite Cardinals are Limit Ordinals*}
(*This proof is modelled upon one assuming nat<=A, with injection
- lam z:cons(u,A). if z=u then 0 else if z : nat then succ(z) else z
+ \<lambda>z\<in>cons(u,A). if z=u then 0 else if z \<in> 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.*)
lemma nat_cons_lepoll: "nat \<lesssim> A ==> cons(u,A) \<lesssim> A"
apply (unfold lepoll_def)
apply (erule exE)
-apply (rule_tac x =
- "lam z:cons (u,A).
- if z=u then f`0
- else if z: range (f) then f`succ (converse (f) `z) else z"
+apply (rule_tac x =
+ "\<lambda>z\<in>cons (u,A).
+ if z=u then f`0
+ 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)
- else 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)
apply (simp add: inj_is_fun [THEN apply_rangeI]
@@ -431,7 +431,7 @@
done
(*Specialized version required below*)
-lemma nat_succ_eqpoll: "nat <= A ==> succ(A) \<approx> A"
+lemma nat_succ_eqpoll: "nat \<subseteq> A ==> succ(A) \<approx> A"
apply (unfold succ_def)
apply (erule subset_imp_lepoll [THEN nat_cons_eqpoll])
done
@@ -447,7 +447,7 @@
done
lemma InfCard_Un:
- "[| InfCard(K); Card(L) |] ==> InfCard(K Un L)"
+ "[| InfCard(K); Card(L) |] ==> InfCard(K \<union> L)"
apply (unfold InfCard_def)
apply (simp add: Card_Un Un_upper1_le [THEN [2] le_trans] Card_is_Ord)
done
@@ -477,7 +477,7 @@
apply (unfold eqpoll_def)
apply (rule exI)
apply (simp add: ordermap_eq_image well_ord_is_wf)
-apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij,
+apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij,
THEN bij_converse_bij])
apply (rule pred_subset)
done
@@ -485,7 +485,7 @@
subsubsection{*Establishing the well-ordering*}
lemma csquare_lam_inj:
- "Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)"
+ "Ord(K) ==> (lam <x,y>:K*K. <x \<union> y, x, y>) \<in> inj(K*K, K*K*K)"
apply (unfold inj_def)
apply (force intro: lam_type Un_least_lt [THEN ltD] ltI)
done
@@ -499,7 +499,7 @@
subsubsection{*Characterising initial segments of the well-ordering*}
lemma csquareD:
- "[| <<x,y>, <z,z>> : csquare_rel(K); x<K; y<K; z<K |] ==> x le z & y le z"
+ "[| <<x,y>, <z,z>> \<in> csquare_rel(K); x<K; y<K; z<K |] ==> x \<le> z & y \<le> z"
apply (unfold csquare_rel_def)
apply (erule rev_mp)
apply (elim ltE)
@@ -508,45 +508,45 @@
apply (simp_all add: lt_def succI2)
done
-lemma pred_csquare_subset:
- "z<K ==> Order.pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)"
+lemma pred_csquare_subset:
+ "z<K ==> Order.pred(K*K, <z,z>, csquare_rel(K)) \<subseteq> succ(z)*succ(z)"
apply (unfold Order.pred_def)
apply (safe del: SigmaI succCI)
apply (erule csquareD [THEN conjE])
-apply (unfold lt_def, auto)
+apply (unfold lt_def, auto)
done
lemma csquare_ltI:
- "[| x<z; y<z; z<K |] ==> <<x,y>, <z,z>> : csquare_rel(K)"
+ "[| x<z; y<z; z<K |] ==> <<x,y>, <z,z>> \<in> csquare_rel(K)"
apply (unfold csquare_rel_def)
apply (subgoal_tac "x<K & y<K")
- prefer 2 apply (blast intro: lt_trans)
+ prefer 2 apply (blast intro: lt_trans)
apply (elim ltE)
apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD)
done
(*Part of the traditional proof. UNUSED since it's harder to prove & apply *)
lemma csquare_or_eqI:
- "[| x le z; y le z; z<K |] ==> <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z"
+ "[| x \<le> z; y \<le> z; z<K |] ==> <<x,y>, <z,z>> \<in> csquare_rel(K) | x=z & y=z"
apply (unfold csquare_rel_def)
apply (subgoal_tac "x<K & y<K")
- prefer 2 apply (blast intro: lt_trans1)
+ prefer 2 apply (blast intro: lt_trans1)
apply (elim ltE)
apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD)
apply (elim succE)
-apply (simp_all add: subset_Un_iff [THEN iff_sym]
+apply (simp_all add: subset_Un_iff [THEN iff_sym]
subset_Un_iff2 [THEN iff_sym] OrdmemD)
done
subsubsection{*The cardinality of initial segments*}
lemma ordermap_z_lt:
- "[| Limit(K); x<K; y<K; z=succ(x Un y) |] ==>
+ "[| Limit(K); x<K; y<K; z=succ(x \<union> y) |] ==>
ordermap(K*K, csquare_rel(K)) ` <x,y> <
ordermap(K*K, csquare_rel(K)) ` <z,z>"
apply (subgoal_tac "z<K & well_ord (K*K, csquare_rel (K))")
prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ
- Limit_is_Ord [THEN well_ord_csquare], clarify)
+ Limit_is_Ord [THEN well_ord_csquare], clarify)
apply (rule csquare_ltI [THEN ordermap_mono, THEN ltI])
apply (erule_tac [4] well_ord_is_wf)
apply (blast intro!: Un_upper1_le Un_upper2_le Ord_ordermap elim!: ltE)+
@@ -554,14 +554,14 @@
(*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
lemma ordermap_csquare_le:
- "[| Limit(K); x<K; y<K; z=succ(x Un y) |]
- ==> | ordermap(K*K, csquare_rel(K)) ` <x,y> | le |succ(z)| |*| |succ(z)|"
+ "[| Limit(K); x<K; y<K; z=succ(x \<union> y) |]
+ ==> | ordermap(K*K, csquare_rel(K)) ` <x,y> | \<le> |succ(z)| |*| |succ(z)|"
apply (unfold cmult_def)
apply (rule well_ord_rmult [THEN well_ord_lepoll_imp_Card_le])
apply (rule Ord_cardinal [THEN well_ord_Memrel])+
apply (subgoal_tac "z<K")
prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ)
-apply (rule ordermap_z_lt [THEN leI, THEN le_imp_lepoll, THEN lepoll_trans],
+apply (rule ordermap_z_lt [THEN leI, THEN le_imp_lepoll, THEN lepoll_trans],
assumption+)
apply (rule ordermap_eqpoll_pred [THEN eqpoll_imp_lepoll, THEN lepoll_trans])
apply (erule Limit_is_Ord [THEN well_ord_csquare])
@@ -573,10 +573,10 @@
apply (erule Ord_succ [THEN Ord_cardinal_eqpoll])+
done
-(*Kunen: "... so the order type <= K" *)
+(*Kunen: "... so the order type is @{text"\<le>"} K" *)
lemma ordertype_csquare_le:
- "[| InfCard(K); ALL y:K. InfCard(y) --> y |*| y = y |]
- ==> ordertype(K*K, csquare_rel(K)) le K"
+ "[| InfCard(K); \<forall>y\<in>K. InfCard(y) \<longrightarrow> y |*| y = y |]
+ ==> ordertype(K*K, csquare_rel(K)) \<le> K"
apply (frule InfCard_is_Card [THEN Card_is_Ord])
apply (rule all_lt_imp_le, assumption)
apply (erule well_ord_csquare [THEN Ord_ordertype])
@@ -587,16 +587,16 @@
apply (safe elim!: ltE)
apply (subgoal_tac "Ord (xa) & Ord (ya)")
prefer 2 apply (blast intro: Ord_in_Ord, clarify)
-(*??WHAT A MESS!*)
+(*??WHAT A MESS!*)
apply (rule InfCard_is_Limit [THEN ordermap_csquare_le, THEN lt_trans1],
- (assumption | rule refl | erule ltI)+)
-apply (rule_tac i = "xa Un ya" and j = nat in Ord_linear2,
+ (assumption | rule refl | erule ltI)+)
+apply (rule_tac i = "xa \<union> ya" and j = nat in Ord_linear2,
simp_all add: Ord_Un Ord_nat)
-prefer 2 (*case nat le (xa Un ya) *)
- apply (simp add: le_imp_subset [THEN nat_succ_eqpoll, THEN cardinal_cong]
+prefer 2 (*case @{term"nat \<le> (xa \<union> ya)"} *)
+ apply (simp add: le_imp_subset [THEN nat_succ_eqpoll, THEN cardinal_cong]
le_succ_iff InfCard_def Card_cardinal Un_least_lt Ord_Un
ltI nat_le_cardinal Ord_cardinal_le [THEN lt_trans1, THEN ltD])
-(*the finite case: xa Un ya < nat *)
+(*the finite case: @{term"xa \<union> ya < nat"} *)
apply (rule_tac j = nat in lt_trans2)
apply (simp add: lt_def nat_cmult_eq_mult nat_succI mult_type
nat_into_Card [THEN Card_cardinal_eq] Ord_nat)
@@ -607,14 +607,14 @@
lemma InfCard_csquare_eq: "InfCard(K) ==> K |*| K = K"
apply (frule InfCard_is_Card [THEN Card_is_Ord])
apply (erule rev_mp)
-apply (erule_tac i=K in trans_induct)
+apply (erule_tac i=K in trans_induct)
apply (rule impI)
apply (rule le_anti_sym)
apply (erule_tac [2] InfCard_is_Card [THEN cmult_square_le])
apply (rule ordertype_csquare_le [THEN [2] le_trans])
-apply (simp add: cmult_def Ord_cardinal_le
+apply (simp add: cmult_def Ord_cardinal_le
well_ord_csquare [THEN Ord_ordertype]
- well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll,
+ well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll,
THEN cardinal_cong], assumption+)
done
@@ -629,9 +629,9 @@
done
lemma InfCard_square_eqpoll: "InfCard(K) ==> K \<times> K \<approx> K"
-apply (rule well_ord_InfCard_square_eq)
- apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel])
-apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq])
+apply (rule well_ord_InfCard_square_eq)
+ apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel])
+apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq])
done
lemma Inf_Card_is_InfCard: "[| ~Finite(i); Card(i) |] ==> InfCard(i)"
@@ -639,7 +639,7 @@
subsubsection{*Toward's Kunen's Corollary 10.13 (1)*}
-lemma InfCard_le_cmult_eq: "[| InfCard(K); L le K; 0<L |] ==> K |*| L = K"
+lemma InfCard_le_cmult_eq: "[| InfCard(K); L \<le> K; 0<L |] ==> K |*| L = K"
apply (rule le_anti_sym)
prefer 2
apply (erule ltE, blast intro: cmult_le_self InfCard_is_Card)
@@ -649,12 +649,12 @@
done
(*Corollary 10.13 (1), for cardinal multiplication*)
-lemma InfCard_cmult_eq: "[| InfCard(K); InfCard(L) |] ==> K |*| L = K Un L"
+lemma InfCard_cmult_eq: "[| InfCard(K); InfCard(L) |] ==> K |*| L = K \<union> L"
apply (rule_tac i = K and j = L in Ord_linear_le)
apply (typecheck add: InfCard_is_Card Card_is_Ord)
apply (rule cmult_commute [THEN ssubst])
apply (rule Un_commute [THEN ssubst])
-apply (simp_all add: InfCard_is_Limit [THEN Limit_has_0] InfCard_le_cmult_eq
+apply (simp_all add: InfCard_is_Limit [THEN Limit_has_0] InfCard_le_cmult_eq
subset_Un_iff2 [THEN iffD1] le_imp_subset)
done
@@ -664,7 +664,7 @@
done
(*Corollary 10.13 (1), for cardinal addition*)
-lemma InfCard_le_cadd_eq: "[| InfCard(K); L le K |] ==> K |+| L = K"
+lemma InfCard_le_cadd_eq: "[| InfCard(K); L \<le> K |] ==> K |+| L = K"
apply (rule le_anti_sym)
prefer 2
apply (erule ltE, blast intro: cadd_le_self InfCard_is_Card)
@@ -673,7 +673,7 @@
apply (simp add: InfCard_cdouble_eq)
done
-lemma InfCard_cadd_eq: "[| InfCard(K); InfCard(L) |] ==> K |+| L = K Un L"
+lemma InfCard_cadd_eq: "[| InfCard(K); InfCard(L) |] ==> K |+| L = K \<union> L"
apply (rule_tac i = K and j = L in Ord_linear_le)
apply (typecheck add: InfCard_is_Card Card_is_Ord)
apply (rule cadd_commute [THEN ssubst])
@@ -704,10 +704,10 @@
(*Allows selective unfolding. Less work than deriving intro/elim rules*)
lemma jump_cardinal_iff:
- "i : jump_cardinal(K) <->
- (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))"
+ "i \<in> jump_cardinal(K) <->
+ (\<exists>r X. r \<subseteq> K*K & X \<subseteq> K & well_ord(X,r) & i = ordertype(X,r))"
apply (unfold jump_cardinal_def)
-apply (blast del: subsetI)
+apply (blast del: subsetI)
done
(*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*)
@@ -715,7 +715,7 @@
apply (rule Ord_jump_cardinal [THEN [2] ltI])
apply (rule jump_cardinal_iff [THEN iffD2])
apply (rule_tac x="Memrel(K)" in exI)
-apply (rule_tac x=K in exI)
+apply (rule_tac x=K in exI)
apply (simp add: ordertype_Memrel well_ord_Memrel)
apply (simp add: Memrel_def subset_iff)
done
@@ -723,10 +723,10 @@
(*The proof by contradiction: the bijection f yields a wellordering of X
whose ordertype is jump_cardinal(K). *)
lemma Card_jump_cardinal_lemma:
- "[| well_ord(X,r); r <= K * K; X <= K;
- f : bij(ordertype(X,r), jump_cardinal(K)) |]
- ==> jump_cardinal(K) : jump_cardinal(K)"
-apply (subgoal_tac "f O ordermap (X,r) : bij (X, jump_cardinal (K))")
+ "[| well_ord(X,r); r \<subseteq> K * K; X \<subseteq> K;
+ f \<in> bij(ordertype(X,r), jump_cardinal(K)) |]
+ ==> jump_cardinal(K) \<in> jump_cardinal(K)"
+apply (subgoal_tac "f O ordermap (X,r) \<in> bij (X, jump_cardinal (K))")
prefer 2 apply (blast intro: comp_bij ordermap_bij)
apply (rule jump_cardinal_iff [THEN iffD2])
apply (intro exI conjI)
@@ -760,13 +760,13 @@
lemma Ord_0_lt_csucc: "Ord(K) ==> 0 < csucc(K)"
by (blast intro: Ord_0_le lt_csucc lt_trans1)
-lemma csucc_le: "[| Card(L); K<L |] ==> csucc(K) le L"
+lemma csucc_le: "[| Card(L); K<L |] ==> csucc(K) \<le> L"
apply (unfold csucc_def)
apply (rule Least_le)
apply (blast intro: Card_is_Ord)+
done
-lemma lt_csucc_iff: "[| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K"
+lemma lt_csucc_iff: "[| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| \<le> K"
apply (rule iffI)
apply (rule_tac [2] Card_lt_imp_lt)
apply (erule_tac [2] lt_trans1)
@@ -774,21 +774,21 @@
apply (rule notI [THEN not_lt_imp_le])
apply (rule Card_cardinal [THEN csucc_le, THEN lt_trans1, THEN lt_irrefl], assumption)
apply (rule Ord_cardinal_le [THEN lt_trans1])
-apply (simp_all add: Ord_cardinal Card_is_Ord)
+apply (simp_all add: Ord_cardinal Card_is_Ord)
done
lemma Card_lt_csucc_iff:
- "[| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K"
+ "[| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' \<le> K"
by (simp add: lt_csucc_iff Card_cardinal_eq Card_is_Ord)
lemma InfCard_csucc: "InfCard(K) ==> InfCard(csucc(K))"
-by (simp add: InfCard_def Card_csucc Card_is_Ord
+by (simp add: InfCard_def Card_csucc Card_is_Ord
lt_csucc [THEN leI, THEN [2] le_trans])
subsubsection{*Removing elements from a finite set decreases its cardinality*}
-lemma Fin_imp_not_cons_lepoll: "A: Fin(U) ==> x~:A --> ~ cons(x,A) \<lesssim> A"
+lemma Fin_imp_not_cons_lepoll: "A: Fin(U) ==> x\<notin>A \<longrightarrow> ~ cons(x,A) \<lesssim> A"
apply (erule Fin_induct)
apply (simp add: lepoll_0_iff)
apply (subgoal_tac "cons (x,cons (xa,y)) = cons (xa,cons (x,y))")
@@ -797,7 +797,7 @@
done
lemma Finite_imp_cardinal_cons [simp]:
- "[| Finite(A); a~:A |] ==> |cons(a,A)| = succ(|A|)"
+ "[| Finite(A); a\<notin>A |] ==> |cons(a,A)| = succ(|A|)"
apply (unfold cardinal_def)
apply (rule Least_equality)
apply (fold cardinal_def)
@@ -827,29 +827,29 @@
apply (simp add: Finite_imp_succ_cardinal_Diff)
done
-lemma Finite_cardinal_in_nat [simp]: "Finite(A) ==> |A| : nat"
+lemma Finite_cardinal_in_nat [simp]: "Finite(A) ==> |A| \<in> nat"
apply (erule Finite_induct)
apply (auto simp add: cardinal_0 Finite_imp_cardinal_cons)
done
lemma card_Un_Int:
- "[|Finite(A); Finite(B)|] ==> |A| #+ |B| = |A Un B| #+ |A Int B|"
-apply (erule Finite_induct, simp)
+ "[|Finite(A); Finite(B)|] ==> |A| #+ |B| = |A \<union> B| #+ |A \<inter> B|"
+apply (erule Finite_induct, simp)
apply (simp add: Finite_Int cons_absorb Un_cons Int_cons_left)
done
-lemma card_Un_disjoint:
- "[|Finite(A); Finite(B); A Int B = 0|] ==> |A Un B| = |A| #+ |B|"
+lemma card_Un_disjoint:
+ "[|Finite(A); Finite(B); A \<inter> B = 0|] ==> |A \<union> B| = |A| #+ |B|"
by (simp add: Finite_Un card_Un_Int)
lemma card_partition [rule_format]:
- "Finite(C) ==>
- Finite (\<Union> C) -->
- (\<forall>c\<in>C. |c| = k) -->
- (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = 0) -->
+ "Finite(C) ==>
+ Finite (\<Union> C) \<longrightarrow>
+ (\<forall>c\<in>C. |c| = k) \<longrightarrow>
+ (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = 0) \<longrightarrow>
k #* |C| = |\<Union> C|"
apply (erule Finite_induct, auto)
-apply (subgoal_tac " x \<inter> \<Union>B = 0")
+apply (subgoal_tac " x \<inter> \<Union>B = 0")
apply (auto simp add: card_Un_disjoint Finite_Union
subset_Finite [of _ "\<Union> (cons(x,F))"])
done
@@ -866,12 +866,12 @@
apply (simp add: nat_cadd_eq_add [symmetric] cadd_def eqpoll_refl)
done
-lemma Ord_subset_natD [rule_format]: "Ord(i) ==> i <= nat --> i : nat | i=nat"
+lemma Ord_subset_natD [rule_format]: "Ord(i) ==> i \<subseteq> nat \<longrightarrow> i \<in> nat | i=nat"
apply (erule trans_induct3, auto)
apply (blast dest!: nat_le_Limit [THEN le_imp_subset])
done
-lemma Ord_nat_subset_into_Card: "[| Ord(i); i <= nat |] ==> Card(i)"
+lemma Ord_nat_subset_into_Card: "[| Ord(i); i \<subseteq> nat |] ==> Card(i)"
by (blast dest: Ord_subset_natD intro: Card_nat nat_into_Card)
lemma Finite_Diff_sing_eq_diff_1: "[| Finite(A); x:A |] ==> |A-{x}| = |A| #- 1"
@@ -889,7 +889,7 @@
done
lemma cardinal_lt_imp_Diff_not_0 [rule_format]:
- "Finite(B) ==> ALL A. |B|<|A| --> A - B ~= 0"
+ "Finite(B) ==> \<forall>A. |B|<|A| \<longrightarrow> A - B \<noteq> 0"
apply (erule Finite_induct, auto)
apply (case_tac "Finite (A)")
apply (subgoal_tac [2] "Finite (cons (x, B))")
@@ -900,100 +900,13 @@
apply (case_tac "x:A")
apply (subgoal_tac [2] "A - cons (x, B) = A - B")
apply auto
-apply (subgoal_tac "|A| le |cons (x, B) |")
+apply (subgoal_tac "|A| \<le> |cons (x, B) |")
prefer 2
- apply (blast dest: Finite_cons [THEN Finite_imp_well_ord]
+ apply (blast dest: Finite_cons [THEN Finite_imp_well_ord]
intro: well_ord_lepoll_imp_Card_le subset_imp_lepoll)
apply (auto simp add: Finite_imp_cardinal_cons)
apply (auto dest!: Finite_cardinal_in_nat simp add: le_iff)
apply (blast intro: lt_trans)
done
-
-ML{*
-val InfCard_def = @{thm InfCard_def};
-val cmult_def = @{thm cmult_def};
-val cadd_def = @{thm cadd_def};
-val jump_cardinal_def = @{thm jump_cardinal_def};
-val csucc_def = @{thm csucc_def};
-
-val sum_commute_eqpoll = @{thm sum_commute_eqpoll};
-val cadd_commute = @{thm cadd_commute};
-val sum_assoc_eqpoll = @{thm sum_assoc_eqpoll};
-val well_ord_cadd_assoc = @{thm well_ord_cadd_assoc};
-val sum_0_eqpoll = @{thm sum_0_eqpoll};
-val cadd_0 = @{thm cadd_0};
-val sum_lepoll_self = @{thm sum_lepoll_self};
-val cadd_le_self = @{thm cadd_le_self};
-val sum_lepoll_mono = @{thm sum_lepoll_mono};
-val cadd_le_mono = @{thm cadd_le_mono};
-val eq_imp_not_mem = @{thm eq_imp_not_mem};
-val sum_succ_eqpoll = @{thm sum_succ_eqpoll};
-val nat_cadd_eq_add = @{thm nat_cadd_eq_add};
-val prod_commute_eqpoll = @{thm prod_commute_eqpoll};
-val cmult_commute = @{thm cmult_commute};
-val prod_assoc_eqpoll = @{thm prod_assoc_eqpoll};
-val well_ord_cmult_assoc = @{thm well_ord_cmult_assoc};
-val sum_prod_distrib_eqpoll = @{thm sum_prod_distrib_eqpoll};
-val well_ord_cadd_cmult_distrib = @{thm well_ord_cadd_cmult_distrib};
-val prod_0_eqpoll = @{thm prod_0_eqpoll};
-val cmult_0 = @{thm cmult_0};
-val prod_singleton_eqpoll = @{thm prod_singleton_eqpoll};
-val cmult_1 = @{thm cmult_1};
-val prod_lepoll_self = @{thm prod_lepoll_self};
-val cmult_le_self = @{thm cmult_le_self};
-val prod_lepoll_mono = @{thm prod_lepoll_mono};
-val cmult_le_mono = @{thm cmult_le_mono};
-val prod_succ_eqpoll = @{thm prod_succ_eqpoll};
-val nat_cmult_eq_mult = @{thm nat_cmult_eq_mult};
-val cmult_2 = @{thm cmult_2};
-val sum_lepoll_prod = @{thm sum_lepoll_prod};
-val lepoll_imp_sum_lepoll_prod = @{thm lepoll_imp_sum_lepoll_prod};
-val nat_cons_lepoll = @{thm nat_cons_lepoll};
-val nat_cons_eqpoll = @{thm nat_cons_eqpoll};
-val nat_succ_eqpoll = @{thm nat_succ_eqpoll};
-val InfCard_nat = @{thm InfCard_nat};
-val InfCard_is_Card = @{thm InfCard_is_Card};
-val InfCard_Un = @{thm InfCard_Un};
-val InfCard_is_Limit = @{thm InfCard_is_Limit};
-val ordermap_eqpoll_pred = @{thm ordermap_eqpoll_pred};
-val ordermap_z_lt = @{thm ordermap_z_lt};
-val InfCard_le_cmult_eq = @{thm InfCard_le_cmult_eq};
-val InfCard_cmult_eq = @{thm InfCard_cmult_eq};
-val InfCard_cdouble_eq = @{thm InfCard_cdouble_eq};
-val InfCard_le_cadd_eq = @{thm InfCard_le_cadd_eq};
-val InfCard_cadd_eq = @{thm InfCard_cadd_eq};
-val Ord_jump_cardinal = @{thm Ord_jump_cardinal};
-val jump_cardinal_iff = @{thm jump_cardinal_iff};
-val K_lt_jump_cardinal = @{thm K_lt_jump_cardinal};
-val Card_jump_cardinal = @{thm Card_jump_cardinal};
-val csucc_basic = @{thm csucc_basic};
-val Card_csucc = @{thm Card_csucc};
-val lt_csucc = @{thm lt_csucc};
-val Ord_0_lt_csucc = @{thm Ord_0_lt_csucc};
-val csucc_le = @{thm csucc_le};
-val lt_csucc_iff = @{thm lt_csucc_iff};
-val Card_lt_csucc_iff = @{thm Card_lt_csucc_iff};
-val InfCard_csucc = @{thm InfCard_csucc};
-val Finite_into_Fin = @{thm Finite_into_Fin};
-val Fin_into_Finite = @{thm Fin_into_Finite};
-val Finite_Fin_iff = @{thm Finite_Fin_iff};
-val Finite_Un = @{thm Finite_Un};
-val Finite_Union = @{thm Finite_Union};
-val Finite_induct = @{thm Finite_induct};
-val Fin_imp_not_cons_lepoll = @{thm Fin_imp_not_cons_lepoll};
-val Finite_imp_cardinal_cons = @{thm Finite_imp_cardinal_cons};
-val Finite_imp_succ_cardinal_Diff = @{thm Finite_imp_succ_cardinal_Diff};
-val Finite_imp_cardinal_Diff = @{thm Finite_imp_cardinal_Diff};
-val nat_implies_well_ord = @{thm nat_implies_well_ord};
-val nat_sum_eqpoll_sum = @{thm nat_sum_eqpoll_sum};
-val Diff_sing_Finite = @{thm Diff_sing_Finite};
-val Diff_Finite = @{thm Diff_Finite};
-val Ord_subset_natD = @{thm Ord_subset_natD};
-val Ord_nat_subset_into_Card = @{thm Ord_nat_subset_into_Card};
-val Finite_cardinal_in_nat = @{thm Finite_cardinal_in_nat};
-val Finite_Diff_sing_eq_diff_1 = @{thm Finite_Diff_sing_eq_diff_1};
-val cardinal_lt_imp_Diff_not_0 = @{thm cardinal_lt_imp_Diff_not_0};
-*}
-
end
--- a/src/ZF/Cardinal_AC.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Cardinal_AC.thy Tue Mar 06 15:15:49 2012 +0000
@@ -29,11 +29,11 @@
by (blast intro: cardinal_cong cardinal_eqE)
lemma cardinal_disjoint_Un:
- "[| |A|=|B|; |C|=|D|; A Int C = 0; B Int D = 0 |]
- ==> |A Un C| = |B Un D|"
+ "[| |A|=|B|; |C|=|D|; A \<inter> C = 0; B \<inter> D = 0 |]
+ ==> |A \<union> C| = |B \<union> D|"
by (simp add: cardinal_eqpoll_iff eqpoll_disjoint_Un)
-lemma lepoll_imp_Card_le: "A lepoll B ==> |A| le |B|"
+lemma lepoll_imp_Card_le: "A lepoll B ==> |A| \<le> |B|"
apply (rule AC_well_ord [THEN exE])
apply (erule well_ord_lepoll_imp_Card_le, assumption)
done
@@ -67,21 +67,21 @@
subsection {*The relationship between cardinality and le-pollence*}
-lemma Card_le_imp_lepoll: "|A| le |B| ==> A lepoll B"
+lemma Card_le_imp_lepoll: "|A| \<le> |B| ==> A lepoll B"
apply (rule cardinal_eqpoll
[THEN eqpoll_sym, THEN eqpoll_imp_lepoll, THEN lepoll_trans])
apply (erule le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_trans])
apply (rule cardinal_eqpoll [THEN eqpoll_imp_lepoll])
done
-lemma le_Card_iff: "Card(K) ==> |A| le K <-> A lepoll K"
-apply (erule Card_cardinal_eq [THEN subst], rule iffI,
+lemma le_Card_iff: "Card(K) ==> |A| \<le> K <-> A lepoll K"
+apply (erule Card_cardinal_eq [THEN subst], rule iffI,
erule Card_le_imp_lepoll)
-apply (erule lepoll_imp_Card_le)
+apply (erule lepoll_imp_Card_le)
done
lemma cardinal_0_iff_0 [simp]: "|A| = 0 <-> A = 0";
-apply auto
+apply auto
apply (drule cardinal_0 [THEN ssubst])
apply (blast intro: eqpoll_0_iff [THEN iffD1] cardinal_eqpoll_iff [THEN iffD1])
done
@@ -90,7 +90,7 @@
apply (cut_tac A = "A" in cardinal_eqpoll)
apply (auto simp add: eqpoll_iff)
apply (blast intro: lesspoll_trans2 lt_Card_imp_lesspoll Card_cardinal)
-apply (force intro: cardinal_lt_imp_lt lesspoll_cardinal_lt lesspoll_trans2
+apply (force intro: cardinal_lt_imp_lt lesspoll_cardinal_lt lesspoll_trans2
simp add: cardinal_idem)
done
@@ -101,17 +101,17 @@
subsection{*Other Applications of AC*}
-lemma surj_implies_inj: "f: surj(X,Y) ==> EX g. g: inj(Y,X)"
+lemma surj_implies_inj: "f: surj(X,Y) ==> \<exists>g. g: inj(Y,X)"
apply (unfold surj_def)
apply (erule CollectE)
apply (rule_tac A1 = Y and B1 = "%y. f-``{y}" in AC_Pi [THEN exE])
apply (fast elim!: apply_Pair)
-apply (blast dest: apply_type Pi_memberD
+apply (blast dest: apply_type Pi_memberD
intro: apply_equality Pi_type f_imp_injective)
done
(*Kunen's Lemma 10.20*)
-lemma surj_implies_cardinal_le: "f: surj(X,Y) ==> |Y| le |X|"
+lemma surj_implies_cardinal_le: "f: surj(X,Y) ==> |Y| \<le> |X|"
apply (rule lepoll_imp_Card_le)
apply (erule surj_implies_inj [THEN exE])
apply (unfold lepoll_def)
@@ -120,7 +120,7 @@
(*Kunen's Lemma 10.21*)
lemma cardinal_UN_le:
- "[| InfCard(K); ALL i:K. |X(i)| le K |] ==> |\<Union>i\<in>K. X(i)| le K"
+ "[| InfCard(K); \<forall>i\<in>K. |X(i)| \<le> K |] ==> |\<Union>i\<in>K. X(i)| \<le> K"
apply (simp add: InfCard_is_Card le_Card_iff)
apply (rule lepoll_trans)
prefer 2
@@ -131,12 +131,12 @@
apply (erule AC_ball_Pi [THEN exE])
apply (rule exI)
(*Lemma needed in both subgoals, for a fixed z*)
-apply (subgoal_tac "ALL z: (\<Union>i\<in>K. X (i)). z: X (LEAST i. z:X (i)) &
- (LEAST i. z:X (i)) : K")
+apply (subgoal_tac "\<forall>z\<in>(\<Union>i\<in>K. X (i)). z: X (LEAST i. z:X (i)) &
+ (LEAST i. z:X (i)) \<in> K")
prefer 2
apply (fast intro!: Least_le [THEN lt_trans1, THEN ltD] ltI
elim!: LeastI Ord_in_Ord)
-apply (rule_tac c = "%z. <LEAST i. z:X (i), f ` (LEAST i. z:X (i)) ` z>"
+apply (rule_tac c = "%z. <LEAST i. z:X (i), f ` (LEAST i. z:X (i)) ` z>"
and d = "%<i,j>. converse (f`i) ` j" in lam_injective)
(*Instantiate the lemma proved above*)
by (blast intro: inj_is_fun [THEN apply_type] dest: apply_type, force)
@@ -144,14 +144,14 @@
(*The same again, using csucc*)
lemma cardinal_UN_lt_csucc:
- "[| InfCard(K); ALL i:K. |X(i)| < csucc(K) |]
+ "[| InfCard(K); \<forall>i\<in>K. |X(i)| < csucc(K) |]
==> |\<Union>i\<in>K. X(i)| < csucc(K)"
by (simp add: Card_lt_csucc_iff cardinal_UN_le InfCard_is_Card Card_cardinal)
(*The same again, for a union of ordinals. In use, j(i) is a bit like rank(i),
the least ordinal j such that i:Vfrom(A,j). *)
lemma cardinal_UN_Ord_lt_csucc:
- "[| InfCard(K); ALL i:K. j(i) < csucc(K) |]
+ "[| InfCard(K); \<forall>i\<in>K. j(i) < csucc(K) |]
==> (\<Union>i\<in>K. j(i)) < csucc(K)"
apply (rule cardinal_UN_lt_csucc [THEN Card_lt_imp_lt], assumption)
apply (blast intro: Ord_cardinal_le [THEN lt_trans1] elim: ltE)
@@ -164,10 +164,10 @@
set need not be a cardinal. Surprisingly complicated proof!
**)
-(*Work backwards along the injection from W into K, given that W~=0.*)
+(*Work backwards along the injection from W into K, given that @{term"W\<noteq>0"}.*)
lemma inj_UN_subset:
- "[| f: inj(A,B); a:A |] ==>
- (\<Union>x\<in>A. C(x)) <= (\<Union>y\<in>B. C(if y: range(f) then converse(f)`y else a))"
+ "[| f: inj(A,B); a:A |] ==>
+ (\<Union>x\<in>A. C(x)) \<subseteq> (\<Union>y\<in>B. C(if y: range(f) then converse(f)`y else a))"
apply (rule UN_least)
apply (rule_tac x1= "f`x" in subset_trans [OF _ UN_upper])
apply (simp add: inj_is_fun [THEN apply_rangeI])
@@ -177,15 +177,15 @@
(*Simpler to require |W|=K; we'd have a bijection; but the theorem would
be weaker.*)
lemma le_UN_Ord_lt_csucc:
- "[| InfCard(K); |W| le K; ALL w:W. j(w) < csucc(K) |]
+ "[| InfCard(K); |W| \<le> K; \<forall>w\<in>W. j(w) < csucc(K) |]
==> (\<Union>w\<in>W. j(w)) < csucc(K)"
apply (case_tac "W=0")
(*solve the easy 0 case*)
- apply (simp add: InfCard_is_Card Card_is_Ord [THEN Card_csucc]
+ apply (simp add: InfCard_is_Card Card_is_Ord [THEN Card_csucc]
Card_is_Ord Ord_0_lt_csucc)
apply (simp add: InfCard_is_Card le_Card_iff lepoll_def)
apply (safe intro!: equalityI)
-apply (erule swap)
+apply (erule swap)
apply (rule lt_subset_trans [OF inj_UN_subset cardinal_UN_Ord_lt_csucc], assumption+)
apply (simp add: inj_converse_fun [THEN apply_type])
apply (blast intro!: Ord_UN elim: ltE)
--- a/src/ZF/Epsilon.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Epsilon.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,23 +9,23 @@
definition
eclose :: "i=>i" where
- "eclose(A) == \<Union>n\<in>nat. nat_rec(n, A, %m r. Union(r))"
+ "eclose(A) == \<Union>n\<in>nat. nat_rec(n, A, %m r. \<Union>(r))"
definition
transrec :: "[i, [i,i]=>i] =>i" where
"transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)"
-
+
definition
rank :: "i=>i" where
"rank(a) == transrec(a, %x f. \<Union>y\<in>x. succ(f`y))"
definition
transrec2 :: "[i, i, [i,i]=>i] =>i" where
- "transrec2(k, a, b) ==
- transrec(k,
- %i r. if(i=0, a,
- if(EX j. i=succ(j),
- b(THE j. i=succ(j), r`(THE j. i=succ(j))),
+ "transrec2(k, a, b) ==
+ transrec(k,
+ %i r. if(i=0, a,
+ if(\<exists>j. i=succ(j),
+ b(THE j. i=succ(j), r`(THE j. i=succ(j))),
\<Union>j<i. r`j)))"
definition
@@ -39,7 +39,7 @@
subsection{*Basic Closure Properties*}
-lemma arg_subset_eclose: "A <= eclose(A)"
+lemma arg_subset_eclose: "A \<subseteq> eclose(A)"
apply (unfold eclose_def)
apply (rule nat_rec_0 [THEN equalityD2, THEN subset_trans])
apply (rule nat_0I [THEN UN_upper])
@@ -56,19 +56,19 @@
apply (erule UnionI, assumption)
done
-(* x : eclose(A) ==> x <= eclose(A) *)
-lemmas eclose_subset =
+(* @{term"x \<in> eclose(A) ==> x \<subseteq> eclose(A)"} *)
+lemmas eclose_subset =
Transset_eclose [unfolded Transset_def, THEN bspec]
-(* [| A : eclose(B); c : A |] ==> c : eclose(B) *)
+(* @{term"[| A \<in> eclose(B); c \<in> A |] ==> c \<in> eclose(B)"} *)
lemmas ecloseD = eclose_subset [THEN subsetD]
lemmas arg_in_eclose_sing = arg_subset_eclose [THEN singleton_subsetD]
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); ALL y:x. P(y) |] ==> P(x)
- |] ==> P(a)
+ [| a: eclose(A); !!x. [| x: eclose(A); \<forall>y\<in>x. P(y) |] ==> P(x)
+ |] ==> P(a)
*)
lemmas eclose_induct =
Transset_induct [OF _ Transset_eclose, induct set: eclose]
@@ -76,37 +76,37 @@
(*Epsilon induction*)
lemma eps_induct:
- "[| !!x. ALL y:x. P(y) ==> P(x) |] ==> P(a)"
-by (rule arg_in_eclose_sing [THEN eclose_induct], blast)
+ "[| !!x. \<forall>y\<in>x. P(y) ==> P(x) |] ==> P(a)"
+by (rule arg_in_eclose_sing [THEN eclose_induct], blast)
subsection{*Leastness of @{term eclose}*}
(** 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. Union(r)) <= X"
+lemma eclose_least_lemma:
+ "[| Transset(X); A<=X; n: nat |] ==> nat_rec(n, A, %m r. \<Union>(r)) \<subseteq> X"
apply (unfold Transset_def)
-apply (erule nat_induct)
+apply (erule nat_induct)
apply (simp add: nat_rec_0)
apply (simp add: nat_rec_succ, blast)
done
-lemma eclose_least:
- "[| Transset(X); A<=X |] ==> eclose(A) <= X"
+lemma eclose_least:
+ "[| Transset(X); A<=X |] ==> eclose(A) \<subseteq> X"
apply (unfold eclose_def)
apply (rule eclose_least_lemma [THEN UN_least], assumption+)
done
(*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
- apply (unfold Transset_def)
+ apply (unfold Transset_def)
apply (blast intro: ecloseD)
apply (blast intro: arg_subset_eclose [THEN subsetD])
done
@@ -117,10 +117,10 @@
done
text{*A transitive set either is empty or contains the empty set.*}
-lemma Transset_0_lemma [rule_format]: "Transset(A) ==> x\<in>A --> 0\<in>A";
-apply (simp add: Transset_def)
-apply (rule_tac a=x in eps_induct, clarify)
-apply (drule bspec, assumption)
+lemma Transset_0_lemma [rule_format]: "Transset(A) ==> x\<in>A \<longrightarrow> 0\<in>A";
+apply (simp add: Transset_def)
+apply (rule_tac a=x in eps_induct, clarify)
+apply (drule bspec, assumption)
apply (case_tac "x=0", auto)
done
@@ -132,28 +132,28 @@
(*Unused...*)
lemma mem_eclose_trans: "[| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"
-by (rule eclose_least [OF Transset_eclose eclose_subset, THEN subsetD],
+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})"
-by (rule eclose_least [OF Transset_eclose singleton_subsetI, THEN subsetD],
+by (rule eclose_least [OF Transset_eclose singleton_subsetI, THEN subsetD],
assumption+)
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"
-by (simp add: lt_def Ord_def under_Memrel)
+by (simp add: lt_def Ord_def under_Memrel)
-(* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *)
+(* @{term"j \<in> eclose(A) ==> Memrel(eclose(A)) -`` j = j"} *)
lemmas under_Memrel_eclose = Transset_eclose [THEN under_Memrel]
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)
@@ -161,13 +161,13 @@
apply (simp add: under_Memrel_eclose mem_eclose_sing_trans [of _ j i])
done
-lemma wfrec_eclose_eq2:
+lemma wfrec_eclose_eq2:
"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
-lemma transrec: "transrec(a,H) = H(a, lam x:a. transrec(x,H))"
+lemma transrec: "transrec(a,H) = H(a, \<lambda>x\<in>a. transrec(x,H))"
apply (unfold transrec_def)
apply (rule wfrec_ssubst)
apply (simp add: wfrec_eclose_eq2 arg_in_eclose_sing under_Memrel_eclose)
@@ -175,44 +175,44 @@
(*Avoids explosions in proofs; resolve it with a meta-level definition.*)
lemma def_transrec:
- "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, lam x:a. f(x))"
+ "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, \<lambda>x\<in>a. f(x))"
apply simp
apply (rule transrec)
done
lemma transrec_type:
- "[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) : B(x) |]
- ==> transrec(a,H) : B(a)"
+ "[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) \<in> B(x) |]
+ ==> transrec(a,H) \<in> B(a)"
apply (rule_tac i = a in arg_in_eclose_sing [THEN eclose_induct])
apply (subst transrec)
-apply (simp add: lam_type)
+apply (simp add: lam_type)
done
-lemma eclose_sing_Ord: "Ord(i) ==> eclose({i}) <= succ(i)"
+lemma eclose_sing_Ord: "Ord(i) ==> eclose({i}) \<subseteq> succ(i)"
apply (erule Ord_is_Transset [THEN Transset_succ, THEN eclose_least])
apply (rule succI1 [THEN singleton_subsetI])
done
-lemma succ_subset_eclose_sing: "succ(i) <= eclose({i})"
-apply (insert arg_subset_eclose [of "{i}"], simp)
-apply (frule eclose_subset, blast)
+lemma succ_subset_eclose_sing: "succ(i) \<subseteq> eclose({i})"
+apply (insert arg_subset_eclose [of "{i}"], simp)
+apply (frule eclose_subset, blast)
done
lemma eclose_sing_Ord_eq: "Ord(i) ==> eclose({i}) = succ(i)"
apply (rule equalityI)
-apply (erule eclose_sing_Ord)
-apply (rule succ_subset_eclose_sing)
+apply (erule eclose_sing_Ord)
+apply (rule succ_subset_eclose_sing)
done
lemma Ord_transrec_type:
assumes jini: "j: i"
and ordi: "Ord(i)"
- and minor: " !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) : B(x)"
- shows "transrec(j,H) : B(j)"
+ and minor: " !!x u. [| x: i; u: Pi(x,B) |] ==> H(x,u) \<in> B(x)"
+ shows "transrec(j,H) \<in> B(j)"
apply (rule transrec_type)
apply (insert jini ordi)
apply (blast intro!: minor
- intro: Ord_trans
+ intro: Ord_trans
dest: Ord_in_Ord [THEN eclose_sing_Ord, THEN subsetD])
done
@@ -223,7 +223,7 @@
by (subst rank_def [THEN def_transrec], simp)
lemma Ord_rank [simp]: "Ord(rank(a))"
-apply (rule_tac a=a in eps_induct)
+apply (rule_tac a=a in eps_induct)
apply (subst rank)
apply (rule Ord_succ [THEN Ord_UN])
apply (erule bspec, assumption)
@@ -247,9 +247,9 @@
apply (erule rank_lt [THEN lt_trans], assumption)
done
-lemma rank_mono: "a<=b ==> rank(a) le rank(b)"
+lemma rank_mono: "a<=b ==> rank(a) \<le> rank(b)"
apply (rule subset_imp_le)
-apply (auto simp add: rank [of a] rank [of b])
+apply (auto simp add: rank [of a] rank [of b])
done
lemma rank_Pow: "rank(Pow(a)) = succ(rank(a))"
@@ -270,7 +270,7 @@
apply (erule rank_lt [THEN leI, THEN succ_leI, THEN le_imp_subset])
done
-lemma rank_Union: "rank(Union(A)) = (\<Union>x\<in>A. rank(x))"
+lemma rank_Union: "rank(\<Union>(A)) = (\<Union>x\<in>A. rank(x))"
apply (rule equalityI)
apply (rule_tac [2] rank_mono [THEN le_imp_subset, THEN UN_least])
apply (erule_tac [2] Union_upper)
@@ -308,13 +308,13 @@
"P(a) ==> (THE x. P(x)) = (if (EX!x. P(x)) then a else 0)"
by (simp add: the_0 the_equality2)
-(*The first premise not only fixs i but ensures f~=0.
- The second premise is now essential. Consider otherwise the relation
- r = {<0,0>,<0,1>,<0,2>,...}. Then f`0 = Union(f``{0}) = Union(nat) = nat,
+(*The first premise not only fixs i but ensures @{term"f\<noteq>0"}.
+ The second premise is now essential. Consider otherwise the relation
+ r = {<0,0>,<0,1>,<0,2>,...}. Then f`0 = \<Union>(f``{0}) = \<Union>(nat) = nat,
whose rank equals that of r.*)
-lemma rank_apply: "[|i : domain(f); function(f)|] ==> rank(f`i) < rank(f)"
-apply clarify
-apply (simp add: function_apply_equality)
+lemma rank_apply: "[|i \<in> domain(f); function(f)|] ==> rank(f`i) < rank(f)"
+apply clarify
+apply (simp add: function_apply_equality)
apply (blast intro: lt_trans rank_lt rank_pair2)
done
@@ -326,7 +326,7 @@
apply (erule arg_into_eclose [THEN eclose_subset])
done
-lemma eclose_mono: "A<=B ==> eclose(A) <= eclose(B)"
+lemma eclose_mono: "A<=B ==> eclose(A) \<subseteq> eclose(B)"
apply (rule Transset_eclose [THEN eclose_least])
apply (erule subset_trans)
apply (rule arg_subset_eclose)
@@ -340,7 +340,7 @@
apply (rule arg_subset_eclose)
done
-(** Transfinite recursion for definitions based on the
+(** Transfinite recursion for definitions based on the
three cases of ordinals **)
lemma transrec2_0 [simp]: "transrec2(0,a,b) = a"
@@ -354,14 +354,14 @@
lemma transrec2_Limit:
"Limit(i) ==> transrec2(i,a,b) = (\<Union>j<i. transrec2(j,a,b))"
apply (rule transrec2_def [THEN def_transrec, THEN trans])
-apply (auto simp add: OUnion_def)
+apply (auto simp add: OUnion_def)
done
lemma def_transrec2:
"(!!x. f(x)==transrec2(x,a,b))
- ==> f(0) = a &
- f(succ(i)) = b(i, f(i)) &
- (Limit(K) --> f(K) = (\<Union>j<K. f(j)))"
+ ==> f(0) = a &
+ f(succ(i)) = b(i, f(i)) &
+ (Limit(K) \<longrightarrow> f(K) = (\<Union>j<K. f(j)))"
by (simp add: transrec2_Limit)
@@ -390,10 +390,10 @@
done
lemma rec_type:
- "[| n: nat;
- a: C(0);
+ "[| n: nat;
+ a: C(0);
!!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) |]
- ==> rec(n,a,b) : C(n)"
+ ==> rec(n,a,b) \<in> C(n)"
by (erule nat_induct, auto)
end
--- a/src/ZF/EquivClass.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/EquivClass.thy Tue Mar 06 15:15:49 2012 +0000
@@ -13,12 +13,12 @@
definition
congruent :: "[i,i=>i]=>o" where
- "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)"
+ "congruent(r,b) == \<forall>y z. <y,z>:r \<longrightarrow> b(y)=b(z)"
definition
congruent2 :: "[i,i,[i,i]=>i]=>o" where
- "congruent2(r1,r2,b) == ALL y1 z1 y2 z2.
- <y1,z1>:r1 --> <y2,z2>:r2 --> b(y1,y2) = b(z1,z2)"
+ "congruent2(r1,r2,b) == \<forall>y1 z1 y2 z2.
+ <y1,z1>:r1 \<longrightarrow> <y2,z2>:r2 \<longrightarrow> b(y1,y2) = b(z1,z2)"
abbreviation
RESPECTS ::"[i=>i, i] => o" (infixr "respects" 80) where
@@ -36,11 +36,11 @@
(** first half: equiv(A,r) ==> converse(r) O r = r **)
lemma sym_trans_comp_subset:
- "[| sym(r); trans(r) |] ==> converse(r) O r <= r"
+ "[| sym(r); trans(r) |] ==> converse(r) O r \<subseteq> r"
by (unfold trans_def sym_def, blast)
lemma refl_comp_subset:
- "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r"
+ "[| refl(A,r); r \<subseteq> A*A |] ==> r \<subseteq> converse(r) O r"
by (unfold refl_def, blast)
lemma equiv_comp_eq:
@@ -54,14 +54,14 @@
"[| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"
apply (unfold equiv_def refl_def sym_def trans_def)
apply (erule equalityE)
-apply (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r", blast+)
+apply (subgoal_tac "\<forall>x y. <x,y> \<in> r \<longrightarrow> <y,x> \<in> r", blast+)
done
(** Equivalence classes **)
(*Lemma for the next result*)
lemma equiv_class_subset:
- "[| sym(r); trans(r); <a,b>: r |] ==> r``{a} <= r``{b}"
+ "[| sym(r); trans(r); <a,b>: r |] ==> r``{a} \<subseteq> r``{b}"
by (unfold trans_def sym_def, blast)
lemma equiv_class_eq:
@@ -77,7 +77,7 @@
(*Lemma for the next result*)
lemma subset_equiv_class:
- "[| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> <a,b>: r"
+ "[| equiv(A,r); r``{b} \<subseteq> r``{a}; b: A |] ==> <a,b>: r"
by (unfold equiv_def refl_def, blast)
lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> <a,b>: r"
@@ -85,10 +85,10 @@
(*thus r``{a} = r``{b} as well*)
lemma equiv_class_nondisjoint:
- "[| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> <a,b>: r"
+ "[| equiv(A,r); x: (r``{a} \<inter> r``{b}) |] ==> <a,b>: r"
by (unfold equiv_def trans_def sym_def, blast)
-lemma equiv_type: "equiv(A,r) ==> r <= A*A"
+lemma equiv_type: "equiv(A,r) ==> r \<subseteq> A*A"
by (unfold equiv_def, blast)
lemma equiv_class_eq_iff:
@@ -113,11 +113,11 @@
by (unfold quotient_def, blast)
lemma Union_quotient:
- "equiv(A,r) ==> Union(A//r) = A"
+ "equiv(A,r) ==> \<Union>(A//r) = A"
by (unfold equiv_def refl_def quotient_def, blast)
lemma quotient_disj:
- "[| equiv(A,r); X: A//r; Y: A//r |] ==> X=Y | (X Int Y <= 0)"
+ "[| equiv(A,r); X: A//r; Y: A//r |] ==> X=Y | (X \<inter> Y \<subseteq> 0)"
apply (unfold quotient_def)
apply (safe intro!: equiv_class_eq, assumption)
apply (unfold equiv_def trans_def sym_def, blast)
@@ -130,17 +130,17 @@
(*Conversion rule*)
lemma UN_equiv_class:
- "[| equiv(A,r); b respects r; a: A |] ==> (UN x:r``{a}. b(x)) = b(a)"
-apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)")
+ "[| equiv(A,r); b respects r; a: A |] ==> (\<Union>x\<in>r``{a}. b(x)) = b(a)"
+apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)")
apply simp
- apply (blast intro: equiv_class_self)
+ apply (blast intro: equiv_class_self)
apply (unfold equiv_def sym_def congruent_def, blast)
done
-(*type checking of UN x:r``{a}. b(x) *)
+(*type checking of @{term"\<Union>x\<in>r``{a}. b(x)"} *)
lemma UN_equiv_class_type:
- "[| equiv(A,r); b respects r; X: A//r; !!x. x : A ==> b(x) : B |]
- ==> (UN x:X. b(x)) : B"
+ "[| equiv(A,r); b respects r; X: A//r; !!x. x \<in> A ==> b(x) \<in> B |]
+ ==> (\<Union>x\<in>X. b(x)) \<in> B"
apply (unfold quotient_def, safe)
apply (simp (no_asm_simp) add: UN_equiv_class)
done
@@ -150,12 +150,12 @@
*)
lemma UN_equiv_class_inject:
"[| equiv(A,r); b respects r;
- (UN x:X. b(x))=(UN y:Y. b(y)); X: A//r; Y: A//r;
+ (\<Union>x\<in>X. b(x))=(\<Union>y\<in>Y. b(y)); X: A//r; Y: A//r;
!!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |]
==> X=Y"
apply (unfold quotient_def, safe)
apply (rule equiv_class_eq, assumption)
-apply (simp add: UN_equiv_class [of A r b])
+apply (simp add: UN_equiv_class [of A r b])
done
@@ -170,7 +170,7 @@
congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
apply (unfold congruent_def, safe)
apply (frule equiv_type [THEN subsetD], assumption)
-apply clarify
+apply clarify
apply (simp add: UN_equiv_class congruent2_implies_congruent)
apply (unfold congruent2_def equiv_def refl_def, blast)
done
@@ -185,10 +185,10 @@
lemma UN_equiv_class_type2:
"[| equiv(A,r); b respects2 r;
X1: A//r; X2: A//r;
- !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B
- |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"
+ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) \<in> B
+ |] ==> (\<Union>x1\<in>X1. \<Union>x2\<in>X2. b(x1,x2)) \<in> B"
apply (unfold quotient_def, safe)
-apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
+apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN
congruent2_implies_congruent quotientI)
done
@@ -196,12 +196,12 @@
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
than the direct proof*)
lemma congruent2I:
- "[| equiv(A1,r1); equiv(A2,r2);
+ "[| equiv(A1,r1); equiv(A2,r2);
!! y z w. [| w \<in> A2; <y,z> \<in> r1 |] ==> b(y,w) = b(z,w);
!! y z w. [| w \<in> A1; <y,z> \<in> r2 |] ==> b(w,y) = b(w,z)
|] ==> congruent2(r1,r2,b)"
apply (unfold congruent2_def equiv_def refl_def, safe)
-apply (blast intro: trans)
+apply (blast intro: trans)
done
lemma congruent2_commuteI:
@@ -209,11 +209,11 @@
and commute: "!! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y)"
and congt: "!! y z w. [| w: A; <y,z>: r |] ==> b(w,y) = b(w,z)"
shows "b respects2 r"
-apply (insert equivA [THEN equiv_type, THEN subsetD])
+apply (insert equivA [THEN equiv_type, THEN subsetD])
apply (rule congruent2I [OF equivA equivA])
apply (rule commute [THEN trans])
apply (rule_tac [3] commute [THEN trans, symmetric])
-apply (rule_tac [5] sym)
+apply (rule_tac [5] sym)
apply (blast intro: congt)+
done
@@ -222,12 +222,12 @@
"[| 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. UN z: Z. b(w,z))"
+ |] ==> congruent(r, %w. \<Union>z\<in>Z. b(w,z))"
apply (simp (no_asm) add: congruent_def)
apply (safe elim!: quotientE)
apply (frule equiv_type [THEN subsetD], assumption)
-apply (simp add: UN_equiv_class [of A r])
-apply (simp add: congruent_def)
+apply (simp add: UN_equiv_class [of A r])
+apply (simp add: congruent_def)
done
end
--- a/src/ZF/Finite.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Finite.thy Tue Mar 06 15:15:49 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) \<subseteq> surj(b,b)
*)
header{*Finite Powerset Operator and Finite Function Space*}
@@ -22,41 +22,41 @@
FiniteFun :: "[i,i]=>i" ("(_ -||>/ _)" [61, 60] 60)
inductive
- domains "Fin(A)" <= "Pow(A)"
+ domains "Fin(A)" \<subseteq> "Pow(A)"
intros
- emptyI: "0 : Fin(A)"
- consI: "[| a: A; b: Fin(A) |] ==> cons(a,b) : Fin(A)"
+ emptyI: "0 \<in> Fin(A)"
+ consI: "[| a: A; b: Fin(A) |] ==> cons(a,b) \<in> Fin(A)"
type_intros empty_subsetI cons_subsetI PowI
type_elims PowD [elim_format]
inductive
- domains "FiniteFun(A,B)" <= "Fin(A*B)"
+ domains "FiniteFun(A,B)" \<subseteq> "Fin(A*B)"
intros
- emptyI: "0 : A -||> B"
- consI: "[| a: A; b: B; h: A -||> B; a ~: domain(h) |]
- ==> cons(<a,b>,h) : A -||> B"
+ emptyI: "0 \<in> A -||> B"
+ consI: "[| a: A; b: B; h: A -||> B; a \<notin> domain(h) |]
+ ==> cons(<a,b>,h) \<in> A -||> B"
type_intros Fin.intros
subsection {* Finite Powerset Operator *}
-lemma Fin_mono: "A<=B ==> Fin(A) <= Fin(B)"
+lemma Fin_mono: "A<=B ==> Fin(A) \<subseteq> Fin(B)"
apply (unfold Fin.defs)
apply (rule lfp_mono)
apply (rule Fin.bnd_mono)+
apply blast
done
-(* A : Fin(B) ==> A <= B *)
+(* @{term"A \<in> Fin(B) ==> A \<subseteq> B"} *)
lemmas FinD = Fin.dom_subset [THEN subsetD, THEN PowD]
(** Induction on finite sets **)
-(*Discharging x~:y entails extra work*)
+(*Discharging @{term"x\<notin>y"} entails extra work*)
lemma Fin_induct [case_names 0 cons, induct set: Fin]:
"[| 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\<notin>y; P(y) |] ==> P(cons(x,y))
|] ==> P(b)"
apply (erule Fin.induct, simp)
apply (case_tac "a:b")
@@ -72,18 +72,18 @@
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 Un c : Fin(A)"
+lemma Fin_UnI [simp]: "[| b: Fin(A); c: Fin(A) |] ==> b \<union> c \<in> Fin(A)"
apply (erule Fin_induct)
apply (simp_all add: Un_cons)
done
(*The union of a set of finite sets is finite.*)
-lemma Fin_UnionI: "C : Fin(Fin(A)) ==> Union(C) : Fin(A)"
+lemma Fin_UnionI: "C \<in> Fin(Fin(A)) ==> \<Union>(C) \<in> Fin(A)"
by (erule Fin_induct, simp_all)
(*Every subset of a finite set is finite.*)
-lemma Fin_subset_lemma [rule_format]: "b: Fin(A) ==> \<forall>z. z<=b --> z: Fin(A)"
+lemma Fin_subset_lemma [rule_format]: "b: Fin(A) ==> \<forall>z. z<=b \<longrightarrow> z: Fin(A)"
apply (erule Fin_induct)
apply (simp add: subset_empty_iff)
apply (simp add: subset_cons_iff distrib_simps, safe)
@@ -93,16 +93,16 @@
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 Int c : Fin(A)"
+lemma Fin_IntI1 [intro,simp]: "b: Fin(A) ==> b \<inter> c \<in> Fin(A)"
by (blast intro: Fin_subset)
-lemma Fin_IntI2 [intro,simp]: "c: Fin(A) ==> b Int c : Fin(A)"
+lemma Fin_IntI2 [intro,simp]: "c: Fin(A) ==> b \<inter> c \<in> 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<=b --> P(b-c)"
+ |] ==> c<=b \<longrightarrow> P(b-c)"
apply (erule Fin_induct, simp)
apply (subst Diff_cons)
apply (simp add: cons_subset_iff Diff_subset [THEN Fin_subset])
@@ -114,11 +114,11 @@
!!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)
+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 \<subseteq> Fin(nat*A)"
apply (induct_tac "n")
apply (simp add: subset_iff)
apply (simp add: succ_def mem_not_refl [THEN cons_fun_eq])
@@ -129,14 +129,14 @@
subsection{*Finite Function Space*}
lemma FiniteFun_mono:
- "[| A<=C; B<=D |] ==> A -||> B <= C -||> D"
+ "[| A<=C; B<=D |] ==> A -||> B \<subseteq> C -||> D"
apply (unfold FiniteFun.defs)
apply (rule lfp_mono)
apply (rule FiniteFun.bnd_mono)+
apply (intro Fin_mono Sigma_mono basic_monos, assumption+)
done
-lemma FiniteFun_mono1: "A<=B ==> A -||> A <= B -||> B"
+lemma FiniteFun_mono1: "A<=B ==> A -||> A \<subseteq> B -||> B"
by (blast dest: FiniteFun_mono)
lemma FiniteFun_is_fun: "h: A -||>B ==> h: domain(h) -> B"
@@ -144,14 +144,14 @@
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) \<in> 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 ==> ALL z. z<=b --> z: A-||>B"
+ "b: A-||>B ==> \<forall>z. z<=b \<longrightarrow> z: A-||>B"
apply (erule FiniteFun.induct)
apply (simp add: subset_empty_iff FiniteFun.intros)
apply (simp add: subset_cons_iff distrib_simps, safe)
@@ -165,29 +165,29 @@
(** Some further results by Sidi O. Ehmety **)
-lemma fun_FiniteFunI [rule_format]: "A:Fin(X) ==> ALL f. f:A->B --> f:A-||>B"
+lemma fun_FiniteFunI [rule_format]: "A:Fin(X) ==> \<forall>f. f:A->B \<longrightarrow> f:A-||>B"
apply (erule Fin.induct)
apply (simp add: FiniteFun.intros, clarify)
apply (case_tac "a:b")
apply (simp add: cons_absorb)
-apply (subgoal_tac "restrict (f,b) : b -||> B")
+apply (subgoal_tac "restrict (f,b) \<in> b -||> B")
prefer 2 apply (blast intro: restrict_type2)
apply (subst fun_cons_restrict_eq, assumption)
apply (simp add: restrict_def lam_def)
-apply (blast intro: apply_funtype FiniteFun.intros
+apply (blast intro: apply_funtype FiniteFun.intros
FiniteFun_mono [THEN [2] rev_subsetD])
done
-lemma lam_FiniteFun: "A: Fin(X) ==> (lam x:A. b(x)) : A -||> {b(x). x:A}"
+lemma lam_FiniteFun: "A: Fin(X) ==> (\<lambda>x\<in>A. b(x)) \<in> 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,B) & (ALL x:domain(f). P(f`x))"
+ "f \<in> FiniteFun(A, {y:B. P(y)})
+ <-> f \<in> FiniteFun(A,B) & (\<forall>x\<in>domain(f). P(f`x))"
apply auto
apply (blast intro: FiniteFun_mono [THEN [2] rev_subsetD])
apply (blast dest: Pair_mem_PiD FiniteFun_is_fun)
-apply (rule_tac A1="domain(f)" in
+apply (rule_tac A1="domain(f)" in
subset_refl [THEN [2] FiniteFun_mono, THEN subsetD])
apply (fast dest: FiniteFun_domain_Fin Fin.dom_subset [THEN subsetD])
apply (rule fun_FiniteFunI)
--- a/src/ZF/Fixedpt.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Fixedpt.thy Tue Mar 06 15:15:49 2012 +0000
@@ -10,15 +10,15 @@
definition
(*monotone operator from Pow(D) to itself*)
bnd_mono :: "[i,i=>i]=>o" where
- "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))"
+ "bnd_mono(D,h) == h(D)<=D & (\<forall>W X. W<=X \<longrightarrow> X<=D \<longrightarrow> h(W) \<subseteq> h(X))"
definition
lfp :: "[i,i=>i]=>i" where
- "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})"
+ "lfp(D,h) == \<Inter>({X: Pow(D). h(X) \<subseteq> X})"
definition
gfp :: "[i,i=>i]=>i" where
- "gfp(D,h) == Union({X: Pow(D). X <= h(X)})"
+ "gfp(D,h) == \<Union>({X: Pow(D). X \<subseteq> h(X)})"
text{*The theorem is proved in the lattice of subsets of @{term D},
namely @{term "Pow(D)"}, with Inter as the greatest lower bound.*}
@@ -27,32 +27,32 @@
lemma bnd_monoI:
"[| h(D)<=D;
- !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) <= h(X)
+ !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) \<subseteq> h(X)
|] ==> bnd_mono(D,h)"
by (unfold bnd_mono_def, clarify, blast)
-lemma bnd_monoD1: "bnd_mono(D,h) ==> h(D) <= D"
+lemma bnd_monoD1: "bnd_mono(D,h) ==> h(D) \<subseteq> D"
apply (unfold bnd_mono_def)
apply (erule conjunct1)
done
-lemma bnd_monoD2: "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) <= h(X)"
+lemma bnd_monoD2: "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) \<subseteq> h(X)"
by (unfold bnd_mono_def, blast)
lemma bnd_mono_subset:
- "[| bnd_mono(D,h); X<=D |] ==> h(X) <= D"
+ "[| bnd_mono(D,h); X<=D |] ==> h(X) \<subseteq> D"
by (unfold bnd_mono_def, clarify, blast)
lemma bnd_mono_Un:
- "[| bnd_mono(D,h); A <= D; B <= D |] ==> h(A) Un h(B) <= h(A Un B)"
+ "[| bnd_mono(D,h); A \<subseteq> D; B \<subseteq> D |] ==> h(A) \<union> h(B) \<subseteq> h(A \<union> B)"
apply (unfold bnd_mono_def)
apply (rule Un_least, blast+)
done
(*unused*)
lemma bnd_mono_UN:
- "[| bnd_mono(D,h); \<forall>i\<in>I. A(i) <= D |]
- ==> (\<Union>i\<in>I. h(A(i))) <= h((\<Union>i\<in>I. A(i)))"
+ "[| bnd_mono(D,h); \<forall>i\<in>I. A(i) \<subseteq> D |]
+ ==> (\<Union>i\<in>I. h(A(i))) \<subseteq> h((\<Union>i\<in>I. A(i)))"
apply (unfold bnd_mono_def)
apply (rule UN_least)
apply (elim conjE)
@@ -63,7 +63,7 @@
(*Useful??*)
lemma bnd_mono_Int:
- "[| bnd_mono(D,h); A <= D; B <= D |] ==> h(A Int B) <= h(A) Int h(B)"
+ "[| bnd_mono(D,h); A \<subseteq> D; B \<subseteq> D |] ==> h(A \<inter> B) \<subseteq> h(A) \<inter> h(B)"
apply (rule Int_greatest)
apply (erule bnd_monoD2, rule Int_lower1, assumption)
apply (erule bnd_monoD2, rule Int_lower2, assumption)
@@ -73,37 +73,37 @@
(*lfp is contained in each pre-fixedpoint*)
lemma lfp_lowerbound:
- "[| h(A) <= A; A<=D |] ==> lfp(D,h) <= A"
+ "[| h(A) \<subseteq> A; A<=D |] ==> lfp(D,h) \<subseteq> A"
by (unfold lfp_def, blast)
(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*)
-lemma lfp_subset: "lfp(D,h) <= D"
+lemma lfp_subset: "lfp(D,h) \<subseteq> D"
by (unfold lfp_def Inter_def, blast)
(*Used in datatype package*)
-lemma def_lfp_subset: "A == lfp(D,h) ==> A <= D"
+lemma def_lfp_subset: "A == lfp(D,h) ==> A \<subseteq> D"
apply simp
apply (rule lfp_subset)
done
lemma lfp_greatest:
- "[| h(D) <= D; !!X. [| h(X) <= X; X<=D |] ==> A<=X |] ==> A <= lfp(D,h)"
+ "[| h(D) \<subseteq> D; !!X. [| h(X) \<subseteq> X; X<=D |] ==> A<=X |] ==> A \<subseteq> lfp(D,h)"
by (unfold lfp_def, blast)
lemma lfp_lemma1:
- "[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) <= A"
+ "[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) \<subseteq> A"
apply (erule bnd_monoD2 [THEN subset_trans])
apply (rule lfp_lowerbound, assumption+)
done
-lemma lfp_lemma2: "bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)"
+lemma lfp_lemma2: "bnd_mono(D,h) ==> h(lfp(D,h)) \<subseteq> lfp(D,h)"
apply (rule bnd_monoD1 [THEN lfp_greatest])
apply (rule_tac [2] lfp_lemma1)
apply (assumption+)
done
lemma lfp_lemma3:
- "bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))"
+ "bnd_mono(D,h) ==> lfp(D,h) \<subseteq> h(lfp(D,h))"
apply (rule lfp_lowerbound)
apply (rule bnd_monoD2, assumption)
apply (rule lfp_lemma2, assumption)
@@ -127,16 +127,16 @@
subsection{*General Induction Rule for Least Fixedpoints*}
lemma Collect_is_pre_fixedpt:
- "[| bnd_mono(D,h); !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) |]
- ==> h(Collect(lfp(D,h),P)) <= Collect(lfp(D,h),P)"
+ "[| bnd_mono(D,h); !!x. x \<in> h(Collect(lfp(D,h),P)) ==> P(x) |]
+ ==> h(Collect(lfp(D,h),P)) \<subseteq> Collect(lfp(D,h),P)"
by (blast intro: lfp_lemma2 [THEN subsetD] bnd_monoD2 [THEN subsetD]
lfp_subset [THEN subsetD])
(*This rule yields an induction hypothesis in which the components of a
data structure may be assumed to be elements of lfp(D,h)*)
lemma induct:
- "[| bnd_mono(D,h); a : lfp(D,h);
- !!x. x : h(Collect(lfp(D,h),P)) ==> P(x)
+ "[| bnd_mono(D,h); a \<in> lfp(D,h);
+ !!x. x \<in> h(Collect(lfp(D,h),P)) ==> P(x)
|] ==> P(a)"
apply (rule Collect_is_pre_fixedpt
[THEN lfp_lowerbound, THEN subsetD, THEN CollectD2])
@@ -147,14 +147,14 @@
(*Definition form, to control unfolding*)
lemma def_induct:
"[| A == lfp(D,h); bnd_mono(D,h); a:A;
- !!x. x : h(Collect(A,P)) ==> P(x)
+ !!x. x \<in> h(Collect(A,P)) ==> P(x)
|] ==> P(a)"
by (rule induct, blast+)
(*This version is useful when "A" is not a subset of D
- second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *)
+ second premise could simply be h(D \<inter> A) \<subseteq> D or !!X. X<=D ==> h(X)<=D *)
lemma lfp_Int_lowerbound:
- "[| h(D Int A) <= A; bnd_mono(D,h) |] ==> lfp(D,h) <= A"
+ "[| h(D \<inter> A) \<subseteq> A; bnd_mono(D,h) |] ==> lfp(D,h) \<subseteq> A"
apply (rule lfp_lowerbound [THEN subset_trans])
apply (erule bnd_mono_subset [THEN Int_greatest], blast+)
done
@@ -164,8 +164,8 @@
lemma lfp_mono:
assumes hmono: "bnd_mono(D,h)"
and imono: "bnd_mono(E,i)"
- and subhi: "!!X. X<=D ==> h(X) <= i(X)"
- shows "lfp(D,h) <= lfp(E,i)"
+ and subhi: "!!X. X<=D ==> h(X) \<subseteq> i(X)"
+ shows "lfp(D,h) \<subseteq> lfp(E,i)"
apply (rule bnd_monoD1 [THEN lfp_greatest])
apply (rule imono)
apply (rule hmono [THEN [2] lfp_Int_lowerbound])
@@ -176,13 +176,13 @@
(*This (unused) version illustrates that monotonicity is not really needed,
but both lfp's must be over the SAME set D; Inter is anti-monotonic!*)
lemma lfp_mono2:
- "[| i(D) <= D; !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(D,i)"
+ "[| i(D) \<subseteq> D; !!X. X<=D ==> h(X) \<subseteq> i(X) |] ==> lfp(D,h) \<subseteq> lfp(D,i)"
apply (rule lfp_greatest, assumption)
apply (rule lfp_lowerbound, blast, assumption)
done
lemma lfp_cong:
- "[|D=D'; !!X. X <= D' ==> h(X) = h'(X)|] ==> lfp(D,h) = lfp(D',h')"
+ "[|D=D'; !!X. X \<subseteq> D' ==> h(X) = h'(X)|] ==> lfp(D,h) = lfp(D',h')"
apply (simp add: lfp_def)
apply (rule_tac t=Inter in subst_context)
apply (rule Collect_cong, simp_all)
@@ -192,44 +192,44 @@
subsection{*Proof of Knaster-Tarski Theorem using @{term gfp}*}
(*gfp contains each post-fixedpoint that is contained in D*)
-lemma gfp_upperbound: "[| A <= h(A); A<=D |] ==> A <= gfp(D,h)"
+lemma gfp_upperbound: "[| A \<subseteq> h(A); A<=D |] ==> A \<subseteq> gfp(D,h)"
apply (unfold gfp_def)
apply (rule PowI [THEN CollectI, THEN Union_upper])
apply (assumption+)
done
-lemma gfp_subset: "gfp(D,h) <= D"
+lemma gfp_subset: "gfp(D,h) \<subseteq> D"
by (unfold gfp_def, blast)
(*Used in datatype package*)
-lemma def_gfp_subset: "A==gfp(D,h) ==> A <= D"
+lemma def_gfp_subset: "A==gfp(D,h) ==> A \<subseteq> D"
apply simp
apply (rule gfp_subset)
done
lemma gfp_least:
- "[| bnd_mono(D,h); !!X. [| X <= h(X); X<=D |] ==> X<=A |] ==>
- gfp(D,h) <= A"
+ "[| bnd_mono(D,h); !!X. [| X \<subseteq> h(X); X<=D |] ==> X<=A |] ==>
+ gfp(D,h) \<subseteq> A"
apply (unfold gfp_def)
apply (blast dest: bnd_monoD1)
done
lemma gfp_lemma1:
- "[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A <= h(gfp(D,h))"
+ "[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A \<subseteq> h(gfp(D,h))"
apply (rule subset_trans, assumption)
apply (erule bnd_monoD2)
apply (rule_tac [2] gfp_subset)
apply (simp add: gfp_upperbound)
done
-lemma gfp_lemma2: "bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))"
+lemma gfp_lemma2: "bnd_mono(D,h) ==> gfp(D,h) \<subseteq> h(gfp(D,h))"
apply (rule gfp_least)
apply (rule_tac [2] gfp_lemma1)
apply (assumption+)
done
lemma gfp_lemma3:
- "bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)"
+ "bnd_mono(D,h) ==> h(gfp(D,h)) \<subseteq> gfp(D,h)"
apply (rule gfp_upperbound)
apply (rule bnd_monoD2, assumption)
apply (rule gfp_lemma2, assumption)
@@ -253,12 +253,12 @@
subsection{*Coinduction Rules for Greatest Fixed Points*}
(*weak version*)
-lemma weak_coinduct: "[| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"
+lemma weak_coinduct: "[| a: X; X \<subseteq> h(X); X \<subseteq> D |] ==> a \<in> gfp(D,h)"
by (blast intro: gfp_upperbound [THEN subsetD])
lemma coinduct_lemma:
- "[| X <= h(X Un gfp(D,h)); X <= D; bnd_mono(D,h) |] ==>
- X Un gfp(D,h) <= h(X Un gfp(D,h))"
+ "[| X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D; bnd_mono(D,h) |] ==>
+ X \<union> gfp(D,h) \<subseteq> h(X \<union> gfp(D,h))"
apply (erule Un_least)
apply (rule gfp_lemma2 [THEN subset_trans], assumption)
apply (rule Un_upper2 [THEN subset_trans])
@@ -268,8 +268,8 @@
(*strong version*)
lemma coinduct:
- "[| bnd_mono(D,h); a: X; X <= h(X Un gfp(D,h)); X <= D |]
- ==> a : gfp(D,h)"
+ "[| bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D |]
+ ==> a \<in> gfp(D,h)"
apply (rule weak_coinduct)
apply (erule_tac [2] coinduct_lemma)
apply (simp_all add: gfp_subset Un_subset_iff)
@@ -277,8 +277,8 @@
(*Definition form, to control unfolding*)
lemma def_coinduct:
- "[| A == gfp(D,h); bnd_mono(D,h); a: X; X <= h(X Un A); X <= D |] ==>
- a : A"
+ "[| A == gfp(D,h); bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> A); X \<subseteq> D |] ==>
+ a \<in> A"
apply simp
apply (rule coinduct, assumption+)
done
@@ -286,15 +286,15 @@
(*The version used in the induction/coinduction package*)
lemma def_Collect_coinduct:
"[| A == gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w)));
- a: X; X <= D; !!z. z: X ==> P(X Un A, z) |] ==>
- a : A"
+ a: X; X \<subseteq> D; !!z. z: X ==> P(X \<union> A, z) |] ==>
+ a \<in> A"
apply (rule def_coinduct, assumption+, blast+)
done
(*Monotonicity of gfp!*)
lemma gfp_mono:
- "[| bnd_mono(D,h); D <= E;
- !!X. X<=D ==> h(X) <= i(X) |] ==> gfp(D,h) <= gfp(E,i)"
+ "[| bnd_mono(D,h); D \<subseteq> E;
+ !!X. X<=D ==> h(X) \<subseteq> i(X) |] ==> gfp(D,h) \<subseteq> gfp(E,i)"
apply (rule gfp_upperbound)
apply (rule gfp_lemma2 [THEN subset_trans], assumption)
apply (blast del: subsetI intro: gfp_subset)
--- a/src/ZF/InfDatatype.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/InfDatatype.thy Tue Mar 06 15:15:49 2012 +0000
@@ -7,68 +7,68 @@
theory InfDatatype imports Datatype_ZF Univ Finite Cardinal_AC begin
-lemmas fun_Limit_VfromE =
+lemmas fun_Limit_VfromE =
Limit_VfromE [OF apply_funtype InfCard_csucc [THEN InfCard_is_Limit]]
lemma fun_Vcsucc_lemma:
- "[| f: D -> Vfrom(A,csucc(K)); |D| le K; InfCard(K) |]
- ==> EX j. f: D -> Vfrom(A,j) & j < csucc(K)"
-apply (rule_tac x = "\<Union>d\<in>D. LEAST i. f`d : Vfrom (A,i) " in exI)
+ "[| f: D -> Vfrom(A,csucc(K)); |D| \<le> K; InfCard(K) |]
+ ==> \<exists>j. f: D -> Vfrom(A,j) & j < csucc(K)"
+apply (rule_tac x = "\<Union>d\<in>D. LEAST i. f`d \<in> Vfrom (A,i) " in exI)
apply (rule conjI)
-apply (rule_tac [2] le_UN_Ord_lt_csucc)
-apply (rule_tac [4] ballI, erule_tac [4] fun_Limit_VfromE, simp_all)
+apply (rule_tac [2] le_UN_Ord_lt_csucc)
+apply (rule_tac [4] ballI, erule_tac [4] fun_Limit_VfromE, simp_all)
prefer 2 apply (fast elim: Least_le [THEN lt_trans1] ltE)
apply (rule Pi_type)
apply (rename_tac [2] d)
apply (erule_tac [2] fun_Limit_VfromE, simp_all)
-apply (subgoal_tac "f`d : Vfrom (A, LEAST i. f`d : Vfrom (A,i))")
+apply (subgoal_tac "f`d \<in> Vfrom (A, LEAST i. f`d \<in> Vfrom (A,i))")
apply (erule Vfrom_mono [OF subset_refl UN_upper, THEN subsetD])
apply assumption
apply (fast elim: LeastI ltE)
done
lemma subset_Vcsucc:
- "[| D <= Vfrom(A,csucc(K)); |D| le K; InfCard(K) |]
- ==> EX j. D <= Vfrom(A,j) & j < csucc(K)"
+ "[| D \<subseteq> Vfrom(A,csucc(K)); |D| \<le> K; InfCard(K) |]
+ ==> \<exists>j. D \<subseteq> Vfrom(A,j) & j < csucc(K)"
by (simp add: subset_iff_id fun_Vcsucc_lemma)
(*Version for arbitrary index sets*)
lemma fun_Vcsucc:
- "[| |D| le K; InfCard(K); D <= Vfrom(A,csucc(K)) |] ==>
- D -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"
+ "[| |D| \<le> K; InfCard(K); D \<subseteq> Vfrom(A,csucc(K)) |] ==>
+ D -> Vfrom(A,csucc(K)) \<subseteq> Vfrom(A,csucc(K))"
apply (safe dest!: fun_Vcsucc_lemma subset_Vcsucc)
apply (rule Vfrom [THEN ssubst])
apply (drule fun_is_rel)
(*This level includes the function, and is below csucc(K)*)
-apply (rule_tac a1 = "succ (succ (j Un ja))" in UN_I [THEN UnI2])
+apply (rule_tac a1 = "succ (succ (j \<union> ja))" in UN_I [THEN UnI2])
apply (blast intro: ltD InfCard_csucc InfCard_is_Limit Limit_has_succ
- Un_least_lt)
+ Un_least_lt)
apply (erule subset_trans [THEN PowI])
apply (fast intro: Pair_in_Vfrom Vfrom_UnI1 Vfrom_UnI2)
done
lemma fun_in_Vcsucc:
- "[| f: D -> Vfrom(A, csucc(K)); |D| le K; InfCard(K);
- D <= Vfrom(A,csucc(K)) |]
+ "[| f: D -> Vfrom(A, csucc(K)); |D| \<le> K; InfCard(K);
+ D \<subseteq> Vfrom(A,csucc(K)) |]
==> f: Vfrom(A,csucc(K))"
by (blast intro: fun_Vcsucc [THEN subsetD])
-(*Remove <= from the rule above*)
+text{*Remove @{text "\<subseteq>"} from the rule above*}
lemmas fun_in_Vcsucc' = fun_in_Vcsucc [OF _ _ _ subsetI]
(** Version where K itself is the index set **)
lemma Card_fun_Vcsucc:
- "InfCard(K) ==> K -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"
+ "InfCard(K) ==> K -> Vfrom(A,csucc(K)) \<subseteq> Vfrom(A,csucc(K))"
apply (frule InfCard_is_Card [THEN Card_is_Ord])
apply (blast del: subsetI
- intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom
- lt_csucc [THEN leI, THEN le_imp_subset, THEN subset_trans])
+ intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom
+ lt_csucc [THEN leI, THEN le_imp_subset, THEN subset_trans])
done
lemma Card_fun_in_Vcsucc:
"[| f: K -> Vfrom(A, csucc(K)); InfCard(K) |] ==> f: Vfrom(A,csucc(K))"
-by (blast intro: Card_fun_Vcsucc [THEN subsetD])
+by (blast intro: Card_fun_Vcsucc [THEN subsetD])
lemma Limit_csucc: "InfCard(K) ==> Limit(csucc(K))"
by (erule InfCard_csucc [THEN InfCard_is_Limit])
@@ -79,7 +79,7 @@
lemmas zero_in_Vcsucc = Limit_csucc [THEN zero_in_VLimit]
lemmas nat_into_Vcsucc = nat_into_VLimit [OF _ Limit_csucc]
-(*For handling Cardinals of the form (nat Un |X|) *)
+(*For handling Cardinals of the form @{term"nat \<union> |X|"} *)
lemmas InfCard_nat_Un_cardinal = InfCard_Un [OF InfCard_nat Card_cardinal]
@@ -91,15 +91,15 @@
(*The new version of Data_Arg.intrs, declared in Datatype.ML*)
lemmas Data_Arg_intros =
SigmaI InlI InrI
- Pair_in_univ Inl_in_univ Inr_in_univ
+ Pair_in_univ Inl_in_univ Inr_in_univ
zero_in_univ A_into_univ nat_into_univ UnCI
(*For most K-branching datatypes with domain Vfrom(A, csucc(K)) *)
lemmas inf_datatype_intros =
InfCard_nat InfCard_nat_Un_cardinal
- Pair_in_Vcsucc Inl_in_Vcsucc Inr_in_Vcsucc
+ Pair_in_Vcsucc Inl_in_Vcsucc Inr_in_Vcsucc
zero_in_Vcsucc A_into_Vfrom nat_into_Vcsucc
- Card_fun_in_Vcsucc fun_in_Vcsucc' UN_I
+ Card_fun_in_Vcsucc fun_in_Vcsucc' UN_I
end
--- a/src/ZF/IntDiv_ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/IntDiv_ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -18,11 +18,11 @@
fun negateSnd (q,r:int) = (q,~r);
- fun divAlg (a,b) = if 0<=a then
- if b>0 then posDivAlg (a,b)
+ fun divAlg (a,b) = if 0<=a then
+ if b>0 then posDivAlg (a,b)
else if a=0 then (0,0)
else negateSnd (negDivAlg (~a,~b))
- else
+ else
if 0<b then negDivAlg (a,b)
else negateSnd (posDivAlg (~a,~b));
*)
@@ -66,7 +66,7 @@
%<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b>
else adjust(b, f ` <a,#2$*b>))"
-(*for the general case b\<noteq>0*)
+(*for the general case @{term"b\<noteq>0"}*)
definition
negateSnd :: "i => i" where
@@ -81,7 +81,7 @@
if #0 $<= b then posDivAlg (<a,b>)
else if a=#0 then <#0,#0>
else negateSnd (negDivAlg (<$-a,$-b>))
- else
+ else
if #0$<b then negDivAlg (<a,b>)
else negateSnd (posDivAlg (<$-a,$-b>))"
@@ -226,7 +226,7 @@
(** strict, in 1st argument; proof is by induction on k>0 **)
lemma zmult_zless_mono2_lemma [rule_format]:
- "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j"
+ "[| i$<j; k \<in> nat |] ==> 0<k \<longrightarrow> $#k $* i $< $#k $* j"
apply (induct_tac "k")
prefer 2
apply (subst int_succ_int_1)
@@ -266,13 +266,13 @@
lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k"
apply (rule zminus_zless_zminus [THEN iffD1])
-apply (simp del: zmult_zminus_right
+apply (simp del: zmult_zminus_right
add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus)
done
lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i"
apply (rule zminus_zless_zminus [THEN iffD1])
-apply (simp del: zmult_zminus
+apply (simp del: zmult_zminus
add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus)
done
@@ -291,16 +291,16 @@
done
-(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
+(** Cancellation laws for k*m < k*n and m*k < n*k, also for @{text"\<le>"} and =,
but not (yet?) for k*m < n*k. **)
lemma zmult_zless_lemma:
- "[| k \<in> int; m \<in> int; n \<in> int |]
+ "[| k \<in> int; m \<in> int; n \<in> int |]
==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
apply (case_tac "k = #0")
apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg)
-apply (auto simp add: not_zless_iff_zle
- not_zle_iff_zless [THEN iff_sym, of "m$*k"]
+apply (auto simp add: not_zless_iff_zle
+ not_zle_iff_zless [THEN iff_sym, of "m$*k"]
not_zle_iff_zless [THEN iff_sym, of m])
apply (auto elim: notE
simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg)
@@ -308,7 +308,7 @@
lemma zmult_zless_cancel2:
"(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
-apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)"
+apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)"
in zmult_zless_lemma)
apply auto
done
@@ -318,11 +318,11 @@
by (simp add: zmult_commute [of k] zmult_zless_cancel2)
lemma zmult_zle_cancel2:
- "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
+ "(m$*k $<= n$*k) <-> ((#0 $< k \<longrightarrow> m$<=n) & (k $< #0 \<longrightarrow> n$<=m))"
by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2)
lemma zmult_zle_cancel1:
- "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))"
+ "(k$*m $<= k$*n) <-> ((#0 $< k \<longrightarrow> m$<=n) & (k $< #0 \<longrightarrow> n$<=m))"
by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1)
lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)"
@@ -350,7 +350,7 @@
subsection{* Uniqueness and monotonicity of quotients and remainders *}
lemma unique_quotient_lemma:
- "[| b$*q' $+ r' $<= b$*q $+ r; #0 $<= r'; #0 $< b; r $< b |]
+ "[| b$*q' $+ r' $<= b$*q $+ r; #0 $<= r'; #0 $< b; r $< b |]
==> q' $<= q"
apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r")
prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls)
@@ -359,18 +359,18 @@
apply (erule zle_zless_trans)
apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
apply (erule zle_zless_trans)
- apply (simp add: );
+ apply (simp add: );
apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)")
- prefer 2
+ prefer 2
apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls)
apply (auto elim: zless_asym
simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls)
done
lemma unique_quotient_lemma_neg:
- "[| b$*q' $+ r' $<= b$*q $+ r; r $<= #0; b $< #0; b $< r' |]
+ "[| b$*q' $+ r' $<= b$*q $+ r; r $<= #0; b $< #0; b $< r' |]
==> q $<= q'"
-apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r"
+apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r"
in unique_quotient_lemma)
apply (auto simp del: zminus_zadd_distrib
simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus)
@@ -378,19 +378,19 @@
lemma unique_quotient:
- "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0;
+ "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;
q \<in> int; q' \<in> int |] ==> q = q'"
apply (simp add: split_ifs quorem_def neq_iff_zless)
apply safe
apply simp_all
apply (blast intro: zle_anti_sym
- dest: zle_eq_refl [THEN unique_quotient_lemma]
+ dest: zle_eq_refl [THEN unique_quotient_lemma]
zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+
done
lemma unique_remainder:
- "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0;
- q \<in> int; q' \<in> int;
+ "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;
+ q \<in> int; q' \<in> int;
r \<in> int; r' \<in> int |] ==> r = r'"
apply (subgoal_tac "q = q'")
prefer 2 apply (blast intro: unique_quotient)
@@ -398,18 +398,18 @@
done
-subsection{*Correctness of posDivAlg,
+subsection{*Correctness of posDivAlg,
the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *}
lemma adjust_eq [simp]:
- "adjust(b, <q,r>) = (let diff = r$-b in
- if #0 $<= diff then <#2$*q $+ #1,diff>
+ "adjust(b, <q,r>) = (let diff = r$-b in
+ if #0 $<= diff then <#2$*q $+ #1,diff>
else <#2$*q,r>)"
by (simp add: Let_def adjust_def)
lemma posDivAlg_termination:
- "[| #0 $< b; ~ a $< b |]
+ "[| #0 $< b; ~ a $< b |]
==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)"
apply (simp (no_asm) add: zless_nat_conj)
apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls)
@@ -418,8 +418,8 @@
lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure]
lemma posDivAlg_eqn:
- "[| #0 $< b; a \<in> int; b \<in> int |] ==>
- posDivAlg(<a,b>) =
+ "[| #0 $< b; a \<in> int; b \<in> int |] ==>
+ posDivAlg(<a,b>) =
(if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))"
apply (rule posDivAlg_unfold [THEN trans])
apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym])
@@ -428,11 +428,11 @@
lemma posDivAlg_induct_lemma [rule_format]:
assumes prem:
- "!!a b. [| a \<in> int; b \<in> int;
- ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)"
- shows "<u,v> \<in> int*int --> P(<u,v>)"
+ "!!a b. [| a \<in> int; b \<in> int;
+ ~ (a $< b | b $<= #0) \<longrightarrow> P(<a, #2 $* b>) |] ==> P(<a,b>)"
+ shows "<u,v> \<in> int*int \<longrightarrow> P(<u,v>)"
apply (rule_tac a = "<u,v>" in wf_induct)
-apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)"
+apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)"
in wf_measure)
apply clarify
apply (rule prem)
@@ -446,7 +446,7 @@
assumes u_int: "u \<in> int"
and v_int: "v \<in> int"
and ih: "!!a b. [| a \<in> int; b \<in> int;
- ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)"
+ ~ (a $< b | b $<= #0) \<longrightarrow> P(a, #2 $* b) |] ==> P(a,b)"
shows "P(u,v)"
apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)")
apply simp
@@ -456,8 +456,8 @@
apply (auto simp add: u_int v_int)
done
-(*FIXME: use intify in integ_of so that we always have integ_of w \<in> int.
- then this rewrite can work for ALL constants!!*)
+(*FIXME: use intify in integ_of so that we always have @{term"integ_of w \<in> int"}.
+ then this rewrite can work for all constants!!*)
lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)"
apply (simp (no_asm) add: int_eq_iff_zle)
done
@@ -483,17 +483,17 @@
(** Inequality reasoning **)
lemma int_0_less_lemma:
- "[| x \<in> int; y \<in> int |]
+ "[| x \<in> int; y \<in> int |]
==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg)
-apply (rule ccontr)
-apply (rule_tac [2] ccontr)
+apply (rule ccontr)
+apply (rule_tac [2] ccontr)
apply (auto simp add: zle_def not_zless_iff_zle)
apply (erule_tac P = "#0$< x$* y" in rev_mp)
apply (erule_tac [2] P = "#0$< x$* y" in rev_mp)
-apply (drule zmult_pos_neg, assumption)
+apply (drule zmult_pos_neg, assumption)
prefer 2
- apply (drule zmult_pos_neg, assumption)
+ apply (drule zmult_pos_neg, assumption)
apply (auto dest: zless_not_sym simp add: zmult_commute)
done
@@ -504,7 +504,7 @@
done
lemma int_0_le_lemma:
- "[| x \<in> int; y \<in> int |]
+ "[| x \<in> int; y \<in> int |]
==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)"
by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff)
@@ -532,7 +532,7 @@
apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
apply assumption+
apply (case_tac "#0 $< ba")
- apply (simp add: posDivAlg_eqn adjust_def integ_of_type
+ apply (simp add: posDivAlg_eqn adjust_def integ_of_type
split add: split_if_asm)
apply clarify
apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
@@ -543,8 +543,8 @@
(*Correctness of posDivAlg: it computes quotients correctly*)
lemma posDivAlg_correct [rule_format]:
- "[| a \<in> int; b \<in> int |]
- ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))"
+ "[| a \<in> int; b \<in> int |]
+ ==> #0 $<= a \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, posDivAlg(<a,b>))"
apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
apply auto
apply (simp_all add: quorem_def)
@@ -567,7 +567,7 @@
subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*}
lemma negDivAlg_termination:
- "[| #0 $< b; a $+ b $< #0 |]
+ "[| #0 $< b; a $+ b $< #0 |]
==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)"
apply (simp (no_asm) add: zless_nat_conj)
apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym]
@@ -577,9 +577,9 @@
lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure]
lemma negDivAlg_eqn:
- "[| #0 $< b; a : int; b : int |] ==>
- negDivAlg(<a,b>) =
- (if #0 $<= a$+b then <#-1,a$+b>
+ "[| #0 $< b; a \<in> int; b \<in> int |] ==>
+ negDivAlg(<a,b>) =
+ (if #0 $<= a$+b then <#-1,a$+b>
else adjust(b, negDivAlg (<a, #2$*b>)))"
apply (rule negDivAlg_unfold [THEN trans])
apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym])
@@ -588,12 +588,12 @@
lemma negDivAlg_induct_lemma [rule_format]:
assumes prem:
- "!!a b. [| a \<in> int; b \<in> int;
- ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |]
+ "!!a b. [| a \<in> int; b \<in> int;
+ ~ (#0 $<= a $+ b | b $<= #0) \<longrightarrow> P(<a, #2 $* b>) |]
==> P(<a,b>)"
- shows "<u,v> \<in> int*int --> P(<u,v>)"
+ shows "<u,v> \<in> int*int \<longrightarrow> P(<u,v>)"
apply (rule_tac a = "<u,v>" in wf_induct)
-apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)"
+apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)"
in wf_measure)
apply clarify
apply (rule prem)
@@ -605,8 +605,8 @@
lemma negDivAlg_induct [consumes 2]:
assumes u_int: "u \<in> int"
and v_int: "v \<in> int"
- and ih: "!!a b. [| a \<in> int; b \<in> int;
- ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |]
+ and ih: "!!a b. [| a \<in> int; b \<in> int;
+ ~ (#0 $<= a $+ b | b $<= #0) \<longrightarrow> P(a, #2 $* b) |]
==> P(a,b)"
shows "P(u,v)"
apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)")
@@ -624,7 +624,7 @@
apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
apply assumption+
apply (case_tac "#0 $< ba")
- apply (simp add: negDivAlg_eqn adjust_def integ_of_type
+ apply (simp add: negDivAlg_eqn adjust_def integ_of_type
split add: split_if_asm)
apply clarify
apply (simp add: int_0_less_mult_iff not_zle_iff_zless)
@@ -637,8 +637,8 @@
(*Correctness of negDivAlg: it computes quotients correctly
It doesn't work if a=0 because the 0/b=0 rather than -1*)
lemma negDivAlg_correct [rule_format]:
- "[| a \<in> int; b \<in> int |]
- ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))"
+ "[| a \<in> int; b \<in> int |]
+ ==> a $< #0 \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, negDivAlg(<a,b>))"
apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
apply auto
apply (simp_all add: quorem_def)
@@ -698,7 +698,7 @@
done
lemma quorem_neg:
- "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|]
+ "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|]
==> quorem (<a,b>, negateSnd(qr))"
apply clarify
apply (auto elim: zless_asym simp add: quorem_def zless_zminus)
@@ -714,7 +714,7 @@
"[|b \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))"
apply (auto simp add: quorem_0 divAlg_def)
apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct
- posDivAlg_type negDivAlg_type)
+ posDivAlg_type negDivAlg_type)
apply (auto simp add: quorem_def neq_iff_zless)
txt{*linear arithmetic from here on*}
apply (auto simp add: zle_def)
@@ -756,7 +756,7 @@
done
-(** Arbitrary definitions for division by zero. Useful to simplify
+(** Arbitrary definitions for division by zero. Useful to simplify
certain equations **)
lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0"
@@ -774,7 +774,7 @@
lemma raw_zmod_zdiv_equality:
"[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (cut_tac a = "a" and b = "b" in divAlg_correct)
apply (auto simp add: quorem_def zdiv_def zmod_def split_def)
done
@@ -808,21 +808,21 @@
(** proving general properties of zdiv and zmod **)
lemma quorem_div_mod:
- "[|b \<noteq> #0; a \<in> int; b \<in> int |]
+ "[|b \<noteq> #0; a \<in> int; b \<in> int |]
==> quorem (<a,b>, <a zdiv b, a zmod b>)"
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
-apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound
+apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound
neg_mod_sign neg_mod_bound)
done
-(*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*)
+(*Surely quorem(<a,b>,<q,r>) implies @{term"a \<in> int"}, but it doesn't matter*)
lemma quorem_div:
- "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |]
+ "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |]
==> a zdiv b = q"
by (blast intro: quorem_div_mod [THEN unique_quotient])
lemma quorem_mod:
- "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |]
+ "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |]
==> a zmod b = r"
by (blast intro: quorem_div_mod [THEN unique_remainder])
@@ -835,7 +835,7 @@
done
lemma zdiv_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zdiv b = #0"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_pos_pos_trivial_raw)
apply auto
done
@@ -849,7 +849,7 @@
done
lemma zdiv_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zdiv b = #0"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_neg_neg_trivial_raw)
apply auto
done
@@ -869,7 +869,7 @@
done
lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_pos_neg_trivial_raw)
apply auto
done
@@ -886,7 +886,7 @@
done
lemma zmod_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zmod b = intify(a)"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_pos_pos_trivial_raw)
apply auto
done
@@ -900,7 +900,7 @@
done
lemma zmod_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zmod b = intify(a)"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_neg_neg_trivial_raw)
apply auto
done
@@ -914,7 +914,7 @@
done
lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b"
-apply (cut_tac a = "intify (a)" and b = "intify (b)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_pos_neg_trivial_raw)
apply auto
done
@@ -927,7 +927,7 @@
lemma zdiv_zminus_zminus_raw:
"[|a \<in> int; b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div])
apply auto
done
@@ -941,7 +941,7 @@
lemma zmod_zminus_zminus_raw:
"[|a \<in> int; b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod])
apply auto
done
@@ -1008,7 +1008,7 @@
(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *)
lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0"
apply (case_tac "a = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (blast intro: quorem_div_mod [THEN self_remainder])
done
@@ -1042,14 +1042,14 @@
(** a positive, b positive **)
-lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |]
+lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |]
==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))"
apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
apply (auto simp add: zle_def)
done
lemma zmod_pos_pos:
- "[| #0 $< a; #0 $<= b |]
+ "[| #0 $< a; #0 $<= b |]
==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))"
apply (simp (no_asm_simp) add: zmod_def divAlg_def)
apply (auto simp add: zle_def)
@@ -1058,14 +1058,14 @@
(** a negative, b positive **)
lemma zdiv_neg_pos:
- "[| a $< #0; #0 $< b |]
+ "[| a $< #0; #0 $< b |]
==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))"
apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
apply (blast dest: zle_zless_trans)
done
lemma zmod_neg_pos:
- "[| a $< #0; #0 $< b |]
+ "[| a $< #0; #0 $< b |]
==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))"
apply (simp (no_asm_simp) add: zmod_def divAlg_def)
apply (blast dest: zle_zless_trans)
@@ -1074,7 +1074,7 @@
(** a positive, b negative **)
lemma zdiv_pos_neg:
- "[| #0 $< a; b $< #0 |]
+ "[| #0 $< a; b $< #0 |]
==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle)
apply auto
@@ -1084,7 +1084,7 @@
done
lemma zmod_pos_neg:
- "[| #0 $< a; b $< #0 |]
+ "[| #0 $< a; b $< #0 |]
==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle)
apply auto
@@ -1096,7 +1096,7 @@
(** a negative, b negative **)
lemma zdiv_neg_neg:
- "[| a $< #0; b $<= #0 |]
+ "[| a $< #0; b $<= #0 |]
==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
apply auto
@@ -1104,7 +1104,7 @@
done
lemma zmod_neg_neg:
- "[| a $< #0; b $<= #0 |]
+ "[| a $< #0; b $<= #0 |]
==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zmod_def divAlg_def)
apply auto
@@ -1198,10 +1198,10 @@
done
lemma zdiv_mono2_lemma:
- "[| b$*q $+ r = b'$*q' $+ r'; #0 $<= b'$*q' $+ r';
- r' $< b'; #0 $<= r; #0 $< b'; b' $<= b |]
+ "[| b$*q $+ r = b'$*q' $+ r'; #0 $<= b'$*q' $+ r';
+ r' $< b'; #0 $<= r; #0 $< b'; b' $<= b |]
==> q $<= q'"
-apply (frule q_pos_lemma, assumption+)
+apply (frule q_pos_lemma, assumption+)
apply (subgoal_tac "b$*q $< b$* (q' $+ #1)")
apply (simp add: zmult_zless_cancel1)
apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans)
@@ -1219,7 +1219,7 @@
lemma zdiv_mono2_raw:
- "[| #0 $<= a; #0 $< b'; b' $<= b; a \<in> int |]
+ "[| #0 $<= a; #0 $< b'; b' $<= b; a \<in> int |]
==> a zdiv b $<= a zdiv b'"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
@@ -1232,7 +1232,7 @@
done
lemma zdiv_mono2:
- "[| #0 $<= a; #0 $< b'; b' $<= b |]
+ "[| #0 $<= a; #0 $< b'; b' $<= b |]
==> a zdiv b $<= a zdiv b'"
apply (cut_tac a = "intify (a)" in zdiv_mono2_raw)
apply auto
@@ -1249,12 +1249,12 @@
lemma zdiv_mono2_neg_lemma:
- "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0;
- r $< b; #0 $<= r'; #0 $< b'; b' $<= b |]
+ "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0;
+ r $< b; #0 $<= r'; #0 $< b'; b' $<= b |]
==> q' $<= q"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
-apply (frule q_neg_lemma, assumption+)
+apply (frule q_neg_lemma, assumption+)
apply (subgoal_tac "b$*q' $< b$* (q $+ #1)")
apply (simp add: zmult_zless_cancel1)
apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1])
@@ -1278,7 +1278,7 @@
done
lemma zdiv_mono2_neg_raw:
- "[| a $< #0; #0 $< b'; b' $<= b; a \<in> int |]
+ "[| a $< #0; #0 $< b'; b' $<= b; a \<in> int |]
==> a zdiv b' $<= a zdiv b"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
@@ -1290,7 +1290,7 @@
apply (simp_all add: pos_mod_sign pos_mod_bound)
done
-lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $<= b |]
+lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $<= b |]
==> a zdiv b' $<= a zdiv b"
apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw)
apply auto
@@ -1303,18 +1303,18 @@
(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **)
lemma zmult1_lemma:
- "[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |]
+ "[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |]
==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)"
apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
-apply (auto intro: raw_zmod_zdiv_equality)
+apply (auto intro: raw_zmod_zdiv_equality)
done
lemma zdiv_zmult1_eq_raw:
- "[|b \<in> int; c \<in> int|]
+ "[|b \<in> int; c \<in> int|]
==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
apply (case_tac "c = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
apply auto
done
@@ -1327,7 +1327,7 @@
lemma zmod_zmult1_eq_raw:
"[|b \<in> int; c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c"
apply (case_tac "c = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
apply auto
done
@@ -1366,12 +1366,12 @@
done
-(** proving (a$+b) zdiv c =
+(** proving (a$+b) zdiv c =
a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **)
lemma zadd1_lemma:
- "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>);
- c \<in> int; c \<noteq> #0 |]
+ "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>);
+ c \<in> int; c \<noteq> #0 |]
==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)"
apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
@@ -1380,32 +1380,32 @@
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
lemma zdiv_zadd1_eq_raw:
- "[|a \<in> int; b \<in> int; c \<in> int|] ==>
+ "[|a \<in> int; b \<in> int; c \<in> int|] ==>
(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
apply (case_tac "c = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,
THEN quorem_div])
done
lemma zdiv_zadd1_eq:
"(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)"
-apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"
in zdiv_zadd1_eq_raw)
apply auto
done
lemma zmod_zadd1_eq_raw:
- "[|a \<in> int; b \<in> int; c \<in> int|]
+ "[|a \<in> int; b \<in> int; c \<in> int|]
==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
apply (case_tac "c = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
-apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod,
THEN quorem_mod])
done
lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
-apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"
+apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)"
in zmod_zadd1_eq_raw)
apply auto
done
@@ -1413,7 +1413,7 @@
lemma zmod_div_trivial_raw:
"[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound
zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial)
done
@@ -1426,8 +1426,8 @@
lemma zmod_mod_trivial_raw:
"[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
-apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound
zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial)
done
@@ -1461,13 +1461,13 @@
lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a"
apply (case_tac "a = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (simp (no_asm_simp) add: zmod_zadd1_eq)
done
lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a"
apply (case_tac "a = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (simp (no_asm_simp) add: zmod_zadd1_eq)
done
@@ -1495,7 +1495,7 @@
"[| #0 $< c; b $< r; r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0"
apply (subgoal_tac "b $* (q zmod c) $<= #0")
prefer 2
- apply (simp add: zmult_le_0_iff pos_mod_sign)
+ apply (simp add: zmult_le_0_iff pos_mod_sign)
apply (blast intro: zless_imp_zle dest: zless_zle_trans)
(*arithmetic*)
apply (drule zadd_zle_mono)
@@ -1507,7 +1507,7 @@
"[| #0 $< c; #0 $<= r; r $< b |] ==> #0 $<= b $* (q zmod c) $+ r"
apply (subgoal_tac "#0 $<= b $* (q zmod c)")
prefer 2
- apply (simp add: int_0_le_mult_iff pos_mod_sign)
+ apply (simp add: int_0_le_mult_iff pos_mod_sign)
apply (blast intro: zless_imp_zle dest: zle_zless_trans)
(*arithmetic*)
apply (drule zadd_zle_mono)
@@ -1527,10 +1527,10 @@
done
lemma zdiv_zmult2_lemma:
- "[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |]
+ "[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |]
==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)"
apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def
- neq_iff_zless int_0_less_mult_iff
+ neq_iff_zless int_0_less_mult_iff
zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2
zdiv_zmult2_aux3 zdiv_zmult2_aux4)
apply (blast dest: zless_trans)+
@@ -1539,7 +1539,7 @@
lemma zdiv_zmult2_eq_raw:
"[|#0 $< c; a \<in> int; b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div])
apply (auto simp add: intify_eq_0_iff_zle)
apply (blast dest: zle_zless_trans)
@@ -1551,10 +1551,10 @@
done
lemma zmod_zmult2_eq_raw:
- "[|#0 $< c; a \<in> int; b \<in> int|]
+ "[|#0 $< c; a \<in> int; b \<in> int|]
==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod])
apply (auto simp add: intify_eq_0_iff_zle)
apply (blast dest: zle_zless_trans)
@@ -1584,7 +1584,7 @@
lemma zdiv_zmult_zmult1_raw:
"[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (auto simp add: neq_iff_zless [of b]
zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
done
@@ -1603,14 +1603,14 @@
subsection{* Distribution of factors over "zmod" *}
lemma zmod_zmult_zmult1_aux1:
- "[| #0 $< b; intify(c) \<noteq> #0 |]
+ "[| #0 $< b; intify(c) \<noteq> #0 |]
==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
apply (subst zmod_zmult2_eq)
apply auto
done
lemma zmod_zmult_zmult1_aux2:
- "[| b $< #0; intify(c) \<noteq> #0 |]
+ "[| b $< #0; intify(c) \<noteq> #0 |]
==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))")
apply (rule_tac [2] zmod_zmult_zmult1_aux1)
@@ -1620,9 +1620,9 @@
lemma zmod_zmult_zmult1_raw:
"[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
apply (case_tac "b = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (case_tac "c = #0")
- apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
+ apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (auto simp add: neq_iff_zless [of b]
zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2)
done
@@ -1726,9 +1726,9 @@
apply auto
done
- lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =
- (if ~b | #0 $<= integ_of w
- then integ_of v zdiv (integ_of w)
+ lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =
+ (if ~b | #0 $<= integ_of w
+ then integ_of v zdiv (integ_of w)
else (integ_of v $+ #1) zdiv (integ_of w))"
apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT)
apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2)
@@ -1770,11 +1770,11 @@
apply auto
done
- lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) =
- (if b then
- if #0 $<= integ_of w
- then #2 $* (integ_of v zmod integ_of w) $+ #1
- else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1
+ lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) =
+ (if b then
+ if #0 $<= integ_of w
+ then #2 $* (integ_of v zmod integ_of w) $+ #1
+ else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1
else #2 $* (integ_of v zmod integ_of w))"
apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT)
apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2)
--- a/src/ZF/Int_ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Int_ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,7 +9,7 @@
definition
intrel :: i where
- "intrel == {p : (nat*nat)*(nat*nat).
+ "intrel == {p \<in> (nat*nat)*(nat*nat).
\<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
definition
@@ -22,7 +22,7 @@
definition
intify :: "i=>i" --{*coercion from ANYTHING to int*} where
- "intify(m) == if m : int then m else $#0"
+ "intify(m) == if m \<in> int then m else $#0"
definition
raw_zminus :: "i=>i" where
@@ -135,7 +135,7 @@
apply (fast elim!: sym int_trans_lemma)
done
-lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : int"
+lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} \<in> int"
by (simp add: int_def)
declare equiv_intrel [THEN eq_equiv_class_iff, simp]
@@ -145,7 +145,7 @@
(** int_of: the injection from nat to int **)
-lemma int_of_type [simp,TC]: "$#m : int"
+lemma int_of_type [simp,TC]: "$#m \<in> int"
by (simp add: int_def quotient_def int_of_def, auto)
lemma int_of_eq [iff]: "($# m = $# n) <-> natify(m)=natify(n)"
@@ -157,10 +157,10 @@
(** intify: coercion from anything to int **)
-lemma intify_in_int [iff,TC]: "intify(x) : int"
+lemma intify_in_int [iff,TC]: "intify(x) \<in> int"
by (simp add: intify_def)
-lemma intify_ident [simp]: "n : int ==> intify(n) = n"
+lemma intify_ident [simp]: "n \<in> int ==> intify(n) = n"
by (simp add: intify_def)
@@ -220,12 +220,12 @@
lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel"
by (auto simp add: congruent_def add_ac)
-lemma raw_zminus_type: "z : int ==> raw_zminus(z) : int"
+lemma raw_zminus_type: "z \<in> int ==> raw_zminus(z) \<in> int"
apply (simp add: int_def raw_zminus_def)
apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent])
done
-lemma zminus_type [TC,iff]: "$-z : int"
+lemma zminus_type [TC,iff]: "$-z \<in> int"
by (simp add: zminus_def raw_zminus_type)
lemma raw_zminus_inject:
@@ -253,7 +253,7 @@
==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
by (simp add: zminus_def raw_zminus image_intrel_int)
-lemma raw_zminus_zminus: "z : int ==> raw_zminus (raw_zminus(z)) = z"
+lemma raw_zminus_zminus: "z \<in> int ==> raw_zminus (raw_zminus(z)) = z"
by (auto simp add: int_def raw_zminus)
lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)"
@@ -262,7 +262,7 @@
lemma zminus_int0 [simp]: "$- ($#0) = $#0"
by (simp add: int_of_def zminus)
-lemma zminus_zminus: "z : int ==> $- ($- z) = z"
+lemma zminus_zminus: "z \<in> int ==> $- ($- z) = z"
by simp
@@ -352,7 +352,7 @@
"[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z"
by (drule zneg_int_of, auto)
-lemma int_cases: "z : int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))"
+lemma int_cases: "z \<in> int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))"
apply (case_tac "znegative (z) ")
prefer 2 apply (blast dest: not_zneg_mag sym)
apply (blast dest: zneg_int_of)
@@ -398,13 +398,13 @@
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) \<in> 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)
done
-lemma zadd_type [iff,TC]: "z $+ w : int"
+lemma zadd_type [iff,TC]: "z $+ w \<in> int"
by (simp add: zadd_def raw_zadd_type)
lemma raw_zadd:
@@ -422,7 +422,7 @@
intrel `` {<x1#+x2, y1#+y2>}"
by (simp add: zadd_def raw_zadd image_intrel_int)
-lemma raw_zadd_int0: "z : int ==> raw_zadd ($#0,z) = z"
+lemma raw_zadd_int0: "z \<in> int ==> raw_zadd ($#0,z) = z"
by (auto simp add: int_def int_of_def raw_zadd)
lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)"
@@ -469,13 +469,13 @@
by (simp add: int_of_add [symmetric] natify_succ)
lemma int_of_diff:
- "[| m\<in>nat; n le m |] ==> $# (m #- n) = ($#m) $- ($#n)"
+ "[| m\<in>nat; n \<le> m |] ==> $# (m #- n) = ($#m) $- ($#n)"
apply (simp add: int_of_def zdiff_def)
apply (frule lt_nat_in_nat)
apply (simp_all add: zadd zminus add_diff_inverse2)
done
-lemma raw_zadd_zminus_inverse: "z : int ==> raw_zadd (z, $- z) = $#0"
+lemma raw_zadd_zminus_inverse: "z \<in> int ==> raw_zadd (z, $- z) = $#0"
by (auto simp add: int_def int_of_def zminus raw_zadd add_commute)
lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0"
@@ -511,13 +511,13 @@
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) \<in> 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)
done
-lemma zmult_type [iff,TC]: "z $* w : int"
+lemma zmult_type [iff,TC]: "z $* w \<in> int"
by (simp add: zmult_def raw_zmult_type)
lemma raw_zmult:
@@ -533,19 +533,19 @@
intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
by (simp add: zmult_def raw_zmult image_intrel_int)
-lemma raw_zmult_int0: "z : int ==> raw_zmult ($#0,z) = $#0"
+lemma raw_zmult_int0: "z \<in> int ==> raw_zmult ($#0,z) = $#0"
by (auto simp add: int_def int_of_def raw_zmult)
lemma zmult_int0 [simp]: "$#0 $* z = $#0"
by (simp add: zmult_def raw_zmult_int0)
-lemma raw_zmult_int1: "z : int ==> raw_zmult ($#1,z) = z"
+lemma raw_zmult_int1: "z \<in> int ==> raw_zmult ($#1,z) = z"
by (auto simp add: int_def int_of_def raw_zmult)
lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)"
by (simp add: zmult_def raw_zmult_int1)
-lemma zmult_int1: "z : int ==> $#1 $* z = z"
+lemma zmult_int1: "z \<in> int ==> $#1 $* z = z"
by simp
lemma raw_zmult_commute:
@@ -603,7 +603,7 @@
(*** Subtraction laws ***)
-lemma zdiff_type [iff,TC]: "z $- w : int"
+lemma zdiff_type [iff,TC]: "z $- w \<in> int"
by (simp add: zdiff_def)
lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z"
@@ -644,10 +644,10 @@
lemma zless_not_refl [iff]: "~ (z$<z)"
by (auto simp add: zless_def znegative_def int_of_def zdiff_def)
-lemma neq_iff_zless: "[| x: int; y: int |] ==> (x ~= y) <-> (x $< y | y $< x)"
+lemma neq_iff_zless: "[| x: int; y: int |] ==> (x \<noteq> 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)"
+lemma zless_imp_intify_neq: "w $< z ==> intify(w) \<noteq> intify(z)"
apply auto
apply (subgoal_tac "~ (intify (w) $< intify (z))")
apply (erule_tac [2] ssubst)
@@ -672,7 +672,7 @@
done
lemma zless_succ_zadd_lemma:
- "w : int ==> w $< w $+ $# succ(n)"
+ "w \<in> 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)
apply (rule_tac x = 0 in exI, auto)
@@ -695,7 +695,7 @@
done
lemma zless_trans_lemma:
- "[| x $< y; y $< z; x: int; y : int; z: int |] ==> x $< z"
+ "[| x $< y; y $< z; x: int; y \<in> 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)
--- a/src/ZF/List_ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/List_ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -93,20 +93,20 @@
definition
(* Function `take' returns the first n elements of a list *)
take :: "[i,i]=>i" where
- "take(n, as) == list_rec(lam n:nat. [],
- %a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
+ "take(n, as) == list_rec(\<lambda>n\<in>nat. [],
+ %a l r. \<lambda>n\<in>nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
definition
nth :: "[i, i]=>i" where
--{*returns the (n+1)th element of a list, or 0 if the
list is too short.*}
- "nth(n, as) == list_rec(lam n:nat. 0,
- %a l r. lam n:nat. nat_case(a, %m. r`m, n), as) ` n"
+ "nth(n, as) == list_rec(\<lambda>n\<in>nat. 0,
+ %a l r. \<lambda>n\<in>nat. nat_case(a, %m. r`m, n), as) ` n"
definition
list_update :: "[i, i, i]=>i" where
- "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
- %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
+ "list_update(xs, i, v) == list_rec(\<lambda>n\<in>nat. Nil,
+ %u us vs. \<lambda>n\<in>nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
consts
filter :: "[i=>o, i] => i"
@@ -119,25 +119,25 @@
primrec
"upt(i, 0) = Nil"
- "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
+ "upt(i, succ(j)) = (if i \<le> j then upt(i, j)@[j] else Nil)"
definition
min :: "[i,i] =>i" where
- "min(x, y) == (if x le y then x else y)"
+ "min(x, y) == (if x \<le> y then x else y)"
definition
max :: "[i, i] =>i" where
- "max(x, y) == (if x le y then y else x)"
+ "max(x, y) == (if x \<le> y then y else x)"
(*** Aspects of the datatype definition ***)
declare list.intros [simp,TC]
(*An elimination rule, for type-checking*)
-inductive_cases ConsE: "Cons(a,l) : list(A)"
+inductive_cases ConsE: "Cons(a,l) \<in> list(A)"
lemma Cons_type_iff [simp]: "Cons(a,l) \<in> list(A) <-> a \<in> A & l \<in> list(A)"
-by (blast elim: ConsE)
+by (blast elim: ConsE)
(*Proving freeness results*)
lemma Cons_iff: "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"
@@ -153,7 +153,7 @@
(** Lemmas to justify using "list" in other recursive type definitions **)
-lemma list_mono: "A<=B ==> list(A) <= list(B)"
+lemma list_mono: "A<=B ==> list(A) \<subseteq> list(B)"
apply (unfold list.defs )
apply (rule lfp_mono)
apply (simp_all add: list.bnd_mono)
@@ -161,7 +161,7 @@
done
(*There is a similar proof by list induction.*)
-lemma list_univ: "list(univ(A)) <= univ(A)"
+lemma list_univ: "list(univ(A)) \<subseteq> univ(A)"
apply (unfold list.defs list.con_defs)
apply (rule lfp_lowerbound)
apply (rule_tac [2] A_subset_univ [THEN univ_mono])
@@ -171,25 +171,25 @@
(*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 \<subseteq> 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))
- |] ==> list_case(c,h,l) : C(l)"
+ |] ==> list_case(c,h,l) \<in> C(l)"
by (erule list.induct, auto)
lemma list_0_triv: "list(0) = {Nil}"
-apply (rule equalityI, auto)
-apply (induct_tac x, auto)
+apply (rule equalityI, auto)
+apply (induct_tac x, auto)
done
(*** List functions ***)
-lemma tl_type: "l: list(A) ==> tl(l) : list(A)"
+lemma tl_type: "l: list(A) ==> tl(l) \<in> list(A)"
apply (induct_tac "l")
apply (simp_all (no_asm_simp) add: list.intros)
done
@@ -208,7 +208,7 @@
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) \<in> list(A)"
apply (induct_tac "i")
apply (simp_all (no_asm_simp) add: tl_type)
done
@@ -222,57 +222,57 @@
"[| 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)"
+ |] ==> list_rec(c,h,l) \<in> 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) \<in> 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) \<in> 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) \<in> nat"
by (simp add: length_list_def)
lemma lt_length_in_nat:
"[|x < length(xs); xs \<in> list(A)|] ==> x \<in> nat"
-by (frule lt_nat_in_nat, typecheck)
+by (frule lt_nat_in_nat, typecheck)
(** app **)
-lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"
+lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys \<in> list(A)"
by (simp add: app_list_def)
(** rev **)
-lemma rev_type [TC]: "xs: list(A) ==> rev(xs) : list(A)"
+lemma rev_type [TC]: "xs: list(A) ==> rev(xs) \<in> list(A)"
by (simp add: rev_list_def)
(** flat **)
-lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) : list(A)"
+lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) \<in> list(A)"
by (simp add: flat_list_def)
(** 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) \<in> Pow(A)"
apply (unfold set_of_list_list_def)
apply (erule list_rec_type, auto)
done
lemma set_of_list_append:
- "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"
+ "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) \<union> set_of_list(ys)"
apply (erule list.induct)
apply (simp_all (no_asm_simp) add: Un_cons)
done
@@ -280,7 +280,7 @@
(** list_add **)
-lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) : nat"
+lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \<in> nat"
by (simp add: list_add_list_def)
@@ -316,7 +316,7 @@
(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
-(* c : list(Collect(B,P)) ==> c : list(B) *)
+(* @{term"c \<in> list(Collect(B,P)) ==> c \<in> 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)"
@@ -350,11 +350,11 @@
(*Lemma for the inductive step of drop_length*)
lemma drop_length_Cons [rule_format]:
"xs: list(A) ==>
- \<forall>x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
+ \<forall>x. \<exists>z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"
by (erule list.induct, simp_all)
lemma drop_length [rule_format]:
- "l: list(A) ==> \<forall>i \<in> length(l). (EX z zs. drop(i,l) = Cons(z,zs))"
+ "l: list(A) ==> \<forall>i \<in> length(l). (\<exists>z zs. drop(i,l) = Cons(z,zs))"
apply (erule list.induct, simp_all, safe)
apply (erule drop_length_Cons)
apply (rule natE)
@@ -435,25 +435,25 @@
done
lemma list_complete_induct_lemma [rule_format]:
- assumes ih:
- "\<And>l. [| l \<in> list(A);
- \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|]
+ assumes ih:
+ "\<And>l. [| l \<in> list(A);
+ \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')|]
==> P(l)"
- shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n --> P(l)"
+ shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n \<longrightarrow> P(l)"
apply (induct_tac n, simp)
-apply (blast intro: ih elim!: leE)
+apply (blast intro: ih elim!: leE)
done
theorem list_complete_induct:
- "[| l \<in> list(A);
- \<And>l. [| l \<in> list(A);
- \<forall>l' \<in> list(A). length(l') < length(l) --> P(l')|]
+ "[| l \<in> list(A);
+ \<And>l. [| l \<in> list(A);
+ \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')|]
==> P(l)
|] ==> P(l)"
-apply (rule list_complete_induct_lemma [of A])
- prefer 4 apply (rule le_refl, simp)
- apply blast
- apply simp
+apply (rule list_complete_induct_lemma [of A])
+ prefer 4 apply (rule le_refl, simp)
+ apply blast
+ apply simp
apply assumption
done
@@ -501,13 +501,13 @@
lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)"
by (induct_tac "xs", auto)
-lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) le length(xs)"
+lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) \<le> length(xs)"
apply (induct_tac "xs", auto)
apply (rule_tac j = "length (l) " in le_trans)
apply (auto simp add: le_iff)
done
-lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) <= set_of_list(xs)"
+lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) \<subseteq> set_of_list(xs)"
by (induct_tac "xs", auto)
lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil"
@@ -527,10 +527,10 @@
lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1"
by (erule list.induct, auto)
-lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) <-> xs ~= Nil"
+lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) <-> xs \<noteq> Nil"
by (erule list.induct, auto)
-lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (EX y ys. xs=Cons(y, ys) & length(ys)=n)"
+lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) <-> (\<exists>y ys. xs=Cons(y, ys) & length(ys)=n)"
by (erule list.induct, auto)
(** more theorems about append **)
@@ -554,7 +554,7 @@
(*TOO SLOW as a default simprule!*)
lemma append_left_is_Nil_iff [rule_format]:
"[| xs:list(A); ys:list(A); zs:list(A) |] ==>
- length(ys)=length(zs) --> (xs@ys=zs <-> (xs=Nil & ys=zs))"
+ length(ys)=length(zs) \<longrightarrow> (xs@ys=zs <-> (xs=Nil & ys=zs))"
apply (erule list.induct)
apply (auto simp add: length_app)
done
@@ -562,14 +562,14 @@
(*TOO SLOW as a default simprule!*)
lemma append_left_is_Nil_iff2 [rule_format]:
"[| xs:list(A); ys:list(A); zs:list(A) |] ==>
- length(ys)=length(zs) --> (zs=ys@xs <-> (xs=Nil & ys=zs))"
+ length(ys)=length(zs) \<longrightarrow> (zs=ys@xs <-> (xs=Nil & ys=zs))"
apply (erule list.induct)
apply (auto simp add: length_app)
done
lemma append_eq_append_iff [rule_format,simp]:
"xs:list(A) ==> \<forall>ys \<in> list(A).
- length(xs)=length(ys) --> (xs@us = ys@vs) <-> (xs=ys & us=vs)"
+ length(xs)=length(ys) \<longrightarrow> (xs@us = ys@vs) <-> (xs=ys & us=vs)"
apply (erule list.induct)
apply (simp (no_asm_simp))
apply clarify
@@ -579,7 +579,7 @@
lemma append_eq_append [rule_format]:
"xs:list(A) ==>
\<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
- length(us) = length(vs) --> (xs@us = ys@vs) --> (xs=ys & us=vs)"
+ length(us) = length(vs) \<longrightarrow> (xs@us = ys@vs) \<longrightarrow> (xs=ys & us=vs)"
apply (induct_tac "xs")
apply (force simp add: length_app, clarify)
apply (erule_tac a = ys in list.cases, simp)
@@ -606,13 +606,13 @@
but the proof requires two more hypotheses: x:A and y:A *)
lemma append1_eq_iff [rule_format,simp]:
"xs:list(A) ==> \<forall>ys \<in> list(A). xs@[x] = ys@[y] <-> (xs = ys & x=y)"
-apply (erule list.induct)
- apply clarify
+apply (erule list.induct)
+ apply clarify
apply (erule list.cases)
apply simp_all
-txt{*Inductive step*}
-apply clarify
-apply (erule_tac a=ys in list.cases, simp_all)
+txt{*Inductive step*}
+apply clarify
+apply (erule_tac a=ys in list.cases, simp_all)
done
@@ -623,15 +623,15 @@
lemma append_right_is_self_iff2 [simp]:
"[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) <-> (xs=Nil)"
apply (rule iffI)
-apply (drule sym, auto)
+apply (drule sym, auto)
done
lemma hd_append [rule_format,simp]:
- "xs:list(A) ==> xs ~= Nil --> hd(xs @ ys) = hd(xs)"
+ "xs:list(A) ==> xs \<noteq> Nil \<longrightarrow> hd(xs @ ys) = hd(xs)"
by (induct_tac "xs", auto)
lemma tl_append [rule_format,simp]:
- "xs:list(A) ==> xs~=Nil --> tl(xs @ ys) = tl(xs)@ys"
+ "xs:list(A) ==> xs\<noteq>Nil \<longrightarrow> tl(xs @ ys) = tl(xs)@ys"
by (induct_tac "xs", auto)
(** rev **)
@@ -649,7 +649,7 @@
lemma rev_list_elim [rule_format]:
"xs:list(A) ==>
- (xs=Nil --> P) --> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] -->P)-->P"
+ (xs=Nil \<longrightarrow> P) \<longrightarrow> (\<forall>ys \<in> list(A). \<forall>y \<in> A. xs =ys@[y] \<longrightarrow>P)\<longrightarrow>P"
by (erule list_append_induct, auto)
@@ -662,13 +662,13 @@
done
lemma drop_all [rule_format,simp]:
- "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n --> drop(n, xs)=Nil"
+ "n:nat ==> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> drop(n, xs)=Nil"
apply (erule nat_induct)
apply (auto elim: list.cases)
done
lemma drop_append [rule_format]:
- "n:nat ==>
+ "n:nat ==>
\<forall>xs \<in> list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)"
apply (induct_tac "n")
apply (auto elim: list.cases)
@@ -696,14 +696,14 @@
by (unfold take_def, auto)
lemma take_all [rule_format,simp]:
- "n:nat ==> \<forall>xs \<in> list(A). length(xs) le n --> take(n, xs) = xs"
+ "n:nat ==> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> take(n, xs) = xs"
apply (erule nat_induct)
-apply (auto elim: list.cases)
+apply (auto elim: list.cases)
done
lemma take_type [rule_format,simp,TC]:
"xs:list(A) ==> \<forall>n \<in> nat. take(n, xs):list(A)"
-apply (erule list.induct, simp, clarify)
+apply (erule list.induct, simp, clarify)
apply (erule natE, auto)
done
@@ -711,12 +711,12 @@
"xs:list(A) ==>
\<forall>ys \<in> list(A). \<forall>n \<in> nat. take(n, xs @ ys) =
take(n, xs) @ take(n #- length(xs), ys)"
-apply (erule list.induct, simp, clarify)
+apply (erule list.induct, simp, clarify)
apply (erule natE, auto)
done
lemma take_take [rule_format]:
- "m : nat ==>
+ "m \<in> nat ==>
\<forall>xs \<in> list(A). \<forall>n \<in> nat. take(n, take(m,xs))= take(min(n, m), xs)"
apply (induct_tac "m", auto)
apply (erule_tac a = xs in list.cases)
@@ -728,38 +728,38 @@
(** nth **)
lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
-by (simp add: nth_def)
+by (simp add: nth_def)
lemma nth_Cons [simp]: "n:nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)"
-by (simp add: nth_def)
+by (simp add: nth_def)
lemma nth_empty [simp]: "nth(n, Nil) = 0"
-by (simp add: nth_def)
+by (simp add: nth_def)
lemma nth_type [rule_format,simp,TC]:
- "xs:list(A) ==> \<forall>n. n < length(xs) --> nth(n,xs) : A"
+ "xs:list(A) ==> \<forall>n. n < length(xs) \<longrightarrow> nth(n,xs) \<in> A"
apply (erule list.induct, simp, clarify)
-apply (subgoal_tac "n \<in> nat")
+apply (subgoal_tac "n \<in> nat")
apply (erule natE, auto dest!: le_in_nat)
done
lemma nth_eq_0 [rule_format]:
- "xs:list(A) ==> \<forall>n \<in> nat. length(xs) le n --> nth(n,xs) = 0"
-apply (erule list.induct, simp, clarify)
+ "xs:list(A) ==> \<forall>n \<in> nat. length(xs) \<le> n \<longrightarrow> nth(n,xs) = 0"
+apply (erule list.induct, simp, clarify)
apply (erule natE, auto)
done
lemma nth_append [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) ==>
\<forall>n \<in> nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs)
else nth(n #- length(xs), ys))"
-apply (induct_tac "xs", simp, clarify)
+apply (induct_tac "xs", simp, clarify)
apply (erule natE, auto)
done
lemma set_of_list_conv_nth:
"xs:list(A)
- ==> set_of_list(xs) = {x:A. EX i:nat. i<length(xs) & x = nth(i,xs)}"
+ ==> set_of_list(xs) = {x:A. \<exists>i\<in>nat. i<length(xs) & x = nth(i,xs)}"
apply (induct_tac "xs", simp_all)
apply (rule equalityI, auto)
apply (rule_tac x = 0 in bexI, auto)
@@ -770,40 +770,40 @@
lemma nth_take_lemma [rule_format]:
"k:nat ==>
- \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k le length(xs) --> k le length(ys) -->
- (\<forall>i \<in> nat. i<k --> nth(i,xs) = nth(i,ys))--> take(k,xs) = take(k,ys))"
+ \<forall>xs \<in> list(A). (\<forall>ys \<in> list(A). k \<le> length(xs) \<longrightarrow> k \<le> length(ys) \<longrightarrow>
+ (\<forall>i \<in> nat. i<k \<longrightarrow> nth(i,xs) = nth(i,ys))\<longrightarrow> take(k,xs) = take(k,ys))"
apply (induct_tac "k")
apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib)
apply clarify
(*Both lists are non-empty*)
-apply (erule_tac a=xs in list.cases, simp)
-apply (erule_tac a=ys in list.cases, clarify)
+apply (erule_tac a=xs in list.cases, simp)
+apply (erule_tac a=ys in list.cases, clarify)
apply (simp (no_asm_use) )
apply clarify
apply (simp (no_asm_simp))
apply (rule conjI, force)
-apply (rename_tac y ys z zs)
-apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)
+apply (rename_tac y ys z zs)
+apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto)
done
lemma nth_equalityI [rule_format]:
"[| xs:list(A); ys:list(A); length(xs) = length(ys);
- \<forall>i \<in> nat. i < length(xs) --> nth(i,xs) = nth(i,ys) |]
+ \<forall>i \<in> nat. i < length(xs) \<longrightarrow> nth(i,xs) = nth(i,ys) |]
==> xs = ys"
-apply (subgoal_tac "length (xs) le length (ys) ")
-apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma)
+apply (subgoal_tac "length (xs) \<le> length (ys) ")
+apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma)
apply (simp_all add: take_all)
done
(*The famous take-lemma*)
lemma take_equalityI [rule_format]:
- "[| xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |]
+ "[| xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys)) |]
==> xs = ys"
-apply (case_tac "length (xs) le length (ys) ")
+apply (case_tac "length (xs) \<le> length (ys) ")
apply (drule_tac x = "length (ys) " in bspec)
apply (drule_tac [3] not_lt_imp_le)
-apply (subgoal_tac [5] "length (ys) le length (xs) ")
+apply (subgoal_tac [5] "length (ys) \<le> length (xs) ")
apply (rule_tac [6] j = "succ (length (ys))" in le_trans)
apply (rule_tac [6] leI)
apply (drule_tac [5] x = "length (xs) " in bspec)
@@ -813,20 +813,20 @@
lemma nth_drop [rule_format]:
"n:nat ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
apply (induct_tac "n", simp_all, clarify)
-apply (erule list.cases, auto)
+apply (erule list.cases, auto)
done
lemma take_succ [rule_format]:
- "xs\<in>list(A)
- ==> \<forall>i. i < length(xs) --> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
+ "xs\<in>list(A)
+ ==> \<forall>i. i < length(xs) \<longrightarrow> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
apply (induct_tac "xs", auto)
-apply (subgoal_tac "i\<in>nat")
+apply (subgoal_tac "i\<in>nat")
apply (erule natE)
-apply (auto simp add: le_in_nat)
+apply (auto simp add: le_in_nat)
done
lemma take_add [rule_format]:
- "[|xs\<in>list(A); j\<in>nat|]
+ "[|xs\<in>list(A); j\<in>nat|]
==> \<forall>i\<in>nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))"
apply (induct_tac "xs", simp_all, clarify)
apply (erule_tac n = i in natE, simp_all)
@@ -861,37 +861,37 @@
(* zip equations *)
lemma list_on_set_of_list: "xs \<in> list(A) ==> xs \<in> list(set_of_list(xs))"
-apply (induct_tac xs, simp_all)
-apply (blast intro: list_mono [THEN subsetD])
+apply (induct_tac xs, simp_all)
+apply (blast intro: list_mono [THEN subsetD])
done
lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil"
apply (simp add: zip_def list_on_set_of_list [of _ A])
-apply (erule list.cases, simp_all)
+apply (erule list.cases, simp_all)
done
lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil"
apply (simp add: zip_def list_on_set_of_list [of _ A])
-apply (erule list.cases, simp_all)
+apply (erule list.cases, simp_all)
done
lemma zip_aux_unique [rule_format]:
- "[|B<=C; xs \<in> list(A)|]
+ "[|B<=C; xs \<in> list(A)|]
==> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
-apply (induct_tac xs)
- apply simp_all
- apply (blast intro: list_mono [THEN subsetD], clarify)
-apply (erule_tac a=ys in list.cases, auto)
-apply (blast intro: list_mono [THEN subsetD])
+apply (induct_tac xs)
+ apply simp_all
+ apply (blast intro: list_mono [THEN subsetD], clarify)
+apply (erule_tac a=ys in list.cases, auto)
+apply (blast intro: list_mono [THEN subsetD])
done
lemma zip_Cons_Cons [simp]:
"[| xs:list(A); ys:list(B); x:A; y:B |] ==>
zip(Cons(x,xs), Cons(y, ys)) = Cons(<x,y>, zip(xs, ys))"
-apply (simp add: zip_def, auto)
-apply (rule zip_aux_unique, auto)
+apply (simp add: zip_def, auto)
+apply (rule zip_aux_unique, auto)
apply (simp add: list_on_set_of_list [of _ B])
-apply (blast intro: list_on_set_of_list list_mono [THEN subsetD])
+apply (blast intro: list_on_set_of_list list_mono [THEN subsetD])
done
lemma zip_type [rule_format,simp,TC]:
@@ -907,53 +907,53 @@
"xs:list(A) ==> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
min(length(xs), length(ys))"
apply (unfold min_def)
-apply (induct_tac "xs", simp_all, clarify)
+apply (induct_tac "xs", simp_all, clarify)
apply (erule_tac a = ys in list.cases, auto)
done
lemma zip_append1 [rule_format]:
"[| ys:list(A); zs:list(B) |] ==>
- \<forall>xs \<in> list(A). zip(xs @ ys, zs) =
+ \<forall>xs \<in> list(A). zip(xs @ ys, zs) =
zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))"
-apply (induct_tac "zs", force, clarify)
-apply (erule_tac a = xs in list.cases, simp_all)
+apply (induct_tac "zs", force, clarify)
+apply (erule_tac a = xs in list.cases, simp_all)
done
lemma zip_append2 [rule_format]:
"[| xs:list(A); zs:list(B) |] ==> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)"
-apply (induct_tac "xs", force, clarify)
+apply (induct_tac "xs", force, clarify)
apply (erule_tac a = ys in list.cases, auto)
done
lemma zip_append [simp]:
"[| length(xs) = length(us); length(ys) = length(vs);
- xs:list(A); us:list(B); ys:list(A); vs:list(B) |]
+ xs:list(A); us:list(B); ys:list(A); vs:list(B) |]
==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0)
lemma zip_rev [rule_format,simp]:
"ys:list(B) ==> \<forall>xs \<in> list(A).
- length(xs) = length(ys) --> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
-apply (induct_tac "ys", force, clarify)
+ length(xs) = length(ys) \<longrightarrow> zip(rev(xs), rev(ys)) = rev(zip(xs, ys))"
+apply (induct_tac "ys", force, clarify)
apply (erule_tac a = xs in list.cases)
apply (auto simp add: length_rev)
done
lemma nth_zip [rule_format,simp]:
"ys:list(B) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
- i < length(xs) --> i < length(ys) -->
+ i < length(xs) \<longrightarrow> i < length(ys) \<longrightarrow>
nth(i,zip(xs, ys)) = <nth(i,xs),nth(i, ys)>"
-apply (induct_tac "ys", force, clarify)
-apply (erule_tac a = xs in list.cases, simp)
+apply (induct_tac "ys", force, clarify)
+apply (erule_tac a = xs in list.cases, simp)
apply (auto elim: natE)
done
lemma set_of_list_zip [rule_format]:
"[| xs:list(A); ys:list(B); i:nat |]
==> set_of_list(zip(xs, ys)) =
- {<x, y>:A*B. EX i:nat. i < min(length(xs), length(ys))
+ {<x, y>:A*B. \<exists>i\<in>nat. i < min(length(xs), length(ys))
& x = nth(i, xs) & y = nth(i, ys)}"
by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
@@ -988,15 +988,15 @@
done
lemma nth_list_update [rule_format]:
- "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs) -->
+ "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs) \<longrightarrow>
nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))"
apply (induct_tac "xs")
apply simp_all
apply clarify
-apply (rename_tac i j)
-apply (erule_tac n=i in natE)
+apply (rename_tac i j)
+apply (erule_tac n=i in natE)
apply (erule_tac [2] n=j in natE)
-apply (erule_tac n=j in natE, simp_all, force)
+apply (erule_tac n=j in natE, simp_all, force)
done
lemma nth_list_update_eq [simp]:
@@ -1005,19 +1005,19 @@
lemma nth_list_update_neq [rule_format,simp]:
- "xs:list(A) ==>
- \<forall>i \<in> nat. \<forall>j \<in> nat. i ~= j --> nth(j, list_update(xs,i,x)) = nth(j,xs)"
+ "xs:list(A) ==>
+ \<forall>i \<in> nat. \<forall>j \<in> nat. i \<noteq> j \<longrightarrow> nth(j, list_update(xs,i,x)) = nth(j,xs)"
apply (induct_tac "xs")
apply (simp (no_asm))
apply clarify
apply (erule natE)
-apply (erule_tac [2] natE, simp_all)
+apply (erule_tac [2] natE, simp_all)
apply (erule natE, simp_all)
done
lemma list_update_overwrite [rule_format,simp]:
"xs:list(A) ==> \<forall>i \<in> nat. i < length(xs)
- --> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
+ \<longrightarrow> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
apply (induct_tac "xs")
apply (simp (no_asm))
apply clarify
@@ -1025,8 +1025,8 @@
done
lemma list_update_same_conv [rule_format]:
- "xs:list(A) ==>
- \<forall>i \<in> nat. i < length(xs) -->
+ "xs:list(A) ==>
+ \<forall>i \<in> nat. i < length(xs) \<longrightarrow>
(list_update(xs, i, x) = xs) <-> (nth(i, xs) = x)"
apply (induct_tac "xs")
apply (simp (no_asm))
@@ -1035,9 +1035,9 @@
done
lemma update_zip [rule_format]:
- "ys:list(B) ==>
+ "ys:list(B) ==>
\<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
- length(xs) = length(ys) -->
+ length(xs) = length(ys) \<longrightarrow>
list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
list_update(ys, i, snd(xy)))"
apply (induct_tac "ys")
@@ -1047,8 +1047,8 @@
done
lemma set_update_subset_cons [rule_format]:
- "xs:list(A) ==>
- \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) <= cons(x, set_of_list(xs))"
+ "xs:list(A) ==>
+ \<forall>i \<in> nat. set_of_list(list_update(xs, i, x)) \<subseteq> cons(x, set_of_list(xs))"
apply (induct_tac "xs")
apply simp
apply (rule ballI)
@@ -1056,8 +1056,8 @@
done
lemma set_of_list_update_subsetI:
- "[| set_of_list(xs) <= A; xs:list(A); x:A; i:nat|]
- ==> set_of_list(list_update(xs, i,x)) <= A"
+ "[| set_of_list(xs) \<subseteq> A; xs:list(A); x:A; i:nat|]
+ ==> set_of_list(list_update(xs, i,x)) \<subseteq> A"
apply (rule subset_trans)
apply (rule set_update_subset_cons, auto)
done
@@ -1071,7 +1071,7 @@
apply (auto simp: lt_Ord intro: le_anti_sym)
done
-lemma upt_conv_Nil [simp]: "[| j le i; j:nat |] ==> upt(i,j) = Nil"
+lemma upt_conv_Nil [simp]: "[| j \<le> 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,7 +1079,7 @@
(*Only needed if upt_Suc is deleted from the simpset*)
lemma upt_succ_append:
- "[| i le j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
+ "[| i \<le> j; j:nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
by simp
lemma upt_conv_Cons:
@@ -1093,7 +1093,7 @@
(*LOOPS as a simprule, since j<=j*)
lemma upt_add_eq_append:
- "[| i le j; j:nat; k:nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
+ "[| i \<le> 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)
@@ -1106,22 +1106,22 @@
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 \<longrightarrow> 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
+apply (auto dest!: not_lt_imp_le
simp add: nth_append less_diff_conv add_commute)
done
lemma take_upt [rule_format,simp]:
"[| m:nat; n:nat |] ==>
- \<forall>i \<in> nat. i #+ m le n --> take(m, upt(i,n)) = upt(i,i#+m)"
+ \<forall>i \<in> nat. i #+ m \<le> n \<longrightarrow> take(m, upt(i,n)) = upt(i,i#+m)"
apply (induct_tac "m")
apply (simp (no_asm_simp) add: take_0)
apply clarify
-apply (subst upt_rec, simp)
+apply (subst upt_rec, simp)
apply (rule sym)
-apply (subst upt_rec, simp)
+apply (subst upt_rec, simp)
apply (simp_all del: upt.simps)
apply (rule_tac j = "succ (i #+ x) " in lt_trans2)
apply auto
@@ -1135,7 +1135,7 @@
lemma nth_map [rule_format,simp]:
"xs:list(A) ==>
- \<forall>n \<in> nat. n < length(xs) --> nth(n, map(f, xs)) = f(nth(n, xs))"
+ \<forall>n \<in> nat. n < length(xs) \<longrightarrow> nth(n, map(f, xs)) = f(nth(n, xs))"
apply (induct_tac "xs", simp)
apply (rule ballI)
apply (induct_tac "n", auto)
@@ -1143,15 +1143,15 @@
lemma nth_map_upt [rule_format]:
"[| m:nat; n:nat |] ==>
- \<forall>i \<in> 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)
+ \<forall>i \<in> nat. i < n #- m \<longrightarrow> 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)
apply (subgoal_tac "i < length (upt (0, x))")
- prefer 2
- apply (simp add: less_diff_conv)
+ prefer 2
+ apply (simp add: less_diff_conv)
apply (rule_tac j = "succ (i #+ y) " in lt_trans2)
- apply simp
- apply simp
+ apply simp
+ apply simp
apply (subgoal_tac "i < length (upt (y, x))")
apply (simp_all add: add_commute less_diff_conv)
done
@@ -1202,29 +1202,29 @@
lemma sublist_Cons:
"[| xs:list(B); x:B |] ==>
- sublist(Cons(x, xs), A) =
- (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) : A})"
+ sublist(Cons(x, xs), A) =
+ (if 0:A then [x] else []) @ sublist(xs, {j:nat. succ(j) \<in> 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)
+apply (simp (no_asm_simp) add: sublist_def)
+apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp)
done
lemma sublist_singleton [simp]:
- "sublist([x], A) = (if 0 : A then [x] else [])"
+ "sublist([x], A) = (if 0 \<in> A then [x] else [])"
by (simp add: sublist_Cons)
lemma sublist_upt_eq_take [rule_format, simp]:
- "xs:list(A) ==> ALL n:nat. sublist(xs,n) = take(n,xs)"
-apply (erule list.induct, simp)
-apply (clarify );
-apply (erule natE)
+ "xs:list(A) ==> \<forall>n\<in>nat. sublist(xs,n) = take(n,xs)"
+apply (erule list.induct, simp)
+apply (clarify );
+apply (erule natE)
apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons)
done
lemma sublist_Int_eq:
- "xs : list(B) ==> sublist(xs, A \<inter> nat) = sublist(xs, A)"
+ "xs \<in> list(B) ==> sublist(xs, A \<inter> nat) = sublist(xs, A)"
apply (erule list.induct)
-apply (simp_all add: sublist_Cons)
+apply (simp_all add: sublist_Cons)
done
text{*Repetition of a List Element*}
--- a/src/ZF/Main_ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Main_ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -27,8 +27,8 @@
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) \<in> A |]
+ ==> F^n (a) \<in> A"
by (induct n rule: nat_induct, simp_all)
lemma iterates_omega_triv:
--- a/src/ZF/Nat_ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Nat_ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,7 +9,7 @@
definition
nat :: i where
- "nat == lfp(Inf, %X. {0} Un {succ(i). i:X})"
+ "nat == lfp(Inf, %X. {0} \<union> {succ(i). i:X})"
definition
quasinat :: "i => o" where
@@ -18,26 +18,26 @@
definition
(*Has an unconditional succ case, which is used in "recursor" below.*)
nat_case :: "[i, i=>i, i]=>i" where
- "nat_case(a,b,k) == THE y. k=0 & y=a | (EX x. k=succ(x) & y=b(x))"
+ "nat_case(a,b,k) == THE y. k=0 & y=a | (\<exists>x. k=succ(x) & y=b(x))"
definition
nat_rec :: "[i, i, [i,i]=>i]=>i" where
- "nat_rec(k,a,b) ==
+ "nat_rec(k,a,b) ==
wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))"
(*Internalized relations on the naturals*)
-
+
definition
Le :: i where
- "Le == {<x,y>:nat*nat. x le y}"
+ "Le == {<x,y>:nat*nat. x \<le> y}"
definition
Lt :: i where
"Lt == {<x, y>:nat*nat. x < y}"
-
+
definition
Ge :: i where
- "Ge == {<x,y>:nat*nat. y le x}"
+ "Ge == {<x,y>:nat*nat. y \<le> x}"
definition
Gt :: i where
@@ -51,33 +51,33 @@
predecessors!*}
-lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"
+lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \<union> {succ(i). i:X})"
apply (rule bnd_monoI)
-apply (cut_tac infinity, blast, blast)
+apply (cut_tac infinity, blast, blast)
done
-(* nat = {0} Un {succ(x). x:nat} *)
+(* @{term"nat = {0} \<union> {succ(x). x:nat}"} *)
lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold]]
(** Type checking of 0 and successor **)
-lemma nat_0I [iff,TC]: "0 : nat"
+lemma nat_0I [iff,TC]: "0 \<in> nat"
apply (subst nat_unfold)
apply (rule singletonI [THEN UnI1])
done
-lemma nat_succI [intro!,TC]: "n : nat ==> succ(n) : nat"
+lemma nat_succI [intro!,TC]: "n \<in> nat ==> succ(n) \<in> nat"
apply (subst nat_unfold)
apply (erule RepFunI [THEN UnI2])
done
-lemma nat_1I [iff,TC]: "1 : nat"
+lemma nat_1I [iff,TC]: "1 \<in> nat"
by (rule nat_0I [THEN nat_succI])
-lemma nat_2I [iff,TC]: "2 : nat"
+lemma nat_2I [iff,TC]: "2 \<in> nat"
by (rule nat_1I [THEN nat_succI])
-lemma bool_subset_nat: "bool <= nat"
+lemma bool_subset_nat: "bool \<subseteq> nat"
by (blast elim!: boolE)
lemmas bool_into_nat = bool_subset_nat [THEN subsetD]
@@ -92,15 +92,15 @@
lemma natE:
"[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"
-by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto)
+by (erule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE], auto)
lemma nat_into_Ord [simp]: "n: nat ==> Ord(n)"
by (erule nat_induct, auto)
-(* i: nat ==> 0 le i; same thing as 0<succ(i) *)
+(* @{term"i: nat ==> 0 \<le> i"}; same thing as @{term"0<succ(i)"} *)
lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le]
-(* i: nat ==> i le i; same thing as i<succ(i) *)
+(* @{term"i: nat ==> i \<le> i"}; same thing as @{term"i<succ(i)"} *)
lemmas nat_le_refl = nat_into_Ord [THEN le_refl]
lemma Ord_nat [iff]: "Ord(nat)"
@@ -108,7 +108,7 @@
apply (erule_tac [2] nat_into_Ord [THEN Ord_is_Transset])
apply (unfold Transset_def)
apply (rule ballI)
-apply (erule nat_induct, auto)
+apply (erule nat_induct, auto)
done
lemma Limit_nat [iff]: "Limit(nat)"
@@ -126,12 +126,12 @@
lemma nat_succ_iff [iff]: "succ(n): nat <-> n: nat"
by (blast dest!: succ_natD)
-lemma nat_le_Limit: "Limit(i) ==> nat le i"
+lemma nat_le_Limit: "Limit(i) ==> nat \<le> i"
apply (rule subset_imp_le)
-apply (simp_all add: Limit_is_Ord)
+apply (simp_all add: Limit_is_Ord)
apply (rule subsetI)
apply (erule nat_induct)
- apply (erule Limit_has_0 [THEN ltD])
+ apply (erule Limit_has_0 [THEN ltD])
apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord)
done
@@ -140,10 +140,10 @@
lemma lt_nat_in_nat: "[| m<n; n: nat |] ==> m: nat"
apply (erule ltE)
-apply (erule Ord_trans, assumption, simp)
+apply (erule Ord_trans, assumption, simp)
done
-lemma le_in_nat: "[| m le n; n:nat |] ==> m:nat"
+lemma le_in_nat: "[| m \<le> n; n:nat |] ==> m:nat"
by (blast dest!: lt_nat_in_nat)
@@ -153,59 +153,59 @@
lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1]
-lemmas complete_induct_rule =
+lemmas complete_induct_rule =
complete_induct [rule_format, case_names less, consumes 1]
-lemma nat_induct_from_lemma [rule_format]:
- "[| n: nat; m: nat;
- !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) |]
- ==> m le n --> P(m) --> P(n)"
-apply (erule nat_induct)
+lemma nat_induct_from_lemma [rule_format]:
+ "[| n: nat; m: nat;
+ !!x. [| x: nat; m \<le> x; P(x) |] ==> P(succ(x)) |]
+ ==> m \<le> n \<longrightarrow> P(m) \<longrightarrow> P(n)"
+apply (erule nat_induct)
apply (simp_all add: distrib_simps le0_iff le_succ_iff)
done
(*Induction starting from m rather than 0*)
-lemma nat_induct_from:
- "[| m le n; m: nat; n: nat;
- P(m);
- !!x. [| x: nat; m le x; P(x) |] ==> P(succ(x)) |]
+lemma nat_induct_from:
+ "[| m \<le> n; m: nat; n: nat;
+ P(m);
+ !!x. [| x: nat; m \<le> 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));
+ "[| 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)
+apply (erule nat_induct, simp)
apply (rule ballI)
apply (rename_tac i j)
-apply (erule_tac n=j in nat_induct, auto)
+apply (erule_tac n=j in nat_induct, auto)
done
(** Induction principle analogous to trancl_induct **)
lemma succ_lt_induct_lemma [rule_format]:
- "m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) -->
- (ALL n:nat. m<n --> P(m,n))"
+ "m: nat ==> P(m,succ(m)) \<longrightarrow> (\<forall>x\<in>nat. P(m,x) \<longrightarrow> P(m,succ(x))) \<longrightarrow>
+ (\<forall>n\<in>nat. m<n \<longrightarrow> P(m,n))"
apply (erule nat_induct)
apply (intro impI, rule nat_induct [THEN ballI])
prefer 4 apply (intro impI, rule nat_induct [THEN ballI])
-apply (auto simp add: le_iff)
+apply (auto simp add: le_iff)
done
lemma succ_lt_induct:
- "[| m<n; n: nat;
- P(m,succ(m));
+ "[| m<n; n: nat;
+ P(m,succ(m));
!!x. [| x: nat; P(m,x) |] ==> P(m,succ(x)) |]
==> P(m,n)"
-by (blast intro: succ_lt_induct_lemma lt_nat_in_nat)
+by (blast intro: succ_lt_induct_lemma lt_nat_in_nat)
subsection{*quasinat: to allow a case-split rule for @{term nat_case}*}
@@ -219,36 +219,36 @@
lemma nat_imp_quasinat: "n \<in> nat ==> quasinat(n)"
by (erule natE, simp_all)
-lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0"
-by (simp add: quasinat_def nat_case_def)
+lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0"
+by (simp add: quasinat_def nat_case_def)
lemma nat_cases_disj: "k=0 | (\<exists>y. k = succ(y)) | ~ quasinat(k)"
-apply (case_tac "k=0", simp)
-apply (case_tac "\<exists>m. k = succ(m)")
-apply (simp_all add: quasinat_def)
+apply (case_tac "k=0", simp)
+apply (case_tac "\<exists>m. k = succ(m)")
+apply (simp_all add: quasinat_def)
done
lemma nat_cases:
"[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P"
-by (insert nat_cases_disj [of k], blast)
+by (insert nat_cases_disj [of k], blast)
(** nat_case **)
lemma nat_case_0 [simp]: "nat_case(a,b,0) = a"
by (simp add: nat_case_def)
-lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)"
+lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)"
by (simp add: nat_case_def)
lemma nat_case_type [TC]:
- "[| 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)
+ "[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) |]
+ ==> nat_case(a,b,n) \<in> C(n)";
+by (erule nat_induct, auto)
lemma split_nat_case:
- "P(nat_case(a,b,k)) <->
- ((k=0 --> P(a)) & (\<forall>x. k=succ(x) --> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))"
-apply (rule nat_cases [of k])
+ "P(nat_case(a,b,k)) <->
+ ((k=0 \<longrightarrow> P(a)) & (\<forall>x. k=succ(x) \<longrightarrow> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))"
+apply (rule nat_cases [of k])
apply (auto simp add: non_nat_case)
done
@@ -260,41 +260,41 @@
lemma nat_rec_0: "nat_rec(0,a,b) = a"
apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
- apply (rule wf_Memrel)
+ apply (rule wf_Memrel)
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))"
apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
- apply (rule wf_Memrel)
+ apply (rule wf_Memrel)
apply (simp add: vimage_singleton_iff)
done
(** The union of two natural numbers is a natural number -- their maximum **)
-lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i Un j: nat"
+lemma Un_nat_type [TC]: "[| i: nat; j: nat |] ==> i \<union> j: nat"
apply (rule Un_least_lt [THEN ltD])
-apply (simp_all add: lt_def)
+apply (simp_all add: lt_def)
done
-lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i Int j: nat"
+lemma Int_nat_type [TC]: "[| i: nat; j: nat |] ==> i \<inter> j: nat"
apply (rule Int_greatest_lt [THEN ltD])
-apply (simp_all add: lt_def)
+apply (simp_all add: lt_def)
done
(*needed to simplify unions over nat*)
-lemma nat_nonempty [simp]: "nat ~= 0"
+lemma nat_nonempty [simp]: "nat \<noteq> 0"
by blast
text{*A natural number is the set of its predecessors*}
lemma nat_eq_Collect_lt: "i \<in> nat ==> {j\<in>nat. j<i} = i"
apply (rule equalityI)
-apply (blast dest: ltD)
+apply (blast dest: ltD)
apply (auto simp add: Ord_mem_iff_lt)
-apply (blast intro: lt_trans)
+apply (blast intro: lt_trans)
done
-lemma Le_iff [iff]: "<x,y> : Le <-> x le y & x : nat & y : nat"
+lemma Le_iff [iff]: "<x,y> \<in> Le <-> x \<le> y & x \<in> nat & y \<in> nat"
by (force simp add: Le_def)
end
--- a/src/ZF/OrdQuant.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/OrdQuant.thy Tue Mar 06 15:15:49 2012 +0000
@@ -11,11 +11,11 @@
definition
(* Ordinal Quantifiers *)
oall :: "[i, i => o] => o" where
- "oall(A, P) == ALL x. x<A --> P(x)"
+ "oall(A, P) == \<forall>x. x<A \<longrightarrow> P(x)"
definition
oex :: "[i, i => o] => o" where
- "oex(A, P) == EX x. x<A & P(x)"
+ "oex(A, P) == \<exists>x. x<A & P(x)"
definition
(* Ordinal Union *)
@@ -48,18 +48,18 @@
(*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize
is proved. Ord_atomize would convert this rule to
x < 0 ==> P(x) == True, which causes dire effects!*)
-lemma [simp]: "(ALL x<0. P(x))"
+lemma [simp]: "(\<forall>x<0. P(x))"
by (simp add: oall_def)
-lemma [simp]: "~(EX x<0. P(x))"
+lemma [simp]: "~(\<exists>x<0. P(x))"
by (simp add: oex_def)
-lemma [simp]: "(ALL x<succ(i). P(x)) <-> (Ord(i) --> P(i) & (ALL x<i. P(x)))"
+lemma [simp]: "(\<forall>x<succ(i). P(x)) <-> (Ord(i) \<longrightarrow> P(i) & (\<forall>x<i. P(x)))"
apply (simp add: oall_def le_iff)
apply (blast intro: lt_Ord2)
done
-lemma [simp]: "(EX x<succ(i). P(x)) <-> (Ord(i) & (P(i) | (EX x<i. P(x))))"
+lemma [simp]: "(\<exists>x<succ(i). P(x)) <-> (Ord(i) & (P(i) | (\<exists>x<i. P(x))))"
apply (simp add: oex_def le_iff)
apply (blast intro: lt_Ord2)
done
@@ -84,12 +84,11 @@
lemma Limit_OUN_eq: "Limit(i) ==> (\<Union>x<i. x) = i"
by (simp add: OUnion_def Limit_Union_eq Limit_is_Ord)
-(* No < version; consider (\<Union>i\<in>nat.i)=nat *)
+(* No < version of this theorem: consider that @{term"(\<Union>i\<in>nat.i)=nat"}! *)
lemma OUN_least:
"(!!x. x<A ==> B(x) \<subseteq> C) ==> (\<Union>x<A. B(x)) \<subseteq> C"
by (simp add: OUnion_def UN_least ltI)
-(* No < version; consider (\<Union>i\<in>nat.i)=nat *)
lemma OUN_least_le:
"[| Ord(i); !!x. x<A ==> b(x) \<le> i |] ==> (\<Union>x<A. b(x)) \<le> i"
by (simp add: OUnion_def UN_least_le ltI Ord_0_le)
@@ -105,34 +104,34 @@
lemma OUN_Union_eq:
"(!!x. x:X ==> Ord(x))
- ==> (\<Union>z < Union(X). C(z)) = (\<Union>x\<in>X. \<Union>z < x. C(z))"
+ ==> (\<Union>z < \<Union>(X). C(z)) = (\<Union>x\<in>X. \<Union>z < x. C(z))"
by (simp add: OUnion_def)
-(*So that rule_format will get rid of ALL x<A...*)
+(*So that rule_format will get rid of this quantifier...*)
lemma atomize_oall [symmetric, rulify]:
- "(!!x. x<A ==> P(x)) == Trueprop (ALL x<A. P(x))"
+ "(!!x. x<A ==> P(x)) == Trueprop (\<forall>x<A. P(x))"
by (simp add: oall_def atomize_all atomize_imp)
subsubsection {*universal quantifier for ordinals*}
lemma oallI [intro!]:
- "[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)"
+ "[| !!x. x<A ==> P(x) |] ==> \<forall>x<A. P(x)"
by (simp add: oall_def)
-lemma ospec: "[| ALL x<A. P(x); x<A |] ==> P(x)"
+lemma ospec: "[| \<forall>x<A. P(x); x<A |] ==> P(x)"
by (simp add: oall_def)
lemma oallE:
- "[| ALL x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q"
+ "[| \<forall>x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q"
by (simp add: oall_def, blast)
lemma rev_oallE [elim]:
- "[| ALL x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q"
+ "[| \<forall>x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q"
by (simp add: oall_def, blast)
-(*Trival rewrite rule; (ALL x<a.P)<->P holds only if a is not 0!*)
-lemma oall_simp [simp]: "(ALL x<a. True) <-> True"
+(*Trival rewrite rule. @{term"(\<forall>x<a.P)<->P"} holds only if a is not 0!*)
+lemma oall_simp [simp]: "(\<forall>x<a. True) <-> True"
by blast
(*Congruence rule for rewriting*)
@@ -145,18 +144,18 @@
subsubsection {*existential quantifier for ordinals*}
lemma oexI [intro]:
- "[| P(x); x<A |] ==> EX x<A. P(x)"
+ "[| P(x); x<A |] ==> \<exists>x<A. P(x)"
apply (simp add: oex_def, blast)
done
-(*Not of the general form for such rules; ~EX has become ALL~ *)
+(*Not of the general form for such rules... *)
lemma oexCI:
- "[| ALL x<A. ~P(x) ==> P(a); a<A |] ==> EX x<A. P(x)"
+ "[| \<forall>x<A. ~P(x) ==> P(a); a<A |] ==> \<exists>x<A. P(x)"
apply (simp add: oex_def, blast)
done
lemma oexE [elim!]:
- "[| EX x<A. P(x); !!x. [| x<A; P(x) |] ==> Q |] ==> Q"
+ "[| \<exists>x<A. P(x); !!x. [| x<A; P(x) |] ==> Q |] ==> Q"
apply (simp add: oex_def, blast)
done
@@ -173,11 +172,11 @@
by (unfold OUnion_def lt_def, blast)
lemma OUN_E [elim!]:
- "[| b : (\<Union>z<i. B(z)); !!a.[| b: B(a); a<i |] ==> R |] ==> R"
+ "[| b \<in> (\<Union>z<i. B(z)); !!a.[| b: B(a); a<i |] ==> R |] ==> R"
apply (unfold OUnion_def lt_def, blast)
done
-lemma OUN_iff: "b : (\<Union>x<i. B(x)) <-> (EX x<i. b : B(x))"
+lemma OUN_iff: "b \<in> (\<Union>x<i. B(x)) <-> (\<exists>x<i. b \<in> B(x))"
by (unfold OUnion_def oex_def lt_def, blast)
lemma OUN_cong [cong]:
@@ -185,7 +184,7 @@
by (simp add: OUnion_def lt_def OUN_iff)
lemma lt_induct:
- "[| i<k; !!x.[| x<k; ALL y<x. P(y) |] ==> P(x) |] ==> P(i)"
+ "[| i<k; !!x.[| x<k; \<forall>y<x. P(y) |] ==> P(x) |] ==> P(i)"
apply (simp add: lt_def oall_def)
apply (erule conjE)
apply (erule Ord_induct, assumption, blast)
@@ -196,11 +195,11 @@
definition
"rall" :: "[i=>o, i=>o] => o" where
- "rall(M, P) == ALL x. M(x) --> P(x)"
+ "rall(M, P) == \<forall>x. M(x) \<longrightarrow> P(x)"
definition
"rex" :: "[i=>o, i=>o] => o" where
- "rex(M, P) == EX x. M(x) & P(x)"
+ "rex(M, P) == \<exists>x. M(x) & P(x)"
syntax
"_rall" :: "[pttrn, i=>o, o] => o" ("(3ALL _[_]./ _)" 10)
@@ -220,18 +219,18 @@
subsubsection{*Relativized universal quantifier*}
-lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> ALL x[M]. P(x)"
+lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> \<forall>x[M]. P(x)"
by (simp add: rall_def)
-lemma rspec: "[| ALL x[M]. P(x); M(x) |] ==> P(x)"
+lemma rspec: "[| \<forall>x[M]. P(x); M(x) |] ==> P(x)"
by (simp add: rall_def)
(*Instantiates x first: better for automatic theorem proving?*)
lemma rev_rallE [elim]:
- "[| ALL x[M]. P(x); ~ M(x) ==> Q; P(x) ==> Q |] ==> Q"
+ "[| \<forall>x[M]. P(x); ~ M(x) ==> Q; P(x) ==> Q |] ==> Q"
by (simp add: rall_def, blast)
-lemma rallE: "[| ALL x[M]. P(x); P(x) ==> Q; ~ M(x) ==> Q |] ==> Q"
+lemma rallE: "[| \<forall>x[M]. P(x); P(x) ==> Q; ~ M(x) ==> Q |] ==> Q"
by blast
(*Trival rewrite rule; (ALL x[M].P)<->P holds only if A is nonempty!*)
@@ -240,24 +239,24 @@
(*Congruence rule for rewriting*)
lemma rall_cong [cong]:
- "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (ALL x[M]. P(x)) <-> (ALL x[M]. P'(x))"
+ "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\<forall>x[M]. P(x)) <-> (\<forall>x[M]. P'(x))"
by (simp add: rall_def)
subsubsection{*Relativized existential quantifier*}
-lemma rexI [intro]: "[| P(x); M(x) |] ==> EX x[M]. P(x)"
+lemma rexI [intro]: "[| P(x); M(x) |] ==> \<exists>x[M]. P(x)"
by (simp add: rex_def, blast)
(*The best argument order when there is only one M(x)*)
-lemma rev_rexI: "[| M(x); P(x) |] ==> EX x[M]. P(x)"
+lemma rev_rexI: "[| M(x); P(x) |] ==> \<exists>x[M]. P(x)"
by blast
-(*Not of the general form for such rules; ~EX has become ALL~ *)
-lemma rexCI: "[| ALL x[M]. ~P(x) ==> P(a); M(a) |] ==> EX x[M]. P(x)"
+(*Not of the general form for such rules... *)
+lemma rexCI: "[| \<forall>x[M]. ~P(x) ==> P(a); M(a) |] ==> \<exists>x[M]. P(x)"
by blast
-lemma rexE [elim!]: "[| EX x[M]. P(x); !!x. [| M(x); P(x) |] ==> Q |] ==> Q"
+lemma rexE [elim!]: "[| \<exists>x[M]. P(x); !!x. [| M(x); P(x) |] ==> Q |] ==> Q"
by (simp add: rex_def, blast)
(*We do not even have (EX x[M]. True) <-> True unless A is nonempty!!*)
@@ -265,7 +264,7 @@
by (simp add: rex_def)
lemma rex_cong [cong]:
- "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (EX x[M]. P(x)) <-> (EX x[M]. P'(x))"
+ "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\<exists>x[M]. P(x)) <-> (\<exists>x[M]. P'(x))"
by (simp add: rex_def cong: conj_cong)
lemma rall_is_ball [simp]: "(\<forall>x[%z. z\<in>A]. P(x)) <-> (\<forall>x\<in>A. P(x))"
@@ -274,68 +273,68 @@
lemma rex_is_bex [simp]: "(\<exists>x[%z. z\<in>A]. P(x)) <-> (\<exists>x\<in>A. P(x))"
by blast
-lemma atomize_rall: "(!!x. M(x) ==> P(x)) == Trueprop (ALL x[M]. P(x))";
+lemma atomize_rall: "(!!x. M(x) ==> P(x)) == Trueprop (\<forall>x[M]. P(x))";
by (simp add: rall_def atomize_all atomize_imp)
declare atomize_rall [symmetric, rulify]
lemma rall_simps1:
- "(ALL x[M]. P(x) & Q) <-> (ALL x[M]. P(x)) & ((ALL x[M]. False) | Q)"
- "(ALL x[M]. P(x) | Q) <-> ((ALL x[M]. P(x)) | Q)"
- "(ALL x[M]. P(x) --> Q) <-> ((EX x[M]. P(x)) --> Q)"
- "(~(ALL x[M]. P(x))) <-> (EX x[M]. ~P(x))"
+ "(\<forall>x[M]. P(x) & Q) <-> (\<forall>x[M]. P(x)) & ((\<forall>x[M]. False) | Q)"
+ "(\<forall>x[M]. P(x) | Q) <-> ((\<forall>x[M]. P(x)) | Q)"
+ "(\<forall>x[M]. P(x) \<longrightarrow> Q) <-> ((\<exists>x[M]. P(x)) \<longrightarrow> Q)"
+ "(~(\<forall>x[M]. P(x))) <-> (\<exists>x[M]. ~P(x))"
by blast+
lemma rall_simps2:
- "(ALL x[M]. P & Q(x)) <-> ((ALL x[M]. False) | P) & (ALL x[M]. Q(x))"
- "(ALL x[M]. P | Q(x)) <-> (P | (ALL x[M]. Q(x)))"
- "(ALL x[M]. P --> Q(x)) <-> (P --> (ALL x[M]. Q(x)))"
+ "(\<forall>x[M]. P & Q(x)) <-> ((\<forall>x[M]. False) | P) & (\<forall>x[M]. Q(x))"
+ "(\<forall>x[M]. P | Q(x)) <-> (P | (\<forall>x[M]. Q(x)))"
+ "(\<forall>x[M]. P \<longrightarrow> Q(x)) <-> (P \<longrightarrow> (\<forall>x[M]. Q(x)))"
by blast+
lemmas rall_simps [simp] = rall_simps1 rall_simps2
lemma rall_conj_distrib:
- "(ALL x[M]. P(x) & Q(x)) <-> ((ALL x[M]. P(x)) & (ALL x[M]. Q(x)))"
+ "(\<forall>x[M]. P(x) & Q(x)) <-> ((\<forall>x[M]. P(x)) & (\<forall>x[M]. Q(x)))"
by blast
lemma rex_simps1:
- "(EX x[M]. P(x) & Q) <-> ((EX x[M]. P(x)) & Q)"
- "(EX x[M]. P(x) | Q) <-> (EX x[M]. P(x)) | ((EX x[M]. True) & Q)"
- "(EX x[M]. P(x) --> Q) <-> ((ALL x[M]. P(x)) --> ((EX x[M]. True) & Q))"
- "(~(EX x[M]. P(x))) <-> (ALL x[M]. ~P(x))"
+ "(\<exists>x[M]. P(x) & Q) <-> ((\<exists>x[M]. P(x)) & Q)"
+ "(\<exists>x[M]. P(x) | Q) <-> (\<exists>x[M]. P(x)) | ((\<exists>x[M]. True) & Q)"
+ "(\<exists>x[M]. P(x) \<longrightarrow> Q) <-> ((\<forall>x[M]. P(x)) \<longrightarrow> ((\<exists>x[M]. True) & Q))"
+ "(~(\<exists>x[M]. P(x))) <-> (\<forall>x[M]. ~P(x))"
by blast+
lemma rex_simps2:
- "(EX x[M]. P & Q(x)) <-> (P & (EX x[M]. Q(x)))"
- "(EX x[M]. P | Q(x)) <-> ((EX x[M]. True) & P) | (EX x[M]. Q(x))"
- "(EX x[M]. P --> Q(x)) <-> (((ALL x[M]. False) | P) --> (EX x[M]. Q(x)))"
+ "(\<exists>x[M]. P & Q(x)) <-> (P & (\<exists>x[M]. Q(x)))"
+ "(\<exists>x[M]. P | Q(x)) <-> ((\<exists>x[M]. True) & P) | (\<exists>x[M]. Q(x))"
+ "(\<exists>x[M]. P \<longrightarrow> Q(x)) <-> (((\<forall>x[M]. False) | P) \<longrightarrow> (\<exists>x[M]. Q(x)))"
by blast+
lemmas rex_simps [simp] = rex_simps1 rex_simps2
lemma rex_disj_distrib:
- "(EX x[M]. P(x) | Q(x)) <-> ((EX x[M]. P(x)) | (EX x[M]. Q(x)))"
+ "(\<exists>x[M]. P(x) | Q(x)) <-> ((\<exists>x[M]. P(x)) | (\<exists>x[M]. Q(x)))"
by blast
subsubsection{*One-point rule for bounded quantifiers*}
-lemma rex_triv_one_point1 [simp]: "(EX x[M]. x=a) <-> ( M(a))"
+lemma rex_triv_one_point1 [simp]: "(\<exists>x[M]. x=a) <-> ( M(a))"
by blast
-lemma rex_triv_one_point2 [simp]: "(EX x[M]. a=x) <-> ( M(a))"
+lemma rex_triv_one_point2 [simp]: "(\<exists>x[M]. a=x) <-> ( M(a))"
by blast
-lemma rex_one_point1 [simp]: "(EX x[M]. x=a & P(x)) <-> ( M(a) & P(a))"
+lemma rex_one_point1 [simp]: "(\<exists>x[M]. x=a & P(x)) <-> ( M(a) & P(a))"
by blast
-lemma rex_one_point2 [simp]: "(EX x[M]. a=x & P(x)) <-> ( M(a) & P(a))"
+lemma rex_one_point2 [simp]: "(\<exists>x[M]. a=x & P(x)) <-> ( M(a) & P(a))"
by blast
-lemma rall_one_point1 [simp]: "(ALL x[M]. x=a --> P(x)) <-> ( M(a) --> P(a))"
+lemma rall_one_point1 [simp]: "(\<forall>x[M]. x=a \<longrightarrow> P(x)) <-> ( M(a) \<longrightarrow> P(a))"
by blast
-lemma rall_one_point2 [simp]: "(ALL x[M]. a=x --> P(x)) <-> ( M(a) --> P(a))"
+lemma rall_one_point2 [simp]: "(\<forall>x[M]. a=x \<longrightarrow> P(x)) <-> ( M(a) \<longrightarrow> P(a))"
by blast
@@ -343,9 +342,9 @@
definition
setclass :: "[i,i] => o" ("##_" [40] 40) where
- "setclass(A) == %x. x : A"
+ "setclass(A) == %x. x \<in> A"
-lemma setclass_iff [simp]: "setclass(A,x) <-> x : A"
+lemma setclass_iff [simp]: "setclass(A,x) <-> x \<in> A"
by (simp add: setclass_def)
lemma rall_setclass_is_ball [simp]: "(\<forall>x[##A]. P(x)) <-> (\<forall>x\<in>A. P(x))"
@@ -368,14 +367,14 @@
text {* Setting up the one-point-rule simproc *}
-simproc_setup defined_rex ("EX x[M]. P(x) & Q(x)") = {*
+simproc_setup defined_rex ("\<exists>x[M]. P(x) & Q(x)") = {*
let
val unfold_rex_tac = unfold_tac @{thms rex_def};
fun prove_rex_tac ss = unfold_rex_tac ss THEN Quantifier1.prove_one_point_ex_tac;
in fn _ => fn ss => Quantifier1.rearrange_bex (prove_rex_tac ss) ss end
*}
-simproc_setup defined_rall ("ALL x[M]. P(x) --> Q(x)") = {*
+simproc_setup defined_rall ("\<forall>x[M]. P(x) \<longrightarrow> Q(x)") = {*
let
val unfold_rall_tac = unfold_tac @{thms rall_def};
fun prove_rall_tac ss = unfold_rall_tac ss THEN Quantifier1.prove_one_point_all_tac;
--- a/src/ZF/Order.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Order.thy Tue Mar 06 15:15:49 2012 +0000
@@ -21,7 +21,7 @@
definition
linear :: "[i,i]=>o" (*Strict total ordering*) where
- "linear(A,r) == (ALL x:A. ALL y:A. <x,y>:r | x=y | <y,x>:r)"
+ "linear(A,r) == (\<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r | x=y | <y,x>:r)"
definition
tot_ord :: "[i,i]=>o" (*Strict total ordering*) where
@@ -46,12 +46,12 @@
definition
mono_map :: "[i,i,i,i]=>i" (*Order-preserving maps*) where
"mono_map(A,r,B,s) ==
- {f: A->B. ALL x:A. ALL y:A. <x,y>:r --> <f`x,f`y>:s}"
+ {f: A->B. \<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r \<longrightarrow> <f`x,f`y>:s}"
definition
ord_iso :: "[i,i,i,i]=>i" (*Order isomorphisms*) where
"ord_iso(A,r,B,s) ==
- {f: bij(A,B). ALL x:A. ALL y:A. <x,y>:r <-> <f`x,f`y>:s}"
+ {f: bij(A,B). \<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r <-> <f`x,f`y>:s}"
definition
pred :: "[i,i,i]=>i" (*Set of predecessors*) where
@@ -64,7 +64,7 @@
definition
first :: "[i, i, i] => o" where
- "first(u, X, R) == u:X & (ALL v:X. v~=u --> <u,v> : R)"
+ "first(u, X, R) == u:X & (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
notation (xsymbols)
@@ -74,7 +74,7 @@
subsection{*Immediate Consequences of the Definitions*}
lemma part_ord_Imp_asym:
- "part_ord(A,r) ==> asym(r Int A*A)"
+ "part_ord(A,r) ==> asym(r \<inter> A*A)"
by (unfold part_ord_def irrefl_def trans_on_def asym_def, blast)
lemma linearE:
@@ -107,7 +107,7 @@
(** Derived rules for pred(A,x,r) **)
-lemma pred_iff: "y : pred(A,x,r) <-> <y,x>:r & y:A"
+lemma pred_iff: "y \<in> pred(A,x,r) <-> <y,x>:r & y:A"
by (unfold pred_def, blast)
lemmas predI = conjI [THEN pred_iff [THEN iffD2]]
@@ -115,14 +115,14 @@
lemma predE: "[| y: pred(A,x,r); [| y:A; <y,x>:r |] ==> P |] ==> P"
by (simp add: pred_def)
-lemma pred_subset_under: "pred(A,x,r) <= r -`` {x}"
+lemma pred_subset_under: "pred(A,x,r) \<subseteq> r -`` {x}"
by (simp add: pred_def, blast)
-lemma pred_subset: "pred(A,x,r) <= A"
+lemma pred_subset: "pred(A,x,r) \<subseteq> A"
by (simp add: pred_def, blast)
lemma pred_pred_eq:
- "pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)"
+ "pred(pred(A,x,r), y, r) = pred(A,x,r) \<inter> pred(A,y,r)"
by (simp add: pred_def, blast)
lemma trans_pred_pred_eq:
@@ -160,30 +160,30 @@
(** Relations restricted to a smaller domain, by Krzysztof Grabczewski **)
-lemma irrefl_Int_iff: "irrefl(A,r Int A*A) <-> irrefl(A,r)"
+lemma irrefl_Int_iff: "irrefl(A,r \<inter> A*A) <-> irrefl(A,r)"
by (unfold irrefl_def, blast)
-lemma trans_on_Int_iff: "trans[A](r Int A*A) <-> trans[A](r)"
+lemma trans_on_Int_iff: "trans[A](r \<inter> A*A) <-> trans[A](r)"
by (unfold trans_on_def, blast)
-lemma part_ord_Int_iff: "part_ord(A,r Int A*A) <-> part_ord(A,r)"
+lemma part_ord_Int_iff: "part_ord(A,r \<inter> A*A) <-> part_ord(A,r)"
apply (unfold part_ord_def)
apply (simp add: irrefl_Int_iff trans_on_Int_iff)
done
-lemma linear_Int_iff: "linear(A,r Int A*A) <-> linear(A,r)"
+lemma linear_Int_iff: "linear(A,r \<inter> A*A) <-> linear(A,r)"
by (unfold linear_def, blast)
-lemma tot_ord_Int_iff: "tot_ord(A,r Int A*A) <-> tot_ord(A,r)"
+lemma tot_ord_Int_iff: "tot_ord(A,r \<inter> A*A) <-> tot_ord(A,r)"
apply (unfold tot_ord_def)
apply (simp add: part_ord_Int_iff linear_Int_iff)
done
-lemma wf_on_Int_iff: "wf[A](r Int A*A) <-> wf[A](r)"
+lemma wf_on_Int_iff: "wf[A](r \<inter> A*A) <-> wf[A](r)"
apply (unfold wf_on_def wf_def, fast) (*10 times faster than blast!*)
done
-lemma well_ord_Int_iff: "well_ord(A,r Int A*A) <-> well_ord(A,r)"
+lemma well_ord_Int_iff: "well_ord(A,r \<inter> A*A) <-> well_ord(A,r)"
apply (unfold well_ord_def)
apply (simp add: tot_ord_Int_iff wf_on_Int_iff)
done
@@ -256,7 +256,7 @@
lemma ord_isoI:
"[| f: bij(A, B);
- !!x y. [| x:A; y:A |] ==> <x, y> : r <-> <f`x, f`y> : s |]
+ !!x y. [| x:A; y:A |] ==> <x, y> \<in> r <-> <f`x, f`y> \<in> s |]
==> f: ord_iso(A,r,B,s)"
by (simp add: ord_iso_def)
@@ -272,12 +272,12 @@
(*Needed? But ord_iso_converse is!*)
lemma ord_iso_apply:
- "[| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> <f`x, f`y> : s"
+ "[| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> <f`x, f`y> \<in> s"
by (simp add: ord_iso_def)
lemma ord_iso_converse:
"[| f: ord_iso(A,r,B,s); <x,y>: s; x:B; y:B |]
- ==> <converse(f) ` x, converse(f) ` y> : r"
+ ==> <converse(f) ` x, converse(f) ` y> \<in> r"
apply (simp add: ord_iso_def, clarify)
apply (erule bspec [THEN bspec, THEN iffD2])
apply (erule asm_rl bij_converse_bij [THEN bij_is_fun, THEN apply_type])+
@@ -323,7 +323,7 @@
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 "<g` (f`x), g` (f`y) > : r")
+apply (subgoal_tac "<g` (f`x), g` (f`y) > \<in> r")
apply (simp add: comp_eq_id_iff [THEN iffD1])
apply (blast intro: apply_funtype)
done
@@ -360,7 +360,7 @@
lemma wf_on_ord_iso:
"[| 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 Int A}" in spec)
+apply (drule_tac x = "{f`z. z:Z \<inter> A}" in spec)
apply (safe intro!: equalityI)
apply (blast dest!: equalityD1 intro: bij_is_fun [THEN apply_type])+
done
@@ -388,18 +388,18 @@
(*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 \<in> 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 *)
apply (drule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption)
-(*Now we also know f`x : pred(A,x,r); contradiction! *)
+(*Now we also know @{term"f`x \<in> pred(A,x,r)"}: contradiction! *)
apply (simp add: well_ord_def pred_def)
done
(*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);
+ "[| well_ord(A,r); f \<in> ord_iso(pred(A,a,r), r, pred(A,c,r), r);
a:A; c:A |] ==> a=c"
apply (frule well_ord_is_trans_on)
apply (frule well_ord_is_linear)
@@ -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 \<in> 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])
@@ -425,26 +425,26 @@
done
lemma ord_iso_restrict_image:
- "[| f : ord_iso(A,r,B,s); C<=A |]
- ==> restrict(f,C) : ord_iso(C, r, f``C, s)"
-apply (simp add: ord_iso_def)
-apply (blast intro: bij_is_inj restrict_bij)
+ "[| f \<in> ord_iso(A,r,B,s); C<=A |]
+ ==> restrict(f,C) \<in> ord_iso(C, r, f``C, s)"
+apply (simp add: ord_iso_def)
+apply (blast intro: bij_is_inj restrict_bij)
done
(*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 |]
- ==> 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)
+ "[| f \<in> ord_iso(A,r,B,s); a:A |]
+ ==> restrict(f, pred(A,a,r)) \<in> 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)
done
(*Tricky; a lot of forward proof!*)
lemma well_ord_iso_preserving:
"[| well_ord(A,r); well_ord(B,s); <a,c>: 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);
+ f \<in> ord_iso(pred(A,a,r), r, pred(B,b,s), s);
+ g \<in> ord_iso(pred(A,c,r), r, pred(B,d,s), s);
a:A; c:A; b:B; d:B |] ==> <b,d>: 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")
@@ -459,7 +459,7 @@
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 |]
- ==> ~ <g`y, f`y> : s"
+ ==> ~ <g`y, f`y> \<in> s"
apply (frule well_ord_iso_subset_lemma)
apply (rule_tac f = "converse (f) " and g = g in ord_iso_trans)
apply auto
@@ -479,7 +479,7 @@
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)")
+apply (subgoal_tac "f`x \<in> B & g`x \<in> B & linear(B,s)")
apply (simp add: linear_def)
apply (blast dest: well_ord_iso_unique_lemma)
apply (blast intro: ord_iso_is_bij bij_is_fun apply_funtype
@@ -488,13 +488,13 @@
subsection{*Towards Kunen's Theorem 6.3: Linearity of the Similarity Relation*}
-lemma ord_iso_map_subset: "ord_iso_map(A,r,B,s) <= A*B"
+lemma ord_iso_map_subset: "ord_iso_map(A,r,B,s) \<subseteq> A*B"
by (unfold ord_iso_map_def, blast)
-lemma domain_ord_iso_map: "domain(ord_iso_map(A,r,B,s)) <= A"
+lemma domain_ord_iso_map: "domain(ord_iso_map(A,r,B,s)) \<subseteq> A"
by (unfold ord_iso_map_def, blast)
-lemma range_ord_iso_map: "range(ord_iso_map(A,r,B,s)) <= B"
+lemma range_ord_iso_map: "range(ord_iso_map(A,r,B,s)) \<subseteq> B"
by (unfold ord_iso_map_def, blast)
lemma converse_ord_iso_map:
@@ -510,14 +510,14 @@
done
lemma ord_iso_map_fun: "well_ord(B,s) ==> ord_iso_map(A,r,B,s)
- : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))"
+ \<in> domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))"
by (simp add: Pi_iff function_ord_iso_map
ord_iso_map_subset [THEN domain_times_range])
lemma ord_iso_map_mono_map:
"[| well_ord(A,r); well_ord(B,s) |]
==> ord_iso_map(A,r,B,s)
- : mono_map(domain(ord_iso_map(A,r,B,s)), r,
+ \<in> mono_map(domain(ord_iso_map(A,r,B,s)), r,
range(ord_iso_map(A,r,B,s)), s)"
apply (unfold mono_map_def)
apply (simp (no_asm_simp) add: ord_iso_map_fun)
@@ -530,7 +530,7 @@
lemma ord_iso_map_ord_iso:
"[| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s)
- : ord_iso(domain(ord_iso_map(A,r,B,s)), r,
+ \<in> ord_iso(domain(ord_iso_map(A,r,B,s)), r,
range(ord_iso_map(A,r,B,s)), s)"
apply (rule well_ord_mono_ord_isoI)
prefer 4
@@ -545,8 +545,8 @@
(*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)) |]
- ==> domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)"
+ a: A; a \<notin> domain(ord_iso_map(A,r,B,s)) |]
+ ==> domain(ord_iso_map(A,r,B,s)) \<subseteq> pred(A, a, r)"
apply (unfold ord_iso_map_def)
apply (safe intro!: predI)
(*Case analysis on xa vs a in r *)
@@ -570,7 +570,7 @@
lemma domain_ord_iso_map_cases:
"[| well_ord(A,r); well_ord(B,s) |]
==> domain(ord_iso_map(A,r,B,s)) = A |
- (EX x:A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))"
+ (\<exists>x\<in>A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))"
apply (frule well_ord_is_wf)
apply (unfold wf_on_def wf_def)
apply (drule_tac x = "A-domain (ord_iso_map (A,r,B,s))" in spec)
@@ -589,7 +589,7 @@
lemma range_ord_iso_map_cases:
"[| well_ord(A,r); well_ord(B,s) |]
==> range(ord_iso_map(A,r,B,s)) = B |
- (EX y:B. range(ord_iso_map(A,r,B,s)) = pred(B,y,s))"
+ (\<exists>y\<in>B. range(ord_iso_map(A,r,B,s)) = pred(B,y,s))"
apply (rule converse_ord_iso_map [THEN subst])
apply (simp add: domain_ord_iso_map_cases)
done
@@ -597,9 +597,9 @@
text{*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*}
theorem well_ord_trichotomy:
"[| well_ord(A,r); well_ord(B,s) |]
- ==> ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) |
- (EX x:A. ord_iso_map(A,r,B,s) : ord_iso(pred(A,x,r), r, B, s)) |
- (EX y:B. ord_iso_map(A,r,B,s) : ord_iso(A, r, pred(B,y,s), s))"
+ ==> ord_iso_map(A,r,B,s) \<in> ord_iso(A, r, B, s) |
+ (\<exists>x\<in>A. ord_iso_map(A,r,B,s) \<in> ord_iso(pred(A,x,r), r, B, s)) |
+ (\<exists>y\<in>B. ord_iso_map(A,r,B,s) \<in> ord_iso(A, r, pred(B,y,s), s))"
apply (frule_tac B = B in domain_ord_iso_map_cases, assumption)
apply (frule_tac B = B in range_ord_iso_map_cases, assumption)
apply (drule ord_iso_map_ord_iso, assumption)
@@ -646,7 +646,7 @@
by (unfold first_def, blast)
lemma well_ord_imp_ex1_first:
- "[| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))"
+ "[| well_ord(A,r); B<=A; B\<noteq>0 |] ==> (EX! b. first(b,B,r))"
apply (unfold well_ord_def wf_on_def wf_def first_def)
apply (elim conjE allE disjE, blast)
apply (erule bexE)
@@ -655,7 +655,7 @@
done
lemma the_first_in:
- "[| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B"
+ "[| well_ord(A,r); B<=A; B\<noteq>0 |] ==> (THE b. first(b,B,r)) \<in> B"
apply (drule well_ord_imp_ex1_first, assumption+)
apply (rule first_is_elem)
apply (erule theI)
@@ -666,7 +666,7 @@
lemma subset_vimage_vimage_iff:
"[| Preorder(r); A \<subseteq> field(r); B \<subseteq> field(r) |] ==>
- r -`` A \<subseteq> r -`` B <-> (ALL a:A. EX b:B. <a, b> : r)"
+ r -`` A \<subseteq> r -`` B <-> (\<forall>a\<in>A. \<exists>b\<in>B. <a, b> \<in> r)"
apply (auto simp: subset_def preorder_on_def refl_def vimage_def image_def)
apply blast
unfolding trans_on_def
@@ -678,12 +678,12 @@
done
lemma subset_vimage1_vimage1_iff:
- "[| Preorder(r); a : field(r); b : field(r) |] ==>
- r -`` {a} \<subseteq> r -`` {b} <-> <a, b> : r"
+ "[| Preorder(r); a \<in> field(r); b \<in> field(r) |] ==>
+ r -`` {a} \<subseteq> r -`` {b} <-> <a, b> \<in> r"
by (simp add: subset_vimage_vimage_iff)
lemma Refl_antisym_eq_Image1_Image1_iff:
- "[| refl(field(r), r); antisym(r); a : field(r); b : field(r) |] ==>
+ "[| refl(field(r), r); antisym(r); a \<in> field(r); b \<in> field(r) |] ==>
r `` {a} = r `` {b} <-> a = b"
apply rule
apply (frule equality_iffD)
@@ -694,13 +694,13 @@
done
lemma Partial_order_eq_Image1_Image1_iff:
- "[| Partial_order(r); a : field(r); b : field(r) |] ==>
+ "[| Partial_order(r); a \<in> field(r); b \<in> field(r) |] ==>
r `` {a} = r `` {b} <-> a = b"
by (simp add: partial_order_on_def preorder_on_def
Refl_antisym_eq_Image1_Image1_iff)
lemma Refl_antisym_eq_vimage1_vimage1_iff:
- "[| refl(field(r), r); antisym(r); a : field(r); b : field(r) |] ==>
+ "[| refl(field(r), r); antisym(r); a \<in> field(r); b \<in> field(r) |] ==>
r -`` {a} = r -`` {b} <-> a = b"
apply rule
apply (frule equality_iffD)
@@ -711,7 +711,7 @@
done
lemma Partial_order_eq_vimage1_vimage1_iff:
- "[| Partial_order(r); a : field(r); b : field(r) |] ==>
+ "[| Partial_order(r); a \<in> field(r); b \<in> field(r) |] ==>
r -`` {a} = r -`` {b} <-> a = b"
by (simp add: partial_order_on_def preorder_on_def
Refl_antisym_eq_vimage1_vimage1_iff)
--- a/src/ZF/OrderArith.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/OrderArith.thy Tue Mar 06 15:15:49 2012 +0000
@@ -12,22 +12,22 @@
radd :: "[i,i,i,i]=>i" where
"radd(A,r,B,s) ==
{z: (A+B) * (A+B).
- (EX x y. z = <Inl(x), Inr(y)>) |
- (EX x' x. z = <Inl(x'), Inl(x)> & <x',x>:r) |
- (EX y' y. z = <Inr(y'), Inr(y)> & <y',y>:s)}"
+ (\<exists>x y. z = <Inl(x), Inr(y)>) |
+ (\<exists>x' x. z = <Inl(x'), Inl(x)> & <x',x>:r) |
+ (\<exists>y' y. z = <Inr(y'), Inr(y)> & <y',y>: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).
- EX x' y' x y. z = <<x',y'>, <x,y>> &
+ \<exists>x' y' x y. z = <<x',y'>, <x,y>> &
(<x',x>: r | (x'=x & <y',y>: s))}"
definition
(*inverse image of a relation*)
rvimage :: "[i,i,i]=>i" where
- "rvimage(A,f,r) == {z: A*A. EX x y. z = <x,y> & <f`x,f`y>: r}"
+ "rvimage(A,f,r) == {z: A*A. \<exists>x y. z = <x,y> & <f`x,f`y>: r}"
definition
measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i" where
@@ -39,19 +39,19 @@
subsubsection{*Rewrite rules. Can be used to obtain introduction rules*}
lemma radd_Inl_Inr_iff [iff]:
- "<Inl(a), Inr(b)> : radd(A,r,B,s) <-> a:A & b:B"
+ "<Inl(a), Inr(b)> \<in> radd(A,r,B,s) <-> a:A & b:B"
by (unfold radd_def, blast)
lemma radd_Inl_iff [iff]:
- "<Inl(a'), Inl(a)> : radd(A,r,B,s) <-> a':A & a:A & <a',a>:r"
+ "<Inl(a'), Inl(a)> \<in> radd(A,r,B,s) <-> a':A & a:A & <a',a>:r"
by (unfold radd_def, blast)
lemma radd_Inr_iff [iff]:
- "<Inr(b'), Inr(b)> : radd(A,r,B,s) <-> b':B & b:B & <b',b>:s"
+ "<Inr(b'), Inr(b)> \<in> radd(A,r,B,s) <-> b':B & b:B & <b',b>:s"
by (unfold radd_def, blast)
lemma radd_Inr_Inl_iff [simp]:
- "<Inr(b), Inl(a)> : radd(A,r,B,s) <-> False"
+ "<Inr(b), Inl(a)> \<in> radd(A,r,B,s) <-> False"
by (unfold radd_def, blast)
declare radd_Inr_Inl_iff [THEN iffD1, dest!]
@@ -59,7 +59,7 @@
subsubsection{*Elimination Rule*}
lemma raddE:
- "[| <p',p> : radd(A,r,B,s);
+ "[| <p',p> \<in> 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); <x',x>: r; x':A; x:A |] ==> Q;
!!y' y. [| p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y:B |] ==> Q
@@ -68,7 +68,7 @@
subsubsection{*Type checking*}
-lemma radd_type: "radd(A,r,B,s) <= (A+B) * (A+B)"
+lemma radd_type: "radd(A,r,B,s) \<subseteq> (A+B) * (A+B)"
apply (unfold radd_def)
apply (rule Collect_subset)
done
@@ -86,10 +86,10 @@
lemma wf_on_radd: "[| wf[A](r); wf[B](s) |] ==> wf[A+B](radd(A,r,B,s))"
apply (rule wf_onI2)
-apply (subgoal_tac "ALL x:A. Inl (x) : Ba")
+apply (subgoal_tac "\<forall>x\<in>A. Inl (x) \<in> Ba")
--{*Proving the lemma, which is needed twice!*}
prefer 2
- apply (erule_tac V = "y : A + B" in thin_rl)
+ apply (erule_tac V = "y \<in> A + B" in thin_rl)
apply (rule_tac ballI)
apply (erule_tac r = r and a = x in wf_on_induct, assumption)
apply blast
@@ -116,7 +116,7 @@
lemma sum_bij:
"[| f: bij(A,C); g: bij(B,D) |]
- ==> (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) : bij(A+B, C+D)"
+ ==> (\<lambda>z\<in>A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \<in> bij(A+B, C+D)"
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)
@@ -125,8 +125,8 @@
lemma sum_ord_iso_cong:
"[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==>
- (lam 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'))"
+ (\<lambda>z\<in>A+B. case(%x. Inl(f`x), %y. Inr(g`y), z))
+ \<in> 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)
(*Do the beta-reductions now*)
@@ -134,9 +134,9 @@
done
(*Could we prove an ord_iso result? Perhaps
- ord_iso(A+B, radd(A,r,B,s), A Un B, r Un s) *)
-lemma sum_disjoint_bij: "A Int B = 0 ==>
- (lam z:A+B. case(%x. x, %y. y, z)) : bij(A+B, A Un B)"
+ ord_iso(A+B, radd(A,r,B,s), A \<union> B, r \<union> s) *)
+lemma sum_disjoint_bij: "A \<inter> B = 0 ==>
+ (\<lambda>z\<in>A+B. case(%x. x, %y. y, z)) \<in> bij(A+B, A \<union> B)"
apply (rule_tac d = "%z. if z:A then Inl (z) else Inr (z) " in lam_bijective)
apply auto
done
@@ -144,16 +144,16 @@
subsubsection{*Associativity*}
lemma sum_assoc_bij:
- "(lam z:(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z))
- : bij((A+B)+C, A+(B+C))"
+ "(\<lambda>z\<in>(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z))
+ \<in> bij((A+B)+C, A+(B+C))"
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:
- "(lam 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),
+ "(\<lambda>z\<in>(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z))
+ \<in> 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)
@@ -163,14 +163,14 @@
subsubsection{*Rewrite rule. Can be used to obtain introduction rules*}
lemma rmult_iff [iff]:
- "<<a',b'>, <a,b>> : rmult(A,r,B,s) <->
+ "<<a',b'>, <a,b>> \<in> rmult(A,r,B,s) <->
(<a',a>: r & a':A & a:A & b': B & b: B) |
(<b',b>: s & a'=a & a:A & b': B & b: B)"
by (unfold rmult_def, blast)
lemma rmultE:
- "[| <<a',b'>, <a,b>> : rmult(A,r,B,s);
+ "[| <<a',b'>, <a,b>> \<in> rmult(A,r,B,s);
[| <a',a>: r; a':A; a:A; b':B; b:B |] ==> Q;
[| <b',b>: s; a:A; a'=a; b':B; b:B |] ==> Q
|] ==> Q"
@@ -178,7 +178,7 @@
subsubsection{*Type checking*}
-lemma rmult_type: "rmult(A,r,B,s) <= (A*B) * (A*B)"
+lemma rmult_type: "rmult(A,r,B,s) \<subseteq> (A*B) * (A*B)"
by (unfold rmult_def, rule Collect_subset)
lemmas field_rmult = rmult_type [THEN field_rel_subset]
@@ -195,7 +195,7 @@
apply (rule wf_onI2)
apply (erule SigmaE)
apply (erule ssubst)
-apply (subgoal_tac "ALL b:B. <x,b>: Ba", blast)
+apply (subgoal_tac "\<forall>b\<in>B. <x,b>: Ba", blast)
apply (erule_tac a = x in wf_on_induct, assumption)
apply (rule ballI)
apply (erule_tac a = b in wf_on_induct, assumption)
@@ -221,7 +221,7 @@
lemma prod_bij:
"[| f: bij(A,C); g: bij(B,D) |]
- ==> (lam <x,y>:A*B. <f`x, g`y>) : bij(A*B, C*D)"
+ ==> (lam <x,y>:A*B. <f`x, g`y>) \<in> bij(A*B, C*D)"
apply (rule_tac d = "%<x,y>. <converse (f) `x, converse (g) `y>"
in lam_bijective)
apply (typecheck add: bij_is_inj inj_is_fun)
@@ -231,20 +231,20 @@
lemma prod_ord_iso_cong:
"[| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |]
==> (lam <x,y>:A*B. <f`x, g`y>)
- : ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))"
+ \<in> 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)
apply (simp_all add: bij_is_fun [THEN apply_type])
apply (blast intro: bij_is_inj [THEN inj_apply_equality])
done
-lemma singleton_prod_bij: "(lam z:A. <x,z>) : bij(A, {x}*A)"
+lemma singleton_prod_bij: "(\<lambda>z\<in>A. <x,z>) \<in> bij(A, {x}*A)"
by (rule_tac d = snd in lam_bijective, auto)
(*Used??*)
lemma singleton_prod_ord_iso:
"well_ord({x},xr) ==>
- (lam z:A. <x,z>) : ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"
+ (\<lambda>z\<in>A. <x,z>) \<in> ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"
apply (rule singleton_prod_bij [THEN ord_isoI])
apply (simp (no_asm_simp))
apply (blast dest: well_ord_is_wf [THEN wf_on_not_refl])
@@ -253,9 +253,9 @@
(*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 ==>
- (lam x:C*B + D. case(%x. x, %y.<a,y>, x))
- : bij(C*B + D, C*B Un {a}*D)"
+ "a\<notin>C ==>
+ (\<lambda>x\<in>C*B + D. case(%x. x, %y.<a,y>, x))
+ \<in> bij(C*B + D, C*B \<union> {a}*D)"
apply (rule subst_elem)
apply (rule id_bij [THEN sum_bij, THEN comp_bij])
apply (rule singleton_prod_bij)
@@ -268,10 +268,10 @@
lemma prod_sum_singleton_ord_iso:
"[| a:A; well_ord(A,r) |] ==>
- (lam x:pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y.<a,y>, x))
- : ord_iso(pred(A,a,r)*B + pred(B,b,s),
+ (\<lambda>x\<in>pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y.<a,y>, x))
+ \<in> 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 Un {a}*pred(B,b,s), rmult(A,r,B,s))"
+ pred(A,a,r)*B \<union> {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])
apply (auto elim!: well_ord_is_wf [THEN wf_on_asym] predE)
@@ -281,25 +281,25 @@
lemma sum_prod_distrib_bij:
"(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x))
- : bij((A+B)*C, (A*C)+(B*C))"
+ \<in> bij((A+B)*C, (A*C)+(B*C))"
by (rule_tac d = "case (%<x,y>.<Inl (x),y>, %<x,y>.<Inr (x),y>) "
in lam_bijective, auto)
lemma sum_prod_distrib_ord_iso:
"(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x))
- : ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t),
+ \<in> 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)
subsubsection{*Associativity*}
lemma prod_assoc_bij:
- "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) : bij((A*B)*C, A*(B*C))"
+ "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) \<in> bij((A*B)*C, A*(B*C))"
by (rule_tac d = "%<x, <y,z>>. <<x,y>, z>" in lam_bijective, auto)
lemma prod_assoc_ord_iso:
"(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>)
- : ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t),
+ \<in> 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,12 +307,12 @@
subsubsection{*Rewrite rule*}
-lemma rvimage_iff: "<a,b> : rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A"
+lemma rvimage_iff: "<a,b> \<in> rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A"
by (unfold rvimage_def, blast)
subsubsection{*Type checking*}
-lemma rvimage_type: "rvimage(A,f,r) <= A*A"
+lemma rvimage_type: "rvimage(A,f,r) \<subseteq> A*A"
by (unfold rvimage_def, rule Collect_subset)
lemmas field_rvimage = rvimage_type [THEN field_rel_subset]
@@ -361,7 +361,7 @@
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 "EX w. w : {w: {f`x. x:Q}. EX x. x: Q & (f`x = w) }")
+apply (subgoal_tac "\<exists>w. w \<in> {w: {f`x. x:Q}. \<exists>x. x: Q & (f`x = w) }")
apply (erule allE)
apply (erule impE)
apply assumption
@@ -373,7 +373,7 @@
@{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))"
apply (rule wf_onI2)
-apply (subgoal_tac "ALL z:A. f`z=f`y --> z: Ba")
+apply (subgoal_tac "\<forall>z\<in>A. f`z=f`y \<longrightarrow> z: Ba")
apply blast
apply (erule_tac a = "f`y" in wf_on_induct)
apply (blast intro!: apply_funtype)
@@ -396,7 +396,7 @@
done
lemma ord_iso_rvimage_eq:
- "f: ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r Int A*A"
+ "f: ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r \<inter> A*A"
by (unfold ord_iso_def rvimage_def, blast)
@@ -440,7 +440,7 @@
lemma wf_imp_subset_rvimage:
- "[|wf(r); r \<subseteq> A*A|] ==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))"
+ "[|wf(r); r \<subseteq> A*A|] ==> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
apply (rule_tac x="wftype(r)" in exI)
apply (rule_tac x="\<lambda>x\<in>A. wfrank(r,x)" in exI)
apply (simp add: Ord_wftype, clarify)
@@ -450,25 +450,25 @@
done
theorem wf_iff_subset_rvimage:
- "relation(r) ==> wf(r) <-> (\<exists>i f A. Ord(i) & r <= rvimage(A, f, Memrel(i)))"
+ "relation(r) ==> wf(r) <-> (\<exists>i f A. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i)))"
by (blast dest!: relation_field_times_field wf_imp_subset_rvimage
intro: wf_rvimage_Ord [THEN wf_subset])
subsection{*Other Results*}
-lemma wf_times: "A Int B = 0 ==> wf(A*B)"
+lemma wf_times: "A \<inter> B = 0 ==> wf(A*B)"
by (simp add: wf_def, blast)
text{*Could also be used to prove @{text wf_radd}*}
lemma wf_Un:
- "[| range(r) Int domain(s) = 0; wf(r); wf(s) |] ==> wf(r Un s)"
+ "[| range(r) \<inter> domain(s) = 0; wf(r); wf(s) |] ==> wf(r \<union> s)"
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 Int domain(s)" in spec)
+apply (drule_tac x="Z \<inter> domain(s)" in spec)
apply simp
apply (blast intro: elim: equalityE)
done
@@ -496,7 +496,7 @@
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]: "<x,y> : measure(A,f) <-> x:A & y:A & f(x)<f(y)"
+lemma measure_iff [iff]: "<x,y> \<in> measure(A,f) <-> x:A & y:A & f(x)<f(y)"
by (simp (no_asm) add: measure_def)
lemma linear_measure:
@@ -521,7 +521,7 @@
apply (blast intro: linear_measure Ordf inj)
done
-lemma measure_type: "measure(A,f) <= A*A"
+lemma measure_type: "measure(A,f) \<subseteq> A*A"
by (auto simp add: measure_def)
subsubsection{*Well-foundedness of Unions*}
@@ -549,7 +549,7 @@
lemma Pow_sum_bij:
"(\<lambda>Z \<in> Pow(A+B). <{x \<in> A. Inl(x) \<in> Z}, {y \<in> B. Inr(y) \<in> Z}>)
\<in> bij(Pow(A+B), Pow(A)*Pow(B))"
-apply (rule_tac d = "%<X,Y>. {Inl (x). x \<in> X} Un {Inr (y). y \<in> Y}"
+apply (rule_tac d = "%<X,Y>. {Inl (x). x \<in> X} \<union> {Inr (y). y \<in> Y}"
in lam_bijective)
apply force+
done
--- a/src/ZF/OrderType.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/OrderType.thy Tue Mar 06 15:15:49 2012 +0000
@@ -11,44 +11,44 @@
Ordinal arithmetic is traditionally defined in terms of order types, as it is
here. But a definition by transfinite recursion would be much simpler!*}
-definition
+definition
ordermap :: "[i,i]=>i" where
- "ordermap(A,r) == lam x:A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
+ "ordermap(A,r) == \<lambda>x\<in>A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
-definition
+definition
ordertype :: "[i,i]=>i" where
"ordertype(A,r) == ordermap(A,r)``A"
-definition
+definition
(*alternative definition of ordinal numbers*)
Ord_alt :: "i => o" where
- "Ord_alt(X) == well_ord(X, Memrel(X)) & (ALL u:X. u=pred(X, u, Memrel(X)))"
+ "Ord_alt(X) == well_ord(X, Memrel(X)) & (\<forall>u\<in>X. u=pred(X, u, Memrel(X)))"
-definition
+definition
(*coercion to ordinal: if not, just 0*)
ordify :: "i=>i" where
"ordify(x) == if Ord(x) then x else 0"
-definition
+definition
(*ordinal multiplication*)
omult :: "[i,i]=>i" (infixl "**" 70) where
"i ** j == ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))"
-definition
+definition
(*ordinal addition*)
raw_oadd :: "[i,i]=>i" where
"raw_oadd(i,j) == ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))"
-definition
+definition
oadd :: "[i,i]=>i" (infixl "++" 65) where
"i ++ j == raw_oadd(ordify(i),ordify(j))"
-definition
+definition
(*ordinal subtraction*)
odiff :: "[i,i]=>i" (infixl "--" 65) where
"i -- j == ordertype(i-j, Memrel(i))"
-
+
notation (xsymbols)
omult (infixl "\<times>\<times>" 70)
@@ -58,7 +58,7 @@
subsection{*Proofs needing the combination of Ordinal.thy and Order.thy*}
-lemma le_well_ord_Memrel: "j le i ==> well_ord(j, Memrel(i))"
+lemma le_well_ord_Memrel: "j \<le> i ==> well_ord(j, Memrel(i))"
apply (rule well_ordI)
apply (rule wf_Memrel [THEN wf_imp_wf_on])
apply (simp add: ltD lt_Ord linear_def
@@ -72,22 +72,22 @@
(*Kunen's Theorem 7.3 (i), page 16; see also Ordinal/Ord_in_Ord
The smaller ordinal is an initial segment of the larger *)
-lemma lt_pred_Memrel:
+lemma lt_pred_Memrel:
"j<i ==> pred(i, j, Memrel(i)) = j"
apply (unfold pred_def lt_def)
apply (simp (no_asm_simp))
apply (blast intro: Ord_trans)
done
-lemma pred_Memrel:
- "x:A ==> pred(A, x, Memrel(A)) = A Int x"
+lemma pred_Memrel:
+ "x:A ==> pred(A, x, Memrel(A)) = A \<inter> x"
by (unfold pred_def Memrel_def, blast)
lemma Ord_iso_implies_eq_lemma:
"[| j<i; f: 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)
+apply (rule well_ord_Memrel [THEN well_ord_iso_predE, of i f j], auto)
apply (unfold ord_iso_def)
(*Combining the two simplifications causes looping*)
apply (simp (no_asm_simp))
@@ -96,7 +96,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)+
@@ -105,8 +105,8 @@
subsection{*Ordermap and ordertype*}
-lemma ordermap_type:
- "ordermap(A,r) : A -> ordertype(A,r)"
+lemma ordermap_type:
+ "ordermap(A,r) \<in> A -> ordertype(A,r)"
apply (unfold ordermap_def ordertype_def)
apply (rule lam_type)
apply (rule lamI [THEN imageI], assumption+)
@@ -115,7 +115,7 @@
subsubsection{*Unfolding of ordermap *}
(*Useful for cardinality reasoning; see CardinalArith.ML*)
-lemma ordermap_eq_image:
+lemma ordermap_eq_image:
"[| wf[A](r); x:A |]
==> ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)"
apply (unfold ordermap_def pred_def)
@@ -127,23 +127,23 @@
(*Useful for rewriting PROVIDED pred is not unfolded until later!*)
lemma ordermap_pred_unfold:
"[| wf[A](r); x:A |]
- ==> ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}"
+ ==> ordermap(A,r) ` x = {ordermap(A,r)`y . y \<in> pred(A,x,r)}"
by (simp add: ordermap_eq_image pred_subset ordermap_type [THEN image_fun])
(*pred-unfolded version. NOT suitable for rewriting -- loops!*)
-lemmas ordermap_unfold = ordermap_pred_unfold [simplified pred_def]
+lemmas ordermap_unfold = ordermap_pred_unfold [simplified pred_def]
-(*The theorem above is
+(*The theorem above is
-[| wf[A](r); x : A |]
-==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y: A . <y,x> : r}}
+[| wf[A](r); x \<in> A |]
+==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y: A . <y,x> \<in> 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 . <y,x> : r}},
+ ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y: A . <y,x> \<in> r}},
might eliminate the need for r to be transitive.
*)
@@ -151,7 +151,7 @@
subsubsection{*Showing that ordermap, ordertype yield ordinals *}
-lemma Ord_ordermap:
+lemma Ord_ordermap:
"[| 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+)
@@ -159,10 +159,10 @@
apply (rule OrdI [OF _ Ord_is_Transset])
apply (unfold pred_def Transset_def)
apply (blast intro: trans_onD
- dest!: ordermap_unfold [THEN equalityD1])+
+ dest!: ordermap_unfold [THEN equalityD1])+
done
-lemma Ord_ordertype:
+lemma Ord_ordertype:
"well_ord(A,r) ==> Ord(ordertype(A,r))"
apply (unfold ordertype_def)
apply (subst image_fun [OF ordermap_type subset_refl])
@@ -178,38 +178,38 @@
lemma ordermap_mono:
"[| <w,x>: r; wf[A](r); w: A; x: A |]
- ==> ordermap(A,r)`w : ordermap(A,r)`x"
+ ==> ordermap(A,r)`w \<in> 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 |]
+lemma converse_ordermap_mono:
+ "[| ordermap(A,r)`w \<in> ordermap(A,r)`x; well_ord(A,r); w: A; x: A |]
==> <w,x>: r"
apply (unfold well_ord_def tot_ord_def, safe)
-apply (erule_tac x=w and y=x in linearE, assumption+)
+apply (erule_tac x=w and y=x in linearE, assumption+)
apply (blast elim!: mem_not_refl [THEN notE])
-apply (blast dest: ordermap_mono intro: mem_asym)
+apply (blast dest: ordermap_mono intro: mem_asym)
done
-lemmas ordermap_surj =
+lemmas ordermap_surj =
ordermap_type [THEN surj_image, unfolded ordertype_def [symmetric]]
-lemma ordermap_bij:
- "well_ord(A,r) ==> ordermap(A,r) : bij(A, ordertype(A,r))"
+lemma ordermap_bij:
+ "well_ord(A,r) ==> ordermap(A,r) \<in> bij(A, ordertype(A,r))"
apply (unfold well_ord_def tot_ord_def bij_def inj_def)
-apply (force intro!: ordermap_type ordermap_surj
- elim: linearE dest: ordermap_mono
+apply (force intro!: ordermap_type ordermap_surj
+ elim: linearE dest: ordermap_mono
simp add: mem_not_refl)
done
subsubsection{*Isomorphisms involving ordertype *}
-lemma ordertype_ord_iso:
+lemma ordertype_ord_iso:
"well_ord(A,r)
- ==> ordermap(A,r) : ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))"
+ ==> ordermap(A,r) \<in> ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))"
apply (unfold ord_iso_def)
-apply (safe elim!: well_ord_is_wf
+apply (safe elim!: well_ord_is_wf
intro!: ordermap_type [THEN apply_type] ordermap_mono ordermap_bij)
apply (blast dest!: converse_ordermap_mono)
done
@@ -223,8 +223,8 @@
done
lemma ordertype_eq_imp_ord_iso:
- "[| ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s) |]
- ==> EX f. f: ord_iso(A,r,B,s)"
+ "[| ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s) |]
+ ==> \<exists>f. f: ord_iso(A,r,B,s)"
apply (rule exI)
apply (rule ordertype_ord_iso [THEN ord_iso_trans], assumption)
apply (erule ssubst)
@@ -234,7 +234,7 @@
subsubsection{*Basic equalities for ordertype *}
(*Ordertype of Memrel*)
-lemma le_ordertype_Memrel: "j le i ==> ordertype(j,Memrel(i)) = j"
+lemma le_ordertype_Memrel: "j \<le> i ==> ordertype(j,Memrel(i)) = j"
apply (rule Ord_iso_implies_eq [symmetric])
apply (erule ltE, assumption)
apply (blast intro: le_well_ord_Memrel Ord_ordertype)
@@ -277,16 +277,16 @@
apply (fast elim!: trans_onD)
done
-lemma ordertype_unfold:
- "ordertype(A,r) = {ordermap(A,r)`y . y : A}"
+lemma ordertype_unfold:
+ "ordertype(A,r) = {ordermap(A,r)`y . y \<in> A}"
apply (unfold ordertype_def)
apply (rule image_fun [OF ordermap_type subset_refl])
done
text{*Theorems by Krzysztof Grabczewski; proofs simplified by lcp *}
-lemma ordertype_pred_subset: "[| well_ord(A,r); x:A |] ==>
- ordertype(pred(A,x,r),r) <= ordertype(A,r)"
+lemma ordertype_pred_subset: "[| well_ord(A,r); x:A |] ==>
+ ordertype(pred(A,x,r),r) \<subseteq> ordertype(A,r)"
apply (simp add: ordertype_unfold well_ord_subset [OF _ pred_subset])
apply (fast intro: ordermap_pred_eq_ordermap elim: predE)
done
@@ -321,13 +321,13 @@
apply (unfold Ord_alt_def)
apply (rule conjI)
apply (erule well_ord_Memrel)
-apply (unfold Ord_def Transset_def pred_def Memrel_def, blast)
+apply (unfold Ord_def Transset_def pred_def Memrel_def, blast)
done
(*proof by lcp*)
-lemma Ord_alt_is_Ord:
+lemma Ord_alt_is_Ord:
"Ord_alt(i) ==> Ord(i)"
-apply (unfold Ord_alt_def Ord_def Transset_def well_ord_def
+apply (unfold Ord_alt_def Ord_def Transset_def well_ord_def
tot_ord_def part_ord_def trans_on_def)
apply (simp add: pred_Memrel)
apply (blast elim!: equalityE)
@@ -340,7 +340,7 @@
text{*Addition with 0 *}
-lemma bij_sum_0: "(lam z:A+0. case(%x. x, %y. y, z)) : bij(A+0, A)"
+lemma bij_sum_0: "(\<lambda>z\<in>A+0. case(%x. x, %y. y, z)) \<in> bij(A+0, A)"
apply (rule_tac d = Inl in lam_bijective, safe)
apply (simp_all (no_asm_simp))
done
@@ -352,7 +352,7 @@
apply force
done
-lemma bij_0_sum: "(lam z:0+A. case(%x. x, %y. y, z)) : bij(0+A, A)"
+lemma bij_0_sum: "(\<lambda>z\<in>0+A. case(%x. x, %y. y, z)) \<in> bij(0+A, A)"
apply (rule_tac d = Inr in lam_bijective, safe)
apply (simp_all (no_asm_simp))
done
@@ -367,9 +367,9 @@
text{*Initial segments of radd. Statements by Grabczewski *}
(*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *)
-lemma pred_Inl_bij:
- "a:A ==> (lam x:pred(A,a,r). Inl(x))
- : bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))"
+lemma pred_Inl_bij:
+ "a:A ==> (\<lambda>x\<in>pred(A,a,r). Inl(x))
+ \<in> 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)
apply auto
@@ -377,24 +377,24 @@
lemma ordertype_pred_Inl_eq:
"[| 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+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])
apply (simp_all add: well_ord_subset [OF _ pred_subset])
apply (simp add: pred_def)
done
-lemma pred_Inr_bij:
- "b:B ==>
- id(A+pred(B,b,s))
- : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"
+lemma pred_Inr_bij:
+ "b:B ==>
+ id(A+pred(B,b,s))
+ \<in> bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"
apply (unfold pred_def id_def)
-apply (rule_tac d = "%z. z" in lam_bijective, auto)
+apply (rule_tac d = "%z. z" in lam_bijective, auto)
done
lemma ordertype_pred_Inr_eq:
"[| 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(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])
prefer 2 apply (force simp add: pred_def id_def, assumption)
@@ -441,7 +441,7 @@
lemma oadd_eq_if_raw_oadd:
- "i++j = (if Ord(i) then (if Ord(j) then raw_oadd(i,j) else i)
+ "i++j = (if Ord(i) then (if Ord(j) then raw_oadd(i,j) else i)
else (if Ord(j) then j else 0))"
by (simp add: oadd_def ordify_def raw_oadd_0_left raw_oadd_0)
@@ -462,15 +462,15 @@
apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel)
done
-(*Thus also we obtain the rule i++j = k ==> i le k *)
-lemma oadd_le_self: "Ord(i) ==> i le i++j"
+(*Thus also we obtain the rule @{term"i++j = k ==> i \<le> k"} *)
+lemma oadd_le_self: "Ord(i) ==> i \<le> i++j"
apply (rule all_lt_imp_le)
-apply (auto simp add: Ord_oadd lt_oadd1)
+apply (auto simp add: Ord_oadd lt_oadd1)
done
text{*Various other results *}
-lemma id_ord_iso_Memrel: "A<=B ==> id(A) : ord_iso(A, Memrel(A), A, Memrel(B))"
+lemma id_ord_iso_Memrel: "A<=B ==> id(A) \<in> ord_iso(A, Memrel(A), A, Memrel(B))"
apply (rule id_bij [THEN ord_isoI])
apply (simp (no_asm_simp))
apply blast
@@ -478,32 +478,32 @@
lemma subset_ord_iso_Memrel:
"[| f: ord_iso(A,Memrel(B),C,r); A<=B |] ==> f: ord_iso(A,Memrel(A),C,r)"
-apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN fun_is_rel])
-apply (frule ord_iso_trans [OF id_ord_iso_Memrel], assumption)
-apply (simp add: right_comp_id)
+apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN fun_is_rel])
+apply (frule ord_iso_trans [OF id_ord_iso_Memrel], assumption)
+apply (simp add: right_comp_id)
done
lemma restrict_ord_iso:
- "[| f \<in> ord_iso(i, Memrel(i), Order.pred(A,a,r), r); a \<in> A; j < i;
+ "[| f \<in> ord_iso(i, Memrel(i), Order.pred(A,a,r), r); a \<in> A; j < i;
trans[A](r) |]
==> restrict(f,j) \<in> ord_iso(j, Memrel(j), Order.pred(A,f`j,r), r)"
-apply (frule ltD)
-apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption)
-apply (frule ord_iso_restrict_pred, assumption)
+apply (frule ltD)
+apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption)
+apply (frule ord_iso_restrict_pred, assumption)
apply (simp add: pred_iff trans_pred_pred_eq lt_pred_Memrel)
-apply (blast intro!: subset_ord_iso_Memrel le_imp_subset [OF leI])
+apply (blast intro!: subset_ord_iso_Memrel le_imp_subset [OF leI])
done
lemma restrict_ord_iso2:
- "[| f \<in> ord_iso(Order.pred(A,a,r), r, i, Memrel(i)); a \<in> A;
+ "[| f \<in> ord_iso(Order.pred(A,a,r), r, i, Memrel(i)); a \<in> A;
j < i; trans[A](r) |]
- ==> converse(restrict(converse(f), j))
+ ==> converse(restrict(converse(f), j))
\<in> ord_iso(Order.pred(A, converse(f)`j, r), r, j, Memrel(j))"
by (blast intro: restrict_ord_iso ord_iso_sym ltI)
lemma ordertype_sum_Memrel:
"[| well_ord(A,r); k<j |]
- ==> ordertype(A+k, radd(A, r, k, Memrel(j))) =
+ ==> ordertype(A+k, radd(A, r, k, Memrel(j))) =
ordertype(A+k, radd(A, r, k, Memrel(k)))"
apply (erule ltE)
apply (rule ord_iso_refl [THEN sum_ord_iso_cong, THEN ordertype_eq])
@@ -528,7 +528,7 @@
prefer 2
apply (frule_tac i = i and j = j in oadd_le_self)
apply (simp (asm_lr) add: oadd_def ordify_def lt_Ord not_lt_iff_le [THEN iff_sym])
-apply (rule Ord_linear_lt, auto)
+apply (rule Ord_linear_lt, auto)
apply (simp_all add: raw_oadd_eq_oadd)
apply (blast dest: oadd_lt_mono2 elim: lt_irrefl lt_asym)+
done
@@ -539,18 +539,18 @@
lemma oadd_inject: "[| i++j = i++k; Ord(j); Ord(k) |] ==> j=k"
apply (simp add: oadd_eq_if_raw_oadd split add: split_if_asm)
apply (simp add: raw_oadd_eq_oadd)
-apply (rule Ord_linear_lt, auto)
+apply (rule Ord_linear_lt, auto)
apply (force dest: oadd_lt_mono2 [of concl: i] simp add: lt_not_refl)+
done
-lemma lt_oadd_disj: "k < i++j ==> k<i | (EX l:j. k = i++l )"
+lemma lt_oadd_disj: "k < i++j ==> k<i | (\<exists>l\<in>j. k = i++l )"
apply (simp add: Ord_in_Ord' [of _ j] oadd_eq_if_raw_oadd
split add: split_if_asm)
prefer 2
apply (simp add: Ord_in_Ord' [of _ j] lt_def)
apply (simp add: ordertype_pred_unfold well_ord_radd well_ord_Memrel raw_oadd_def)
apply (erule ltD [THEN RepFunE])
-apply (force simp add: ordertype_pred_Inl_eq well_ord_Memrel ltI
+apply (force simp add: ordertype_pred_Inl_eq well_ord_Memrel ltI
lt_pred_Memrel le_ordertype_Memrel leI
ordertype_pred_Inr_eq ordertype_sum_Memrel)
done
@@ -562,7 +562,7 @@
apply (simp add: oadd_eq_if_raw_oadd Ord_raw_oadd raw_oadd_0 raw_oadd_0_left, clarify)
apply (simp add: raw_oadd_def)
apply (rule ordertype_eq [THEN trans])
-apply (rule sum_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym]
+apply (rule sum_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym]
ord_iso_refl])
apply (simp_all add: Ord_ordertype well_ord_radd well_ord_Memrel)
apply (rule sum_assoc_ord_iso [THEN ordertype_eq, THEN trans])
@@ -571,11 +571,11 @@
apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel)+
done
-lemma oadd_unfold: "[| Ord(i); Ord(j) |] ==> i++j = i Un (\<Union>k\<in>j. {i++k})"
+lemma oadd_unfold: "[| Ord(i); Ord(j) |] ==> i++j = i \<union> (\<Union>k\<in>j. {i++k})"
apply (rule subsetI [THEN equalityI])
apply (erule ltI [THEN lt_oadd_disj, THEN disjE])
-apply (blast intro: Ord_oadd)
-apply (blast elim!: ltE, blast)
+apply (blast intro: Ord_oadd)
+apply (blast elim!: ltE, blast)
apply (force intro: lt_oadd1 oadd_lt_mono2 simp add: Ord_mem_iff_lt)
done
@@ -597,13 +597,13 @@
lemma oadd_UN:
"[| !!x. x:A ==> Ord(j(x)); a:A |]
==> i ++ (\<Union>x\<in>A. j(x)) = (\<Union>x\<in>A. i++j(x))"
-by (blast intro: ltI Ord_UN Ord_oadd lt_oadd1 [THEN ltD]
- oadd_lt_mono2 [THEN ltD]
+by (blast intro: ltI Ord_UN Ord_oadd lt_oadd1 [THEN ltD]
+ oadd_lt_mono2 [THEN ltD]
elim!: ltE dest!: ltI [THEN lt_oadd_disj])
lemma oadd_Limit: "Limit(j) ==> i++j = (\<Union>k\<in>j. i++k)"
apply (frule Limit_has_0 [THEN ltD])
-apply (simp add: Limit_is_Ord [THEN Ord_in_Ord] oadd_UN [symmetric]
+apply (simp add: Limit_is_Ord [THEN Ord_in_Ord] oadd_UN [symmetric]
Union_eq_UN [symmetric] Limit_Union_eq)
done
@@ -626,12 +626,12 @@
Limit_is_Ord [of j, THEN Ord_in_Ord], auto)
apply (rule_tac x="succ(y)" in bexI)
apply (simp add: ltI Limit_is_Ord [of j, THEN Ord_in_Ord])
-apply (simp add: Limit_def lt_def)
+apply (simp add: Limit_def lt_def)
done
text{*Order/monotonicity properties of ordinal addition *}
-lemma oadd_le_self2: "Ord(i) ==> i le j++i"
+lemma oadd_le_self2: "Ord(i) ==> i \<le> j++i"
apply (erule_tac i = i in trans_induct3)
apply (simp (no_asm_simp) add: Ord_0_le)
apply (simp (no_asm_simp) add: oadd_succ succ_leI)
@@ -643,7 +643,7 @@
apply (simp add: Union_eq_UN [symmetric] Limit_Union_eq le_refl Limit_is_Ord)
done
-lemma oadd_le_mono1: "k le j ==> k++i le j++i"
+lemma oadd_le_mono1: "k \<le> j ==> k++i \<le> j++i"
apply (frule lt_Ord)
apply (frule le_Ord2)
apply (simp add: oadd_eq_if_raw_oadd, clarify)
@@ -655,31 +655,31 @@
apply (rule le_implies_UN_le_UN, blast)
done
-lemma oadd_lt_mono: "[| i' le i; j'<j |] ==> i'++j' < i++j"
+lemma oadd_lt_mono: "[| i' \<le> i; j'<j |] ==> i'++j' < i++j"
by (blast intro: lt_trans1 oadd_le_mono1 oadd_lt_mono2 Ord_succD elim: ltE)
-lemma oadd_le_mono: "[| i' le i; j' le j |] ==> i'++j' le i++j"
+lemma oadd_le_mono: "[| i' \<le> i; j' \<le> j |] ==> i'++j' \<le> i++j"
by (simp del: oadd_succ add: oadd_succ [symmetric] le_Ord2 oadd_lt_mono)
-lemma oadd_le_iff2: "[| Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k"
+lemma oadd_le_iff2: "[| Ord(j); Ord(k) |] ==> i++j \<le> i++k <-> j \<le> k"
by (simp del: oadd_succ add: oadd_lt_iff2 oadd_succ [symmetric] Ord_succ)
lemma oadd_lt_self: "[| Ord(i); 0<j |] ==> i < i++j"
-apply (rule lt_trans2)
-apply (erule le_refl)
-apply (simp only: lt_Ord2 oadd_1 [of i, symmetric])
+apply (rule lt_trans2)
+apply (erule le_refl)
+apply (simp only: lt_Ord2 oadd_1 [of i, symmetric])
apply (blast intro: succ_leI oadd_le_mono)
done
text{*Every ordinal is exceeded by some limit ordinal.*}
lemma Ord_imp_greater_Limit: "Ord(i) ==> \<exists>k. i<k & Limit(k)"
-apply (rule_tac x="i ++ nat" in exI)
+apply (rule_tac x="i ++ nat" in exI)
apply (blast intro: oadd_LimitI oadd_lt_self Limit_nat [THEN Limit_has_0])
done
lemma Ord2_imp_greater_Limit: "[|Ord(i); Ord(j)|] ==> \<exists>k. i<k & j<k & Limit(k)"
-apply (insert Ord_Un [of i j, THEN Ord_imp_greater_Limit])
-apply (simp add: Un_least_lt_iff)
+apply (insert Ord_Un [of i j, THEN Ord_imp_greater_Limit])
+apply (simp add: Un_least_lt_iff)
done
@@ -689,7 +689,7 @@
It's probably simpler to define the difference recursively!*}
lemma bij_sum_Diff:
- "A<=B ==> (lam y:B. if(y:A, Inl(y), Inr(y))) : bij(B, A+(B-A))"
+ "A<=B ==> (\<lambda>y\<in>B. if(y:A, Inl(y), Inr(y))) \<in> 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)
@@ -698,8 +698,8 @@
done
lemma ordertype_sum_Diff:
- "i le j ==>
- ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) =
+ "i \<le> j ==>
+ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) =
ordertype(j, Memrel(j))"
apply (safe dest!: le_subset_iff [THEN iffD1])
apply (rule bij_sum_Diff [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq])
@@ -711,15 +711,15 @@
apply (blast intro: lt_trans2 lt_trans)
done
-lemma Ord_odiff [simp,TC]:
+lemma Ord_odiff [simp,TC]:
"[| Ord(i); Ord(j) |] ==> Ord(i--j)"
apply (unfold odiff_def)
apply (blast intro: Ord_ordertype Diff_subset well_ord_subset well_ord_Memrel)
done
-lemma raw_oadd_ordertype_Diff:
- "i le j
+lemma raw_oadd_ordertype_Diff:
+ "i \<le> j
==> raw_oadd(i,j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))"
apply (simp add: raw_oadd_def odiff_def)
apply (safe dest!: le_subset_iff [THEN iffD1])
@@ -729,7 +729,7 @@
apply (blast intro: well_ord_radd Diff_subset well_ord_subset well_ord_Memrel)+
done
-lemma oadd_odiff_inverse: "i le j ==> i ++ (j--i) = j"
+lemma oadd_odiff_inverse: "i \<le> j ==> i ++ (j--i) = j"
by (simp add: lt_Ord le_Ord2 oadd_def ordify_def raw_oadd_ordertype_Diff
ordertype_sum_Diff ordertype_Memrel lt_Ord2 [THEN Ord_succD])
@@ -741,7 +741,7 @@
apply (blast intro: Ord_ordertype Ord_oadd Ord_odiff)+
done
-lemma odiff_lt_mono2: "[| i<j; k le i |] ==> i--k < j--k"
+lemma odiff_lt_mono2: "[| i<j; k \<le> i |] ==> i--k < j--k"
apply (rule_tac i = k in oadd_lt_cancel2)
apply (simp add: oadd_odiff_inverse)
apply (subst oadd_odiff_inverse)
@@ -752,7 +752,7 @@
subsection{*Ordinal Multiplication*}
-lemma Ord_omult [simp,TC]:
+lemma Ord_omult [simp,TC]:
"[| Ord(i); Ord(j) |] ==> Ord(i**j)"
apply (unfold omult_def)
apply (blast intro: Ord_ordertype well_ord_rmult well_ord_Memrel)
@@ -760,67 +760,67 @@
subsubsection{*A useful unfolding law *}
-lemma pred_Pair_eq:
- "[| a:A; b:B |] ==> pred(A*B, <a,b>, rmult(A,r,B,s)) =
- pred(A,a,r)*B Un ({a} * pred(B,b,s))"
+lemma pred_Pair_eq:
+ "[| a:A; b:B |] ==> pred(A*B, <a,b>, rmult(A,r,B,s)) =
+ pred(A,a,r)*B \<union> ({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) |] ==>
- ordertype(pred(A*B, <a,b>, rmult(A,r,B,s)), rmult(A,r,B,s)) =
- ordertype(pred(A,a,r)*B + pred(B,b,s),
+ "[| a:A; b:B; well_ord(A,r); well_ord(B,s) |] ==>
+ ordertype(pred(A*B, <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))"
apply (simp (no_asm_simp) add: pred_Pair_eq)
apply (rule ordertype_eq [symmetric])
apply (rule prod_sum_singleton_ord_iso)
apply (simp_all add: pred_subset well_ord_rmult [THEN well_ord_subset])
-apply (blast intro: pred_subset well_ord_rmult [THEN well_ord_subset]
+apply (blast intro: pred_subset well_ord_rmult [THEN well_ord_subset]
elim!: predE)
done
-lemma ordertype_pred_Pair_lemma:
+lemma ordertype_pred_Pair_lemma:
"[| i'<i; j'<j |]
- ==> ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))),
- rmult(i,Memrel(i),j,Memrel(j))) =
+ ==> ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))),
+ rmult(i,Memrel(i),j,Memrel(j))) =
raw_oadd (j**i', j')"
apply (unfold raw_oadd_def omult_def)
-apply (simp add: ordertype_pred_Pair_eq lt_pred_Memrel ltD lt_Ord2
+apply (simp add: ordertype_pred_Pair_eq lt_pred_Memrel ltD lt_Ord2
well_ord_Memrel)
apply (rule trans)
- apply (rule_tac [2] ordertype_ord_iso
+ apply (rule_tac [2] ordertype_ord_iso
[THEN sum_ord_iso_cong, THEN ordertype_eq])
apply (rule_tac [3] ord_iso_refl)
apply (rule id_bij [THEN ord_isoI, THEN ordertype_eq])
apply (elim SigmaE sumE ltE ssubst)
apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel
- Ord_ordertype lt_Ord lt_Ord2)
+ Ord_ordertype lt_Ord lt_Ord2)
apply (blast intro: Ord_trans)+
done
-lemma lt_omult:
+lemma lt_omult:
"[| Ord(i); Ord(j); k<j**i |]
- ==> EX j' i'. k = j**i' ++ j' & j'<j & i'<i"
+ ==> \<exists>j' i'. k = j**i' ++ j' & j'<j & i'<i"
apply (unfold omult_def)
apply (simp add: ordertype_pred_unfold well_ord_rmult well_ord_Memrel)
apply (safe elim!: ltE)
-apply (simp add: ordertype_pred_Pair_lemma ltI raw_oadd_eq_oadd
+apply (simp add: ordertype_pred_Pair_lemma ltI raw_oadd_eq_oadd
omult_def [symmetric] Ord_in_Ord' [of _ i] Ord_in_Ord' [of _ j])
apply (blast intro: ltI)
done
-lemma omult_oadd_lt:
+lemma omult_oadd_lt:
"[| j'<j; i'<i |] ==> j**i' ++ j' < j**i"
apply (unfold omult_def)
apply (rule ltI)
prefer 2
apply (simp add: Ord_ordertype well_ord_rmult well_ord_Memrel lt_Ord2)
apply (simp add: ordertype_pred_unfold well_ord_rmult well_ord_Memrel lt_Ord2)
-apply (rule bexI [of _ i'])
-apply (rule bexI [of _ j'])
+apply (rule bexI [of _ i'])
+apply (rule bexI [of _ j'])
apply (simp add: ordertype_pred_Pair_lemma ltI omult_def [symmetric])
apply (simp add: lt_Ord lt_Ord2 raw_oadd_eq_oadd)
-apply (simp_all add: lt_def)
+apply (simp_all add: lt_def)
done
lemma omult_unfold:
@@ -828,7 +828,7 @@
apply (rule subsetI [THEN equalityI])
apply (rule lt_omult [THEN exE])
apply (erule_tac [3] ltI)
-apply (simp_all add: Ord_omult)
+apply (simp_all add: Ord_omult)
apply (blast elim!: ltE)
apply (blast intro: omult_oadd_lt [THEN ltD] ltI)
done
@@ -851,7 +851,7 @@
lemma omult_1 [simp]: "Ord(i) ==> i**1 = i"
apply (unfold omult_def)
-apply (rule_tac s1="Memrel(i)"
+apply (rule_tac s1="Memrel(i)"
in ord_isoI [THEN ordertype_eq, THEN trans])
apply (rule_tac c = snd and d = "%z.<0,z>" in lam_bijective)
apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel)
@@ -859,7 +859,7 @@
lemma omult_1_left [simp]: "Ord(i) ==> 1**i = i"
apply (unfold omult_def)
-apply (rule_tac s1="Memrel(i)"
+apply (rule_tac s1="Memrel(i)"
in ord_isoI [THEN ordertype_eq, THEN trans])
apply (rule_tac c = fst and d = "%z.<z,0>" in lam_bijective)
apply (auto elim!: fst_type well_ord_Memrel ordertype_Memrel)
@@ -872,14 +872,14 @@
apply (simp add: oadd_eq_if_raw_oadd)
apply (simp add: omult_def raw_oadd_def)
apply (rule ordertype_eq [THEN trans])
-apply (rule prod_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym]
+apply (rule prod_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym]
ord_iso_refl])
-apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel
+apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel
Ord_ordertype)
apply (rule sum_prod_distrib_ord_iso [THEN ordertype_eq, THEN trans])
apply (rule_tac [2] ordertype_eq)
apply (rule_tac [2] sum_ord_iso_cong [OF ordertype_ord_iso ordertype_ord_iso])
-apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel
+apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel
Ord_ordertype)
done
@@ -888,14 +888,14 @@
text{*Associative law *}
-lemma omult_assoc:
+lemma omult_assoc:
"[| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)"
apply (unfold omult_def)
apply (rule ordertype_eq [THEN trans])
-apply (rule prod_ord_iso_cong [OF ord_iso_refl
+apply (rule prod_ord_iso_cong [OF ord_iso_refl
ordertype_ord_iso [THEN ord_iso_sym]])
apply (blast intro: well_ord_rmult well_ord_Memrel)+
-apply (rule prod_assoc_ord_iso
+apply (rule prod_assoc_ord_iso
[THEN ord_iso_sym, THEN ordertype_eq, THEN trans])
apply (rule_tac [2] ordertype_eq)
apply (rule_tac [2] prod_ord_iso_cong [OF ordertype_ord_iso ord_iso_refl])
@@ -905,13 +905,13 @@
text{*Ordinal multiplication with limit ordinals *}
-lemma omult_UN:
+lemma omult_UN:
"[| Ord(i); !!x. x:A ==> Ord(j(x)) |]
==> i ** (\<Union>x\<in>A. j(x)) = (\<Union>x\<in>A. i**j(x))"
by (simp (no_asm_simp) add: Ord_UN omult_unfold, blast)
lemma omult_Limit: "[| Ord(i); Limit(j) |] ==> i**j = (\<Union>k\<in>j. i**k)"
-by (simp add: Limit_is_Ord [THEN Ord_in_Ord] omult_UN [symmetric]
+by (simp add: Limit_is_Ord [THEN Ord_in_Ord] omult_UN [symmetric]
Union_eq_UN [symmetric] Limit_Union_eq)
@@ -923,10 +923,10 @@
apply (force simp add: omult_unfold)
done
-lemma omult_le_self: "[| Ord(i); 0<j |] ==> i le i**j"
+lemma omult_le_self: "[| Ord(i); 0<j |] ==> i \<le> i**j"
by (blast intro: all_lt_imp_le Ord_omult lt_omult1 lt_Ord2)
-lemma omult_le_mono1: "[| k le j; Ord(i) |] ==> k**i le j**i"
+lemma omult_le_mono1: "[| k \<le> j; Ord(i) |] ==> k**i \<le> j**i"
apply (frule lt_Ord)
apply (frule le_Ord2)
apply (erule trans_induct3)
@@ -943,20 +943,20 @@
apply (force simp add: Ord_omult)
done
-lemma omult_le_mono2: "[| k le j; Ord(i) |] ==> i**k le i**j"
+lemma omult_le_mono2: "[| k \<le> j; Ord(i) |] ==> i**k \<le> i**j"
apply (rule subset_imp_le)
apply (safe elim!: ltE dest!: Ord_succD intro!: Ord_omult)
apply (simp add: omult_unfold)
-apply (blast intro: Ord_trans)
+apply (blast intro: Ord_trans)
done
-lemma omult_le_mono: "[| i' le i; j' le j |] ==> i'**j' le i**j"
+lemma omult_le_mono: "[| i' \<le> i; j' \<le> j |] ==> i'**j' \<le> i**j"
by (blast intro: le_trans omult_le_mono1 omult_le_mono2 Ord_succD elim: ltE)
-lemma omult_lt_mono: "[| i' le i; j'<j; 0<i |] ==> i'**j' < i**j"
+lemma omult_lt_mono: "[| i' \<le> i; j'<j; 0<i |] ==> i'**j' < i**j"
by (blast intro: lt_trans1 omult_le_mono1 omult_lt_mono2 Ord_succD elim: ltE)
-lemma omult_le_self2: "[| Ord(i); 0<j |] ==> i le j**i"
+lemma omult_le_self2: "[| Ord(i); 0<j |] ==> i \<le> j**i"
apply (frule lt_Ord2)
apply (erule_tac i = i in trans_induct3)
apply (simp (no_asm_simp))
@@ -977,32 +977,32 @@
lemma omult_inject: "[| i**j = i**k; 0<i; Ord(j); Ord(k) |] ==> j=k"
apply (rule Ord_linear_lt)
prefer 4 apply assumption
-apply auto
+apply auto
apply (force dest: omult_lt_mono2 simp add: lt_not_refl)+
done
subsection{*The Relation @{term Lt}*}
lemma wf_Lt: "wf(Lt)"
-apply (rule wf_subset)
-apply (rule wf_Memrel)
-apply (auto simp add: Lt_def Memrel_def lt_def)
+apply (rule wf_subset)
+apply (rule wf_Memrel)
+apply (auto simp add: Lt_def Memrel_def lt_def)
done
lemma irrefl_Lt: "irrefl(A,Lt)"
by (auto simp add: Lt_def irrefl_def)
lemma trans_Lt: "trans[A](Lt)"
-apply (simp add: Lt_def trans_on_def)
-apply (blast intro: lt_trans)
+apply (simp add: Lt_def trans_on_def)
+apply (blast intro: lt_trans)
done
lemma part_ord_Lt: "part_ord(A,Lt)"
by (simp add: part_ord_def irrefl_Lt trans_Lt)
lemma linear_Lt: "linear(nat,Lt)"
-apply (auto dest!: not_lt_imp_le simp add: Lt_def linear_def le_iff)
-apply (drule lt_asym, auto)
+apply (auto dest!: not_lt_imp_le simp add: Lt_def linear_def le_iff)
+apply (drule lt_asym, auto)
done
lemma tot_ord_Lt: "tot_ord(nat,Lt)"
--- a/src/ZF/Ordinal.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Ordinal.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,15 +9,15 @@
definition
Memrel :: "i=>i" where
- "Memrel(A) == {z: A*A . EX x y. z=<x,y> & x:y }"
+ "Memrel(A) == {z: A*A . \<exists>x y. z=<x,y> & x:y }"
definition
Transset :: "i=>o" where
- "Transset(i) == ALL x:i. x<=i"
+ "Transset(i) == \<forall>x\<in>i. x<=i"
definition
Ord :: "i=>o" where
- "Ord(i) == Transset(i) & (ALL x:i. Transset(x))"
+ "Ord(i) == Transset(i) & (\<forall>x\<in>i. Transset(x))"
definition
lt :: "[i,i] => o" (infixl "<" 50) (*less-than on ordinals*) where
@@ -25,7 +25,7 @@
definition
Limit :: "i=>o" where
- "Limit(i) == Ord(i) & 0<i & (ALL y. y<i --> succ(y)<i)"
+ "Limit(i) == Ord(i) & 0<i & (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
abbreviation
le (infixl "le" 50) where
@@ -45,17 +45,17 @@
lemma Transset_iff_Pow: "Transset(A) <-> A<=Pow(A)"
by (unfold Transset_def, blast)
-lemma Transset_iff_Union_succ: "Transset(A) <-> Union(succ(A)) = A"
+lemma Transset_iff_Union_succ: "Transset(A) <-> \<Union>(succ(A)) = A"
apply (unfold Transset_def)
apply (blast elim!: equalityE)
done
-lemma Transset_iff_Union_subset: "Transset(A) <-> Union(A) <= A"
+lemma Transset_iff_Union_subset: "Transset(A) <-> \<Union>(A) \<subseteq> A"
by (unfold Transset_def, blast)
subsubsection{*Consequences of Downwards Closure*}
-lemma Transset_doubleton_D:
+lemma Transset_doubleton_D:
"[| Transset(C); {a,b}: C |] ==> a:C & b: C"
by (unfold Transset_def, blast)
@@ -66,11 +66,11 @@
done
lemma Transset_includes_domain:
- "[| Transset(C); A*B <= C; b: B |] ==> A <= C"
+ "[| Transset(C); A*B \<subseteq> C; b: B |] ==> A \<subseteq> C"
by (blast dest: Transset_Pair_D)
lemma Transset_includes_range:
- "[| Transset(C); A*B <= C; a: A |] ==> B <= C"
+ "[| Transset(C); A*B \<subseteq> C; a: A |] ==> B \<subseteq> C"
by (blast dest: Transset_Pair_D)
subsubsection{*Closure Properties*}
@@ -78,12 +78,12 @@
lemma Transset_0: "Transset(0)"
by (unfold Transset_def, blast)
-lemma Transset_Un:
- "[| Transset(i); Transset(j) |] ==> Transset(i Un j)"
+lemma Transset_Un:
+ "[| Transset(i); Transset(j) |] ==> Transset(i \<union> j)"
by (unfold Transset_def, blast)
-lemma Transset_Int:
- "[| Transset(i); Transset(j) |] ==> Transset(i Int j)"
+lemma Transset_Int:
+ "[| Transset(i); Transset(j) |] ==> Transset(i \<inter> j)"
by (unfold Transset_def, blast)
lemma Transset_succ: "Transset(i) ==> Transset(succ(i))"
@@ -92,36 +92,36 @@
lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))"
by (unfold Transset_def, blast)
-lemma Transset_Union: "Transset(A) ==> Transset(Union(A))"
+lemma Transset_Union: "Transset(A) ==> Transset(\<Union>(A))"
by (unfold Transset_def, blast)
-lemma Transset_Union_family:
- "[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))"
+lemma Transset_Union_family:
+ "[| !!i. i:A ==> Transset(i) |] ==> Transset(\<Union>(A))"
by (unfold Transset_def, blast)
-lemma Transset_Inter_family:
- "[| !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))"
+lemma Transset_Inter_family:
+ "[| !!i. i:A ==> Transset(i) |] ==> Transset(\<Inter>(A))"
by (unfold Inter_def Transset_def, blast)
lemma Transset_UN:
"(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (\<Union>x\<in>A. B(x))"
-by (rule Transset_Union_family, auto)
+by (rule Transset_Union_family, auto)
lemma Transset_INT:
"(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (\<Inter>x\<in>A. B(x))"
-by (rule Transset_Inter_family, auto)
+by (rule Transset_Inter_family, auto)
subsection{*Lemmas for Ordinals*}
lemma OrdI:
"[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i)"
-by (simp add: Ord_def)
+by (simp add: Ord_def)
lemma Ord_is_Transset: "Ord(i) ==> Transset(i)"
-by (simp add: Ord_def)
+by (simp add: Ord_def)
-lemma Ord_contains_Transset:
+lemma Ord_contains_Transset:
"[| Ord(i); j:i |] ==> Transset(j) "
by (unfold Ord_def, blast)
@@ -145,7 +145,7 @@
lemma Ord_trans: "[| i:j; j:k; Ord(k) |] ==> i:k"
by (blast dest: OrdmemD)
-lemma Ord_succ_subsetI: "[| i:j; Ord(j) |] ==> succ(i) <= j"
+lemma Ord_succ_subsetI: "[| i:j; Ord(j) |] ==> succ(i) \<subseteq> j"
by (blast dest: OrdmemD)
@@ -162,18 +162,18 @@
lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)"
by (blast intro: Ord_succ dest!: Ord_succD)
-lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Un j)"
+lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \<union> j)"
apply (unfold Ord_def)
apply (blast intro!: Transset_Un)
done
-lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i Int j)"
+lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \<inter> j)"
apply (unfold Ord_def)
apply (blast intro!: Transset_Int)
done
(*There is no set of all ordinals, for then it would contain itself*)
-lemma ON_class: "~ (ALL i. i:X <-> Ord(i))"
+lemma ON_class: "~ (\<forall>i. i:X <-> Ord(i))"
apply (rule notI)
apply (frule_tac x = X in spec)
apply (safe elim!: mem_irrefl)
@@ -205,7 +205,7 @@
lemma lt_Ord2: "j<i ==> Ord(i)"
by (erule ltE, assumption)
-(* "ja le j ==> Ord(j)" *)
+(* @{term"ja \<le> j ==> Ord(j)"} *)
lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD]
(* i<0 ==> R *)
@@ -231,36 +231,36 @@
done
-(** le is less than or equals; recall i le j abbrevs i<succ(j) !! **)
+(** Recall that @{term"i \<le> j"} abbreviates @{term"i<succ(j)"} !! **)
-lemma le_iff: "i le j <-> i<j | (i=j & Ord(j))"
+lemma le_iff: "i \<le> j <-> i<j | (i=j & Ord(j))"
by (unfold lt_def, blast)
(*Equivalently, i<j ==> i < succ(j)*)
-lemma leI: "i<j ==> i le j"
+lemma leI: "i<j ==> i \<le> j"
by (simp (no_asm_simp) add: le_iff)
-lemma le_eqI: "[| i=j; Ord(j) |] ==> i le j"
+lemma le_eqI: "[| i=j; Ord(j) |] ==> i \<le> j"
by (simp (no_asm_simp) add: le_iff)
lemmas le_refl = refl [THEN le_eqI]
-lemma le_refl_iff [iff]: "i le i <-> Ord(i)"
+lemma le_refl_iff [iff]: "i \<le> i <-> Ord(i)"
by (simp (no_asm_simp) add: lt_not_refl le_iff)
-lemma leCI: "(~ (i=j & Ord(j)) ==> i<j) ==> i le j"
+lemma leCI: "(~ (i=j & Ord(j)) ==> i<j) ==> i \<le> j"
by (simp add: le_iff, blast)
lemma leE:
- "[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"
+ "[| i \<le> j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"
by (simp add: le_iff, blast)
-lemma le_anti_sym: "[| i le j; j le i |] ==> i=j"
+lemma le_anti_sym: "[| i \<le> j; j \<le> i |] ==> i=j"
apply (simp add: le_iff)
apply (blast elim: lt_asym)
done
-lemma le0_iff [simp]: "i le 0 <-> i=0"
+lemma le0_iff [simp]: "i \<le> 0 <-> i=0"
by (blast elim!: leE)
lemmas le0D = le0_iff [THEN iffD1, dest!]
@@ -268,22 +268,22 @@
subsection{*Natural Deduction Rules for Memrel*}
(*The lemmas MemrelI/E give better speed than [iff] here*)
-lemma Memrel_iff [simp]: "<a,b> : Memrel(A) <-> a:b & a:A & b:A"
+lemma Memrel_iff [simp]: "<a,b> \<in> Memrel(A) <-> a:b & a:A & b:A"
by (unfold Memrel_def, blast)
-lemma MemrelI [intro!]: "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)"
+lemma MemrelI [intro!]: "[| a: b; a: A; b: A |] ==> <a,b> \<in> Memrel(A)"
by auto
lemma MemrelE [elim!]:
- "[| <a,b> : Memrel(A);
- [| a: A; b: A; a:b |] ==> P |]
+ "[| <a,b> \<in> Memrel(A);
+ [| a: A; b: A; a:b |] ==> P |]
==> P"
by auto
-lemma Memrel_type: "Memrel(A) <= A*A"
+lemma Memrel_type: "Memrel(A) \<subseteq> A*A"
by (unfold Memrel_def, blast)
-lemma Memrel_mono: "A<=B ==> Memrel(A) <= Memrel(B)"
+lemma Memrel_mono: "A<=B ==> Memrel(A) \<subseteq> Memrel(B)"
by (unfold Memrel_def, blast)
lemma Memrel_0 [simp]: "Memrel(0) = 0"
@@ -299,33 +299,33 @@
Proof idea: show A<=B by applying the foundation axiom to A-B *)
lemma wf_Memrel: "wf(Memrel(A))"
apply (unfold wf_def)
-apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast)
+apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast)
done
text{*The premise @{term "Ord(i)"} does not suffice.*}
-lemma trans_Memrel:
+lemma trans_Memrel:
"Ord(i) ==> trans(Memrel(i))"
by (unfold Ord_def Transset_def trans_def, blast)
text{*However, the following premise is strong enough.*}
-lemma Transset_trans_Memrel:
+lemma Transset_trans_Memrel:
"\<forall>j\<in>i. Transset(j) ==> trans(Memrel(i))"
by (unfold Transset_def trans_def, blast)
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*)
-lemma Transset_Memrel_iff:
- "Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A"
+lemma Transset_Memrel_iff:
+ "Transset(A) ==> <a,b> \<in> Memrel(A) <-> a:b & b:A"
by (unfold Transset_def, blast)
subsection{*Transfinite Induction*}
(*Epsilon induction over a transitive set*)
-lemma Transset_induct:
- "[| i: k; Transset(k);
- !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) |]
+lemma Transset_induct:
+ "[| i: k; Transset(k);
+ !!x.[| x: k; \<forall>y\<in>x. P(y) |] ==> P(x) |]
==> P(i)"
-apply (simp add: Transset_def)
+apply (simp add: Transset_def)
apply (erule wf_Memrel [THEN wf_induct2], blast+)
done
@@ -336,11 +336,11 @@
(*Induction over the class of ordinals -- a useful corollary of Ord_induct*)
lemma trans_induct [consumes 1]:
- "[| Ord(i);
- !!x.[| Ord(x); ALL y:x. P(y) |] ==> P(x) |]
+ "[| Ord(i);
+ !!x.[| Ord(x); \<forall>y\<in>x. P(y) |] ==> P(x) |]
==> P(i)"
apply (rule Ord_succ [THEN succI1 [THEN Ord_induct]], assumption)
-apply (blast intro: Ord_succ [THEN Ord_in_Ord])
+apply (blast intro: Ord_succ [THEN Ord_in_Ord])
done
lemmas trans_induct_rule = trans_induct [rule_format, consumes 1]
@@ -352,36 +352,36 @@
subsubsection{*Proving That < is a Linear Ordering on the Ordinals*}
lemma Ord_linear [rule_format]:
- "Ord(i) ==> (ALL j. Ord(j) --> i:j | i=j | j:i)"
+ "Ord(i) ==> (\<forall>j. Ord(j) \<longrightarrow> i:j | i=j | j:i)"
apply (erule trans_induct)
apply (rule impI [THEN allI])
-apply (erule_tac i=j in trans_induct)
-apply (blast dest: Ord_trans)
+apply (erule_tac i=j in trans_induct)
+apply (blast dest: Ord_trans)
done
(*The trichotomy law for ordinals!*)
lemma Ord_linear_lt:
"[| Ord(i); Ord(j); i<j ==> P; i=j ==> P; j<i ==> P |] ==> P"
-apply (simp add: lt_def)
+apply (simp add: lt_def)
apply (rule_tac i1=i and j1=j in Ord_linear [THEN disjE], blast+)
done
lemma Ord_linear2:
- "[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P"
+ "[| Ord(i); Ord(j); i<j ==> P; j \<le> i ==> P |] ==> P"
apply (rule_tac i = i and j = j in Ord_linear_lt)
apply (blast intro: leI le_eqI sym ) +
done
lemma Ord_linear_le:
- "[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P"
+ "[| Ord(i); Ord(j); i \<le> j ==> P; j \<le> i ==> P |] ==> P"
apply (rule_tac i = i and j = j in Ord_linear_lt)
apply (blast intro: leI le_eqI ) +
done
-lemma le_imp_not_lt: "j le i ==> ~ i<j"
+lemma le_imp_not_lt: "j \<le> i ==> ~ i<j"
by (blast elim!: leE elim: lt_asym)
-lemma not_lt_imp_le: "[| ~ i<j; Ord(i); Ord(j) |] ==> j le i"
+lemma not_lt_imp_le: "[| ~ i<j; Ord(i); Ord(j) |] ==> j \<le> i"
by (rule_tac i = i and j = j in Ord_linear2, auto)
subsubsection{*Some Rewrite Rules for <, le*}
@@ -389,160 +389,160 @@
lemma Ord_mem_iff_lt: "Ord(j) ==> i:j <-> i<j"
by (unfold lt_def, blast)
-lemma not_lt_iff_le: "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i"
+lemma not_lt_iff_le: "[| Ord(i); Ord(j) |] ==> ~ i<j <-> j \<le> i"
by (blast dest: le_imp_not_lt not_lt_imp_le)
-lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i"
+lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i \<le> j <-> j<i"
by (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym])
(*This is identical to 0<succ(i) *)
-lemma Ord_0_le: "Ord(i) ==> 0 le i"
+lemma Ord_0_le: "Ord(i) ==> 0 \<le> i"
by (erule not_lt_iff_le [THEN iffD1], auto)
-lemma Ord_0_lt: "[| Ord(i); i~=0 |] ==> 0<i"
+lemma Ord_0_lt: "[| Ord(i); i\<noteq>0 |] ==> 0<i"
apply (erule not_le_iff_lt [THEN iffD1])
apply (rule Ord_0, blast)
done
-lemma Ord_0_lt_iff: "Ord(i) ==> i~=0 <-> 0<i"
+lemma Ord_0_lt_iff: "Ord(i) ==> i\<noteq>0 <-> 0<i"
by (blast intro: Ord_0_lt)
subsection{*Results about Less-Than or Equals*}
-(** For ordinals, j<=i (subset) implies j le i (less-than or equals) **)
+(** For ordinals, @{term"j\<subseteq>i"} implies @{term"j \<le> i"} (less-than or equals) **)
-lemma zero_le_succ_iff [iff]: "0 le succ(x) <-> Ord(x)"
+lemma zero_le_succ_iff [iff]: "0 \<le> succ(x) <-> Ord(x)"
by (blast intro: Ord_0_le elim: ltE)
-lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j le i"
+lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j \<le> i"
apply (rule not_lt_iff_le [THEN iffD1], assumption+)
apply (blast elim: ltE mem_irrefl)
done
-lemma le_imp_subset: "i le j ==> i<=j"
+lemma le_imp_subset: "i \<le> j ==> i<=j"
by (blast dest: OrdmemD elim: ltE leE)
-lemma le_subset_iff: "j le i <-> j<=i & Ord(i) & Ord(j)"
+lemma le_subset_iff: "j \<le> i <-> j<=i & Ord(i) & Ord(j)"
by (blast dest: subset_imp_le le_imp_subset elim: ltE)
-lemma le_succ_iff: "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"
+lemma le_succ_iff: "i \<le> succ(j) <-> i \<le> j | i=succ(j) & Ord(i)"
apply (simp (no_asm) add: le_iff)
apply blast
done
(*Just a variant of subset_imp_le*)
-lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i"
+lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j \<le> i"
by (blast intro: not_lt_imp_le dest: lt_irrefl)
subsubsection{*Transitivity Laws*}
-lemma lt_trans1: "[| i le j; j<k |] ==> i<k"
+lemma lt_trans1: "[| i \<le> j; j<k |] ==> i<k"
by (blast elim!: leE intro: lt_trans)
-lemma lt_trans2: "[| i<j; j le k |] ==> i<k"
+lemma lt_trans2: "[| i<j; j \<le> k |] ==> i<k"
by (blast elim!: leE intro: lt_trans)
-lemma le_trans: "[| i le j; j le k |] ==> i le k"
+lemma le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> k"
by (blast intro: lt_trans1)
-lemma succ_leI: "i<j ==> succ(i) le j"
-apply (rule not_lt_iff_le [THEN iffD1])
+lemma succ_leI: "i<j ==> succ(i) \<le> j"
+apply (rule not_lt_iff_le [THEN iffD1])
apply (blast elim: ltE leE lt_asym)+
done
(*Identical to succ(i) < succ(j) ==> i<j *)
-lemma succ_leE: "succ(i) le j ==> i<j"
+lemma succ_leE: "succ(i) \<le> j ==> i<j"
apply (rule not_le_iff_lt [THEN iffD1])
apply (blast elim: ltE leE lt_asym)+
done
-lemma succ_le_iff [iff]: "succ(i) le j <-> i<j"
+lemma succ_le_iff [iff]: "succ(i) \<le> j <-> i<j"
by (blast intro: succ_leI succ_leE)
-lemma succ_le_imp_le: "succ(i) le succ(j) ==> i le j"
+lemma succ_le_imp_le: "succ(i) \<le> succ(j) ==> i \<le> j"
by (blast dest!: succ_leE)
-lemma lt_subset_trans: "[| i <= j; j<k; Ord(i) |] ==> i<k"
-apply (rule subset_imp_le [THEN lt_trans1])
+lemma lt_subset_trans: "[| i \<subseteq> j; j<k; Ord(i) |] ==> i<k"
+apply (rule subset_imp_le [THEN lt_trans1])
apply (blast intro: elim: ltE) +
done
lemma lt_imp_0_lt: "j<i ==> 0<i"
-by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord])
+by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord])
lemma succ_lt_iff: "succ(i) < j <-> i<j & succ(i) \<noteq> j"
-apply auto
-apply (blast intro: lt_trans le_refl dest: lt_Ord)
-apply (frule lt_Ord)
-apply (rule not_le_iff_lt [THEN iffD1])
+apply auto
+apply (blast intro: lt_trans le_refl dest: lt_Ord)
+apply (frule lt_Ord)
+apply (rule not_le_iff_lt [THEN iffD1])
apply (blast intro: lt_Ord2)
- apply blast
-apply (simp add: lt_Ord lt_Ord2 le_iff)
-apply (blast dest: lt_asym)
+ apply blast
+apply (simp add: lt_Ord lt_Ord2 le_iff)
+apply (blast dest: lt_asym)
done
lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \<in> succ(j) <-> i\<in>j"
-apply (insert succ_le_iff [of i j])
-apply (simp add: lt_def)
+apply (insert succ_le_iff [of i j])
+apply (simp add: lt_def)
done
subsubsection{*Union and Intersection*}
-lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i le i Un j"
+lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i \<le> i \<union> j"
by (rule Un_upper1 [THEN subset_imp_le], auto)
-lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j le i Un j"
+lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j \<le> i \<union> j"
by (rule Un_upper2 [THEN subset_imp_le], auto)
(*Replacing k by succ(k') yields the similar rule for le!*)
-lemma Un_least_lt: "[| i<k; j<k |] ==> i Un j < k"
+lemma Un_least_lt: "[| i<k; j<k |] ==> i \<union> j < k"
apply (rule_tac i = i and j = j in Ord_linear_le)
-apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord)
+apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord)
done
-lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k"
+lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i \<union> j < k <-> i<k & j<k"
apply (safe intro!: Un_least_lt)
apply (rule_tac [2] Un_upper2_le [THEN lt_trans1])
-apply (rule Un_upper1_le [THEN lt_trans1], auto)
+apply (rule Un_upper1_le [THEN lt_trans1], auto)
done
lemma Un_least_mem_iff:
- "[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k"
-apply (insert Un_least_lt_iff [of i j k])
+ "[| Ord(i); Ord(j); Ord(k) |] ==> i \<union> j \<in> k <-> i:k & j:k"
+apply (insert Un_least_lt_iff [of i j k])
apply (simp add: lt_def)
done
(*Replacing k by succ(k') yields the similar rule for le!*)
-lemma Int_greatest_lt: "[| i<k; j<k |] ==> i Int j < k"
+lemma Int_greatest_lt: "[| i<k; j<k |] ==> i \<inter> j < k"
apply (rule_tac i = i and j = j in Ord_linear_le)
-apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord)
+apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord)
done
lemma Ord_Un_if:
"[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)"
by (simp add: not_lt_iff_le le_imp_subset leI
- subset_Un_iff [symmetric] subset_Un_iff2 [symmetric])
+ subset_Un_iff [symmetric] subset_Un_iff2 [symmetric])
lemma succ_Un_distrib:
"[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)"
-by (simp add: Ord_Un_if lt_Ord le_Ord2)
+by (simp add: Ord_Un_if lt_Ord le_Ord2)
lemma lt_Un_iff:
"[| Ord(i); Ord(j) |] ==> k < i \<union> j <-> k < i | k < j";
-apply (simp add: Ord_Un_if not_lt_iff_le)
-apply (blast intro: leI lt_trans2)+
+apply (simp add: Ord_Un_if not_lt_iff_le)
+apply (blast intro: leI lt_trans2)+
done
lemma le_Un_iff:
"[| Ord(i); Ord(j) |] ==> k \<le> i \<union> j <-> k \<le> i | k \<le> j";
-by (simp add: succ_Un_distrib lt_Un_iff [symmetric])
+by (simp add: succ_Un_distrib lt_Un_iff [symmetric])
-lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i Un j"
-by (simp add: lt_Un_iff lt_Ord2)
+lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i \<union> j"
+by (simp add: lt_Un_iff lt_Ord2)
-lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i Un j"
-by (simp add: lt_Un_iff lt_Ord2)
+lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i \<union> j"
+by (simp add: lt_Un_iff lt_Ord2)
(*See also Transset_iff_Union_succ*)
lemma Ord_Union_succ_eq: "Ord(i) ==> \<Union>(succ(i)) = i"
@@ -551,7 +551,7 @@
subsection{*Results about Limits*}
-lemma Ord_Union [intro,simp,TC]: "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))"
+lemma Ord_Union [intro,simp,TC]: "[| !!i. i:A ==> Ord(i) |] ==> Ord(\<Union>(A))"
apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI])
apply (blast intro: Ord_contains_Transset)+
done
@@ -561,21 +561,21 @@
by (rule Ord_Union, blast)
lemma Ord_Inter [intro,simp,TC]:
- "[| !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))"
+ "[| !!i. i:A ==> Ord(i) |] ==> Ord(\<Inter>(A))"
apply (rule Transset_Inter_family [THEN OrdI])
-apply (blast intro: Ord_is_Transset)
-apply (simp add: Inter_def)
-apply (blast intro: Ord_contains_Transset)
+apply (blast intro: Ord_is_Transset)
+apply (simp add: Inter_def)
+apply (blast intro: Ord_contains_Transset)
done
lemma Ord_INT [intro,simp,TC]:
"[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(\<Inter>x\<in>A. B(x))"
-by (rule Ord_Inter, blast)
+by (rule Ord_Inter, blast)
-(* No < version; consider (\<Union>i\<in>nat.i)=nat *)
+(* No < version of this theorem: consider that @{term"(\<Union>i\<in>nat.i)=nat"}! *)
lemma UN_least_le:
- "[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (\<Union>x\<in>A. b(x)) le i"
+ "[| Ord(i); !!x. x:A ==> b(x) \<le> i |] ==> (\<Union>x\<in>A. b(x)) \<le> i"
apply (rule le_imp_subset [THEN UN_least, THEN subset_imp_le])
apply (blast intro: Ord_UN elim: ltE)+
done
@@ -589,10 +589,10 @@
lemma UN_upper_lt:
"[| a\<in>A; i < b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"
-by (unfold lt_def, blast)
+by (unfold lt_def, blast)
lemma UN_upper_le:
- "[| a: A; i le b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i le (\<Union>x\<in>A. b(x))"
+ "[| a: A; i \<le> b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i \<le> (\<Union>x\<in>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)+
@@ -603,28 +603,28 @@
lemma Union_upper_le:
"[| j: J; i\<le>j; Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"
-apply (subst Union_eq_UN)
+apply (subst Union_eq_UN)
apply (rule UN_upper_le, auto)
done
lemma le_implies_UN_le_UN:
- "[| !!x. x:A ==> c(x) le d(x) |] ==> (\<Union>x\<in>A. c(x)) le (\<Union>x\<in>A. d(x))"
+ "[| !!x. x:A ==> c(x) \<le> d(x) |] ==> (\<Union>x\<in>A. c(x)) \<le> (\<Union>x\<in>A. d(x))"
apply (rule UN_least_le)
apply (rule_tac [2] UN_upper_le)
-apply (blast intro: Ord_UN le_Ord2)+
+apply (blast intro: Ord_UN le_Ord2)+
done
lemma Ord_equality: "Ord(i) ==> (\<Union>y\<in>i. succ(y)) = i"
by (blast intro: Ord_trans)
(*Holds for all transitive sets, not just ordinals*)
-lemma Ord_Union_subset: "Ord(i) ==> Union(i) <= i"
+lemma Ord_Union_subset: "Ord(i) ==> \<Union>(i) \<subseteq> i"
by (blast intro: Ord_trans)
subsection{*Limit Ordinals -- General Properties*}
-lemma Limit_Union_eq: "Limit(i) ==> Union(i) = i"
+lemma Limit_Union_eq: "Limit(i) ==> \<Union>(i) = i"
apply (unfold Limit_def)
apply (fast intro!: ltI elim!: ltE elim: Ord_trans)
done
@@ -639,7 +639,7 @@
apply (erule conjunct2 [THEN conjunct1])
done
-lemma Limit_nonzero: "Limit(i) ==> i ~= 0"
+lemma Limit_nonzero: "Limit(i) ==> i \<noteq> 0"
by (drule Limit_has_0, blast)
lemma Limit_has_succ: "[| Limit(i); j<i |] ==> succ(j) < i"
@@ -648,7 +648,7 @@
lemma Limit_succ_lt_iff [simp]: "Limit(i) ==> succ(j) < i <-> (j<i)"
apply (safe intro!: Limit_has_succ)
apply (frule lt_Ord)
-apply (blast intro: lt_trans)
+apply (blast intro: lt_trans)
done
lemma zero_not_Limit [iff]: "~ Limit(0)"
@@ -659,16 +659,16 @@
lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"
apply (unfold Limit_def, simp add: lt_Ord2, clarify)
-apply (drule_tac i=y in ltD)
+apply (drule_tac i=y in ltD)
apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
done
-lemma non_succ_LimitI:
- "[| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)"
+lemma non_succ_LimitI:
+ "[| 0<i; \<forall>y. succ(y) \<noteq> i |] ==> Limit(i)"
apply (unfold Limit_def)
apply (safe del: subsetI)
apply (rule_tac [2] not_le_iff_lt [THEN iffD1])
-apply (simp_all add: lt_Ord lt_Ord2)
+apply (simp_all add: lt_Ord lt_Ord2)
apply (blast elim: leE lt_asym)
done
@@ -681,28 +681,28 @@
lemma not_succ_Limit [simp]: "~ Limit(succ(i))"
by blast
-lemma Limit_le_succD: "[| Limit(i); i le succ(j) |] ==> i le j"
+lemma Limit_le_succD: "[| Limit(i); i \<le> succ(j) |] ==> i \<le> j"
by (blast elim!: leE)
subsubsection{*Traditional 3-Way Case Analysis on Ordinals*}
-lemma Ord_cases_disj: "Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"
+lemma Ord_cases_disj: "Ord(i) ==> i=0 | (\<exists>j. Ord(j) & i=succ(j)) | Limit(i)"
by (blast intro!: non_succ_LimitI Ord_0_lt)
lemma Ord_cases:
- "[| Ord(i);
- i=0 ==> P;
- !!j. [| Ord(j); i=succ(j) |] ==> P;
- Limit(i) ==> P
+ "[| Ord(i);
+ i=0 ==> P;
+ !!j. [| Ord(j); i=succ(j) |] ==> P;
+ Limit(i) ==> P
|] ==> P"
-by (drule Ord_cases_disj, blast)
+by (drule Ord_cases_disj, blast)
lemma trans_induct3 [case_names 0 succ limit, consumes 1]:
- "[| Ord(i);
- P(0);
- !!x. [| Ord(x); P(x) |] ==> P(succ(x));
- !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x)
+ "[| Ord(i);
+ P(0);
+ !!x. [| Ord(x); P(x) |] ==> P(succ(x));
+ !!x. [| Limit(x); \<forall>y\<in>x. P(y) |] ==> P(x)
|] ==> P(i)"
apply (erule trans_induct)
apply (erule Ord_cases, blast+)
@@ -714,16 +714,16 @@
union is a limit ordinal.*}
lemma Ord_set_cases:
"\<forall>i\<in>I. Ord(i) ==> I=0 \<or> \<Union>(I) \<in> I \<or> (\<Union>(I) \<notin> I \<and> Limit(\<Union>(I)))"
-apply (clarify elim!: not_emptyE)
-apply (cases "\<Union>(I)" rule: Ord_cases)
+apply (clarify elim!: not_emptyE)
+apply (cases "\<Union>(I)" rule: Ord_cases)
apply (blast intro: Ord_Union)
apply (blast intro: subst_elem)
- apply auto
+ apply auto
apply (clarify elim!: equalityE succ_subsetE)
apply (simp add: Union_subset_iff)
apply (subgoal_tac "B = succ(j)", blast)
-apply (rule le_anti_sym)
- apply (simp add: le_subset_iff)
+apply (rule le_anti_sym)
+ apply (simp add: le_subset_iff)
apply (simp add: ltI)
done
--- a/src/ZF/Perm.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Perm.thy Tue Mar 06 15:15:49 2012 +0000
@@ -15,28 +15,28 @@
definition
(*composition of relations and functions; NOT Suppes's relative product*)
comp :: "[i,i]=>i" (infixr "O" 60) where
- "r O s == {xz : domain(s)*range(r) .
- EX x y z. xz=<x,z> & <x,y>:s & <y,z>:r}"
+ "r O s == {xz \<in> domain(s)*range(r) .
+ \<exists>x y z. xz=<x,z> & <x,y>:s & <y,z>:r}"
definition
(*the identity function for A*)
id :: "i=>i" where
- "id(A) == (lam x:A. x)"
+ "id(A) == (\<lambda>x\<in>A. x)"
definition
(*one-to-one functions from A to B*)
inj :: "[i,i]=>i" where
- "inj(A,B) == { f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x}"
+ "inj(A,B) == { f: A->B. \<forall>w\<in>A. \<forall>x\<in>A. f`w=f`x \<longrightarrow> w=x}"
definition
(*onto functions from A to B*)
surj :: "[i,i]=>i" where
- "surj(A,B) == { f: A->B . ALL y:B. EX x:A. f`x=y}"
+ "surj(A,B) == { f: A->B . \<forall>y\<in>B. \<exists>x\<in>A. f`x=y}"
definition
(*one-to-one and onto functions*)
bij :: "[i,i]=>i" where
- "bij(A,B) == inj(A,B) Int surj(A,B)"
+ "bij(A,B) == inj(A,B) \<inter> surj(A,B)"
subsection{*Surjective Function Space*}
@@ -46,7 +46,7 @@
apply (erule CollectD1)
done
-lemma fun_is_surj: "f : Pi(A,B) ==> f: surj(A,range(f))"
+lemma fun_is_surj: "f \<in> Pi(A,B) ==> f: surj(A,range(f))"
apply (unfold surj_def)
apply (blast intro: apply_equality range_of_fun domain_type)
done
@@ -67,13 +67,13 @@
"[| !!x. x:A ==> c(x): B;
!!y. y:B ==> d(y): A;
!!y. y:B ==> c(d(y)) = y
- |] ==> (lam x:A. c(x)) : surj(A,B)"
+ |] ==> (\<lambda>x\<in>A. c(x)) \<in> surj(A,B)"
apply (rule_tac d = d in f_imp_surjective)
apply (simp_all add: lam_type)
done
text{*Cantor's theorem revisited*}
-lemma cantor_surj: "f ~: surj(A,Pow(A))"
+lemma cantor_surj: "f \<notin> surj(A,Pow(A))"
apply (unfold surj_def, safe)
apply (cut_tac cantor)
apply (best del: subsetI)
@@ -99,7 +99,7 @@
text{* A function with a left inverse is an injection *}
-lemma f_imp_injective: "[| f: A->B; ALL x:A. d(f`x)=x |] ==> f: inj(A,B)"
+lemma f_imp_injective: "[| f: A->B; \<forall>x\<in>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
@@ -107,7 +107,7 @@
lemma lam_injective:
"[| !!x. x:A ==> c(x): B;
!!x. x:A ==> d(c(x)) = x |]
- ==> (lam x:A. c(x)) : inj(A,B)"
+ ==> (\<lambda>x\<in>A. c(x)) \<in> inj(A,B)"
apply (rule_tac d = d in f_imp_injective)
apply (simp_all add: lam_type)
done
@@ -132,13 +132,13 @@
!!y. y:B ==> d(y): A;
!!x. x:A ==> d(c(x)) = x;
!!y. y:B ==> c(d(y)) = y
- |] ==> (lam x:A. c(x)) : bij(A,B)"
+ |] ==> (\<lambda>x\<in>A. c(x)) \<in> bij(A,B)"
apply (unfold bij_def)
apply (blast intro!: lam_injective lam_surjective)
done
-lemma RepFun_bijective: "(ALL y : x. EX! y'. f(y') = f(y))
- ==> (lam z:{f(y). y:x}. THE y. f(y) = z) : bij({f(y). y:x}, x)"
+lemma RepFun_bijective: "(\<forall>y\<in>x. EX! y'. f(y') = f(y))
+ ==> (\<lambda>z\<in>{f(y). y:x}. THE y. f(y) = z) \<in> bij({f(y). y:x}, x)"
apply (rule_tac d = f in lam_bijective)
apply (auto simp add: the_equality2)
done
@@ -146,7 +146,7 @@
subsection{*Identity Function*}
-lemma idI [intro!]: "a:A ==> <a,a> : id(A)"
+lemma idI [intro!]: "a:A ==> <a,a> \<in> id(A)"
apply (unfold id_def)
apply (erule lamI)
done
@@ -154,7 +154,7 @@
lemma idE [elim!]: "[| p: id(A); !!x.[| x:A; p=<x,x> |] ==> P |] ==> P"
by (simp add: id_def lam_def, blast)
-lemma id_type: "id(A) : A->A"
+lemma id_type: "id(A) \<in> A->A"
apply (unfold id_def)
apply (rule lam_type, assumption)
done
@@ -164,7 +164,7 @@
apply (simp (no_asm_simp))
done
-lemma id_mono: "A<=B ==> id(A) <= id(B)"
+lemma id_mono: "A<=B ==> id(A) \<subseteq> id(B)"
apply (unfold id_def)
apply (erule lam_mono)
done
@@ -186,7 +186,7 @@
apply (blast intro: id_inj id_surj)
done
-lemma subset_iff_id: "A <= B <-> id(A) : A->B"
+lemma subset_iff_id: "A \<subseteq> B <-> id(A) \<in> A->B"
apply (unfold id_def)
apply (force intro!: lam_type dest: apply_type)
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) \<in> range(f)->A"
apply (unfold inj_def)
apply (simp (no_asm_simp) add: Pi_iff function_def)
apply (erule CollectE)
@@ -259,19 +259,19 @@
subsection{*Composition of Two Relations*}
-text{*The inductive definition package could derive these theorems for @term{r O s}*}
+text{*The inductive definition package could derive these theorems for @{term"r O s"}*}
-lemma compI [intro]: "[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s"
+lemma compI [intro]: "[| <a,b>:s; <b,c>:r |] ==> <a,c> \<in> r O s"
by (unfold comp_def, blast)
lemma compE [elim!]:
- "[| xz : r O s;
+ "[| xz \<in> r O s;
!!x y z. [| xz=<x,z>; <x,y>:s; <y,z>:r |] ==> P |]
==> P"
by (unfold comp_def, blast)
lemma compEpair:
- "[| <a,c> : r O s;
+ "[| <a,c> \<in> r O s;
!!y. [| <a,y>:s; <y,c>:r |] ==> P |]
==> P"
by (erule compE, simp)
@@ -283,35 +283,35 @@
subsection{*Domain and Range -- see Suppes, Section 3.1*}
text{*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*}
-lemma range_comp: "range(r O s) <= range(r)"
+lemma range_comp: "range(r O s) \<subseteq> range(r)"
by blast
-lemma range_comp_eq: "domain(r) <= range(s) ==> range(r O s) = range(r)"
+lemma range_comp_eq: "domain(r) \<subseteq> range(s) ==> range(r O s) = range(r)"
by (rule range_comp [THEN equalityI], blast)
-lemma domain_comp: "domain(r O s) <= domain(s)"
+lemma domain_comp: "domain(r O s) \<subseteq> domain(s)"
by blast
-lemma domain_comp_eq: "range(s) <= domain(r) ==> domain(r O s) = domain(s)"
+lemma domain_comp_eq: "range(s) \<subseteq> domain(r) ==> domain(r O s) = domain(s)"
by (rule domain_comp [THEN equalityI], blast)
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 \<in> 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 \<in> bij(A,range(f))"
by (auto simp add: bij_def intro: inj_inj_range inj_is_fun fun_is_surj)
subsection{*Other Results*}
-lemma comp_mono: "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"
+lemma comp_mono: "[| r'<=r; s'<=s |] ==> (r' O s') \<subseteq> (r O s)"
by blast
text{*composition preserves relations*}
-lemma comp_rel: "[| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"
+lemma comp_rel: "[| s<=A*B; r<=B*C |] ==> (r O s) \<subseteq> A*C"
by blast
text{*associative law for composition*}
@@ -319,14 +319,14 @@
by blast
(*left identity of composition; provable inclusions are
- id(A) O r <= r
- and [| r<=A*B; B<=C |] ==> r <= id(C) O r *)
+ id(A) O r \<subseteq> r
+ and [| r<=A*B; B<=C |] ==> r \<subseteq> id(C) O r *)
lemma left_comp_id: "r<=A*B ==> id(B) O r = r"
by blast
(*right identity of composition; provable inclusions are
- r O id(A) <= r
- and [| r<=A*B; A<=C |] ==> r <= r O id(C) *)
+ r O id(A) \<subseteq> r
+ and [| r<=A*B; A<=C |] ==> r \<subseteq> r O id(C) *)
lemma right_comp_id: "r<=A*B ==> r O id(A) = r"
by blast
@@ -337,7 +337,7 @@
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) \<in> 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
@@ -353,8 +353,8 @@
text{*Simplifies compositions of lambda-abstractions*}
lemma comp_lam:
"[| !!x. x:A ==> b(x): B |]
- ==> (lam y:B. c(y)) O (lam x:A. b(x)) = (lam x:A. c(b(x)))"
-apply (subgoal_tac "(lam x:A. b(x)) : A -> B")
+ ==> (\<lambda>y\<in>B. c(y)) O (\<lambda>x\<in>A. b(x)) = (\<lambda>x\<in>A. c(b(x)))"
+apply (subgoal_tac "(\<lambda>x\<in>A. b(x)) \<in> A -> B")
apply (rule fun_extension)
apply (blast intro: comp_fun lam_funtype)
apply (rule lam_funtype)
@@ -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) \<in> 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) \<in> 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) \<in> bij(A,C)"
apply (unfold bij_def)
apply (blast intro: comp_inj comp_surj)
done
@@ -420,7 +420,7 @@
subsubsection{*Inverses of Composition*}
text{*left inverse of composition; one inclusion is
- @term{f: A->B ==> id(A) <= converse(f) O f} *}
+ @{term "f: A->B ==> id(A) \<subseteq> 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)
@@ -428,7 +428,7 @@
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) \<subseteq> id(B)"} *}
lemma right_comp_inverse:
"f: surj(A,B) ==> f O converse(f) = id(B)"
apply (simp add: surj_def, clarify)
@@ -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) <-> (ALL y:B. f`(g`y)=y)"
+ "[| f: A->B; g: B->A |] ==> f O g = id(B) <-> (\<forall>y\<in>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 \<in> 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 \<in> 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 \<in> bij(A,B)"
by (simp add: fg_imp_bijective comp_eq_id_iff
left_inverse_lemma right_inverse_lemma)
@@ -471,16 +471,16 @@
text{*Theorem by KG, proof by LCP*}
lemma inj_disjoint_Un:
- "[| f: inj(A,B); g: inj(C,D); B Int D = 0 |]
- ==> (lam a: A Un C. if a:A then f`a else g`a) : inj(A Un C, B Un D)"
+ "[| f: inj(A,B); g: inj(C,D); B \<inter> D = 0 |]
+ ==> (\<lambda>a\<in>A \<union> C. if a:A then f`a else g`a) \<in> inj(A \<union> C, B \<union> 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 Int C = 0 |]
- ==> (f Un g) : surj(A Un C, B Un D)"
+ "[| f: surj(A,B); g: surj(C,D); A \<inter> C = 0 |]
+ ==> (f \<union> g) \<in> surj(A \<union> C, B \<union> D)"
apply (simp add: surj_def fun_disjoint_Un)
apply (blast dest!: domain_of_fun
intro!: fun_disjoint_apply1 fun_disjoint_apply2)
@@ -489,8 +489,8 @@
text{*A simple, high-level proof; the version for injections follows from it,
using @term{f:inj(A,B) <-> f:bij(A,range(f))} *}
lemma bij_disjoint_Un:
- "[| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |]
- ==> (f Un g) : bij(A Un C, B Un D)"
+ "[| f: bij(A,B); g: bij(C,D); A \<inter> C = 0; B \<inter> D = 0 |]
+ ==> (f \<union> g) \<in> bij(A \<union> C, B \<union> D)"
apply (rule invertible_imp_bijective)
apply (subst converse_Un)
apply (auto intro: fun_disjoint_Un bij_is_fun bij_converse_bij)
@@ -505,7 +505,7 @@
apply (blast intro: apply_equality apply_Pair Pi_type)
done
-lemma restrict_image [simp]: "restrict(f,A) `` B = f `` (A Int B)"
+lemma restrict_image [simp]: "restrict(f,A) `` B = f `` (A \<inter> B)"
by (auto simp add: restrict_def)
lemma restrict_inj:
@@ -534,7 +534,7 @@
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) \<in> 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,8 +542,8 @@
lemma inj_extend:
- "[| f: inj(A,B); a~:A; b~:B |]
- ==> cons(<a,b>,f) : inj(cons(a,A), cons(b,B))"
+ "[| f: inj(A,B); a\<notin>A; b\<notin>B |]
+ ==> cons(<a,b>,f) \<in> inj(cons(a,A), cons(b,B))"
apply (unfold inj_def)
apply (force intro: apply_type simp add: fun_extend)
done
--- a/src/ZF/QPair.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/QPair.thy Tue Mar 06 15:15:49 2012 +0000
@@ -26,11 +26,11 @@
definition
qfst :: "i => i" where
- "qfst(p) == THE a. EX b. p=<a;b>"
+ "qfst(p) == THE a. \<exists>b. p=<a;b>"
definition
qsnd :: "i => i" where
- "qsnd(p) == THE b. EX a. p=<a;b>"
+ "qsnd(p) == THE b. \<exists>a. p=<a;b>"
definition
qsplit :: "[[i, i] => 'a, i] => 'a::{}" (*for pattern-matching*) where
@@ -38,7 +38,7 @@
definition
qconverse :: "i => i" where
- "qconverse(r) == {z. w:r, EX x y. w=<x;y> & z=<y;x>}"
+ "qconverse(r) == {z. w:r, \<exists>x y. w=<x;y> & z=<y;x>}"
definition
QSigma :: "[i, i => i] => i" where
@@ -55,7 +55,7 @@
definition
qsum :: "[i,i]=>i" (infixr "<+>" 65) where
- "A <+> B == ({0} <*> A) Un ({1} <*> B)"
+ "A <+> B == ({0} <*> A) \<union> ({1} <*> B)"
definition
QInl :: "i=>i" where
@@ -94,7 +94,7 @@
subsubsection{*QSigma: Disjoint union of a family of sets
Generalizes Cartesian product*}
-lemma QSigmaI [intro!]: "[| a:A; b:B(a) |] ==> <a;b> : QSigma(A,B)"
+lemma QSigmaI [intro!]: "[| a:A; b:B(a) |] ==> <a;b> \<in> QSigma(A,B)"
by (simp add: QSigma_def)
@@ -110,10 +110,10 @@
"[| <a;b>: QSigma(A,B); [| a:A; b:B(a) |] ==> P |] ==> P"
by (simp add: QSigma_def)
-lemma QSigmaD1: "<a;b> : QSigma(A,B) ==> a : A"
+lemma QSigmaD1: "<a;b> \<in> QSigma(A,B) ==> a \<in> A"
by blast
-lemma QSigmaD2: "<a;b> : QSigma(A,B) ==> b : B(a)"
+lemma QSigmaD2: "<a;b> \<in> QSigma(A,B) ==> b \<in> B(a)"
by blast
lemma QSigma_cong:
@@ -136,10 +136,10 @@
lemma qsnd_conv [simp]: "qsnd(<a;b>) = 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) \<in> 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) \<in> B(qfst(p))"
by auto
lemma QPair_qfst_qsnd_eq: "a: QSigma(A,B) ==> <qfst(a); qsnd(a)> = a"
@@ -156,11 +156,11 @@
lemma qsplit_type [elim!]:
"[| p:QSigma(A,B);
!!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x;y>)
- |] ==> qsplit(%x y. c(x,y), p) : C(p)"
+ |] ==> qsplit(%x y. c(x,y), p) \<in> C(p)"
by auto
lemma expand_qsplit:
- "u: A<*>B ==> R(qsplit(c,u)) <-> (ALL x:A. ALL y:B. u = <x;y> --> R(c(x,y)))"
+ "u: A<*>B ==> R(qsplit(c,u)) <-> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x;y> \<longrightarrow> R(c(x,y)))"
apply (simp add: qsplit_def, auto)
done
@@ -186,11 +186,11 @@
lemma qconverseI [intro!]: "<a;b>:r ==> <b;a>:qconverse(r)"
by (simp add: qconverse_def, blast)
-lemma qconverseD [elim!]: "<a;b> : qconverse(r) ==> <b;a> : r"
+lemma qconverseD [elim!]: "<a;b> \<in> qconverse(r) ==> <b;a> \<in> r"
by (simp add: qconverse_def, blast)
lemma qconverseE [elim!]:
- "[| yx : qconverse(r);
+ "[| yx \<in> qconverse(r);
!!x y. [| yx=<y;x>; <x;y>:r |] ==> P
|] ==> P"
by (simp add: qconverse_def, blast)
@@ -198,7 +198,7 @@
lemma qconverse_qconverse: "r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r"
by blast
-lemma qconverse_type: "r <= A <*> B ==> qconverse(r) <= B <*> A"
+lemma qconverse_type: "r \<subseteq> A <*> B ==> qconverse(r) \<subseteq> B <*> A"
by blast
lemma qconverse_prod: "qconverse(A <*> B) = B <*> A"
@@ -214,10 +214,10 @@
(** Introduction rules for the injections **)
-lemma QInlI [intro!]: "a : A ==> QInl(a) : A <+> B"
+lemma QInlI [intro!]: "a \<in> A ==> QInl(a) \<in> A <+> B"
by (simp add: qsum_defs, blast)
-lemma QInrI [intro!]: "b : B ==> QInr(b) : A <+> B"
+lemma QInrI [intro!]: "b \<in> B ==> QInr(b) \<in> A <+> B"
by (simp add: qsum_defs, blast)
(** Elimination rules **)
@@ -263,10 +263,10 @@
(** <+> is itself injective... who cares?? **)
lemma qsum_iff:
- "u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"
+ "u: A <+> B <-> (\<exists>x. x:A & u=QInl(x)) | (\<exists>y. y:B & u=QInr(y))"
by blast
-lemma qsum_subset_iff: "A <+> B <= C <+> D <-> A<=C & B<=D"
+lemma qsum_subset_iff: "A <+> B \<subseteq> C <+> D <-> A<=C & B<=D"
by blast
lemma qsum_equal_iff: "A <+> B = C <+> D <-> A=C & B=D"
@@ -287,7 +287,7 @@
"[| 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)"
+ |] ==> qcase(c,d,u) \<in> C(u)"
by (simp add: qsum_defs, auto)
(** Rules for the Part primitive **)
@@ -301,26 +301,26 @@
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) Un Part(C,QInr) = C"
+lemma Part_qsum_equality: "C \<subseteq> A <+> B ==> Part(C,QInl) \<union> Part(C,QInr) = C"
by blast
subsubsection{*Monotonicity*}
-lemma QPair_mono: "[| a<=c; b<=d |] ==> <a;b> <= <c;d>"
+lemma QPair_mono: "[| a<=c; b<=d |] ==> <a;b> \<subseteq> <c;d>"
by (simp add: QPair_def sum_mono)
lemma QSigma_mono [rule_format]:
- "[| A<=C; ALL x:A. B(x) <= D(x) |] ==> QSigma(A,B) <= QSigma(C,D)"
+ "[| A<=C; \<forall>x\<in>A. B(x) \<subseteq> D(x) |] ==> QSigma(A,B) \<subseteq> QSigma(C,D)"
by blast
-lemma QInl_mono: "a<=b ==> QInl(a) <= QInl(b)"
+lemma QInl_mono: "a<=b ==> QInl(a) \<subseteq> QInl(b)"
by (simp add: QInl_def subset_refl [THEN QPair_mono])
-lemma QInr_mono: "a<=b ==> QInr(a) <= QInr(b)"
+lemma QInr_mono: "a<=b ==> QInr(a) \<subseteq> QInr(b)"
by (simp add: QInr_def subset_refl [THEN QPair_mono])
-lemma qsum_mono: "[| A<=C; B<=D |] ==> A <+> B <= C <+> D"
+lemma qsum_mono: "[| A<=C; B<=D |] ==> A <+> B \<subseteq> C <+> D"
by blast
end
--- a/src/ZF/QUniv.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/QUniv.thy Tue Mar 06 15:15:49 2012 +0000
@@ -8,13 +8,13 @@
theory QUniv imports Univ QPair begin
(*Disjoint sums as a datatype*)
-rep_datatype
+rep_datatype
elimination sumE
induction TrueI
case_eqns case_Inl case_Inr
(*Variant disjoint sums as a datatype*)
-rep_datatype
+rep_datatype
elimination qsumE
induction TrueI
case_eqns qcase_QInl qcase_QInr
@@ -27,46 +27,46 @@
subsection{*Properties involving Transset and Sum*}
lemma Transset_includes_summands:
- "[| Transset(C); A+B <= C |] ==> A <= C & B <= C"
-apply (simp add: sum_def Un_subset_iff)
+ "[| Transset(C); A+B \<subseteq> C |] ==> A \<subseteq> C & B \<subseteq> C"
+apply (simp add: sum_def Un_subset_iff)
apply (blast dest: Transset_includes_range)
done
lemma Transset_sum_Int_subset:
- "Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)"
-apply (simp add: sum_def Int_Un_distrib2)
+ "Transset(C) ==> (A+B) \<inter> C \<subseteq> (A \<inter> C) + (B \<inter> C)"
+apply (simp add: sum_def Int_Un_distrib2)
apply (blast dest: Transset_Pair_D)
done
subsection{*Introduction and Elimination Rules*}
-lemma qunivI: "X <= univ(eclose(A)) ==> X : quniv(A)"
+lemma qunivI: "X \<subseteq> univ(eclose(A)) ==> X \<in> quniv(A)"
by (simp add: quniv_def)
-lemma qunivD: "X : quniv(A) ==> X <= univ(eclose(A))"
+lemma qunivD: "X \<in> quniv(A) ==> X \<subseteq> univ(eclose(A))"
by (simp add: quniv_def)
-lemma quniv_mono: "A<=B ==> quniv(A) <= quniv(B)"
+lemma quniv_mono: "A<=B ==> quniv(A) \<subseteq> quniv(B)"
apply (unfold quniv_def)
apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono])
done
subsection{*Closure Properties*}
-lemma univ_eclose_subset_quniv: "univ(eclose(A)) <= quniv(A)"
-apply (simp add: quniv_def Transset_iff_Pow [symmetric])
+lemma univ_eclose_subset_quniv: "univ(eclose(A)) \<subseteq> quniv(A)"
+apply (simp add: quniv_def Transset_iff_Pow [symmetric])
apply (rule Transset_eclose [THEN Transset_univ])
done
(*Key property for proving A_subset_quniv; requires eclose in def of quniv*)
-lemma univ_subset_quniv: "univ(A) <= quniv(A)"
+lemma univ_subset_quniv: "univ(A) \<subseteq> quniv(A)"
apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans])
apply (rule univ_eclose_subset_quniv)
done
lemmas univ_into_quniv = univ_subset_quniv [THEN subsetD]
-lemma Pow_univ_subset_quniv: "Pow(univ(A)) <= quniv(A)"
+lemma Pow_univ_subset_quniv: "Pow(univ(A)) \<subseteq> quniv(A)"
apply (unfold quniv_def)
apply (rule arg_subset_eclose [THEN univ_mono, THEN Pow_mono])
done
@@ -85,24 +85,24 @@
(*** univ(A) closure for Quine-inspired pairs and injections ***)
(*Quine ordered pairs*)
-lemma QPair_subset_univ:
- "[| a <= univ(A); b <= univ(A) |] ==> <a;b> <= univ(A)"
+lemma QPair_subset_univ:
+ "[| a \<subseteq> univ(A); b \<subseteq> univ(A) |] ==> <a;b> \<subseteq> univ(A)"
by (simp add: QPair_def sum_subset_univ)
subsection{*Quine Disjoint Sum*}
-lemma QInl_subset_univ: "a <= univ(A) ==> QInl(a) <= univ(A)"
+lemma QInl_subset_univ: "a \<subseteq> univ(A) ==> QInl(a) \<subseteq> univ(A)"
apply (unfold QInl_def)
apply (erule empty_subsetI [THEN QPair_subset_univ])
done
-lemmas naturals_subset_nat =
+lemmas naturals_subset_nat =
Ord_nat [THEN Ord_is_Transset, unfolded Transset_def, THEN bspec]
lemmas naturals_subset_univ =
subset_trans [OF naturals_subset_nat nat_subset_univ]
-lemma QInr_subset_univ: "a <= univ(A) ==> QInr(a) <= univ(A)"
+lemma QInr_subset_univ: "a \<subseteq> univ(A) ==> QInr(a) \<subseteq> univ(A)"
apply (unfold QInr_def)
apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ])
done
@@ -110,39 +110,39 @@
subsection{*Closure for Quine-Inspired Products and Sums*}
(*Quine ordered pairs*)
-lemma QPair_in_quniv:
- "[| a: quniv(A); b: quniv(A) |] ==> <a;b> : quniv(A)"
-by (simp add: quniv_def QPair_def sum_subset_univ)
+lemma QPair_in_quniv:
+ "[| a: quniv(A); b: quniv(A) |] ==> <a;b> \<in> quniv(A)"
+by (simp add: quniv_def QPair_def sum_subset_univ)
-lemma QSigma_quniv: "quniv(A) <*> quniv(A) <= quniv(A)"
+lemma QSigma_quniv: "quniv(A) <*> quniv(A) \<subseteq> quniv(A)"
by (blast intro: QPair_in_quniv)
lemmas QSigma_subset_quniv = subset_trans [OF QSigma_mono QSigma_quniv]
(*The opposite inclusion*)
-lemma quniv_QPair_D:
- "<a;b> : quniv(A) ==> a: quniv(A) & b: quniv(A)"
+lemma quniv_QPair_D:
+ "<a;b> \<in> quniv(A) ==> a: quniv(A) & b: quniv(A)"
apply (unfold quniv_def QPair_def)
-apply (rule Transset_includes_summands [THEN conjE])
+apply (rule Transset_includes_summands [THEN conjE])
apply (rule Transset_eclose [THEN Transset_univ])
-apply (erule PowD, blast)
+apply (erule PowD, blast)
done
lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE]
-lemma quniv_QPair_iff: "<a;b> : quniv(A) <-> a: quniv(A) & b: quniv(A)"
+lemma quniv_QPair_iff: "<a;b> \<in> quniv(A) <-> a: quniv(A) & b: quniv(A)"
by (blast intro: QPair_in_quniv dest: quniv_QPair_D)
subsection{*Quine Disjoint Sum*}
-lemma QInl_in_quniv: "a: quniv(A) ==> QInl(a) : quniv(A)"
+lemma QInl_in_quniv: "a: quniv(A) ==> QInl(a) \<in> quniv(A)"
by (simp add: QInl_def zero_in_quniv QPair_in_quniv)
-lemma QInr_in_quniv: "b: quniv(A) ==> QInr(b) : quniv(A)"
+lemma QInr_in_quniv: "b: quniv(A) ==> QInr(b) \<in> quniv(A)"
by (simp add: QInr_def one_in_quniv QPair_in_quniv)
-lemma qsum_quniv: "quniv(C) <+> quniv(C) <= quniv(C)"
+lemma qsum_quniv: "quniv(C) <+> quniv(C) \<subseteq> quniv(C)"
by (blast intro: QInl_in_quniv QInr_in_quniv)
lemmas qsum_subset_quniv = subset_trans [OF qsum_mono qsum_quniv]
@@ -162,9 +162,9 @@
(*Intersecting <a;b> with Vfrom...*)
-lemma QPair_Int_Vfrom_succ_subset:
- "Transset(X) ==>
- <a;b> Int Vfrom(X, succ(i)) <= <a Int Vfrom(X,i); b Int Vfrom(X,i)>"
+lemma QPair_Int_Vfrom_succ_subset:
+ "Transset(X) ==>
+ <a;b> \<inter> Vfrom(X, succ(i)) \<subseteq> <a \<inter> Vfrom(X,i); b \<inter> Vfrom(X,i)>"
by (simp add: QPair_def sum_def Int_Un_distrib2 Un_mono
product_Int_Vfrom_subset [THEN subset_trans]
Sigma_mono [OF Int_lower1 subset_refl])
@@ -175,28 +175,28 @@
(*Rule for level i -- preserving the level, not decreasing it*)
-lemma QPair_Int_Vfrom_subset:
- "Transset(X) ==>
- <a;b> Int Vfrom(X,i) <= <a Int Vfrom(X,i); b Int Vfrom(X,i)>"
+lemma QPair_Int_Vfrom_subset:
+ "Transset(X) ==>
+ <a;b> \<inter> Vfrom(X,i) \<subseteq> <a \<inter> Vfrom(X,i); b \<inter> Vfrom(X,i)>"
apply (unfold QPair_def)
apply (erule Transset_Vfrom [THEN Transset_sum_Int_subset])
done
-(*[| a Int Vset(i) <= c; b Int Vset(i) <= d |] ==> <a;b> Int Vset(i) <= <c;d>*)
+(*@{term"[| a \<inter> Vset(i) \<subseteq> c; b \<inter> Vset(i) \<subseteq> d |] ==> <a;b> \<inter> Vset(i) \<subseteq> <c;d>"}*)
lemmas QPair_Int_Vset_subset_trans =
subset_trans [OF Transset_0 [THEN QPair_Int_Vfrom_subset] QPair_mono]
lemma QPair_Int_Vset_subset_UN:
- "Ord(i) ==> <a;b> Int Vset(i) <= (\<Union>j\<in>i. <a Int Vset(j); b Int Vset(j)>)"
+ "Ord(i) ==> <a;b> \<inter> Vset(i) \<subseteq> (\<Union>j\<in>i. <a \<inter> Vset(j); b \<inter> Vset(j)>)"
apply (erule Ord_cases)
(*0 case*)
apply (simp add: Vfrom_0)
(*succ(j) case*)
-apply (erule ssubst)
+apply (erule ssubst)
apply (rule Transset_0 [THEN QPair_Int_Vfrom_succ_subset, THEN subset_trans])
apply (rule succI1 [THEN UN_upper])
(*Limit(i) case*)
-apply (simp del: UN_simps
+apply (simp del: UN_simps
add: Limit_Vfrom_eq Int_UN_distrib UN_mono QPair_Int_Vset_subset_trans)
done
--- a/src/ZF/Sum.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Sum.thy Tue Mar 06 15:15:49 2012 +0000
@@ -10,7 +10,7 @@
text{*And the "Part" primitive for simultaneous recursive type definitions*}
definition sum :: "[i,i]=>i" (infixr "+" 65) where
- "A+B == {0}*A Un {1}*B"
+ "A+B == {0}*A \<union> {1}*B"
definition Inl :: "i=>i" where
"Inl(a) == <0,a>"
@@ -23,29 +23,29 @@
(*operator for selecting out the various summands*)
definition Part :: "[i,i=>i] => i" where
- "Part(A,h) == {x: A. EX z. x = h(z)}"
+ "Part(A,h) == {x: A. \<exists>z. x = h(z)}"
subsection{*Rules for the @{term Part} Primitive*}
lemma Part_iff:
- "a : Part(A,h) <-> a:A & (EX y. a=h(y))"
+ "a \<in> Part(A,h) <-> a:A & (\<exists>y. a=h(y))"
apply (unfold Part_def)
apply (rule separation)
done
lemma Part_eqI [intro]:
- "[| a : A; a=h(b) |] ==> a : Part(A,h)"
+ "[| a \<in> A; a=h(b) |] ==> a \<in> 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
+ "[| a \<in> Part(A,h); !!z. [| a \<in> A; a=h(z) |] ==> P
|] ==> P"
apply (unfold Part_def, blast)
done
-lemma Part_subset: "Part(A,h) <= A"
+lemma Part_subset: "Part(A,h) \<subseteq> A"
apply (unfold Part_def)
apply (rule Collect_subset)
done
@@ -60,10 +60,10 @@
(** Introduction rules for the injections **)
-lemma InlI [intro!,simp,TC]: "a : A ==> Inl(a) : A+B"
+lemma InlI [intro!,simp,TC]: "a \<in> A ==> Inl(a) \<in> A+B"
by (unfold sum_defs, blast)
-lemma InrI [intro!,simp,TC]: "b : B ==> Inr(b) : A+B"
+lemma InrI [intro!,simp,TC]: "b \<in> B ==> Inr(b) \<in> A+B"
by (unfold sum_defs, blast)
(** Elimination rules **)
@@ -106,7 +106,7 @@
lemma InrD: "Inr(b): A+B ==> b: B"
by blast
-lemma sum_iff: "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"
+lemma sum_iff: "u: A+B <-> (\<exists>x. x:A & u=Inl(x)) | (\<exists>y. y:B & u=Inr(y))"
by blast
lemma Inl_in_sum_iff [simp]: "(Inl(x) \<in> A+B) <-> (x \<in> A)";
@@ -115,7 +115,7 @@
lemma Inr_in_sum_iff [simp]: "(Inr(y) \<in> A+B) <-> (y \<in> B)";
by auto
-lemma sum_subset_iff: "A+B <= C+D <-> A<=C & B<=D"
+lemma sum_subset_iff: "A+B \<subseteq> C+D <-> A<=C & B<=D"
by blast
lemma sum_equal_iff: "A+B = C+D <-> A=C & B=D"
@@ -137,13 +137,13 @@
"[| 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)"
+ |] ==> case(c,d,u) \<in> C(u)"
by auto
lemma expand_case: "u: A+B ==>
R(case(c,d,u)) <->
- ((ALL x:A. u = Inl(x) --> R(c(x))) &
- (ALL y:B. u = Inr(y) --> R(d(y))))"
+ ((\<forall>x\<in>A. u = Inl(x) \<longrightarrow> R(c(x))) &
+ (\<forall>y\<in>B. u = Inr(y) \<longrightarrow> R(d(y))))"
by auto
lemma case_cong:
@@ -176,7 +176,7 @@
lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y: B}"
by blast
-lemma PartD1: "a : Part(A,h) ==> a : A"
+lemma PartD1: "a \<in> Part(A,h) ==> a \<in> A"
by (simp add: Part_def)
lemma Part_id: "Part(A,%x. x) = A"
@@ -185,7 +185,7 @@
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) Un Part(C,Inr) = C"
+lemma Part_sum_equality: "C \<subseteq> A+B ==> Part(C,Inl) \<union> Part(C,Inr) = C"
by blast
end
--- a/src/ZF/Trancl.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Trancl.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,36 +9,36 @@
definition
refl :: "[i,i]=>o" where
- "refl(A,r) == (ALL x: A. <x,x> : r)"
+ "refl(A,r) == (\<forall>x\<in>A. <x,x> \<in> r)"
definition
irrefl :: "[i,i]=>o" where
- "irrefl(A,r) == ALL x: A. <x,x> ~: r"
+ "irrefl(A,r) == \<forall>x\<in>A. <x,x> \<notin> r"
definition
sym :: "i=>o" where
- "sym(r) == ALL x y. <x,y>: r --> <y,x>: r"
+ "sym(r) == \<forall>x y. <x,y>: r \<longrightarrow> <y,x>: r"
definition
asym :: "i=>o" where
- "asym(r) == ALL x y. <x,y>:r --> ~ <y,x>:r"
+ "asym(r) == \<forall>x y. <x,y>:r \<longrightarrow> ~ <y,x>:r"
definition
antisym :: "i=>o" where
- "antisym(r) == ALL x y.<x,y>:r --> <y,x>:r --> x=y"
+ "antisym(r) == \<forall>x y.<x,y>:r \<longrightarrow> <y,x>:r \<longrightarrow> x=y"
definition
trans :: "i=>o" where
- "trans(r) == ALL x y z. <x,y>: r --> <y,z>: r --> <x,z>: r"
+ "trans(r) == \<forall>x y z. <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
definition
trans_on :: "[i,i]=>o" ("trans[_]'(_')") where
- "trans[A](r) == ALL x:A. ALL y:A. ALL z:A.
- <x,y>: r --> <y,z>: r --> <x,z>: r"
+ "trans[A](r) == \<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A.
+ <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
definition
rtrancl :: "i=>i" ("(_^*)" [100] 100) (*refl/transitive closure*) where
- "r^* == lfp(field(r)*field(r), %s. id(field(r)) Un (r O s))"
+ "r^* == lfp(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
definition
trancl :: "i=>i" ("(_^+)" [100] 100) (*transitive closure*) where
@@ -46,7 +46,7 @@
definition
equiv :: "[i,i]=>o" where
- "equiv(A,r) == r <= A*A & refl(A,r) & sym(r) & trans(r)"
+ "equiv(A,r) == r \<subseteq> A*A & refl(A,r) & sym(r) & trans(r)"
subsection{*General properties of relations*}
@@ -54,17 +54,17 @@
subsubsection{*irreflexivity*}
lemma irreflI:
- "[| !!x. x:A ==> <x,x> ~: r |] ==> irrefl(A,r)"
-by (simp add: irrefl_def)
+ "[| !!x. x:A ==> <x,x> \<notin> r |] ==> irrefl(A,r)"
+by (simp add: irrefl_def)
-lemma irreflE: "[| irrefl(A,r); x:A |] ==> <x,x> ~: r"
+lemma irreflE: "[| irrefl(A,r); x:A |] ==> <x,x> \<notin> r"
by (simp add: irrefl_def)
subsubsection{*symmetry*}
lemma symI:
"[| !!x y.<x,y>: r ==> <y,x>: r |] ==> sym(r)"
-by (unfold sym_def, blast)
+by (unfold sym_def, blast)
lemma symE: "[| sym(r); <x,y>: r |] ==> <y,x>: r"
by (unfold sym_def, blast)
@@ -73,7 +73,7 @@
lemma antisymI:
"[| !!x y.[| <x,y>: r; <y,x>: r |] ==> x=y |] ==> antisym(r)"
-by (simp add: antisym_def, blast)
+by (simp add: antisym_def, blast)
lemma antisymE: "[| antisym(r); <x,y>: r; <y,x>: r |] ==> x=y"
by (simp add: antisym_def, blast)
@@ -83,62 +83,62 @@
lemma transD: "[| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r"
by (unfold trans_def, blast)
-lemma trans_onD:
+lemma trans_onD:
"[| trans[A](r); <a,b>:r; <b,c>:r; a:A; b:A; c:A |] ==> <a,c>:r"
by (unfold trans_on_def, blast)
lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
by (unfold trans_def trans_on_def, blast)
-lemma trans_on_imp_trans: "[|trans[A](r); r <= A*A|] ==> trans(r)";
+lemma trans_on_imp_trans: "[|trans[A](r); r \<subseteq> A*A|] ==> trans(r)";
by (simp add: trans_on_def trans_def, blast)
subsection{*Transitive closure of a relation*}
lemma rtrancl_bnd_mono:
- "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"
+ "bnd_mono(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
by (rule bnd_monoI, blast+)
-lemma rtrancl_mono: "r<=s ==> r^* <= s^*"
+lemma rtrancl_mono: "r<=s ==> r^* \<subseteq> s^*"
apply (unfold rtrancl_def)
apply (rule lfp_mono)
apply (rule rtrancl_bnd_mono)+
-apply blast
+apply blast
done
-(* r^* = id(field(r)) Un ( r O r^* ) *)
+(* @{term"r^* = id(field(r)) \<union> ( r O r^* )"} *)
lemmas rtrancl_unfold =
rtrancl_bnd_mono [THEN rtrancl_def [THEN def_lfp_unfold]]
(** The relation rtrancl **)
-(* r^* <= field(r) * field(r) *)
+(* @{term"r^* \<subseteq> field(r) * field(r)"} *)
lemmas rtrancl_type = rtrancl_def [THEN def_lfp_subset]
lemma relation_rtrancl: "relation(r^*)"
-apply (simp add: relation_def)
-apply (blast dest: rtrancl_type [THEN subsetD])
+apply (simp add: relation_def)
+apply (blast dest: rtrancl_type [THEN subsetD])
done
(*Reflexivity of rtrancl*)
-lemma rtrancl_refl: "[| a: field(r) |] ==> <a,a> : r^*"
+lemma rtrancl_refl: "[| a: field(r) |] ==> <a,a> \<in> r^*"
apply (rule rtrancl_unfold [THEN ssubst])
apply (erule idI [THEN UnI1])
done
(*Closure under composition with r *)
-lemma rtrancl_into_rtrancl: "[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^*"
+lemma rtrancl_into_rtrancl: "[| <a,b> \<in> r^*; <b,c> \<in> r |] ==> <a,c> \<in> r^*"
apply (rule rtrancl_unfold [THEN ssubst])
apply (rule compI [THEN UnI2], assumption, assumption)
done
(*rtrancl of r contains all pairs in r *)
-lemma r_into_rtrancl: "<a,b> : r ==> <a,b> : r^*"
+lemma r_into_rtrancl: "<a,b> \<in> r ==> <a,b> \<in> r^*"
by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+)
(*The premise ensures that r consists entirely of pairs*)
-lemma r_subset_rtrancl: "relation(r) ==> r <= r^*"
+lemma r_subset_rtrancl: "relation(r) ==> r \<subseteq> r^*"
by (simp add: relation_def, blast intro: r_into_rtrancl)
lemma rtrancl_field: "field(r^*) = field(r)"
@@ -148,22 +148,22 @@
(** standard induction rule **)
lemma rtrancl_full_induct [case_names initial step, consumes 1]:
- "[| <a,b> : r^*;
- !!x. x: field(r) ==> P(<x,x>);
- !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |]
+ "[| <a,b> \<in> r^*;
+ !!x. x: field(r) ==> P(<x,x>);
+ !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |]
==> P(<a,b>)"
-by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast)
+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
caused expensive case splits!*)
lemma rtrancl_induct [case_names initial step, induct set: rtrancl]:
- "[| <a,b> : r^*;
- P(a);
- !!y z.[| <a,y> : r^*; <y,z> : r; P(y) |] ==> P(z)
+ "[| <a,b> \<in> r^*;
+ P(a);
+ !!y z.[| <a,y> \<in> r^*; <y,z> \<in> r; P(y) |] ==> P(z)
|] ==> P(b)"
(*by induction on this formula*)
-apply (subgoal_tac "ALL y. <a,b> = <a,y> --> P (y) ")
+apply (subgoal_tac "\<forall>y. <a,b> = <a,y> \<longrightarrow> P (y) ")
(*now solve first subgoal: this formula is sufficient*)
apply (erule spec [THEN mp], rule refl)
(*now do the induction*)
@@ -175,19 +175,19 @@
apply (unfold trans_def)
apply (intro allI impI)
apply (erule_tac b = z in rtrancl_induct, assumption)
-apply (blast intro: rtrancl_into_rtrancl)
+apply (blast intro: rtrancl_into_rtrancl)
done
lemmas rtrancl_trans = trans_rtrancl [THEN transD]
(*elimination of rtrancl -- by induction on a special formula*)
lemma rtranclE:
- "[| <a,b> : r^*; (a=b) ==> P;
- !!y.[| <a,y> : r^*; <y,b> : r |] ==> P |]
+ "[| <a,b> \<in> r^*; (a=b) ==> P;
+ !!y.[| <a,y> \<in> r^*; <y,b> \<in> r |] ==> P |]
==> P"
-apply (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r) ")
+apply (subgoal_tac "a = b | (\<exists>y. <a,y> \<in> r^* & <y,b> \<in> r) ")
(*see HOL/trancl*)
-apply blast
+apply blast
apply (erule rtrancl_induct, blast+)
done
@@ -207,44 +207,44 @@
(** Conversions between trancl and rtrancl **)
-lemma trancl_into_rtrancl: "<a,b> : r^+ ==> <a,b> : r^*"
+lemma trancl_into_rtrancl: "<a,b> \<in> r^+ ==> <a,b> \<in> r^*"
apply (unfold trancl_def)
apply (blast intro: rtrancl_into_rtrancl)
done
(*r^+ contains all pairs in r *)
-lemma r_into_trancl: "<a,b> : r ==> <a,b> : r^+"
+lemma r_into_trancl: "<a,b> \<in> r ==> <a,b> \<in> r^+"
apply (unfold trancl_def)
apply (blast intro!: rtrancl_refl)
done
(*The premise ensures that r consists entirely of pairs*)
-lemma r_subset_trancl: "relation(r) ==> r <= r^+"
+lemma r_subset_trancl: "relation(r) ==> r \<subseteq> r^+"
by (simp add: relation_def, blast intro: r_into_trancl)
(*intro rule by definition: from r^* and r *)
-lemma rtrancl_into_trancl1: "[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+"
+lemma rtrancl_into_trancl1: "[| <a,b> \<in> r^*; <b,c> \<in> r |] ==> <a,c> \<in> r^+"
by (unfold trancl_def, blast)
(*intro rule from r and r^* *)
lemma rtrancl_into_trancl2:
- "[| <a,b> : r; <b,c> : r^* |] ==> <a,c> : r^+"
+ "[| <a,b> \<in> r; <b,c> \<in> r^* |] ==> <a,c> \<in> r^+"
apply (erule rtrancl_induct)
apply (erule r_into_trancl)
-apply (blast intro: r_into_trancl trancl_trans)
+apply (blast intro: r_into_trancl trancl_trans)
done
(*Nice induction rule for trancl*)
lemma trancl_induct [case_names initial step, induct set: trancl]:
- "[| <a,b> : r^+;
- !!y. [| <a,y> : r |] ==> P(y);
- !!y z.[| <a,y> : r^+; <y,z> : r; P(y) |] ==> P(z)
+ "[| <a,b> \<in> r^+;
+ !!y. [| <a,y> \<in> r |] ==> P(y);
+ !!y z.[| <a,y> \<in> r^+; <y,z> \<in> r; P(y) |] ==> P(z)
|] ==> P(b)"
apply (rule compEpair)
apply (unfold trancl_def, assumption)
(*by induction on this formula*)
-apply (subgoal_tac "ALL z. <y,z> : r --> P (z) ")
+apply (subgoal_tac "\<forall>z. <y,z> \<in> r \<longrightarrow> P (z) ")
(*now solve first subgoal: this formula is sufficient*)
apply blast
apply (erule rtrancl_induct)
@@ -253,40 +253,40 @@
(*elimination of r^+ -- NOT an induction rule*)
lemma tranclE:
- "[| <a,b> : r^+;
- <a,b> : r ==> P;
- !!y.[| <a,y> : r^+; <y,b> : r |] ==> P
+ "[| <a,b> \<in> r^+;
+ <a,b> \<in> r ==> P;
+ !!y.[| <a,y> \<in> r^+; <y,b> \<in> r |] ==> P
|] ==> P"
-apply (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r) ")
-apply blast
+apply (subgoal_tac "<a,b> \<in> r | (\<exists>y. <a,y> \<in> r^+ & <y,b> \<in> r) ")
+apply blast
apply (rule compEpair)
apply (unfold trancl_def, assumption)
apply (erule rtranclE)
apply (blast intro: rtrancl_into_trancl1)+
done
-lemma trancl_type: "r^+ <= field(r)*field(r)"
+lemma trancl_type: "r^+ \<subseteq> field(r)*field(r)"
apply (unfold trancl_def)
apply (blast elim: rtrancl_type [THEN subsetD, THEN SigmaE2])
done
lemma relation_trancl: "relation(r^+)"
-apply (simp add: relation_def)
-apply (blast dest: trancl_type [THEN subsetD])
+apply (simp add: relation_def)
+apply (blast dest: trancl_type [THEN subsetD])
done
lemma trancl_subset_times: "r \<subseteq> A * A ==> r^+ \<subseteq> A * A"
by (insert trancl_type [of r], blast)
-lemma trancl_mono: "r<=s ==> r^+ <= s^+"
+lemma trancl_mono: "r<=s ==> r^+ \<subseteq> s^+"
by (unfold trancl_def, intro comp_mono rtrancl_mono)
lemma trancl_eq_r: "[|relation(r); trans(r)|] ==> r^+ = r"
apply (rule equalityI)
- prefer 2 apply (erule r_subset_trancl, clarify)
-apply (frule trancl_type [THEN subsetD], clarify)
+ prefer 2 apply (erule r_subset_trancl, clarify)
+apply (frule trancl_type [THEN subsetD], clarify)
apply (erule trancl_induct, assumption)
-apply (blast dest: transD)
+apply (blast dest: transD)
done
@@ -296,21 +296,21 @@
apply (rule equalityI, auto)
prefer 2
apply (frule rtrancl_type [THEN subsetD])
- apply (blast intro: r_into_rtrancl )
+ apply (blast intro: r_into_rtrancl )
txt{*converse direction*}
-apply (frule rtrancl_type [THEN subsetD], clarify)
+apply (frule rtrancl_type [THEN subsetD], clarify)
apply (erule rtrancl_induct)
apply (simp add: rtrancl_refl rtrancl_field)
apply (blast intro: rtrancl_trans)
done
-lemma rtrancl_subset: "[| R <= S; S <= R^* |] ==> S^* = R^*"
+lemma rtrancl_subset: "[| R \<subseteq> S; S \<subseteq> R^* |] ==> S^* = R^*"
apply (drule rtrancl_mono)
apply (drule rtrancl_mono, simp_all, blast)
done
lemma rtrancl_Un_rtrancl:
- "[| relation(r); relation(s) |] ==> (r^* Un s^*)^* = (r Un s)^*"
+ "[| relation(r); relation(s) |] ==> (r^* \<union> s^*)^* = (r \<union> s)^*"
apply (rule rtrancl_subset)
apply (blast dest: r_subset_rtrancl)
apply (blast intro: rtrancl_mono [THEN subsetD])
@@ -362,8 +362,8 @@
done
lemma converse_trancl_induct [case_names initial step, consumes 1]:
-"[| <a, b>:r^+; !!y. <y, b> :r ==> P(y);
- !!y z. [| <y, z> : r; <z, b> : r^+; P(z) |] ==> P(y) |]
+"[| <a, b>:r^+; !!y. <y, b> :r ==> P(y);
+ !!y z. [| <y, z> \<in> r; <z, b> \<in> r^+; P(z) |] ==> P(y) |]
==> P(a)"
apply (drule converseI)
apply (simp (no_asm_use) add: trancl_converse [symmetric])
--- a/src/ZF/Univ.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Univ.thy Tue Mar 06 15:15:49 2012 +0000
@@ -15,7 +15,7 @@
definition
Vfrom :: "[i,i]=>i" where
- "Vfrom(A,i) == transrec(i, %x f. A Un (\<Union>y\<in>x. Pow(f`y)))"
+ "Vfrom(A,i) == transrec(i, %x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))"
abbreviation
Vset :: "i=>i" where
@@ -24,13 +24,13 @@
definition
Vrec :: "[i, [i,i]=>i] =>i" where
- "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
- H(z, lam w:Vset(x). g`rank(w)`w)) ` a"
+ "Vrec(a,H) == transrec(rank(a), %x g. \<lambda>z\<in>Vset(succ(x)).
+ H(z, \<lambda>w\<in>Vset(x). g`rank(w)`w)) ` a"
definition
Vrecursor :: "[[i,i]=>i, i] =>i" where
- "Vrecursor(H,a) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
- H(lam w:Vset(x). g`rank(w)`w, z)) ` a"
+ "Vrecursor(H,a) == transrec(rank(a), %x g. \<lambda>z\<in>Vset(succ(x)).
+ H(\<lambda>w\<in>Vset(x). g`rank(w)`w, z)) ` a"
definition
univ :: "i=>i" where
@@ -40,30 +40,30 @@
subsection{*Immediate Consequences of the Definition of @{term "Vfrom(A,i)"}*}
text{*NOT SUITABLE FOR REWRITING -- RECURSIVE!*}
-lemma Vfrom: "Vfrom(A,i) = A Un (\<Union>j\<in>i. Pow(Vfrom(A,j)))"
+lemma Vfrom: "Vfrom(A,i) = A \<union> (\<Union>j\<in>i. Pow(Vfrom(A,j)))"
by (subst Vfrom_def [THEN def_transrec], simp)
subsubsection{* Monotonicity *}
lemma Vfrom_mono [rule_format]:
- "A<=B ==> \<forall>j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"
+ "A<=B ==> \<forall>j. i<=j \<longrightarrow> Vfrom(A,i) \<subseteq> Vfrom(B,j)"
apply (rule_tac a=i in eps_induct)
apply (rule impI [THEN allI])
apply (subst Vfrom [of A])
apply (subst Vfrom [of B])
apply (erule Un_mono)
-apply (erule UN_mono, blast)
+apply (erule UN_mono, blast)
done
lemma VfromI: "[| a \<in> Vfrom(A,j); j<i |] ==> a \<in> Vfrom(A,i)"
-by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]])
+by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]])
subsubsection{* A fundamental equality: Vfrom does not require ordinals! *}
-lemma Vfrom_rank_subset1: "Vfrom(A,x) <= Vfrom(A,rank(x))"
+lemma Vfrom_rank_subset1: "Vfrom(A,x) \<subseteq> Vfrom(A,rank(x))"
proof (induct x rule: eps_induct)
fix x
assume "\<forall>y\<in>x. Vfrom(A,y) \<subseteq> Vfrom(A,rank(y))"
@@ -72,7 +72,7 @@
blast intro!: rank_lt [THEN ltD])
qed
-lemma Vfrom_rank_subset2: "Vfrom(A,rank(x)) <= Vfrom(A,x)"
+lemma Vfrom_rank_subset2: "Vfrom(A,rank(x)) \<subseteq> Vfrom(A,x)"
apply (rule_tac a=x in eps_induct)
apply (subst Vfrom)
apply (subst Vfrom, rule subset_refl [THEN Un_mono])
@@ -99,19 +99,19 @@
lemma zero_in_Vfrom: "y:x ==> 0 \<in> Vfrom(A,x)"
by (subst Vfrom, blast)
-lemma i_subset_Vfrom: "i <= Vfrom(A,i)"
+lemma i_subset_Vfrom: "i \<subseteq> Vfrom(A,i)"
apply (rule_tac a=i in eps_induct)
apply (subst Vfrom, blast)
done
-lemma A_subset_Vfrom: "A <= Vfrom(A,i)"
+lemma A_subset_Vfrom: "A \<subseteq> Vfrom(A,i)"
apply (subst Vfrom)
apply (rule Un_upper1)
done
lemmas A_into_Vfrom = A_subset_Vfrom [THEN subsetD]
-lemma subset_mem_Vfrom: "a <= Vfrom(A,i) ==> a \<in> Vfrom(A,succ(i))"
+lemma subset_mem_Vfrom: "a \<subseteq> Vfrom(A,i) ==> a \<in> Vfrom(A,succ(i))"
by (subst Vfrom, blast)
subsubsection{* Finite sets and ordered pairs *}
@@ -126,13 +126,13 @@
lemma Pair_in_Vfrom:
"[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i) |] ==> <a,b> \<in> Vfrom(A,succ(succ(i)))"
apply (unfold Pair_def)
-apply (blast intro: doubleton_in_Vfrom)
+apply (blast intro: doubleton_in_Vfrom)
done
-lemma succ_in_Vfrom: "a <= Vfrom(A,i) ==> succ(a) \<in> Vfrom(A,succ(succ(i)))"
+lemma succ_in_Vfrom: "a \<subseteq> Vfrom(A,i) ==> succ(a) \<in> Vfrom(A,succ(succ(i)))"
apply (intro subset_mem_Vfrom succ_subsetI, assumption)
-apply (erule subset_trans)
-apply (rule Vfrom_mono [OF subset_refl subset_succI])
+apply (erule subset_trans)
+apply (rule Vfrom_mono [OF subset_refl subset_succI])
done
subsection{* 0, Successor and Limit Equations for @{term Vfrom} *}
@@ -140,9 +140,9 @@
lemma Vfrom_0: "Vfrom(A,0) = A"
by (subst Vfrom, blast)
-lemma Vfrom_succ_lemma: "Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"
+lemma Vfrom_succ_lemma: "Ord(i) ==> Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))"
apply (rule Vfrom [THEN trans])
-apply (rule equalityI [THEN subst_context,
+apply (rule equalityI [THEN subst_context,
OF _ succI1 [THEN RepFunI, THEN Union_upper]])
apply (rule UN_least)
apply (rule subset_refl [THEN Vfrom_mono, THEN Pow_mono])
@@ -150,7 +150,7 @@
apply (erule Ord_succ)
done
-lemma Vfrom_succ: "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"
+lemma Vfrom_succ: "Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))"
apply (rule_tac x1 = "succ (i)" in Vfrom_rank_eq [THEN subst])
apply (rule_tac x1 = i in Vfrom_rank_eq [THEN subst])
apply (subst rank_succ)
@@ -158,8 +158,8 @@
done
(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces
- the conclusion to be Vfrom(A,Union(X)) = A Un (\<Union>y\<in>X. Vfrom(A,y)) *)
-lemma Vfrom_Union: "y:X ==> Vfrom(A,Union(X)) = (\<Union>y\<in>X. Vfrom(A,y))"
+ the conclusion to be Vfrom(A,\<Union>(X)) = A \<union> (\<Union>y\<in>X. Vfrom(A,y)) *)
+lemma Vfrom_Union: "y:X ==> Vfrom(A,\<Union>(X)) = (\<Union>y\<in>X. Vfrom(A,y))"
apply (subst Vfrom)
apply (rule equalityI)
txt{*first inclusion*}
@@ -179,11 +179,11 @@
subsection{* @{term Vfrom} applied to Limit Ordinals *}
(*NB. limit ordinals are non-empty:
- Vfrom(A,0) = A = A Un (\<Union>y\<in>0. Vfrom(A,y)) *)
+ Vfrom(A,0) = A = A \<union> (\<Union>y\<in>0. Vfrom(A,y)) *)
lemma Limit_Vfrom_eq:
"Limit(i) ==> Vfrom(A,i) = (\<Union>y\<in>i. Vfrom(A,y))"
apply (rule Limit_has_0 [THEN ltD, THEN Vfrom_Union, THEN subst], assumption)
-apply (simp add: Limit_Union_eq)
+apply (simp add: Limit_Union_eq)
done
lemma Limit_VfromE:
@@ -193,7 +193,7 @@
apply (rule classical)
apply (rule Limit_Vfrom_eq [THEN equalityD1, THEN subsetD, THEN UN_E])
prefer 2 apply assumption
- apply blast
+ apply blast
apply (blast intro: ltI Limit_is_Ord)
done
@@ -201,12 +201,12 @@
"[| a \<in> Vfrom(A,i); Limit(i) |] ==> {a} \<in> Vfrom(A,i)"
apply (erule Limit_VfromE, assumption)
apply (erule singleton_in_Vfrom [THEN VfromI])
-apply (blast intro: Limit_has_succ)
+apply (blast intro: Limit_has_succ)
done
-lemmas Vfrom_UnI1 =
+lemmas Vfrom_UnI1 =
Un_upper1 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD]]
-lemmas Vfrom_UnI2 =
+lemmas Vfrom_UnI2 =
Un_upper2 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD]]
text{*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*}
@@ -223,12 +223,12 @@
txt{*Infer that a, b occur at ordinals x,xa < i.*}
apply (erule Limit_VfromE, assumption)
apply (erule Limit_VfromE, assumption)
-txt{*Infer that succ(succ(x Un xa)) < i *}
+txt{*Infer that @{term"succ(succ(x \<union> xa)) < i"} *}
apply (blast intro: VfromI [OF Pair_in_Vfrom]
Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt)
done
-lemma product_VLimit: "Limit(i) ==> Vfrom(A,i) * Vfrom(A,i) <= Vfrom(A,i)"
+lemma product_VLimit: "Limit(i) ==> Vfrom(A,i) * Vfrom(A,i) \<subseteq> Vfrom(A,i)"
by (blast intro: Pair_in_VLimit)
lemmas Sigma_subset_VLimit =
@@ -259,7 +259,7 @@
apply (blast intro: one_in_VLimit Pair_in_VLimit)
done
-lemma sum_VLimit: "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)"
+lemma sum_VLimit: "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) \<subseteq> Vfrom(C,i)"
by (blast intro!: Inl_in_VLimit Inr_in_VLimit)
lemmas sum_subset_VLimit = subset_trans [OF sum_mono sum_VLimit]
@@ -283,11 +283,11 @@
apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]])
done
-lemma Transset_Pair_subset: "[| <a,b> <= C; Transset(C) |] ==> a: C & b: C"
+lemma Transset_Pair_subset: "[| <a,b> \<subseteq> C; Transset(C) |] ==> a: C & b: C"
by (unfold Pair_def Transset_def, blast)
lemma Transset_Pair_subset_VLimit:
- "[| <a,b> <= Vfrom(A,i); Transset(A); Limit(i) |]
+ "[| <a,b> \<subseteq> Vfrom(A,i); Transset(A); Limit(i) |]
==> <a,b> \<in> Vfrom(A,i)"
apply (erule Transset_Pair_subset [THEN conjE])
apply (erule Transset_Vfrom)
@@ -295,14 +295,14 @@
done
lemma Union_in_Vfrom:
- "[| X \<in> Vfrom(A,j); Transset(A) |] ==> Union(X) \<in> Vfrom(A, succ(j))"
+ "[| X \<in> Vfrom(A,j); Transset(A) |] ==> \<Union>(X) \<in> Vfrom(A, succ(j))"
apply (drule Transset_Vfrom)
apply (rule subset_mem_Vfrom)
apply (unfold Transset_def, blast)
done
lemma Union_in_VLimit:
- "[| X \<in> Vfrom(A,i); Limit(i); Transset(A) |] ==> Union(X) \<in> Vfrom(A,i)"
+ "[| X \<in> Vfrom(A,i); Limit(i); Transset(A) |] ==> \<Union>(X) \<in> Vfrom(A,i)"
apply (rule Limit_VfromE, assumption+)
apply (blast intro: Limit_has_succ VfromI Union_in_Vfrom)
done
@@ -317,15 +317,15 @@
lemma in_VLimit:
"[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i);
!!x y j. [| j<i; 1:j; x \<in> Vfrom(A,j); y \<in> Vfrom(A,j) |]
- ==> EX k. h(x,y) \<in> Vfrom(A,k) & k<i |]
+ ==> \<exists>k. h(x,y) \<in> Vfrom(A,k) & k<i |]
==> h(a,b) \<in> Vfrom(A,i)"
txt{*Infer that a, b occur at ordinals x,xa < i.*}
apply (erule Limit_VfromE, assumption)
apply (erule Limit_VfromE, assumption, atomize)
-apply (drule_tac x=a in spec)
-apply (drule_tac x=b in spec)
-apply (drule_tac x="x Un xa Un 2" in spec)
-apply (simp add: Un_least_lt_iff lt_Ord Vfrom_UnI1 Vfrom_UnI2)
+apply (drule_tac x=a in spec)
+apply (drule_tac x=b in spec)
+apply (drule_tac x="x \<union> xa \<union> 2" in spec)
+apply (simp add: Un_least_lt_iff lt_Ord Vfrom_UnI1 Vfrom_UnI2)
apply (blast intro: Limit_has_0 Limit_has_succ VfromI)
done
@@ -414,15 +414,15 @@
subsubsection{* Characterisation of the elements of @{term "Vset(i)"} *}
-lemma VsetD [rule_format]: "Ord(i) ==> \<forall>b. b \<in> Vset(i) --> rank(b) < i"
+lemma VsetD [rule_format]: "Ord(i) ==> \<forall>b. b \<in> Vset(i) \<longrightarrow> rank(b) < i"
apply (erule trans_induct)
apply (subst Vset, safe)
apply (subst rank)
-apply (blast intro: ltI UN_succ_least_lt)
+apply (blast intro: ltI UN_succ_least_lt)
done
lemma VsetI_lemma [rule_format]:
- "Ord(i) ==> \<forall>b. rank(b) \<in> i --> b \<in> Vset(i)"
+ "Ord(i) ==> \<forall>b. rank(b) \<in> i \<longrightarrow> b \<in> Vset(i)"
apply (erule trans_induct)
apply (rule allI)
apply (subst Vset)
@@ -447,30 +447,30 @@
lemma rank_Vset: "Ord(i) ==> rank(Vset(i)) = i"
apply (subst rank)
apply (rule equalityI, safe)
-apply (blast intro: VsetD [THEN ltD])
-apply (blast intro: VsetD [THEN ltD] Ord_trans)
+apply (blast intro: VsetD [THEN ltD])
+apply (blast intro: VsetD [THEN ltD] Ord_trans)
apply (blast intro: i_subset_Vfrom [THEN subsetD]
Ord_in_Ord [THEN rank_of_Ord, THEN ssubst])
done
lemma Finite_Vset: "i \<in> nat ==> Finite(Vset(i))";
apply (erule nat_induct)
- apply (simp add: Vfrom_0)
-apply (simp add: Vset_succ)
+ apply (simp add: Vfrom_0)
+apply (simp add: Vset_succ)
done
subsubsection{* Reasoning about Sets in Terms of Their Elements' Ranks *}
-lemma arg_subset_Vset_rank: "a <= Vset(rank(a))"
+lemma arg_subset_Vset_rank: "a \<subseteq> Vset(rank(a))"
apply (rule subsetI)
apply (erule rank_lt [THEN VsetI])
done
lemma Int_Vset_subset:
- "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"
-apply (rule subset_trans)
+ "[| !!i. Ord(i) ==> a \<inter> Vset(i) \<subseteq> b |] ==> a \<subseteq> b"
+apply (rule subset_trans)
apply (rule Int_greatest [OF subset_refl arg_subset_Vset_rank])
-apply (blast intro: Ord_rank)
+apply (blast intro: Ord_rank)
done
subsubsection{* Set Up an Environment for Simplification *}
@@ -490,7 +490,7 @@
subsubsection{* Recursion over Vset Levels! *}
text{*NOT SUITABLE FOR REWRITING: recursive!*}
-lemma Vrec: "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"
+lemma Vrec: "Vrec(a,H) = H(a, \<lambda>x\<in>Vset(rank(a)). Vrec(x,H))"
apply (unfold Vrec_def)
apply (subst transrec, simp)
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
@@ -499,14 +499,14 @@
text{*This form avoids giant explosions in proofs. NOTE USE OF == *}
lemma def_Vrec:
"[| !!x. h(x)==Vrec(x,H) |] ==>
- h(a) = H(a, lam x: Vset(rank(a)). h(x))"
-apply simp
+ h(a) = H(a, \<lambda>x\<in>Vset(rank(a)). h(x))"
+apply simp
apply (rule Vrec)
done
text{*NOT SUITABLE FOR REWRITING: recursive!*}
lemma Vrecursor:
- "Vrecursor(H,a) = H(lam x:Vset(rank(a)). Vrecursor(H,x), a)"
+ "Vrecursor(H,a) = H(\<lambda>x\<in>Vset(rank(a)). Vrecursor(H,x), a)"
apply (unfold Vrecursor_def)
apply (subst transrec, simp)
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
@@ -514,7 +514,7 @@
text{*This form avoids giant explosions in proofs. NOTE USE OF == *}
lemma def_Vrecursor:
- "h == Vrecursor(H) ==> h(a) = H(lam x: Vset(rank(a)). h(x), a)"
+ "h == Vrecursor(H) ==> h(a) = H(\<lambda>x\<in>Vset(rank(a)). h(x), a)"
apply simp
apply (rule Vrecursor)
done
@@ -522,7 +522,7 @@
subsection{* The Datatype Universe: @{term "univ(A)"} *}
-lemma univ_mono: "A<=B ==> univ(A) <= univ(B)"
+lemma univ_mono: "A<=B ==> univ(A) \<subseteq> univ(B)"
apply (unfold univ_def)
apply (erule Vfrom_mono)
apply (rule subset_refl)
@@ -540,28 +540,28 @@
apply (rule Limit_nat [THEN Limit_Vfrom_eq])
done
-lemma subset_univ_eq_Int: "c <= univ(A) ==> c = (\<Union>i\<in>nat. c Int Vfrom(A,i))"
+lemma subset_univ_eq_Int: "c \<subseteq> univ(A) ==> c = (\<Union>i\<in>nat. c \<inter> Vfrom(A,i))"
apply (rule subset_UN_iff_eq [THEN iffD1])
apply (erule univ_eq_UN [THEN subst])
done
lemma univ_Int_Vfrom_subset:
- "[| a <= univ(X);
- !!i. i:nat ==> a Int Vfrom(X,i) <= b |]
- ==> a <= b"
+ "[| a \<subseteq> univ(X);
+ !!i. i:nat ==> a \<inter> Vfrom(X,i) \<subseteq> b |]
+ ==> a \<subseteq> b"
apply (subst subset_univ_eq_Int, assumption)
-apply (rule UN_least, simp)
+apply (rule UN_least, simp)
done
lemma univ_Int_Vfrom_eq:
- "[| a <= univ(X); b <= univ(X);
- !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i)
+ "[| a \<subseteq> univ(X); b \<subseteq> univ(X);
+ !!i. i:nat ==> a \<inter> Vfrom(X,i) = b \<inter> Vfrom(X,i)
|] ==> a = b"
apply (rule equalityI)
apply (rule univ_Int_Vfrom_subset, assumption)
-apply (blast elim: equalityCE)
+apply (blast elim: equalityCE)
apply (rule univ_Int_Vfrom_subset, assumption)
-apply (blast elim: equalityCE)
+apply (blast elim: equalityCE)
done
subsection{* Closure Properties for @{term "univ(A)"}*}
@@ -571,10 +571,10 @@
apply (rule nat_0I [THEN zero_in_Vfrom])
done
-lemma zero_subset_univ: "{0} <= univ(A)"
+lemma zero_subset_univ: "{0} \<subseteq> univ(A)"
by (blast intro: zero_in_univ)
-lemma A_subset_univ: "A <= univ(A)"
+lemma A_subset_univ: "A \<subseteq> univ(A)"
apply (unfold univ_def)
apply (rule A_subset_Vfrom)
done
@@ -601,12 +601,12 @@
done
lemma Union_in_univ:
- "[| X: univ(A); Transset(A) |] ==> Union(X) \<in> univ(A)"
+ "[| X: univ(A); Transset(A) |] ==> \<Union>(X) \<in> univ(A)"
apply (unfold univ_def)
apply (blast intro: Union_in_VLimit Limit_nat)
done
-lemma product_univ: "univ(A)*univ(A) <= univ(A)"
+lemma product_univ: "univ(A)*univ(A) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (rule Limit_nat [THEN product_VLimit])
done
@@ -614,7 +614,7 @@
subsubsection{* The Natural Numbers *}
-lemma nat_subset_univ: "nat <= univ(A)"
+lemma nat_subset_univ: "nat \<subseteq> univ(A)"
apply (unfold univ_def)
apply (rule i_subset_Vfrom)
done
@@ -633,7 +633,7 @@
lemma two_in_univ: "2 \<in> univ(A)"
by (blast intro: nat_into_univ)
-lemma bool_subset_univ: "bool <= univ(A)"
+lemma bool_subset_univ: "bool \<subseteq> univ(A)"
apply (unfold bool_def)
apply (blast intro!: zero_in_univ one_in_univ)
done
@@ -653,7 +653,7 @@
apply (erule Inr_in_VLimit [OF _ Limit_nat])
done
-lemma sum_univ: "univ(C)+univ(C) <= univ(C)"
+lemma sum_univ: "univ(C)+univ(C) \<subseteq> univ(C)"
apply (unfold univ_def)
apply (rule Limit_nat [THEN sum_VLimit])
done
@@ -663,7 +663,7 @@
lemma Sigma_subset_univ:
"[|A \<subseteq> univ(D); \<And>x. x \<in> A \<Longrightarrow> B(x) \<subseteq> univ(D)|] ==> Sigma(A,B) \<subseteq> univ(D)"
apply (simp add: univ_def)
-apply (blast intro: Sigma_subset_VLimit del: subsetI)
+apply (blast intro: Sigma_subset_VLimit del: subsetI)
done
@@ -677,14 +677,14 @@
subsubsection{* Closure under Finite Powerset *}
lemma Fin_Vfrom_lemma:
- "[| b: Fin(Vfrom(A,i)); Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j<i"
+ "[| b: Fin(Vfrom(A,i)); Limit(i) |] ==> \<exists>j. b \<subseteq> Vfrom(A,j) & j<i"
apply (erule Fin_induct)
apply (blast dest!: Limit_has_0, safe)
apply (erule Limit_VfromE, assumption)
apply (blast intro!: Un_least_lt intro: Vfrom_UnI1 Vfrom_UnI2)
done
-lemma Fin_VLimit: "Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)"
+lemma Fin_VLimit: "Limit(i) ==> Fin(Vfrom(A,i)) \<subseteq> Vfrom(A,i)"
apply (rule subsetI)
apply (drule Fin_Vfrom_lemma, safe)
apply (rule Vfrom [THEN ssubst])
@@ -693,7 +693,7 @@
lemmas Fin_subset_VLimit = subset_trans [OF Fin_mono Fin_VLimit]
-lemma Fin_univ: "Fin(univ(A)) <= univ(A)"
+lemma Fin_univ: "Fin(univ(A)) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (rule Limit_nat [THEN Fin_VLimit])
done
@@ -701,7 +701,7 @@
subsubsection{* Closure under Finite Powers: Functions from a Natural Number *}
lemma nat_fun_VLimit:
- "[| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)"
+ "[| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
apply (erule nat_fun_subset_Fin [THEN subset_trans])
apply (blast del: subsetI
intro: subset_refl Fin_subset_VLimit Sigma_subset_VLimit nat_subset_VLimit)
@@ -709,7 +709,7 @@
lemmas nat_fun_subset_VLimit = subset_trans [OF Pi_mono nat_fun_VLimit]
-lemma nat_fun_univ: "n: nat ==> n -> univ(A) <= univ(A)"
+lemma nat_fun_univ: "n: nat ==> n -> univ(A) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (erule nat_fun_VLimit [OF _ Limit_nat])
done
@@ -719,36 +719,36 @@
text{*General but seldom-used version; normally the domain is fixed*}
lemma FiniteFun_VLimit1:
- "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) <= Vfrom(A,i)"
+ "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
apply (rule FiniteFun.dom_subset [THEN subset_trans])
apply (blast del: subsetI
intro: Fin_subset_VLimit Sigma_subset_VLimit subset_refl)
done
-lemma FiniteFun_univ1: "univ(A) -||> univ(A) <= univ(A)"
+lemma FiniteFun_univ1: "univ(A) -||> univ(A) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (rule Limit_nat [THEN FiniteFun_VLimit1])
done
text{*Version for a fixed domain*}
lemma FiniteFun_VLimit:
- "[| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)"
-apply (rule subset_trans)
+ "[| W \<subseteq> Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
+apply (rule subset_trans)
apply (erule FiniteFun_mono [OF _ subset_refl])
apply (erule FiniteFun_VLimit1)
done
lemma FiniteFun_univ:
- "W <= univ(A) ==> W -||> univ(A) <= univ(A)"
+ "W \<subseteq> univ(A) ==> W -||> univ(A) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (erule FiniteFun_VLimit [OF _ Limit_nat])
done
lemma FiniteFun_in_univ:
- "[| f: W -||> univ(A); W <= univ(A) |] ==> f \<in> univ(A)"
+ "[| f: W -||> univ(A); W \<subseteq> univ(A) |] ==> f \<in> univ(A)"
by (erule FiniteFun_univ [THEN subsetD], assumption)
-text{*Remove <= from the rule above*}
+text{*Remove @{text "\<subseteq>"} from the rule above*}
lemmas FiniteFun_in_univ' = FiniteFun_in_univ [OF _ subsetI]
@@ -760,16 +760,16 @@
lemma doubleton_in_Vfrom_D:
"[| {a,b} \<in> Vfrom(X,succ(i)); Transset(X) |]
==> a \<in> Vfrom(X,i) & b \<in> Vfrom(X,i)"
-by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD],
+by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD],
assumption, fast)
text{*This weaker version says a, b exist at the same level*}
lemmas Vfrom_doubleton_D = Transset_Vfrom [THEN Transset_doubleton_D]
-(** Using only the weaker theorem would prove <a,b> \<in> Vfrom(X,i)
- implies a, b \<in> Vfrom(X,i), which is useless for induction.
- Using only the stronger theorem would prove <a,b> \<in> Vfrom(X,succ(succ(i)))
- implies a, b \<in> Vfrom(X,i), leaving the succ(i) case untreated.
+(** Using only the weaker theorem would prove <a,b> : Vfrom(X,i)
+ implies a, b : Vfrom(X,i), which is useless for induction.
+ Using only the stronger theorem would prove <a,b> : Vfrom(X,succ(succ(i)))
+ implies a, b : Vfrom(X,i), leaving the succ(i) case untreated.
The combination gives a reduction by precisely one level, which is
most convenient for proofs.
**)
@@ -783,13 +783,13 @@
lemma product_Int_Vfrom_subset:
"Transset(X) ==>
- (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"
+ (a*b) \<inter> Vfrom(X, succ(i)) \<subseteq> (a \<inter> Vfrom(X,i)) * (b \<inter> Vfrom(X,i))"
by (blast dest!: Pair_in_Vfrom_D)
ML
{*
-val rank_ss = @{simpset} addsimps [@{thm VsetI}]
+val rank_ss = @{simpset} addsimps [@{thm VsetI}]
addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}]));
*}
--- a/src/ZF/WF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/WF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -21,16 +21,16 @@
definition
wf :: "i=>o" where
(*r is a well-founded relation*)
- "wf(r) == ALL Z. Z=0 | (EX x:Z. ALL y. <y,x>:r --> ~ y:Z)"
+ "wf(r) == \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. <y,x>:r \<longrightarrow> ~ y:Z)"
definition
wf_on :: "[i,i]=>o" ("wf[_]'(_')") where
(*r is well-founded on A*)
- "wf_on(A,r) == wf(r Int A*A)"
+ "wf_on(A,r) == wf(r \<inter> A*A)"
definition
is_recfun :: "[i, i, [i,i]=>i, i] =>o" where
- "is_recfun(r,a,H,f) == (f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))))"
+ "is_recfun(r,a,H,f) == (f = (\<lambda>x\<in>r-``{a}. H(x, restrict(f, r-``{x}))))"
definition
the_recfun :: "[i, i, [i,i]=>i] =>i" where
@@ -47,7 +47,7 @@
definition
wfrec_on :: "[i, i, i, [i,i]=>i] =>i" ("wfrec[_]'(_,_,_')") where
- "wfrec[A](r,a,H) == wfrec(r Int A*A, a, H)"
+ "wfrec[A](r,a,H) == wfrec(r \<inter> A*A, a, H)"
subsection{*Well-Founded Relations*}
@@ -55,9 +55,9 @@
subsubsection{*Equivalences between @{term wf} and @{term wf_on}*}
lemma wf_imp_wf_on: "wf(r) ==> wf[A](r)"
-by (unfold wf_def wf_on_def, force)
+by (unfold wf_def wf_on_def, force)
-lemma wf_on_imp_wf: "[|wf[A](r); r <= A*A|] ==> wf(r)";
+lemma wf_on_imp_wf: "[|wf[A](r); r \<subseteq> A*A|] ==> wf(r)";
by (simp add: wf_on_def subset_Int_iff)
lemma wf_on_field_imp_wf: "wf[field(r)](r) ==> wf(r)"
@@ -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; ALL x:Z. EX y:Z. <y,x>:r |] ==> False"
+ assumes prem: "!!Z u. [| Z<=A; u:Z; \<forall>x\<in>Z. \<exists>y\<in>Z. <y,x>:r |] ==> False"
shows "wf[A](r)"
apply (unfold wf_on_def wf_def)
apply (rule equals0I [THEN disjCI, THEN allI])
@@ -89,14 +89,14 @@
text{*If @{term r} allows well-founded induction over @{term A}
then we have @{term "wf[A](r)"}. Premise is equivalent to
- @{prop "!!B. ALL x:A. (ALL y. <y,x>: r --> y:B) --> x:B ==> A<=B"} *}
+ @{prop "!!B. \<forall>x\<in>A. (\<forall>y. <y,x>: r \<longrightarrow> y:B) \<longrightarrow> x:B ==> A<=B"} *}
lemma wf_onI2:
- assumes prem: "!!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A |]
+ assumes prem: "!!y B. [| \<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y:B) \<longrightarrow> x:B; y:A |]
==> y:B"
shows "wf[A](r)"
apply (rule wf_onI)
apply (rule_tac c=u in prem [THEN DiffE])
- prefer 3 apply blast
+ prefer 3 apply blast
apply fast+
done
@@ -107,11 +107,11 @@
@{term "P(z)"} does not hold...*}
lemma wf_induct [induct set: wf]:
"[| wf(r);
- !!x.[| ALL y. <y,x>: r --> P(y) |] ==> P(x) |]
+ !!x.[| \<forall>y. <y,x>: r \<longrightarrow> P(y) |] ==> P(x) |]
==> P(a)"
-apply (unfold wf_def)
-apply (erule_tac x = "{z : domain(r). ~ P(z)}" in allE)
-apply blast
+apply (unfold wf_def)
+apply (erule_tac x = "{z \<in> domain(r). ~ P(z)}" in allE)
+apply blast
done
lemmas wf_induct_rule = wf_induct [rule_format, induct set: wf]
@@ -119,20 +119,20 @@
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; ALL y. <y,x>: r --> P(y) |] ==> P(x) |]
+ !!x.[| x: A; \<forall>y. <y,x>: r \<longrightarrow> P(y) |] ==> P(x) |]
==> P(a)"
apply (erule_tac P="a:A" in rev_mp)
-apply (erule_tac a=a in wf_induct, blast)
+apply (erule_tac a=a in wf_induct, blast)
done
-lemma field_Int_square: "field(r Int A*A) <= A"
+lemma field_Int_square: "field(r \<inter> A*A) \<subseteq> A"
by blast
lemma wf_on_induct [consumes 2, induct set: wf_on]:
"[| wf[A](r); a:A;
- !!x.[| x: A; ALL y:A. <y,x>: r --> P(y) |] ==> P(x)
+ !!x.[| x: A; \<forall>y\<in>A. <y,x>: r \<longrightarrow> P(y) |] ==> P(x)
|] ==> P(a)"
-apply (unfold wf_on_def)
+apply (unfold wf_on_def)
apply (erule wf_induct2, assumption)
apply (rule field_Int_square, blast)
done
@@ -141,71 +141,71 @@
wf_on_induct [rule_format, consumes 2, induct set: wf_on]
-text{*If @{term r} allows well-founded induction
+text{*If @{term r} allows well-founded induction
then we have @{term "wf(r)"}.*}
lemma wfI:
"[| field(r)<=A;
- !!y B. [| ALL x:A. (ALL y:A. <y,x>:r --> y:B) --> x:B; y:A|]
+ !!y B. [| \<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y:B) \<longrightarrow> x:B; y:A|]
==> y:B |]
==> wf(r)"
apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf])
apply (rule wf_onI2)
- prefer 2 apply blast
-apply blast
+ prefer 2 apply blast
+apply blast
done
subsection{*Basic Properties of Well-Founded Relations*}
-lemma wf_not_refl: "wf(r) ==> <a,a> ~: r"
+lemma wf_not_refl: "wf(r) ==> <a,a> \<notin> r"
by (erule_tac a=a in wf_induct, blast)
-lemma wf_not_sym [rule_format]: "wf(r) ==> ALL x. <a,x>:r --> <x,a> ~: r"
+lemma wf_not_sym [rule_format]: "wf(r) ==> \<forall>x. <a,x>:r \<longrightarrow> <x,a> \<notin> r"
by (erule_tac a=a in wf_induct, blast)
-(* [| wf(r); <a,x> : r; ~P ==> <x,a> : r |] ==> P *)
+(* @{term"[| wf(r); <a,x> \<in> r; ~P ==> <x,a> \<in> r |] ==> P"} *)
lemmas wf_asym = wf_not_sym [THEN swap]
-lemma wf_on_not_refl: "[| wf[A](r); a: A |] ==> <a,a> ~: r"
+lemma wf_on_not_refl: "[| wf[A](r); a: A |] ==> <a,a> \<notin> r"
by (erule_tac a=a in wf_on_induct, assumption, blast)
lemma wf_on_not_sym [rule_format]:
- "[| wf[A](r); a:A |] ==> ALL b:A. <a,b>:r --> <b,a>~:r"
+ "[| wf[A](r); a:A |] ==> \<forall>b\<in>A. <a,b>:r \<longrightarrow> <b,a>\<notin>r"
apply (erule_tac a=a in wf_on_induct, assumption, blast)
done
lemma wf_on_asym:
- "[| wf[A](r); ~Z ==> <a,b> : r;
- <b,a> ~: r ==> Z; ~Z ==> a : A; ~Z ==> b : A |] ==> Z"
-by (blast dest: wf_on_not_sym)
+ "[| wf[A](r); ~Z ==> <a,b> \<in> r;
+ <b,a> \<notin> r ==> Z; ~Z ==> a \<in> A; ~Z ==> b \<in> A |] ==> Z"
+by (blast dest: wf_on_not_sym)
(*Needed to prove well_ordI. Could also reason that wf[A](r) means
- wf(r Int A*A); thus wf( (r Int A*A)^+ ) and use wf_not_refl *)
+ wf(r \<inter> A*A); thus wf( (r \<inter> A*A)^+ ) and use wf_not_refl *)
lemma wf_on_chain3:
"[| wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a:A; b:A; c:A |] ==> P"
-apply (subgoal_tac "ALL y:A. ALL z:A. <a,y>:r --> <y,z>:r --> <z,a>:r --> P",
- blast)
+apply (subgoal_tac "\<forall>y\<in>A. \<forall>z\<in>A. <a,y>:r \<longrightarrow> <y,z>:r \<longrightarrow> <z,a>:r \<longrightarrow> P",
+ blast)
apply (erule_tac a=a in wf_on_induct, assumption, blast)
done
-text{*transitive closure of a WF relation is WF provided
+text{*transitive closure of a WF relation is WF provided
@{term A} is downward closed*}
lemma wf_on_trancl:
- "[| wf[A](r); r-``A <= A |] ==> wf[A](r^+)"
+ "[| wf[A](r); r-``A \<subseteq> A |] ==> wf[A](r^+)"
apply (rule wf_onI2)
apply (frule bspec [THEN mp], assumption+)
apply (erule_tac a = y in wf_on_induct, assumption)
-apply (blast elim: tranclE, blast)
+apply (blast elim: tranclE, blast)
done
lemma wf_trancl: "wf(r) ==> wf(r^+)"
apply (simp add: wf_iff_wf_on_field)
-apply (rule wf_on_subset_A)
+apply (rule wf_on_subset_A)
apply (erule wf_on_trancl)
- apply blast
+ apply blast
apply (rule trancl_type [THEN field_rel_subset])
done
@@ -228,15 +228,15 @@
lemma apply_recfun:
"[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))"
-apply (unfold is_recfun_def)
+apply (unfold is_recfun_def)
txt{*replace f only on the left-hand side*}
apply (erule_tac P = "%x.?t(x) = ?u" in ssubst)
-apply (simp add: underI)
+apply (simp add: underI)
done
lemma is_recfun_equal [rule_format]:
"[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |]
- ==> <x,a>:r --> <x,b>:r --> f`x=g`x"
+ ==> <x,a>:r \<longrightarrow> <x,b>:r \<longrightarrow> f`x=g`x"
apply (frule_tac f = f in is_recfun_type)
apply (frule_tac f = g in is_recfun_type)
apply (simp add: is_recfun_def)
@@ -245,7 +245,7 @@
apply (elim ssubst)
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
apply (rule_tac t = "%z. H (?x,z) " in subst_context)
-apply (subgoal_tac "ALL y : r-``{x}. ALL z. <y,z>:f <-> <y,z>:g")
+apply (subgoal_tac "\<forall>y\<in>r-``{x}. \<forall>z. <y,z>:f <-> <y,z>:g")
apply (blast dest: transD)
apply (simp add: apply_iff)
apply (blast dest: transD intro: sym)
@@ -283,13 +283,13 @@
lemma unfold_the_recfun:
"[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))"
apply (rule_tac a=a in wf_induct, assumption)
-apply (rename_tac a1)
-apply (rule_tac f = "lam y: r-``{a1}. wftrec (r,y,H)" in is_the_recfun)
+apply (rename_tac a1)
+apply (rule_tac f = "\<lambda>y\<in>r-``{a1}. wftrec (r,y,H)" in is_the_recfun)
apply typecheck
apply (unfold is_recfun_def wftrec_def)
--{*Applying the substitution: must keep the quantified assumption!*}
-apply (rule lam_cong [OF refl])
-apply (drule underD)
+apply (rule lam_cong [OF refl])
+apply (drule underD)
apply (fold is_recfun_def)
apply (rule_tac t = "%z. H(?x,z)" in subst_context)
apply (rule fun_extension)
@@ -298,9 +298,9 @@
apply blast
apply (blast dest: transD)
apply (frule spec [THEN mp], assumption)
-apply (subgoal_tac "<xa,a1> : r")
+apply (subgoal_tac "<xa,a1> \<in> r")
apply (drule_tac x1 = xa in spec [THEN mp], assumption)
-apply (simp add: vimage_singleton_iff
+apply (simp add: vimage_singleton_iff
apply_recfun is_recfun_cut)
apply (blast dest: transD)
done
@@ -316,7 +316,7 @@
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
lemma wftrec:
"[| wf(r); trans(r) |] ==>
- wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"
+ wftrec(r,a,H) = H(a, \<lambda>x\<in>r-``{a}. wftrec(r,x,H))"
apply (unfold wftrec_def)
apply (subst unfold_the_recfun [unfolded is_recfun_def])
apply (simp_all add: vimage_singleton_iff [THEN iff_sym] the_recfun_cut)
@@ -327,8 +327,8 @@
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
lemma wfrec:
- "wf(r) ==> wfrec(r,a,H) = H(a, lam x:r-``{a}. wfrec(r,x,H))"
-apply (unfold wfrec_def)
+ "wf(r) ==> wfrec(r,a,H) = H(a, \<lambda>x\<in>r-``{a}. wfrec(r,x,H))"
+apply (unfold wfrec_def)
apply (erule wf_trancl [THEN wftrec, THEN ssubst])
apply (rule trans_trancl)
apply (rule vimage_pair_mono [THEN restrict_lam_eq, THEN subst_context])
@@ -339,24 +339,24 @@
(*This form avoids giant explosions in proofs. NOTE USE OF == *)
lemma def_wfrec:
"[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==>
- h(a) = H(a, lam x: r-``{a}. h(x))"
+ h(a) = H(a, \<lambda>x\<in>r-``{a}. h(x))"
apply simp
-apply (elim wfrec)
+apply (elim wfrec)
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)
- |] ==> wfrec(r,a,H) : B(a)"
+ !!x u. [| x: A; u: Pi(r-``{x}, B) |] ==> H(x,u) \<in> B(x)
+ |] ==> wfrec(r,a,H) \<in> B(a)"
apply (rule_tac a = a in wf_induct2, assumption+)
apply (subst wfrec, assumption)
-apply (simp add: lam_type underD)
+apply (simp add: lam_type underD)
done
lemma wfrec_on:
"[| wf[A](r); a: A |] ==>
- wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))"
+ wfrec[A](r,a,H) = H(a, \<lambda>x\<in>(r-``{a}) \<inter> A. wfrec[A](r,x,H))"
apply (unfold wf_on_def wfrec_on_def)
apply (erule wfrec [THEN trans])
apply (simp add: vimage_Int_square cons_subset_iff)
@@ -364,7 +364,7 @@
text{*Minimal-element characterization of well-foundedness*}
lemma wf_eq_minimal:
- "wf(r) <-> (ALL Q x. x:Q --> (EX z:Q. ALL y. <y,z>:r --> y~:Q))"
+ "wf(r) <-> (\<forall>Q x. x:Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. <y,z>:r \<longrightarrow> y\<notin>Q))"
by (unfold wf_def, blast)
end
--- a/src/ZF/ZF.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/ZF.thy Tue Mar 06 15:15:49 2012 +0000
@@ -200,10 +200,10 @@
0 Pow Inf Union PrimReplace mem
defs (* Bounded Quantifiers *)
- Ball_def: "Ball(A, P) == \<forall>x. x\<in>A --> P(x)"
+ Ball_def: "Ball(A, P) == \<forall>x. x\<in>A \<longrightarrow> P(x)"
Bex_def: "Bex(A, P) == \<exists>x. x\<in>A & P(x)"
- subset_def: "A <= B == \<forall>x\<in>A. x\<in>B"
+ subset_def: "A \<subseteq> B == \<forall>x\<in>A. x\<in>B"
axiomatization where
@@ -212,18 +212,18 @@
Axioms for Union, Pow and Replace state existence only,
uniqueness is derivable using extensionality. *)
- extension: "A = B <-> A <= B & B <= A" and
- Union_iff: "A \<in> Union(C) <-> (\<exists>B\<in>C. A\<in>B)" and
- Pow_iff: "A \<in> Pow(B) <-> A <= B" and
+ extension: "A = B <-> A \<subseteq> B & B \<subseteq> A" and
+ Union_iff: "A \<in> \<Union>(C) <-> (\<exists>B\<in>C. A\<in>B)" and
+ Pow_iff: "A \<in> Pow(B) <-> A \<subseteq> B" and
(*We may name this set, though it is not uniquely defined.*)
infinity: "0\<in>Inf & (\<forall>y\<in>Inf. succ(y): Inf)" and
(*This formulation facilitates case analysis on A.*)
- foundation: "A=0 | (\<exists>x\<in>A. \<forall>y\<in>x. y~:A)" and
+ foundation: "A=0 | (\<exists>x\<in>A. \<forall>y\<in>x. y\<notin>A)" and
(*Schema axiom since predicate P is a higher-order variable*)
- replacement: "(\<forall>x\<in>A. \<forall>y z. P(x,y) & P(x,z) --> y=z) ==>
+ replacement: "(\<forall>x\<in>A. \<forall>y z. P(x,y) & P(x,z) \<longrightarrow> y=z) ==>
b \<in> PrimReplace(A,P) <-> (\<exists>x\<in>A. P(x,b))"
@@ -248,18 +248,18 @@
set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*)
Upair_def: "Upair(a,b) == {y. x\<in>Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
- cons_def: "cons(a,A) == Upair(a,a) Un A"
+ cons_def: "cons(a,A) == Upair(a,a) \<union> A"
succ_def: "succ(i) == cons(i, i)"
(* Difference, general intersection, binary union and small intersection *)
Diff_def: "A - B == { x\<in>A . ~(x\<in>B) }"
- Inter_def: "Inter(A) == { x\<in>Union(A) . \<forall>y\<in>A. x\<in>y}"
- Un_def: "A Un B == Union(Upair(A,B))"
- Int_def: "A Int B == Inter(Upair(A,B))"
+ Inter_def: "\<Inter>(A) == { x\<in>\<Union>(A) . \<forall>y\<in>A. x\<in>y}"
+ Un_def: "A \<union> B == \<Union>(Upair(A,B))"
+ Int_def: "A \<inter> B == \<Inter>(Upair(A,B))"
(* definite descriptions *)
- the_def: "The(P) == Union({y . x \<in> {0}, P(y)})"
+ the_def: "The(P) == \<Union>({y . x \<in> {0}, P(y)})"
if_def: "if(P,a,b) == THE z. P & z=a | ~P & z=b"
(* this "symmetric" definition works better than {{a}, {a,b}} *)
@@ -276,21 +276,21 @@
domain_def: "domain(r) == {x. w\<in>r, \<exists>y. w=<x,y>}"
range_def: "range(r) == domain(converse(r))"
- field_def: "field(r) == domain(r) Un range(r)"
+ field_def: "field(r) == domain(r) \<union> range(r)"
relation_def: "relation(r) == \<forall>z\<in>r. \<exists>x y. z = <x,y>"
function_def: "function(r) ==
- \<forall>x y. <x,y>:r --> (\<forall>y'. <x,y'>:r --> y=y')"
- image_def: "r `` A == {y : range(r) . \<exists>x\<in>A. <x,y> : r}"
+ \<forall>x y. <x,y>:r \<longrightarrow> (\<forall>y'. <x,y'>:r \<longrightarrow> y=y')"
+ image_def: "r `` A == {y \<in> range(r) . \<exists>x\<in>A. <x,y> \<in> r}"
vimage_def: "r -`` A == converse(r)``A"
(* Abstraction, application and Cartesian product of a family of sets *)
lam_def: "Lambda(A,b) == {<x,b(x)> . x\<in>A}"
- apply_def: "f`a == Union(f``{a})"
+ apply_def: "f`a == \<Union>(f``{a})"
Pi_def: "Pi(A,B) == {f\<in>Pow(Sigma(A,B)). A<=domain(f) & function(f)}"
(* Restrict the relation r to the domain A *)
- restrict_def: "restrict(r,A) == {z : r. \<exists>x\<in>A. \<exists>y. z = <x,y>}"
+ restrict_def: "restrict(r,A) == {z \<in> r. \<exists>x\<in>A. \<exists>y. z = <x,y>}"
subsection {* Substitution*}
@@ -311,19 +311,19 @@
by (simp add: Ball_def)
(*Instantiates x first: better for automatic theorem proving?*)
-lemma rev_ballE [elim]:
- "[| \<forall>x\<in>A. P(x); x~:A ==> Q; P(x) ==> Q |] ==> Q"
-by (simp add: Ball_def, blast)
+lemma rev_ballE [elim]:
+ "[| \<forall>x\<in>A. P(x); x\<notin>A ==> Q; P(x) ==> Q |] ==> Q"
+by (simp add: Ball_def, blast)
-lemma ballE: "[| \<forall>x\<in>A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"
+lemma ballE: "[| \<forall>x\<in>A. P(x); P(x) ==> Q; x\<notin>A ==> Q |] ==> Q"
by blast
(*Used in the datatype package*)
lemma rev_bspec: "[| x: A; \<forall>x\<in>A. P(x) |] ==> P(x)"
by (simp add: Ball_def)
-(*Trival rewrite rule; (\<forall>x\<in>A.P)<->P holds only if A is nonempty!*)
-lemma ball_triv [simp]: "(\<forall>x\<in>A. P) <-> ((\<exists>x. x\<in>A) --> P)"
+(*Trival rewrite rule; @{term"(\<forall>x\<in>A.P)<->P"} holds only if A is nonempty!*)
+lemma ball_triv [simp]: "(\<forall>x\<in>A. P) <-> ((\<exists>x. x\<in>A) \<longrightarrow> P)"
by (simp add: Ball_def)
(*Congruence rule for rewriting*)
@@ -344,23 +344,23 @@
lemma bexI [intro]: "[| P(x); x: A |] ==> \<exists>x\<in>A. P(x)"
by (simp add: Bex_def, blast)
-(*The best argument order when there is only one x\<in>A*)
+(*The best argument order when there is only one @{term"x\<in>A"}*)
lemma rev_bexI: "[| x\<in>A; P(x) |] ==> \<exists>x\<in>A. P(x)"
by blast
-(*Not of the general form for such rules; ~\<exists>has become ALL~ *)
+(*Not of the general form for such rules. The existential quanitifer becomes universal. *)
lemma bexCI: "[| \<forall>x\<in>A. ~P(x) ==> P(a); a: A |] ==> \<exists>x\<in>A. P(x)"
by blast
lemma bexE [elim!]: "[| \<exists>x\<in>A. P(x); !!x. [| x\<in>A; P(x) |] ==> Q |] ==> Q"
by (simp add: Bex_def, blast)
-(*We do not even have (\<exists>x\<in>A. True) <-> True unless A is nonempty!!*)
+(*We do not even have @{term"(\<exists>x\<in>A. True) <-> True"} unless @{term"A" is nonempty!!*)
lemma bex_triv [simp]: "(\<exists>x\<in>A. P) <-> ((\<exists>x. x\<in>A) & P)"
by (simp add: Bex_def)
lemma bex_cong [cong]:
- "[| A=A'; !!x. x\<in>A' ==> P(x) <-> P'(x) |]
+ "[| A=A'; !!x. x\<in>A' ==> P(x) <-> P'(x) |]
==> (\<exists>x\<in>A. P(x)) <-> (\<exists>x\<in>A'. P'(x))"
by (simp add: Bex_def cong: conj_cong)
@@ -369,39 +369,39 @@
subsection{*Rules for subsets*}
lemma subsetI [intro!]:
- "(!!x. x\<in>A ==> x\<in>B) ==> A <= B"
-by (simp add: subset_def)
+ "(!!x. x\<in>A ==> x\<in>B) ==> A \<subseteq> B"
+by (simp add: subset_def)
(*Rule in Modus Ponens style [was called subsetE] *)
-lemma subsetD [elim]: "[| A <= B; c\<in>A |] ==> c\<in>B"
+lemma subsetD [elim]: "[| A \<subseteq> B; c\<in>A |] ==> c\<in>B"
apply (unfold subset_def)
apply (erule bspec, assumption)
done
(*Classical elimination rule*)
lemma subsetCE [elim]:
- "[| A <= B; c~:A ==> P; c\<in>B ==> P |] ==> P"
-by (simp add: subset_def, blast)
+ "[| A \<subseteq> B; c\<notin>A ==> P; c\<in>B ==> P |] ==> P"
+by (simp add: subset_def, blast)
(*Sometimes useful with premises in this order*)
lemma rev_subsetD: "[| c\<in>A; A<=B |] ==> c\<in>B"
by blast
-lemma contra_subsetD: "[| A <= B; c ~: B |] ==> c ~: A"
+lemma contra_subsetD: "[| A \<subseteq> B; c \<notin> B |] ==> c \<notin> A"
by blast
-lemma rev_contra_subsetD: "[| c ~: B; A <= B |] ==> c ~: A"
+lemma rev_contra_subsetD: "[| c \<notin> B; A \<subseteq> B |] ==> c \<notin> A"
by blast
-lemma subset_refl [simp]: "A <= A"
+lemma subset_refl [simp]: "A \<subseteq> A"
by blast
lemma subset_trans: "[| A<=B; B<=C |] ==> A<=C"
by blast
(*Useful for proving A<=B by rewriting in some cases*)
-lemma subset_iff:
- "A<=B <-> (\<forall>x. x\<in>A --> x\<in>B)"
+lemma subset_iff:
+ "A<=B <-> (\<forall>x. x\<in>A \<longrightarrow> x\<in>B)"
apply (unfold subset_def Ball_def)
apply (rule iff_refl)
done
@@ -410,8 +410,8 @@
subsection{*Rules for equality*}
(*Anti-symmetry of the subset relation*)
-lemma equalityI [intro]: "[| A <= B; B <= A |] ==> A = B"
-by (rule extension [THEN iffD2], rule conjI)
+lemma equalityI [intro]: "[| A \<subseteq> B; B \<subseteq> A |] ==> A = B"
+by (rule extension [THEN iffD2], rule conjI)
lemma equality_iffI: "(!!x. x\<in>A <-> x\<in>B) ==> A = B"
@@ -421,75 +421,75 @@
lemmas equalityD2 = extension [THEN iffD1, THEN conjunct2]
lemma equalityE: "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"
-by (blast dest: equalityD1 equalityD2)
+by (blast dest: equalityD1 equalityD2)
lemma equalityCE:
- "[| A = B; [| c\<in>A; c\<in>B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"
-by (erule equalityE, blast)
+ "[| A = B; [| c\<in>A; c\<in>B |] ==> P; [| c\<notin>A; c\<notin>B |] ==> P |] ==> P"
+by (erule equalityE, blast)
lemma equality_iffD:
- "A = B ==> (!!x. x : A <-> x : B)"
+ "A = B ==> (!!x. x \<in> A <-> x \<in> B)"
by auto
subsection{*Rules for Replace -- the derived form of replacement*}
-lemma Replace_iff:
- "b : {y. x\<in>A, P(x,y)} <-> (\<exists>x\<in>A. P(x,b) & (\<forall>y. P(x,y) --> y=b))"
+lemma Replace_iff:
+ "b \<in> {y. x\<in>A, P(x,y)} <-> (\<exists>x\<in>A. P(x,b) & (\<forall>y. P(x,y) \<longrightarrow> y=b))"
apply (unfold Replace_def)
apply (rule replacement [THEN iff_trans], blast+)
done
(*Introduction; there must be a unique y such that P(x,y), namely y=b. *)
-lemma ReplaceI [intro]:
- "[| P(x,b); x: A; !!y. P(x,y) ==> y=b |] ==>
- b : {y. x\<in>A, P(x,y)}"
-by (rule Replace_iff [THEN iffD2], blast)
+lemma ReplaceI [intro]:
+ "[| P(x,b); x: A; !!y. P(x,y) ==> y=b |] ==>
+ b \<in> {y. x\<in>A, P(x,y)}"
+by (rule Replace_iff [THEN iffD2], blast)
(*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *)
-lemma ReplaceE:
- "[| b : {y. x\<in>A, P(x,y)};
- !!x. [| x: A; P(x,b); \<forall>y. P(x,y)-->y=b |] ==> R
+lemma ReplaceE:
+ "[| b \<in> {y. x\<in>A, P(x,y)};
+ !!x. [| x: A; P(x,b); \<forall>y. P(x,y)\<longrightarrow>y=b |] ==> R
|] ==> R"
by (rule Replace_iff [THEN iffD1, THEN bexE], simp+)
(*As above but without the (generally useless) 3rd assumption*)
-lemma ReplaceE2 [elim!]:
- "[| b : {y. x\<in>A, P(x,y)};
- !!x. [| x: A; P(x,b) |] ==> R
+lemma ReplaceE2 [elim!]:
+ "[| b \<in> {y. x\<in>A, P(x,y)};
+ !!x. [| x: A; P(x,b) |] ==> R
|] ==> R"
-by (erule ReplaceE, blast)
+by (erule ReplaceE, blast)
lemma Replace_cong [cong]:
- "[| A=B; !!x y. x\<in>B ==> P(x,y) <-> Q(x,y) |] ==>
+ "[| A=B; !!x y. x\<in>B ==> P(x,y) <-> Q(x,y) |] ==>
Replace(A,P) = Replace(B,Q)"
-apply (rule equality_iffI)
-apply (simp add: Replace_iff)
+apply (rule equality_iffI)
+apply (simp add: Replace_iff)
done
subsection{*Rules for RepFun*}
-lemma RepFunI: "a \<in> A ==> f(a) : {f(x). x\<in>A}"
+lemma RepFunI: "a \<in> A ==> f(a) \<in> {f(x). x\<in>A}"
by (simp add: RepFun_def Replace_iff, blast)
(*Useful for coinduction proofs*)
-lemma RepFun_eqI [intro]: "[| b=f(a); a \<in> A |] ==> b : {f(x). x\<in>A}"
+lemma RepFun_eqI [intro]: "[| b=f(a); a \<in> A |] ==> b \<in> {f(x). x\<in>A}"
apply (erule ssubst)
apply (erule RepFunI)
done
lemma RepFunE [elim!]:
- "[| b : {f(x). x\<in>A};
- !!x.[| x\<in>A; b=f(x) |] ==> P |] ==>
+ "[| b \<in> {f(x). x\<in>A};
+ !!x.[| x\<in>A; b=f(x) |] ==> P |] ==>
P"
-by (simp add: RepFun_def Replace_iff, blast)
+by (simp add: RepFun_def Replace_iff, blast)
-lemma RepFun_cong [cong]:
+lemma RepFun_cong [cong]:
"[| A=B; !!x. x\<in>B ==> f(x)=g(x) |] ==> RepFun(A,f) = RepFun(B,g)"
by (simp add: RepFun_def)
-lemma RepFun_iff [simp]: "b : {f(x). x\<in>A} <-> (\<exists>x\<in>A. b=f(x))"
+lemma RepFun_iff [simp]: "b \<in> {f(x). x\<in>A} <-> (\<exists>x\<in>A. b=f(x))"
by (unfold Bex_def, blast)
lemma triv_RepFun [simp]: "{x. x\<in>A} = A"
@@ -499,23 +499,23 @@
subsection{*Rules for Collect -- forming a subset by separation*}
(*Separation is derivable from Replacement*)
-lemma separation [simp]: "a : {x\<in>A. P(x)} <-> a\<in>A & P(a)"
+lemma separation [simp]: "a \<in> {x\<in>A. P(x)} <-> a\<in>A & P(a)"
by (unfold Collect_def, blast)
-lemma CollectI [intro!]: "[| a\<in>A; P(a) |] ==> a : {x\<in>A. P(x)}"
+lemma CollectI [intro!]: "[| a\<in>A; P(a) |] ==> a \<in> {x\<in>A. P(x)}"
by simp
-lemma CollectE [elim!]: "[| a : {x\<in>A. P(x)}; [| a\<in>A; P(a) |] ==> R |] ==> R"
+lemma CollectE [elim!]: "[| a \<in> {x\<in>A. P(x)}; [| a\<in>A; P(a) |] ==> R |] ==> R"
by simp
-lemma CollectD1: "a : {x\<in>A. P(x)} ==> a\<in>A"
+lemma CollectD1: "a \<in> {x\<in>A. P(x)} ==> a\<in>A"
by (erule CollectE, assumption)
-lemma CollectD2: "a : {x\<in>A. P(x)} ==> P(a)"
+lemma CollectD2: "a \<in> {x\<in>A. P(x)} ==> P(a)"
by (erule CollectE, assumption)
lemma Collect_cong [cong]:
- "[| A=B; !!x. x\<in>B ==> P(x) <-> Q(x) |]
+ "[| A=B; !!x. x\<in>B ==> P(x) <-> Q(x) |]
==> Collect(A, %x. P(x)) = Collect(B, %x. Q(x))"
by (simp add: Collect_def)
@@ -525,17 +525,17 @@
declare Union_iff [simp]
(*The order of the premises presupposes that C is rigid; A may be flexible*)
-lemma UnionI [intro]: "[| B: C; A: B |] ==> A: Union(C)"
+lemma UnionI [intro]: "[| B: C; A: B |] ==> A: \<Union>(C)"
by (simp, blast)
-lemma UnionE [elim!]: "[| A \<in> Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R"
+lemma UnionE [elim!]: "[| A \<in> \<Union>(C); !!B.[| A: B; B: C |] ==> R |] ==> R"
by (simp, blast)
subsection{*Rules for Unions of families*}
-(* \<Union>x\<in>A. B(x) abbreviates Union({B(x). x\<in>A}) *)
+(* @{term"\<Union>x\<in>A. B(x)"} abbreviates @{term"\<Union>({B(x). x\<in>A})"} *)
-lemma UN_iff [simp]: "b : (\<Union>x\<in>A. B(x)) <-> (\<exists>x\<in>A. b \<in> B(x))"
+lemma UN_iff [simp]: "b \<in> (\<Union>x\<in>A. B(x)) <-> (\<exists>x\<in>A. b \<in> B(x))"
by (simp add: Bex_def, blast)
(*The order of the premises presupposes that A is rigid; b may be flexible*)
@@ -543,16 +543,16 @@
by (simp, blast)
-lemma UN_E [elim!]:
- "[| b : (\<Union>x\<in>A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R"
-by blast
+lemma UN_E [elim!]:
+ "[| b \<in> (\<Union>x\<in>A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R"
+by blast
-lemma UN_cong:
+lemma UN_cong:
"[| A=B; !!x. x\<in>B ==> C(x)=D(x) |] ==> (\<Union>x\<in>A. C(x)) = (\<Union>x\<in>B. D(x))"
-by simp
+by simp
-(*No "Addcongs [UN_cong]" because \<Union>is a combination of constants*)
+(*No "Addcongs [UN_cong]" because @{term\<Union>} is a combination of constants*)
(* UN_E appears before UnionE so that it is tried first, to avoid expensive
calls to hyp_subst_tac. Cannot include UN_I as it is unsafe: would enlarge
@@ -561,9 +561,9 @@
subsection{*Rules for the empty set*}
-(*The set {x\<in>0. False} is empty; by foundation it equals 0
+(*The set @{term"{x\<in>0. False}"} is empty; by foundation it equals 0
See Suppes, page 21.*)
-lemma not_mem_empty [simp]: "a ~: 0"
+lemma not_mem_empty [simp]: "a \<notin> 0"
apply (cut_tac foundation)
apply (best dest: equalityD2)
done
@@ -571,69 +571,69 @@
lemmas emptyE [elim!] = not_mem_empty [THEN notE]
-lemma empty_subsetI [simp]: "0 <= A"
-by blast
+lemma empty_subsetI [simp]: "0 \<subseteq> A"
+by blast
lemma equals0I: "[| !!y. y\<in>A ==> False |] ==> A=0"
by blast
-lemma equals0D [dest]: "A=0 ==> a ~: A"
+lemma equals0D [dest]: "A=0 ==> a \<notin> A"
by blast
declare sym [THEN equals0D, dest]
-lemma not_emptyI: "a\<in>A ==> A ~= 0"
+lemma not_emptyI: "a\<in>A ==> A \<noteq> 0"
by blast
-lemma not_emptyE: "[| A ~= 0; !!x. x\<in>A ==> R |] ==> R"
+lemma not_emptyE: "[| A \<noteq> 0; !!x. x\<in>A ==> R |] ==> R"
by blast
subsection{*Rules for Inter*}
(*Not obviously useful for proving InterI, InterD, InterE*)
-lemma Inter_iff: "A \<in> Inter(C) <-> (\<forall>x\<in>C. A: x) & C\<noteq>0"
+lemma Inter_iff: "A \<in> \<Inter>(C) <-> (\<forall>x\<in>C. A: x) & C\<noteq>0"
by (simp add: Inter_def Ball_def, blast)
(* Intersection is well-behaved only if the family is non-empty! *)
-lemma InterI [intro!]:
- "[| !!x. x: C ==> A: x; C\<noteq>0 |] ==> A \<in> Inter(C)"
+lemma InterI [intro!]:
+ "[| !!x. x: C ==> A: x; C\<noteq>0 |] ==> A \<in> \<Inter>(C)"
by (simp add: Inter_iff)
(*A "destruct" rule -- every B in C contains A as an element, but
A\<in>B can hold when B\<in>C does not! This rule is analogous to "spec". *)
-lemma InterD [elim, Pure.elim]: "[| A \<in> Inter(C); B \<in> C |] ==> A \<in> B"
+lemma InterD [elim, Pure.elim]: "[| A \<in> \<Inter>(C); B \<in> C |] ==> A \<in> B"
by (unfold Inter_def, blast)
-(*"Classical" elimination rule -- does not require exhibiting B\<in>C *)
-lemma InterE [elim]:
- "[| A \<in> Inter(C); B~:C ==> R; A\<in>B ==> R |] ==> R"
-by (simp add: Inter_def, blast)
-
+(*"Classical" elimination rule -- does not require exhibiting @{term"B\<in>C"} *)
+lemma InterE [elim]:
+ "[| A \<in> \<Inter>(C); B\<notin>C ==> R; A\<in>B ==> R |] ==> R"
+by (simp add: Inter_def, blast)
+
subsection{*Rules for Intersections of families*}
-(* \<Inter>x\<in>A. B(x) abbreviates Inter({B(x). x\<in>A}) *)
+(* @{term"\<Inter>x\<in>A. B(x)"} abbreviates @{term"\<Inter>({B(x). x\<in>A})"} *)
-lemma INT_iff: "b : (\<Inter>x\<in>A. B(x)) <-> (\<forall>x\<in>A. b \<in> B(x)) & A\<noteq>0"
+lemma INT_iff: "b \<in> (\<Inter>x\<in>A. B(x)) <-> (\<forall>x\<in>A. b \<in> B(x)) & A\<noteq>0"
by (force simp add: Inter_def)
lemma INT_I: "[| !!x. x: A ==> b: B(x); A\<noteq>0 |] ==> b: (\<Inter>x\<in>A. B(x))"
by blast
-lemma INT_E: "[| b : (\<Inter>x\<in>A. B(x)); a: A |] ==> b \<in> B(a)"
+lemma INT_E: "[| b \<in> (\<Inter>x\<in>A. B(x)); a: A |] ==> b \<in> B(a)"
by blast
lemma INT_cong:
"[| A=B; !!x. x\<in>B ==> C(x)=D(x) |] ==> (\<Inter>x\<in>A. C(x)) = (\<Inter>x\<in>B. D(x))"
by simp
-(*No "Addcongs [INT_cong]" because \<Inter>is a combination of constants*)
+(*No "Addcongs [INT_cong]" because @{term\<Inter>} is a combination of constants*)
subsection{*Rules for Powersets*}
-lemma PowI: "A <= B ==> A \<in> Pow(B)"
+lemma PowI: "A \<subseteq> B ==> A \<in> Pow(B)"
by (erule Pow_iff [THEN iffD2])
lemma PowD: "A \<in> Pow(B) ==> A<=B"
@@ -641,16 +641,16 @@
declare Pow_iff [iff]
-lemmas Pow_bottom = empty_subsetI [THEN PowI] (* 0 \<in> Pow(B) *)
-lemmas Pow_top = subset_refl [THEN PowI] (* A \<in> Pow(A) *)
+lemmas Pow_bottom = empty_subsetI [THEN PowI] --{* @{term"0 \<in> Pow(B)"} *}
+lemmas Pow_top = subset_refl [THEN PowI] --{* @{term"A \<in> Pow(A)"} *}
subsection{*Cantor's Theorem: There is no surjection from a set to its powerset.*}
-(*The search is undirected. Allowing redundant introduction rules may
+(*The search is undirected. Allowing redundant introduction rules may
make it diverge. Variable b represents ANY map, such as
(lam x\<in>A.b(x)): A->Pow(A). *)
-lemma cantor: "\<exists>S \<in> Pow(A). \<forall>x\<in>A. b(x) ~= S"
+lemma cantor: "\<exists>S \<in> Pow(A). \<forall>x\<in>A. b(x) \<noteq> S"
by (best elim!: equalityCE del: ReplaceI RepFun_eqI)
end
--- a/src/ZF/Zorn.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/Zorn.thy Tue Mar 06 15:15:49 2012 +0000
@@ -28,11 +28,11 @@
definition
increasing :: "i=>i" where
- "increasing(A) == {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A --> x<=f`x}"
+ "increasing(A) == {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A \<longrightarrow> x<=f`x}"
text{*Lemma for the inductive definition below*}
-lemma Union_in_Pow: "Y \<in> Pow(Pow(A)) ==> Union(Y) \<in> Pow(A)"
+lemma Union_in_Pow: "Y \<in> Pow(Pow(A)) ==> \<Union>(Y) \<in> Pow(A)"
by blast
@@ -45,12 +45,12 @@
"TFin" :: "[i,i]=>i"
inductive
- domains "TFin(S,next)" <= "Pow(S)"
+ domains "TFin(S,next)" \<subseteq> "Pow(S)"
intros
nextI: "[| x \<in> TFin(S,next); next \<in> increasing(S) |]
==> next`x \<in> TFin(S,next)"
- Pow_UnionI: "Y \<in> Pow(TFin(S,next)) ==> Union(Y) \<in> TFin(S,next)"
+ Pow_UnionI: "Y \<in> Pow(TFin(S,next)) ==> \<Union>(Y) \<in> TFin(S,next)"
monos Pow_mono
con_defs increasing_def
@@ -59,11 +59,11 @@
subsection{*Mathematical Preamble *}
-lemma Union_lemma0: "(\<forall>x\<in>C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)"
+lemma Union_lemma0: "(\<forall>x\<in>C. x<=A | B<=x) ==> \<Union>(C)<=A | B<=\<Union>(C)"
by blast
lemma Inter_lemma0:
- "[| c \<in> C; \<forall>x\<in>C. A<=x | x<=B |] ==> A <= Inter(C) | Inter(C) <= B"
+ "[| c \<in> C; \<forall>x\<in>C. A<=x | x<=B |] ==> A \<subseteq> \<Inter>(C) | \<Inter>(C) \<subseteq> B"
by blast
@@ -74,7 +74,7 @@
apply (erule CollectD1)
done
-lemma increasingD2: "[| f \<in> increasing(A); x<=A |] ==> x <= f`x"
+lemma increasingD2: "[| f \<in> increasing(A); x<=A |] ==> x \<subseteq> f`x"
by (unfold increasing_def, blast)
lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI]
@@ -86,7 +86,7 @@
lemma TFin_induct:
"[| n \<in> TFin(S,next);
!!x. [| x \<in> TFin(S,next); P(x); next \<in> increasing(S) |] ==> P(next`x);
- !!Y. [| Y <= TFin(S,next); \<forall>y\<in>Y. P(y) |] ==> P(Union(Y))
+ !!Y. [| Y \<subseteq> TFin(S,next); \<forall>y\<in>Y. P(y) |] ==> P(\<Union>(Y))
|] ==> P(n)"
by (erule TFin.induct, blast+)
@@ -99,7 +99,7 @@
text{*Lemma 1 of section 3.1*}
lemma TFin_linear_lemma1:
"[| n \<in> TFin(S,next); m \<in> TFin(S,next);
- \<forall>x \<in> TFin(S,next) . x<=m --> x=m | next`x<=m |]
+ \<forall>x \<in> TFin(S,next) . x<=m \<longrightarrow> x=m | next`x<=m |]
==> n<=m | next`m<=n"
apply (erule TFin_induct)
apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*)
@@ -111,7 +111,7 @@
Requires @{term "next \<in> increasing(S)"} in the second induction step.*}
lemma TFin_linear_lemma2:
"[| m \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> \<forall>n \<in> TFin(S,next). n<=m --> n=m | next`n <= m"
+ ==> \<forall>n \<in> TFin(S,next). n<=m \<longrightarrow> n=m | next`n \<subseteq> m"
apply (erule TFin_induct)
apply (rule impI [THEN ballI])
txt{*case split using @{text TFin_linear_lemma1}*}
@@ -136,13 +136,13 @@
text{*a more convenient form for Lemma 2*}
lemma TFin_subsetD:
"[| n<=m; m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> n=m | next`n <= m"
+ ==> n=m | next`n \<subseteq> m"
by (blast dest: TFin_linear_lemma2 [rule_format])
text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*}
lemma TFin_subset_linear:
"[| m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> n <= m | m<=n"
+ ==> n \<subseteq> m | m<=n"
apply (rule disjE)
apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
apply (assumption+, erule disjI2)
@@ -153,7 +153,7 @@
text{*Lemma 3 of section 3.3*}
lemma equal_next_upper:
- "[| n \<in> TFin(S,next); m \<in> TFin(S,next); m = next`m |] ==> n <= m"
+ "[| n \<in> TFin(S,next); m \<in> TFin(S,next); m = next`m |] ==> n \<subseteq> m"
apply (erule TFin_induct)
apply (drule TFin_subsetD)
apply (assumption+, force, blast)
@@ -162,7 +162,7 @@
text{*Property 3.3 of section 3.3*}
lemma equal_next_Union:
"[| m \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> m = next`m <-> m = Union(TFin(S,next))"
+ ==> m = next`m <-> m = \<Union>(TFin(S,next))"
apply (rule iffI)
apply (rule Union_upper [THEN equalityI])
apply (rule_tac [2] equal_next_upper [THEN Union_least])
@@ -181,17 +181,17 @@
text{** Defining the "next" operation for Hausdorff's Theorem **}
-lemma chain_subset_Pow: "chain(A) <= Pow(A)"
+lemma chain_subset_Pow: "chain(A) \<subseteq> Pow(A)"
apply (unfold chain_def)
apply (rule Collect_subset)
done
-lemma super_subset_chain: "super(A,c) <= chain(A)"
+lemma super_subset_chain: "super(A,c) \<subseteq> chain(A)"
apply (unfold super_def)
apply (rule Collect_subset)
done
-lemma maxchain_subset_chain: "maxchain(A) <= chain(A)"
+lemma maxchain_subset_chain: "maxchain(A) \<subseteq> chain(A)"
apply (unfold maxchain_def)
apply (rule Collect_subset)
done
@@ -255,12 +255,12 @@
apply (rule AC_Pi_Pow [THEN exE])
apply (rule Hausdorff_next_exists [THEN bexE], assumption)
apply (rename_tac ch "next")
-apply (subgoal_tac "Union (TFin (S,next)) \<in> chain (S) ")
+apply (subgoal_tac "\<Union>(TFin (S,next)) \<in> chain (S) ")
prefer 2
apply (blast intro!: TFin_chain_lemma4 subset_refl [THEN TFin_UnionI])
-apply (rule_tac x = "Union (TFin (S,next))" in exI)
+apply (rule_tac x = "\<Union>(TFin (S,next))" in exI)
apply (rule classical)
-apply (subgoal_tac "next ` Union (TFin (S,next)) = Union (TFin (S,next))")
+apply (subgoal_tac "next ` Union(TFin (S,next)) = \<Union>(TFin (S,next))")
apply (rule_tac [2] equal_next_Union [THEN iffD2, symmetric])
apply (rule_tac [2] subset_refl [THEN TFin_UnionI])
prefer 2 apply assumption
@@ -280,11 +280,11 @@
"[| c \<in> chain(A); z \<in> A; \<forall>x \<in> c. x<=z |] ==> cons(z,c) \<in> chain(A)"
by (unfold chain_def, blast)
-lemma Zorn: "\<forall>c \<in> chain(S). Union(c) \<in> S ==> \<exists>y \<in> S. \<forall>z \<in> S. y<=z --> y=z"
+lemma Zorn: "\<forall>c \<in> chain(S). \<Union>(c) \<in> S ==> \<exists>y \<in> S. \<forall>z \<in> S. y<=z \<longrightarrow> y=z"
apply (rule Hausdorff [THEN exE])
apply (simp add: maxchain_def)
apply (rename_tac c)
-apply (rule_tac x = "Union (c)" in bexI)
+apply (rule_tac x = "\<Union>(c)" in bexI)
prefer 2 apply blast
apply safe
apply (rename_tac z)
@@ -299,7 +299,7 @@
text {* Alternative version of Zorn's Lemma *}
theorem Zorn2:
- "\<forall>c \<in> chain(S). \<exists>y \<in> S. \<forall>x \<in> c. x <= y ==> \<exists>y \<in> S. \<forall>z \<in> S. y<=z --> y=z"
+ "\<forall>c \<in> chain(S). \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y ==> \<exists>y \<in> S. \<forall>z \<in> S. y<=z \<longrightarrow> y=z"
apply (cut_tac Hausdorff maxchain_subset_chain)
apply (erule exE)
apply (drule subsetD, assumption)
@@ -321,20 +321,20 @@
text{*Lemma 5*}
lemma TFin_well_lemma5:
- "[| n \<in> TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) \<in> Z |]
- ==> \<forall>m \<in> Z. n <= m"
+ "[| n \<in> TFin(S,next); Z \<subseteq> TFin(S,next); z:Z; ~ \<Inter>(Z) \<in> Z |]
+ ==> \<forall>m \<in> Z. n \<subseteq> m"
apply (erule TFin_induct)
prefer 2 apply blast txt{*second induction step is easy*}
apply (rule ballI)
apply (rule bspec [THEN TFin_subsetD, THEN disjE], auto)
-apply (subgoal_tac "m = Inter (Z) ")
+apply (subgoal_tac "m = \<Inter>(Z) ")
apply blast+
done
text{*Well-ordering of @{term "TFin(S,next)"} *}
-lemma well_ord_TFin_lemma: "[| Z <= TFin(S,next); z \<in> Z |] ==> Inter(Z) \<in> Z"
+lemma well_ord_TFin_lemma: "[| Z \<subseteq> TFin(S,next); z \<in> Z |] ==> \<Inter>(Z) \<in> Z"
apply (rule classical)
-apply (subgoal_tac "Z = {Union (TFin (S,next))}")
+apply (subgoal_tac "Z = {\<Union>(TFin (S,next))}")
apply (simp (no_asm_simp) add: Inter_singleton)
apply (erule equal_singleton)
apply (rule Union_upper [THEN equalityI])
@@ -350,7 +350,7 @@
txt{*Prove the well-foundedness goal*}
apply (rule wf_onI)
apply (frule well_ord_TFin_lemma, assumption)
-apply (drule_tac x = "Inter (Z) " in bspec, assumption)
+apply (drule_tac x = "\<Inter>(Z) " in bspec, assumption)
apply blast
txt{*Now prove the linearity goal*}
apply (intro ballI)
@@ -394,25 +394,25 @@
"[| ch \<in> (\<Pi> X \<in> Pow(S)-{0}. X);
next \<in> increasing(S);
\<forall>X \<in> Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) |]
- ==> (\<lambda> x \<in> S. Union({y \<in> TFin(S,next). x \<notin> y}))
+ ==> (\<lambda> x \<in> S. \<Union>({y \<in> TFin(S,next). x \<notin> y}))
\<in> inj(S, TFin(S,next) - {S})"
apply (rule_tac d = "%y. ch` (S-y) " in lam_injective)
apply (rule DiffI)
apply (rule Collect_subset [THEN TFin_UnionI])
apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE)
-apply (subgoal_tac "x \<notin> Union ({y \<in> TFin (S,next) . x \<notin> y}) ")
+apply (subgoal_tac "x \<notin> \<Union>({y \<in> TFin (S,next) . x \<notin> y}) ")
prefer 2 apply (blast elim: equalityE)
-apply (subgoal_tac "Union ({y \<in> TFin (S,next) . x \<notin> y}) \<noteq> S")
+apply (subgoal_tac "\<Union>({y \<in> TFin (S,next) . x \<notin> y}) \<noteq> S")
prefer 2 apply (blast elim: equalityE)
-txt{*For proving @{text "x \<in> next`Union(...)"}.
+txt{*For proving @{text "x \<in> next`\<Union>(...)"}.
Abrial and Laffitte's justification appears to be faulty.*}
-apply (subgoal_tac "~ next ` Union ({y \<in> TFin (S,next) . x \<notin> y})
- <= Union ({y \<in> TFin (S,next) . x \<notin> y}) ")
+apply (subgoal_tac "~ next ` Union({y \<in> TFin (S,next) . x \<notin> y})
+ \<subseteq> \<Union>({y \<in> TFin (S,next) . x \<notin> y}) ")
prefer 2
apply (simp del: Union_iff
add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]
Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2])
-apply (subgoal_tac "x \<in> next ` Union ({y \<in> TFin (S,next) . x \<notin> y}) ")
+apply (subgoal_tac "x \<in> next ` Union({y \<in> TFin (S,next) . x \<notin> y}) ")
prefer 2
apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI)
txt{*End of the lemmas!*}
@@ -436,7 +436,7 @@
definition Chain :: "i => i" where
- "Chain(r) = {A : Pow(field(r)). ALL a:A. ALL b:A. <a, b> : r | <b, a> : r}"
+ "Chain(r) = {A \<in> Pow(field(r)). \<forall>a\<in>A. \<forall>b\<in>A. <a, b> \<in> r | <b, a> \<in> r}"
lemma mono_Chain:
"r \<subseteq> s ==> Chain(r) \<subseteq> Chain(s)"
@@ -445,17 +445,17 @@
theorem Zorn_po:
assumes po: "Partial_order(r)"
- and u: "ALL C:Chain(r). EX u:field(r). ALL a:C. <a, u> : r"
- shows "EX m:field(r). ALL a:field(r). <m, a> : r --> a = m"
+ and u: "\<forall>C\<in>Chain(r). \<exists>u\<in>field(r). \<forall>a\<in>C. <a, u> \<in> r"
+ shows "\<exists>m\<in>field(r). \<forall>a\<in>field(r). <m, a> \<in> r \<longrightarrow> a = m"
proof -
have "Preorder(r)" using po by (simp add: partial_order_on_def)
--{* Mirror r in the set of subsets below (wrt r) elements of A (?). *}
- let ?B = "lam x:field(r). r -`` {x}" let ?S = "?B `` field(r)"
- have "ALL C:chain(?S). EX U:?S. ALL A:C. A \<subseteq> U"
+ let ?B = "\<lambda>x\<in>field(r). r -`` {x}" let ?S = "?B `` field(r)"
+ have "\<forall>C\<in>chain(?S). \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U"
proof (clarsimp simp: chain_def Subset_rel_def bex_image_simp)
fix C
- assume 1: "C \<subseteq> ?S" and 2: "ALL A:C. ALL B:C. A \<subseteq> B | B \<subseteq> A"
- let ?A = "{x : field(r). EX M:C. M = ?B`x}"
+ assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B | B \<subseteq> A"
+ let ?A = "{x \<in> field(r). \<exists>M\<in>C. M = ?B`x}"
have "C = ?B `` ?A" using 1
apply (auto simp: image_def)
apply rule
@@ -472,46 +472,46 @@
apply (thin_tac "C \<subseteq> ?X")
apply (fast elim: lamE)
done
- have "?A : Chain(r)"
+ have "?A \<in> Chain(r)"
proof (simp add: Chain_def subsetI, intro conjI ballI impI)
fix a b
- assume "a : field(r)" "r -`` {a} : C" "b : field(r)" "r -`` {b} : C"
+ assume "a \<in> field(r)" "r -`` {a} \<in> C" "b \<in> field(r)" "r -`` {b} \<in> C"
hence "r -`` {a} \<subseteq> r -`` {b} | r -`` {b} \<subseteq> r -`` {a}" using 2 by auto
- then show "<a, b> : r | <b, a> : r"
- using `Preorder(r)` `a : field(r)` `b : field(r)`
+ then show "<a, b> \<in> r | <b, a> \<in> r"
+ using `Preorder(r)` `a \<in> field(r)` `b \<in> field(r)`
by (simp add: subset_vimage1_vimage1_iff)
qed
- then obtain u where uA: "u : field(r)" "ALL a:?A. <a, u> : r"
+ then obtain u where uA: "u \<in> field(r)" "\<forall>a\<in>?A. <a, u> \<in> r"
using u
apply auto
apply (drule bspec) apply assumption
apply auto
done
- have "ALL A:C. A \<subseteq> r -`` {u}"
+ have "\<forall>A\<in>C. A \<subseteq> r -`` {u}"
proof (auto intro!: vimageI)
fix a B
- assume aB: "B : C" "a : B"
- with 1 obtain x where "x : field(r)" "B = r -`` {x}"
+ assume aB: "B \<in> C" "a \<in> B"
+ with 1 obtain x where "x \<in> field(r)" "B = r -`` {x}"
apply -
apply (drule subsetD) apply assumption
apply (erule imageE)
apply (erule lamE)
apply simp
done
- then show "<a, u> : r" using uA aB `Preorder(r)`
+ then show "<a, u> \<in> r" using uA aB `Preorder(r)`
by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+
qed
- then show "EX U:field(r). ALL A:C. A \<subseteq> r -`` {U}"
- using `u : field(r)` ..
+ then show "\<exists>U\<in>field(r). \<forall>A\<in>C. A \<subseteq> r -`` {U}"
+ using `u \<in> field(r)` ..
qed
from Zorn2 [OF this]
- obtain m B where "m : field(r)" "B = r -`` {m}"
- "ALL x:field(r). B \<subseteq> r -`` {x} --> B = r -`` {x}"
+ obtain m B where "m \<in> field(r)" "B = r -`` {m}"
+ "\<forall>x\<in>field(r). B \<subseteq> r -`` {x} \<longrightarrow> B = r -`` {x}"
by (auto elim!: lamE simp: ball_image_simp)
- then have "ALL a:field(r). <m, a> : r --> a = m"
- using po `Preorder(r)` `m : field(r)`
+ then have "\<forall>a\<in>field(r). <m, a> \<in> r \<longrightarrow> a = m"
+ using po `Preorder(r)` `m \<in> field(r)`
by (auto simp: subset_vimage1_vimage1_iff Partial_order_eq_vimage1_vimage1_iff)
- then show ?thesis using `m : field(r)` by blast
+ then show ?thesis using `m \<in> field(r)` by blast
qed
end
--- a/src/ZF/equalities.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/equalities.thy Tue Mar 06 15:15:49 2012 +0000
@@ -10,14 +10,14 @@
text{*These cover union, intersection, converse, domain, range, etc. Philippe
de Groote proved many of the inclusions.*}
-lemma in_mono: "A\<subseteq>B ==> x\<in>A --> x\<in>B"
+lemma in_mono: "A\<subseteq>B ==> x\<in>A \<longrightarrow> x\<in>B"
by blast
lemma the_eq_0 [simp]: "(THE x. False) = 0"
by (blast intro: the_0)
subsection{*Bounded Quantifiers*}
-text {* \medskip
+text {* \medskip
The following are not added to the default simpset because
(a) they duplicate the body and (b) there are no similar rules for @{text Int}.*}
@@ -46,10 +46,10 @@
by (unfold converse_def, blast)
lemma converseE [elim!]:
- "[| yx \<in> converse(r);
+ "[| yx \<in> converse(r);
!!x y. [| yx=<y,x>; <x,y>\<in>r |] ==> P |]
==> P"
-by (unfold converse_def, blast)
+by (unfold converse_def, blast)
lemma converse_converse: "r\<subseteq>Sigma(A,B) ==> converse(converse(r)) = r"
by blast
@@ -79,7 +79,7 @@
lemma cons_subset_iff [iff]: "cons(a,B)\<subseteq>C <-> a\<in>C & B\<subseteq>C"
by blast
-(*A safe special case of subset elimination, adding no new variables
+(*A safe special case of subset elimination, adding no new variables
[| cons(a,B) \<subseteq> C; [| a \<in> C; B \<subseteq> C |] ==> R |] ==> R *)
lemmas cons_subsetE = cons_subset_iff [THEN iffD1, THEN conjE]
@@ -90,7 +90,7 @@
by blast
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*)
-lemma cons_eq: "{a} Un B = cons(a,B)"
+lemma cons_eq: "{a} \<union> B = cons(a,B)"
by blast
lemma cons_commute: "cons(a, cons(b, C)) = cons(b, cons(a, C))"
@@ -102,7 +102,7 @@
lemma cons_Diff: "a: B ==> cons(a, B-{a}) = B"
by blast
-lemma Diff_cons_eq: "cons(a,B) - C = (if a\<in>C then B-C else cons(a,B-C))"
+lemma Diff_cons_eq: "cons(a,B) - C = (if a\<in>C then B-C else cons(a,B-C))"
by auto
lemma equal_singleton [rule_format]: "[| a: C; \<forall>y\<in>C. y=b |] ==> C = {b}"
@@ -125,14 +125,14 @@
lemma subset_succI: "i \<subseteq> succ(i)"
by blast
-(*But if j is an ordinal or is transitive, then i\<in>j implies i\<subseteq>j!
- See ordinal/Ord_succ_subsetI*)
+(*But if j is an ordinal or is transitive, then @{term"i\<in>j"} implies @{term"i\<subseteq>j"}!
+ See @{text"Ord_succ_subsetI}*)
lemma succ_subsetI: "[| i\<in>j; i\<subseteq>j |] ==> succ(i)\<subseteq>j"
by (unfold succ_def, blast)
lemma succ_subsetE:
"[| succ(i) \<subseteq> j; [| i\<in>j; i\<subseteq>j |] ==> P |] ==> P"
-by (unfold succ_def, blast)
+by (unfold succ_def, blast)
lemma succ_subset_iff: "succ(a) \<subseteq> B <-> (a \<subseteq> B & a \<in> B)"
by (unfold succ_def, blast)
@@ -142,34 +142,34 @@
(** Intersection is the greatest lower bound of two sets **)
-lemma Int_subset_iff: "C \<subseteq> A Int B <-> C \<subseteq> A & C \<subseteq> B"
+lemma Int_subset_iff: "C \<subseteq> A \<inter> B <-> C \<subseteq> A & C \<subseteq> B"
by blast
-lemma Int_lower1: "A Int B \<subseteq> A"
+lemma Int_lower1: "A \<inter> B \<subseteq> A"
by blast
-lemma Int_lower2: "A Int B \<subseteq> B"
+lemma Int_lower2: "A \<inter> B \<subseteq> B"
by blast
-lemma Int_greatest: "[| C\<subseteq>A; C\<subseteq>B |] ==> C \<subseteq> A Int B"
+lemma Int_greatest: "[| C\<subseteq>A; C\<subseteq>B |] ==> C \<subseteq> A \<inter> B"
by blast
-lemma Int_cons: "cons(a,B) Int C \<subseteq> cons(a, B Int C)"
+lemma Int_cons: "cons(a,B) \<inter> C \<subseteq> cons(a, B \<inter> C)"
by blast
-lemma Int_absorb [simp]: "A Int A = A"
+lemma Int_absorb [simp]: "A \<inter> A = A"
by blast
-lemma Int_left_absorb: "A Int (A Int B) = A Int B"
+lemma Int_left_absorb: "A \<inter> (A \<inter> B) = A \<inter> B"
by blast
-lemma Int_commute: "A Int B = B Int A"
+lemma Int_commute: "A \<inter> B = B \<inter> A"
by blast
-lemma Int_left_commute: "A Int (B Int C) = B Int (A Int C)"
+lemma Int_left_commute: "A \<inter> (B \<inter> C) = B \<inter> (A \<inter> C)"
by blast
-lemma Int_assoc: "(A Int B) Int C = A Int (B Int C)"
+lemma Int_assoc: "(A \<inter> B) \<inter> C = A \<inter> (B \<inter> C)"
by blast
(*Intersection is an AC-operator*)
@@ -181,27 +181,27 @@
lemma Int_absorb2: "A \<subseteq> B ==> A \<inter> B = A"
by blast
-lemma Int_Un_distrib: "A Int (B Un C) = (A Int B) Un (A Int C)"
+lemma Int_Un_distrib: "A \<inter> (B \<union> C) = (A \<inter> B) \<union> (A \<inter> C)"
by blast
-lemma Int_Un_distrib2: "(B Un C) Int A = (B Int A) Un (C Int A)"
+lemma Int_Un_distrib2: "(B \<union> C) \<inter> A = (B \<inter> A) \<union> (C \<inter> A)"
by blast
-lemma subset_Int_iff: "A\<subseteq>B <-> A Int B = A"
+lemma subset_Int_iff: "A\<subseteq>B <-> A \<inter> B = A"
by (blast elim!: equalityE)
-lemma subset_Int_iff2: "A\<subseteq>B <-> B Int A = A"
+lemma subset_Int_iff2: "A\<subseteq>B <-> B \<inter> A = A"
by (blast elim!: equalityE)
-lemma Int_Diff_eq: "C\<subseteq>A ==> (A-B) Int C = C-B"
+lemma Int_Diff_eq: "C\<subseteq>A ==> (A-B) \<inter> C = C-B"
by blast
lemma Int_cons_left:
- "cons(a,A) Int B = (if a \<in> B then cons(a, A Int B) else A Int B)"
+ "cons(a,A) \<inter> B = (if a \<in> B then cons(a, A \<inter> B) else A \<inter> B)"
by auto
lemma Int_cons_right:
- "A Int cons(a, B) = (if a \<in> A then cons(a, A Int B) else A Int B)"
+ "A \<inter> cons(a, B) = (if a \<in> A then cons(a, A \<inter> B) else A \<inter> B)"
by auto
lemma cons_Int_distrib: "cons(x, A \<inter> B) = cons(x, A) \<inter> cons(x, B)"
@@ -211,34 +211,34 @@
(** Union is the least upper bound of two sets *)
-lemma Un_subset_iff: "A Un B \<subseteq> C <-> A \<subseteq> C & B \<subseteq> C"
+lemma Un_subset_iff: "A \<union> B \<subseteq> C <-> A \<subseteq> C & B \<subseteq> C"
by blast
-lemma Un_upper1: "A \<subseteq> A Un B"
+lemma Un_upper1: "A \<subseteq> A \<union> B"
by blast
-lemma Un_upper2: "B \<subseteq> A Un B"
+lemma Un_upper2: "B \<subseteq> A \<union> B"
by blast
-lemma Un_least: "[| A\<subseteq>C; B\<subseteq>C |] ==> A Un B \<subseteq> C"
+lemma Un_least: "[| A\<subseteq>C; B\<subseteq>C |] ==> A \<union> B \<subseteq> C"
by blast
-lemma Un_cons: "cons(a,B) Un C = cons(a, B Un C)"
+lemma Un_cons: "cons(a,B) \<union> C = cons(a, B \<union> C)"
by blast
-lemma Un_absorb [simp]: "A Un A = A"
+lemma Un_absorb [simp]: "A \<union> A = A"
by blast
-lemma Un_left_absorb: "A Un (A Un B) = A Un B"
+lemma Un_left_absorb: "A \<union> (A \<union> B) = A \<union> B"
by blast
-lemma Un_commute: "A Un B = B Un A"
+lemma Un_commute: "A \<union> B = B \<union> A"
by blast
-lemma Un_left_commute: "A Un (B Un C) = B Un (A Un C)"
+lemma Un_left_commute: "A \<union> (B \<union> C) = B \<union> (A \<union> C)"
by blast
-lemma Un_assoc: "(A Un B) Un C = A Un (B Un C)"
+lemma Un_assoc: "(A \<union> B) \<union> C = A \<union> (B \<union> C)"
by blast
(*Union is an AC-operator*)
@@ -250,19 +250,19 @@
lemma Un_absorb2: "B \<subseteq> A ==> A \<union> B = A"
by blast
-lemma Un_Int_distrib: "(A Int B) Un C = (A Un C) Int (B Un C)"
+lemma Un_Int_distrib: "(A \<inter> B) \<union> C = (A \<union> C) \<inter> (B \<union> C)"
by blast
-lemma subset_Un_iff: "A\<subseteq>B <-> A Un B = B"
+lemma subset_Un_iff: "A\<subseteq>B <-> A \<union> B = B"
by (blast elim!: equalityE)
-lemma subset_Un_iff2: "A\<subseteq>B <-> B Un A = B"
+lemma subset_Un_iff2: "A\<subseteq>B <-> B \<union> A = B"
by (blast elim!: equalityE)
-lemma Un_empty [iff]: "(A Un B = 0) <-> (A = 0 & B = 0)"
+lemma Un_empty [iff]: "(A \<union> B = 0) <-> (A = 0 & B = 0)"
by blast
-lemma Un_eq_Union: "A Un B = Union({A, B})"
+lemma Un_eq_Union: "A \<union> B = \<Union>({A, B})"
by blast
subsection{*Set Difference*}
@@ -270,16 +270,16 @@
lemma Diff_subset: "A-B \<subseteq> A"
by blast
-lemma Diff_contains: "[| C\<subseteq>A; C Int B = 0 |] ==> C \<subseteq> A-B"
+lemma Diff_contains: "[| C\<subseteq>A; C \<inter> B = 0 |] ==> C \<subseteq> A-B"
by blast
-lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C) <-> B\<subseteq>A-C & c ~: B"
+lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C) <-> B\<subseteq>A-C & c \<notin> B"
by blast
lemma Diff_cancel: "A - A = 0"
by blast
-lemma Diff_triv: "A Int B = 0 ==> A - B = A"
+lemma Diff_triv: "A \<inter> B = 0 ==> A - B = A"
by blast
lemma empty_Diff [simp]: "0 - A = 0"
@@ -299,46 +299,46 @@
lemma Diff_cons2: "A - cons(a,B) = A - {a} - B"
by blast
-lemma Diff_disjoint: "A Int (B-A) = 0"
+lemma Diff_disjoint: "A \<inter> (B-A) = 0"
by blast
-lemma Diff_partition: "A\<subseteq>B ==> A Un (B-A) = B"
+lemma Diff_partition: "A\<subseteq>B ==> A \<union> (B-A) = B"
by blast
-lemma subset_Un_Diff: "A \<subseteq> B Un (A - B)"
+lemma subset_Un_Diff: "A \<subseteq> B \<union> (A - B)"
by blast
lemma double_complement: "[| A\<subseteq>B; B\<subseteq>C |] ==> B-(C-A) = A"
by blast
-lemma double_complement_Un: "(A Un B) - (B-A) = A"
+lemma double_complement_Un: "(A \<union> B) - (B-A) = A"
by blast
-lemma Un_Int_crazy:
- "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"
+lemma Un_Int_crazy:
+ "(A \<inter> B) \<union> (B \<inter> C) \<union> (C \<inter> A) = (A \<union> B) \<inter> (B \<union> C) \<inter> (C \<union> A)"
apply blast
done
-lemma Diff_Un: "A - (B Un C) = (A-B) Int (A-C)"
+lemma Diff_Un: "A - (B \<union> C) = (A-B) \<inter> (A-C)"
by blast
-lemma Diff_Int: "A - (B Int C) = (A-B) Un (A-C)"
+lemma Diff_Int: "A - (B \<inter> C) = (A-B) \<union> (A-C)"
by blast
-lemma Un_Diff: "(A Un B) - C = (A - C) Un (B - C)"
+lemma Un_Diff: "(A \<union> B) - C = (A - C) \<union> (B - C)"
by blast
-lemma Int_Diff: "(A Int B) - C = A Int (B - C)"
+lemma Int_Diff: "(A \<inter> B) - C = A \<inter> (B - C)"
by blast
-lemma Diff_Int_distrib: "C Int (A-B) = (C Int A) - (C Int B)"
+lemma Diff_Int_distrib: "C \<inter> (A-B) = (C \<inter> A) - (C \<inter> B)"
by blast
-lemma Diff_Int_distrib2: "(A-B) Int C = (A Int C) - (B Int C)"
+lemma Diff_Int_distrib2: "(A-B) \<inter> C = (A \<inter> C) - (B \<inter> C)"
by blast
(*Halmos, Naive Set Theory, page 16.*)
-lemma Un_Int_assoc_iff: "(A Int B) Un C = A Int (B Un C) <-> C\<subseteq>A"
+lemma Un_Int_assoc_iff: "(A \<inter> B) \<union> C = A \<inter> (B \<union> C) <-> C\<subseteq>A"
by (blast elim!: equalityE)
@@ -346,42 +346,42 @@
(** Big Union is the least upper bound of a set **)
-lemma Union_subset_iff: "Union(A) \<subseteq> C <-> (\<forall>x\<in>A. x \<subseteq> C)"
+lemma Union_subset_iff: "\<Union>(A) \<subseteq> C <-> (\<forall>x\<in>A. x \<subseteq> C)"
by blast
-lemma Union_upper: "B\<in>A ==> B \<subseteq> Union(A)"
+lemma Union_upper: "B\<in>A ==> B \<subseteq> \<Union>(A)"
by blast
-lemma Union_least: "[| !!x. x\<in>A ==> x\<subseteq>C |] ==> Union(A) \<subseteq> C"
+lemma Union_least: "[| !!x. x\<in>A ==> x\<subseteq>C |] ==> \<Union>(A) \<subseteq> C"
by blast
-lemma Union_cons [simp]: "Union(cons(a,B)) = a Un Union(B)"
+lemma Union_cons [simp]: "\<Union>(cons(a,B)) = a \<union> \<Union>(B)"
by blast
-lemma Union_Un_distrib: "Union(A Un B) = Union(A) Un Union(B)"
+lemma Union_Un_distrib: "\<Union>(A \<union> B) = \<Union>(A) \<union> \<Union>(B)"
by blast
-lemma Union_Int_subset: "Union(A Int B) \<subseteq> Union(A) Int Union(B)"
+lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>(A) \<inter> \<Union>(B)"
by blast
-lemma Union_disjoint: "Union(C) Int A = 0 <-> (\<forall>B\<in>C. B Int A = 0)"
+lemma Union_disjoint: "\<Union>(C) \<inter> A = 0 <-> (\<forall>B\<in>C. B \<inter> A = 0)"
by (blast elim!: equalityE)
-lemma Union_empty_iff: "Union(A) = 0 <-> (\<forall>B\<in>A. B=0)"
+lemma Union_empty_iff: "\<Union>(A) = 0 <-> (\<forall>B\<in>A. B=0)"
by blast
-lemma Int_Union2: "Union(B) Int A = (\<Union>C\<in>B. C Int A)"
+lemma Int_Union2: "\<Union>(B) \<inter> A = (\<Union>C\<in>B. C \<inter> A)"
by blast
(** Big Intersection is the greatest lower bound of a nonempty set **)
-lemma Inter_subset_iff: "A\<noteq>0 ==> C \<subseteq> Inter(A) <-> (\<forall>x\<in>A. C \<subseteq> x)"
+lemma Inter_subset_iff: "A\<noteq>0 ==> C \<subseteq> \<Inter>(A) <-> (\<forall>x\<in>A. C \<subseteq> x)"
by blast
-lemma Inter_lower: "B\<in>A ==> Inter(A) \<subseteq> B"
+lemma Inter_lower: "B\<in>A ==> \<Inter>(A) \<subseteq> B"
by blast
-lemma Inter_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>x |] ==> C \<subseteq> Inter(A)"
+lemma Inter_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>x |] ==> C \<subseteq> \<Inter>(A)"
by blast
(** Intersection of a family of sets **)
@@ -392,31 +392,31 @@
lemma INT_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>B(x) |] ==> C \<subseteq> (\<Inter>x\<in>A. B(x))"
by force
-lemma Inter_0 [simp]: "Inter(0) = 0"
+lemma Inter_0 [simp]: "\<Inter>(0) = 0"
by (unfold Inter_def, blast)
lemma Inter_Un_subset:
- "[| z\<in>A; z\<in>B |] ==> Inter(A) Un Inter(B) \<subseteq> Inter(A Int B)"
+ "[| z\<in>A; z\<in>B |] ==> \<Inter>(A) \<union> \<Inter>(B) \<subseteq> \<Inter>(A \<inter> B)"
by blast
(* A good challenge: Inter is ill-behaved on the empty set *)
lemma Inter_Un_distrib:
- "[| A\<noteq>0; B\<noteq>0 |] ==> Inter(A Un B) = Inter(A) Int Inter(B)"
+ "[| A\<noteq>0; B\<noteq>0 |] ==> \<Inter>(A \<union> B) = \<Inter>(A) \<inter> \<Inter>(B)"
by blast
-lemma Union_singleton: "Union({b}) = b"
+lemma Union_singleton: "\<Union>({b}) = b"
by blast
-lemma Inter_singleton: "Inter({b}) = b"
+lemma Inter_singleton: "\<Inter>({b}) = b"
by blast
lemma Inter_cons [simp]:
- "Inter(cons(a,B)) = (if B=0 then a else a Int Inter(B))"
+ "\<Inter>(cons(a,B)) = (if B=0 then a else a \<inter> \<Inter>(B))"
by force
subsection{*Unions and Intersections of Families*}
-lemma subset_UN_iff_eq: "A \<subseteq> (\<Union>i\<in>I. B(i)) <-> A = (\<Union>i\<in>I. A Int B(i))"
+lemma subset_UN_iff_eq: "A \<subseteq> (\<Union>i\<in>I. B(i)) <-> A = (\<Union>i\<in>I. A \<inter> B(i))"
by (blast elim!: equalityE)
lemma UN_subset_iff: "(\<Union>x\<in>A. B(x)) \<subseteq> C <-> (\<forall>x\<in>A. B(x) \<subseteq> C)"
@@ -428,10 +428,10 @@
lemma UN_least: "[| !!x. x\<in>A ==> B(x)\<subseteq>C |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> C"
by blast
-lemma Union_eq_UN: "Union(A) = (\<Union>x\<in>A. x)"
+lemma Union_eq_UN: "\<Union>(A) = (\<Union>x\<in>A. x)"
by blast
-lemma Inter_eq_INT: "Inter(A) = (\<Inter>x\<in>A. x)"
+lemma Inter_eq_INT: "\<Inter>(A) = (\<Inter>x\<in>A. x)"
by (unfold Inter_def, blast)
lemma UN_0 [simp]: "(\<Union>i\<in>0. A(i)) = 0"
@@ -440,31 +440,31 @@
lemma UN_singleton: "(\<Union>x\<in>A. {x}) = A"
by blast
-lemma UN_Un: "(\<Union>i\<in> A Un B. C(i)) = (\<Union>i\<in> A. C(i)) Un (\<Union>i\<in>B. C(i))"
+lemma UN_Un: "(\<Union>i\<in> A \<union> B. C(i)) = (\<Union>i\<in> A. C(i)) \<union> (\<Union>i\<in>B. C(i))"
by blast
-lemma INT_Un: "(\<Inter>i\<in>I Un J. A(i)) =
- (if I=0 then \<Inter>j\<in>J. A(j)
- else if J=0 then \<Inter>i\<in>I. A(i)
- else ((\<Inter>i\<in>I. A(i)) Int (\<Inter>j\<in>J. A(j))))"
+lemma INT_Un: "(\<Inter>i\<in>I \<union> J. A(i)) =
+ (if I=0 then \<Inter>j\<in>J. A(j)
+ else if J=0 then \<Inter>i\<in>I. A(i)
+ else ((\<Inter>i\<in>I. A(i)) \<inter> (\<Inter>j\<in>J. A(j))))"
by (simp, blast intro!: equalityI)
lemma UN_UN_flatten: "(\<Union>x \<in> (\<Union>y\<in>A. B(y)). C(x)) = (\<Union>y\<in>A. \<Union>x\<in> B(y). C(x))"
by blast
(*Halmos, Naive Set Theory, page 35.*)
-lemma Int_UN_distrib: "B Int (\<Union>i\<in>I. A(i)) = (\<Union>i\<in>I. B Int A(i))"
+lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A(i)) = (\<Union>i\<in>I. B \<inter> A(i))"
by blast
-lemma Un_INT_distrib: "I\<noteq>0 ==> B Un (\<Inter>i\<in>I. A(i)) = (\<Inter>i\<in>I. B Un A(i))"
+lemma Un_INT_distrib: "I\<noteq>0 ==> B \<union> (\<Inter>i\<in>I. A(i)) = (\<Inter>i\<in>I. B \<union> A(i))"
by auto
lemma Int_UN_distrib2:
- "(\<Union>i\<in>I. A(i)) Int (\<Union>j\<in>J. B(j)) = (\<Union>i\<in>I. \<Union>j\<in>J. A(i) Int B(j))"
+ "(\<Union>i\<in>I. A(i)) \<inter> (\<Union>j\<in>J. B(j)) = (\<Union>i\<in>I. \<Union>j\<in>J. A(i) \<inter> B(j))"
by blast
-lemma Un_INT_distrib2: "[| I\<noteq>0; J\<noteq>0 |] ==>
- (\<Inter>i\<in>I. A(i)) Un (\<Inter>j\<in>J. B(j)) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A(i) Un B(j))"
+lemma Un_INT_distrib2: "[| I\<noteq>0; J\<noteq>0 |] ==>
+ (\<Inter>i\<in>I. A(i)) \<union> (\<Inter>j\<in>J. B(j)) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A(i) \<union> B(j))"
by auto
lemma UN_constant [simp]: "(\<Union>y\<in>A. c) = (if A=0 then 0 else c)"
@@ -480,33 +480,33 @@
by (auto simp add: Inter_def)
lemma INT_Union_eq:
- "0 ~: A ==> (\<Inter>x\<in> Union(A). B(x)) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B(x))"
-apply (subgoal_tac "\<forall>x\<in>A. x~=0")
+ "0 \<notin> A ==> (\<Inter>x\<in> \<Union>(A). B(x)) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B(x))"
+apply (subgoal_tac "\<forall>x\<in>A. x\<noteq>0")
prefer 2 apply blast
-apply (force simp add: Inter_def ball_conj_distrib)
+apply (force simp add: Inter_def ball_conj_distrib)
done
lemma INT_UN_eq:
- "(\<forall>x\<in>A. B(x) ~= 0)
+ "(\<forall>x\<in>A. B(x) \<noteq> 0)
==> (\<Inter>z\<in> (\<Union>x\<in>A. B(x)). C(z)) = (\<Inter>x\<in>A. \<Inter>z\<in> B(x). C(z))"
apply (subst INT_Union_eq, blast)
apply (simp add: Inter_def)
done
-(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:
+(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:
Union of a family of unions **)
lemma UN_Un_distrib:
- "(\<Union>i\<in>I. A(i) Un B(i)) = (\<Union>i\<in>I. A(i)) Un (\<Union>i\<in>I. B(i))"
+ "(\<Union>i\<in>I. A(i) \<union> B(i)) = (\<Union>i\<in>I. A(i)) \<union> (\<Union>i\<in>I. B(i))"
by blast
lemma INT_Int_distrib:
- "I\<noteq>0 ==> (\<Inter>i\<in>I. A(i) Int B(i)) = (\<Inter>i\<in>I. A(i)) Int (\<Inter>i\<in>I. B(i))"
+ "I\<noteq>0 ==> (\<Inter>i\<in>I. A(i) \<inter> B(i)) = (\<Inter>i\<in>I. A(i)) \<inter> (\<Inter>i\<in>I. B(i))"
by (blast elim!: not_emptyE)
lemma UN_Int_subset:
- "(\<Union>z\<in>I Int J. A(z)) \<subseteq> (\<Union>z\<in>I. A(z)) Int (\<Union>z\<in>J. A(z))"
+ "(\<Union>z\<in>I \<inter> J. A(z)) \<subseteq> (\<Union>z\<in>I. A(z)) \<inter> (\<Union>z\<in>J. A(z))"
by blast
(** Devlin, page 12, exercise 5: Complements **)
@@ -521,17 +521,17 @@
(** Unions and Intersections with General Sum **)
(*Not suitable for rewriting: LOOPS!*)
-lemma Sigma_cons1: "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)"
+lemma Sigma_cons1: "Sigma(cons(a,B), C) = ({a}*C(a)) \<union> Sigma(B,C)"
by blast
(*Not suitable for rewriting: LOOPS!*)
-lemma Sigma_cons2: "A * cons(b,B) = A*{b} Un A*B"
+lemma Sigma_cons2: "A * cons(b,B) = A*{b} \<union> A*B"
by blast
-lemma Sigma_succ1: "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)"
+lemma Sigma_succ1: "Sigma(succ(A), B) = ({A}*B(A)) \<union> Sigma(A,B)"
by blast
-lemma Sigma_succ2: "A * succ(B) = A*{B} Un A*B"
+lemma Sigma_succ2: "A * succ(B) = A*{B} \<union> A*B"
by blast
lemma SUM_UN_distrib1:
@@ -543,27 +543,27 @@
by blast
lemma SUM_Un_distrib1:
- "(\<Sigma> i\<in>I Un J. C(i)) = (\<Sigma> i\<in>I. C(i)) Un (\<Sigma> j\<in>J. C(j))"
+ "(\<Sigma> i\<in>I \<union> J. C(i)) = (\<Sigma> i\<in>I. C(i)) \<union> (\<Sigma> j\<in>J. C(j))"
by blast
lemma SUM_Un_distrib2:
- "(\<Sigma> i\<in>I. A(i) Un B(i)) = (\<Sigma> i\<in>I. A(i)) Un (\<Sigma> i\<in>I. B(i))"
+ "(\<Sigma> i\<in>I. A(i) \<union> B(i)) = (\<Sigma> i\<in>I. A(i)) \<union> (\<Sigma> i\<in>I. B(i))"
by blast
(*First-order version of the above, for rewriting*)
-lemma prod_Un_distrib2: "I * (A Un B) = I*A Un I*B"
+lemma prod_Un_distrib2: "I * (A \<union> B) = I*A \<union> I*B"
by (rule SUM_Un_distrib2)
lemma SUM_Int_distrib1:
- "(\<Sigma> i\<in>I Int J. C(i)) = (\<Sigma> i\<in>I. C(i)) Int (\<Sigma> j\<in>J. C(j))"
+ "(\<Sigma> i\<in>I \<inter> J. C(i)) = (\<Sigma> i\<in>I. C(i)) \<inter> (\<Sigma> j\<in>J. C(j))"
by blast
lemma SUM_Int_distrib2:
- "(\<Sigma> i\<in>I. A(i) Int B(i)) = (\<Sigma> i\<in>I. A(i)) Int (\<Sigma> i\<in>I. B(i))"
+ "(\<Sigma> i\<in>I. A(i) \<inter> B(i)) = (\<Sigma> i\<in>I. A(i)) \<inter> (\<Sigma> i\<in>I. B(i))"
by blast
(*First-order version of the above, for rewriting*)
-lemma prod_Int_distrib2: "I * (A Int B) = I*A Int I*B"
+lemma prod_Int_distrib2: "I * (A \<inter> B) = I*A \<inter> I*B"
by (rule SUM_Int_distrib2)
(*Cf Aczel, Non-Well-Founded Sets, page 115*)
@@ -575,12 +575,12 @@
by blast
lemma Int_Sigma_eq:
- "(\<Sigma> x \<in> A'. B'(x)) Int (\<Sigma> x \<in> A. B(x)) = (\<Sigma> x \<in> A' Int A. B'(x) Int B(x))"
+ "(\<Sigma> x \<in> A'. B'(x)) \<inter> (\<Sigma> x \<in> A. B(x)) = (\<Sigma> x \<in> A' \<inter> A. B'(x) \<inter> B(x))"
by blast
(** Domain **)
-lemma domain_iff: "a: domain(r) <-> (EX y. <a,y>\<in> r)"
+lemma domain_iff: "a: domain(r) <-> (\<exists>y. <a,y>\<in> r)"
by (unfold domain_def, blast)
lemma domainI [intro]: "<a,b>\<in> r ==> a: domain(r)"
@@ -602,10 +602,10 @@
lemma domain_cons [simp]: "domain(cons(<a,b>,r)) = cons(a, domain(r))"
by blast
-lemma domain_Un_eq [simp]: "domain(A Un B) = domain(A) Un domain(B)"
+lemma domain_Un_eq [simp]: "domain(A \<union> B) = domain(A) \<union> domain(B)"
by blast
-lemma domain_Int_subset: "domain(A Int B) \<subseteq> domain(A) Int domain(B)"
+lemma domain_Int_subset: "domain(A \<inter> B) \<subseteq> domain(A) \<inter> domain(B)"
by blast
lemma domain_Diff_subset: "domain(A) - domain(B) \<subseteq> domain(A - B)"
@@ -614,7 +614,7 @@
lemma domain_UN: "domain(\<Union>x\<in>A. B(x)) = (\<Union>x\<in>A. domain(B(x)))"
by blast
-lemma domain_Union: "domain(Union(A)) = (\<Union>x\<in>A. domain(x))"
+lemma domain_Union: "domain(\<Union>(A)) = (\<Union>x\<in>A. domain(x))"
by blast
@@ -643,10 +643,10 @@
lemma range_cons [simp]: "range(cons(<a,b>,r)) = cons(b, range(r))"
by blast
-lemma range_Un_eq [simp]: "range(A Un B) = range(A) Un range(B)"
+lemma range_Un_eq [simp]: "range(A \<union> B) = range(A) \<union> range(B)"
by blast
-lemma range_Int_subset: "range(A Int B) \<subseteq> range(A) Int range(B)"
+lemma range_Int_subset: "range(A \<inter> B) \<subseteq> range(A) \<inter> range(B)"
by blast
lemma range_Diff_subset: "range(A) - range(B) \<subseteq> range(A - B)"
@@ -667,18 +667,18 @@
lemma fieldI2: "<a,b>\<in> r ==> b \<in> field(r)"
by (unfold field_def, blast)
-lemma fieldCI [intro]:
+lemma fieldCI [intro]:
"(~ <c,a>\<in>r ==> <a,b>\<in> r) ==> a \<in> field(r)"
apply (unfold field_def, blast)
done
-lemma fieldE [elim!]:
- "[| a \<in> field(r);
- !!x. <a,x>\<in> r ==> P;
+lemma fieldE [elim!]:
+ "[| a \<in> field(r);
+ !!x. <a,x>\<in> r ==> P;
!!x. <x,a>\<in> r ==> P |] ==> P"
by (unfold field_def, blast)
-lemma field_subset: "field(A*B) \<subseteq> A Un B"
+lemma field_subset: "field(A*B) \<subseteq> A \<union> B"
by blast
lemma domain_subset_field: "domain(r) \<subseteq> field(r)"
@@ -698,7 +698,7 @@
by blast
lemma relation_field_times_field: "relation(r) ==> r \<subseteq> field(r)*field(r)"
-by (simp add: relation_def, blast)
+by (simp add: relation_def, blast)
lemma field_of_prod: "field(A*A) = A"
by blast
@@ -709,10 +709,10 @@
lemma field_cons [simp]: "field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"
by blast
-lemma field_Un_eq [simp]: "field(A Un B) = field(A) Un field(B)"
+lemma field_Un_eq [simp]: "field(A \<union> B) = field(A) \<union> field(B)"
by blast
-lemma field_Int_subset: "field(A Int B) \<subseteq> field(A) Int field(B)"
+lemma field_Int_subset: "field(A \<inter> B) \<subseteq> field(A) \<inter> field(B)"
by blast
lemma field_Diff_subset: "field(A) - field(B) \<subseteq> field(A - B)"
@@ -722,18 +722,18 @@
by blast
(** The Union of a set of relations is a relation -- Lemma for fun_Union **)
-lemma rel_Union: "(\<forall>x\<in>S. EX A B. x \<subseteq> A*B) ==>
- Union(S) \<subseteq> domain(Union(S)) * range(Union(S))"
+lemma rel_Union: "(\<forall>x\<in>S. \<exists>A B. x \<subseteq> A*B) ==>
+ \<Union>(S) \<subseteq> domain(\<Union>(S)) * range(\<Union>(S))"
by blast
(** The Union of 2 relations is a relation (Lemma for fun_Un) **)
-lemma rel_Un: "[| r \<subseteq> A*B; s \<subseteq> C*D |] ==> (r Un s) \<subseteq> (A Un C) * (B Un D)"
+lemma rel_Un: "[| r \<subseteq> A*B; s \<subseteq> C*D |] ==> (r \<union> s) \<subseteq> (A \<union> C) * (B \<union> D)"
by blast
-lemma domain_Diff_eq: "[| <a,c> \<in> r; c~=b |] ==> domain(r-{<a,b>}) = domain(r)"
+lemma domain_Diff_eq: "[| <a,c> \<in> r; c\<noteq>b |] ==> domain(r-{<a,b>}) = domain(r)"
by blast
-lemma range_Diff_eq: "[| <c,b> \<in> r; c~=a |] ==> range(r-{<a,b>}) = range(r)"
+lemma range_Diff_eq: "[| <c,b> \<in> r; c\<noteq>a |] ==> range(r-{<a,b>}) = range(r)"
by blast
@@ -748,7 +748,7 @@
lemma imageI [intro]: "[| <a,b>\<in> r; a\<in>A |] ==> b \<in> r``A"
by (unfold image_def, blast)
-lemma imageE [elim!]:
+lemma imageE [elim!]:
"[| b: r``A; !!x.[| <x,b>\<in> r; x\<in>A |] ==> P |] ==> P"
by (unfold image_def, blast)
@@ -758,7 +758,7 @@
lemma image_0 [simp]: "r``0 = 0"
by blast
-lemma image_Un [simp]: "r``(A Un B) = (r``A) Un (r``B)"
+lemma image_Un [simp]: "r``(A \<union> B) = (r``A) \<union> (r``B)"
by blast
lemma image_UN: "r `` (\<Union>x\<in>A. B(x)) = (\<Union>x\<in>A. r `` B(x))"
@@ -768,13 +768,13 @@
"{z \<in> Sigma(A,B). P(z)} `` C = (\<Union>x \<in> A. {y \<in> B(x). x \<in> C & P(<x,y>)})"
by blast
-lemma image_Int_subset: "r``(A Int B) \<subseteq> (r``A) Int (r``B)"
+lemma image_Int_subset: "r``(A \<inter> B) \<subseteq> (r``A) \<inter> (r``B)"
by blast
-lemma image_Int_square_subset: "(r Int A*A)``B \<subseteq> (r``B) Int A"
+lemma image_Int_square_subset: "(r \<inter> A*A)``B \<subseteq> (r``B) \<inter> A"
by blast
-lemma image_Int_square: "B\<subseteq>A ==> (r Int A*A)``B = (r``B) Int A"
+lemma image_Int_square: "B\<subseteq>A ==> (r \<inter> A*A)``B = (r``B) \<inter> A"
by blast
@@ -782,16 +782,16 @@
lemma image_0_left [simp]: "0``A = 0"
by blast
-lemma image_Un_left: "(r Un s)``A = (r``A) Un (s``A)"
+lemma image_Un_left: "(r \<union> s)``A = (r``A) \<union> (s``A)"
by blast
-lemma image_Int_subset_left: "(r Int s)``A \<subseteq> (r``A) Int (s``A)"
+lemma image_Int_subset_left: "(r \<inter> s)``A \<subseteq> (r``A) \<inter> (s``A)"
by blast
subsection{*Inverse Image of a Set under a Function or Relation*}
-lemma vimage_iff:
+lemma vimage_iff:
"a \<in> r-``B <-> (\<exists>y\<in>B. <a,y>\<in>r)"
by (unfold vimage_def image_def converse_def, blast)
@@ -801,7 +801,7 @@
lemma vimageI [intro]: "[| <a,b>\<in> r; b\<in>B |] ==> a \<in> r-``B"
by (unfold vimage_def, blast)
-lemma vimageE [elim!]:
+lemma vimageE [elim!]:
"[| a: r-``B; !!x.[| <a,x>\<in> r; x\<in>B |] ==> P |] ==> P"
apply (unfold vimage_def, blast)
done
@@ -814,10 +814,10 @@
lemma vimage_0 [simp]: "r-``0 = 0"
by blast
-lemma vimage_Un [simp]: "r-``(A Un B) = (r-``A) Un (r-``B)"
+lemma vimage_Un [simp]: "r-``(A \<union> B) = (r-``A) \<union> (r-``B)"
by blast
-lemma vimage_Int_subset: "r-``(A Int B) \<subseteq> (r-``A) Int (r-``B)"
+lemma vimage_Int_subset: "r-``(A \<inter> B) \<subseteq> (r-``A) \<inter> (r-``B)"
by blast
(*NOT suitable for rewriting*)
@@ -825,7 +825,7 @@
by blast
lemma function_vimage_Int:
- "function(f) ==> f-``(A Int B) = (f-``A) Int (f-``B)"
+ "function(f) ==> f-``(A \<inter> B) = (f-``A) \<inter> (f-``B)"
by (unfold function_def, blast)
lemma function_vimage_Diff: "function(f) ==> f-``(A-B) = (f-``A) - (f-``B)"
@@ -834,10 +834,10 @@
lemma function_image_vimage: "function(f) ==> f `` (f-`` A) \<subseteq> A"
by (unfold function_def, blast)
-lemma vimage_Int_square_subset: "(r Int A*A)-``B \<subseteq> (r-``B) Int A"
+lemma vimage_Int_square_subset: "(r \<inter> A*A)-``B \<subseteq> (r-``B) \<inter> A"
by blast
-lemma vimage_Int_square: "B\<subseteq>A ==> (r Int A*A)-``B = (r-``B) Int A"
+lemma vimage_Int_square: "B\<subseteq>A ==> (r \<inter> A*A)-``B = (r-``B) \<inter> A"
by blast
@@ -846,19 +846,19 @@
lemma vimage_0_left [simp]: "0-``A = 0"
by blast
-lemma vimage_Un_left: "(r Un s)-``A = (r-``A) Un (s-``A)"
+lemma vimage_Un_left: "(r \<union> s)-``A = (r-``A) \<union> (s-``A)"
by blast
-lemma vimage_Int_subset_left: "(r Int s)-``A \<subseteq> (r-``A) Int (s-``A)"
+lemma vimage_Int_subset_left: "(r \<inter> s)-``A \<subseteq> (r-``A) \<inter> (s-``A)"
by blast
(** Converse **)
-lemma converse_Un [simp]: "converse(A Un B) = converse(A) Un converse(B)"
+lemma converse_Un [simp]: "converse(A \<union> B) = converse(A) \<union> converse(B)"
by blast
-lemma converse_Int [simp]: "converse(A Int B) = converse(A) Int converse(B)"
+lemma converse_Int [simp]: "converse(A \<inter> B) = converse(A) \<inter> converse(B)"
by blast
lemma converse_Diff [simp]: "converse(A - B) = converse(A) - converse(B)"
@@ -879,28 +879,28 @@
lemma Pow_0 [simp]: "Pow(0) = {0}"
by blast
-lemma Pow_insert: "Pow (cons(a,A)) = Pow(A) Un {cons(a,X) . X: Pow(A)}"
+lemma Pow_insert: "Pow (cons(a,A)) = Pow(A) \<union> {cons(a,X) . X: Pow(A)}"
apply (rule equalityI, safe)
apply (erule swap)
-apply (rule_tac a = "x-{a}" in RepFun_eqI, auto)
+apply (rule_tac a = "x-{a}" in RepFun_eqI, auto)
done
-lemma Un_Pow_subset: "Pow(A) Un Pow(B) \<subseteq> Pow(A Un B)"
+lemma Un_Pow_subset: "Pow(A) \<union> Pow(B) \<subseteq> Pow(A \<union> B)"
by blast
lemma UN_Pow_subset: "(\<Union>x\<in>A. Pow(B(x))) \<subseteq> Pow(\<Union>x\<in>A. B(x))"
by blast
-lemma subset_Pow_Union: "A \<subseteq> Pow(Union(A))"
+lemma subset_Pow_Union: "A \<subseteq> Pow(\<Union>(A))"
by blast
-lemma Union_Pow_eq [simp]: "Union(Pow(A)) = A"
+lemma Union_Pow_eq [simp]: "\<Union>(Pow(A)) = A"
by blast
-lemma Union_Pow_iff: "Union(A) \<in> Pow(B) <-> A \<in> Pow(Pow(B))"
+lemma Union_Pow_iff: "\<Union>(A) \<in> Pow(B) <-> A \<in> Pow(Pow(B))"
by blast
-lemma Pow_Int_eq [simp]: "Pow(A Int B) = Pow(A) Int Pow(B)"
+lemma Pow_Int_eq [simp]: "Pow(A \<inter> B) = Pow(A) \<inter> Pow(B)"
by blast
lemma Pow_INT_eq: "A\<noteq>0 ==> Pow(\<Inter>x\<in>A. B(x)) = (\<Inter>x\<in>A. Pow(B(x)))"
@@ -924,20 +924,20 @@
lemma Collect_subset: "Collect(A,P) \<subseteq> A"
by blast
-lemma Collect_Un: "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)"
+lemma Collect_Un: "Collect(A \<union> B, P) = Collect(A,P) \<union> Collect(B,P)"
by blast
-lemma Collect_Int: "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)"
+lemma Collect_Int: "Collect(A \<inter> B, P) = Collect(A,P) \<inter> Collect(B,P)"
by blast
lemma Collect_Diff: "Collect(A - B, P) = Collect(A,P) - Collect(B,P)"
by blast
-lemma Collect_cons: "{x\<in>cons(a,B). P(x)} =
+lemma Collect_cons: "{x\<in>cons(a,B). P(x)} =
(if P(a) then cons(a, {x\<in>B. P(x)}) else {x\<in>B. P(x)})"
by (simp, blast)
-lemma Int_Collect_self_eq: "A Int Collect(A,P) = Collect(A,P)"
+lemma Int_Collect_self_eq: "A \<inter> Collect(A,P) = Collect(A,P)"
by blast
lemma Collect_Collect_eq [simp]:
@@ -945,27 +945,27 @@
by blast
lemma Collect_Int_Collect_eq:
- "Collect(A,P) Int Collect(A,Q) = Collect(A, %x. P(x) & Q(x))"
+ "Collect(A,P) \<inter> Collect(A,Q) = Collect(A, %x. P(x) & Q(x))"
by blast
lemma Collect_Union_eq [simp]:
"Collect(\<Union>x\<in>A. B(x), P) = (\<Union>x\<in>A. Collect(B(x), P))"
by blast
-lemma Collect_Int_left: "{x\<in>A. P(x)} Int B = {x \<in> A Int B. P(x)}"
+lemma Collect_Int_left: "{x\<in>A. P(x)} \<inter> B = {x \<in> A \<inter> B. P(x)}"
by blast
-lemma Collect_Int_right: "A Int {x\<in>B. P(x)} = {x \<in> A Int B. P(x)}"
+lemma Collect_Int_right: "A \<inter> {x\<in>B. P(x)} = {x \<in> A \<inter> B. P(x)}"
by blast
-lemma Collect_disj_eq: "{x\<in>A. P(x) | Q(x)} = Collect(A, P) Un Collect(A, Q)"
+lemma Collect_disj_eq: "{x\<in>A. P(x) | Q(x)} = Collect(A, P) \<union> Collect(A, Q)"
by blast
-lemma Collect_conj_eq: "{x\<in>A. P(x) & Q(x)} = Collect(A, P) Int Collect(A, Q)"
+lemma Collect_conj_eq: "{x\<in>A. P(x) & Q(x)} = Collect(A, P) \<inter> Collect(A, Q)"
by blast
-lemmas subset_SIs = subset_refl cons_subsetI subset_consI
- Union_least UN_least Un_least
+lemmas subset_SIs = subset_refl cons_subsetI subset_consI
+ Union_least UN_least Un_least
Inter_greatest Int_greatest RepFun_subset
Un_upper1 Un_upper2 Int_lower1 Int_lower2
@@ -979,6 +979,6 @@
val ZF_cs = claset_of (@{context} delrules [@{thm equalityI}]);
*}
-
+
end
--- a/src/ZF/ex/Group.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/ex/Group.thy Tue Mar 06 15:15:49 2012 +0000
@@ -412,7 +412,7 @@
qed
text{*Basis for homomorphism proofs: we assume two groups @{term G} and
- @term{H}, with a homomorphism @{term h} between them*}
+ @{term H}, with a homomorphism @{term h} between them*}
locale group_hom = G: group G + H: group H
for G (structure) and H (structure) and h +
assumes homh: "h \<in> hom(G,H)"
--- a/src/ZF/func.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/func.thy Tue Mar 06 15:15:49 2012 +0000
@@ -9,7 +9,7 @@
subsection{*The Pi Operator: Dependent Function Space*}
-lemma subset_Sigma_imp_relation: "r <= Sigma(A,B) ==> relation(r)"
+lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) ==> relation(r)"
by (simp add: relation_def, blast)
lemma relation_converse_converse [simp]:
@@ -25,7 +25,7 @@
(*For upward compatibility with the former definition*)
lemma Pi_iff_old:
- "f: Pi(A,B) <-> f<=Sigma(A,B) & (ALL x:A. EX! y. <x,y>: f)"
+ "f: Pi(A,B) <-> f<=Sigma(A,B) & (\<forall>x\<in>A. EX! y. <x,y>: f)"
by (unfold Pi_def function_def, blast)
lemma fun_is_function: "f: Pi(A,B) ==> function(f)"
@@ -40,7 +40,7 @@
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 \<subseteq> Sigma(A,B)"
by (unfold Pi_def, blast)
lemma Pi_cong:
@@ -69,15 +69,15 @@
done
(*Applying a function outside its domain yields 0*)
-lemma apply_0: "a ~: domain(f) ==> f`a = 0"
+lemma apply_0: "a \<notin> domain(f) ==> f`a = 0"
by (unfold apply_def, blast)
-lemma Pi_memberD: "[| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>"
+lemma Pi_memberD: "[| f: Pi(A,B); c: f |] ==> \<exists>x\<in>A. c = <x,f`x>"
apply (frule fun_is_rel)
apply (blast dest: apply_equality)
done
-lemma function_apply_Pair: "[| function(f); a : domain(f)|] ==> <a,f`a>: f"
+lemma function_apply_Pair: "[| function(f); a \<in> domain(f)|] ==> <a,f`a>: f"
apply (simp add: function_def, clarify)
apply (subgoal_tac "f`a = y", blast)
apply (simp add: apply_def, blast)
@@ -89,11 +89,11 @@
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 \<in> 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 \<in> B"
by (blast dest: apply_type)
lemma apply_iff: "f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b"
@@ -102,28 +102,28 @@
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 \<in> B(x) |] ==> f \<in> 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,B) & (ALL x: A. P(x, f`x))"
+ "(f \<in> Pi(A, %x. {y:B(x). P(x,y)}))
+ <-> f \<in> Pi(A,B) & (\<forall>x\<in>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 \<in> Pi(A,B); !!x. x:A ==> B(x)<=C(x) |] ==> f \<in> Pi(A,C)"
by (blast intro: Pi_type dest: apply_type)
(** Elimination of membership in a function **)
-lemma domain_type: "[| <a,b> : f; f: Pi(A,B) |] ==> a : A"
+lemma domain_type: "[| <a,b> \<in> f; f: Pi(A,B) |] ==> a \<in> A"
by (blast dest: fun_is_rel)
-lemma range_type: "[| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)"
+lemma range_type: "[| <a,b> \<in> f; f: Pi(A,B) |] ==> b \<in> B(a)"
by (blast dest: fun_is_rel)
lemma Pair_mem_PiD: "[| <a,b>: f; f: Pi(A,B) |] ==> a:A & b:B(a) & f`a = b"
@@ -131,39 +131,39 @@
subsection{*Lambda Abstraction*}
-lemma lamI: "a:A ==> <a,b(a)> : (lam x:A. b(x))"
+lemma lamI: "a:A ==> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
apply (unfold lam_def)
apply (erule RepFunI)
done
lemma lamE:
- "[| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P
+ "[| p: (\<lambda>x\<in>A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P
|] ==> P"
by (simp add: lam_def, blast)
-lemma lamD: "[| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)"
+lemma lamD: "[| <a,c>: (\<lambda>x\<in>A. b(x)) |] ==> c = b(a)"
by (simp add: lam_def)
lemma lam_type [TC]:
- "[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A. b(x)) : Pi(A,B)"
+ "[| !!x. x:A ==> b(x): B(x) |] ==> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
by (simp add: lam_def Pi_def function_def, blast)
-lemma lam_funtype: "(lam x:A. b(x)) : A -> {b(x). x:A}"
+lemma lam_funtype: "(\<lambda>x\<in>A. b(x)) \<in> A -> {b(x). x:A}"
by (blast intro: lam_type)
-lemma function_lam: "function (lam x:A. b(x))"
+lemma function_lam: "function (\<lambda>x\<in>A. b(x))"
by (simp add: function_def lam_def)
-lemma relation_lam: "relation (lam x:A. b(x))"
+lemma relation_lam: "relation (\<lambda>x\<in>A. b(x))"
by (simp add: relation_def lam_def)
-lemma beta_if [simp]: "(lam x:A. b(x)) ` a = (if a : A then b(a) else 0)"
+lemma beta_if [simp]: "(\<lambda>x\<in>A. b(x)) ` a = (if a \<in> A then b(a) else 0)"
by (simp add: apply_def lam_def, blast)
-lemma beta: "a : A ==> (lam x:A. b(x)) ` a = b(a)"
+lemma beta: "a \<in> A ==> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
by (simp add: apply_def lam_def, blast)
-lemma lam_empty [simp]: "(lam x:0. b(x)) = 0"
+lemma lam_empty [simp]: "(\<lambda>x\<in>0. b(x)) = 0"
by (simp add: lam_def)
lemma domain_lam [simp]: "domain(Lambda(A,b)) = A"
@@ -175,13 +175,13 @@
by (simp only: lam_def cong add: RepFun_cong)
lemma lam_theI:
- "(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)"
-apply (rule_tac x = "lam x: A. THE y. Q (x,y)" in exI)
+ "(!!x. x:A ==> EX! y. Q(x,y)) ==> \<exists>f. \<forall>x\<in>A. Q(x, f`x)"
+apply (rule_tac x = "\<lambda>x\<in>A. THE y. Q (x,y)" in exI)
apply simp
apply (blast intro: theI)
done
-lemma lam_eqE: "[| (lam x:A. f(x)) = (lam x:A. g(x)); a:A |] ==> f(a)=g(a)"
+lemma lam_eqE: "[| (\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x)); a:A |] ==> f(a)=g(a)"
by (fast intro!: lamI elim: equalityE lamE)
@@ -190,7 +190,7 @@
by (unfold Pi_def function_def, blast)
(*The singleton function*)
-lemma singleton_fun [simp]: "{<a,b>} : {a} -> {b}"
+lemma singleton_fun [simp]: "{<a,b>} \<in> {a} -> {b}"
by (unfold Pi_def function_def, blast)
lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)"
@@ -207,26 +207,26 @@
(*Semi-extensionality!*)
lemma fun_subset:
- "[| f : Pi(A,B); g: Pi(C,D); A<=C;
+ "[| f \<in> 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);
+ "[| f \<in> 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) ==> (lam x:A. f`x) = f"
+lemma eta [simp]: "f \<in> Pi(A,B) ==> (\<lambda>x\<in>A. f`x) = f"
apply (rule fun_extension)
apply (auto simp add: lam_type apply_type beta)
done
lemma fun_extension_iff:
- "[| f:Pi(A,B); g:Pi(A,C) |] ==> (ALL a:A. f`a = g`a) <-> f=g"
+ "[| f:Pi(A,B); g:Pi(A,C) |] ==> (\<forall>a\<in>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 \<subseteq> g <-> (f = g)"
by (blast dest: apply_Pair
intro: fun_extension apply_equality [symmetric])
@@ -234,7 +234,7 @@
(*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
lemma Pi_lamE:
assumes major: "f: Pi(A,B)"
- and minor: "!!b. [| ALL x:A. b(x):B(x); f = (lam x:A. b(x)) |] ==> P"
+ and minor: "!!b. [| \<forall>x\<in>A. b(x):B(x); f = (\<lambda>x\<in>A. b(x)) |] ==> P"
shows "P"
apply (rule minor)
apply (rule_tac [2] eta [symmetric])
@@ -244,12 +244,12 @@
subsection{*Images of Functions*}
-lemma image_lam: "C <= A ==> (lam x:A. b(x)) `` C = {b(x). x:C}"
+lemma image_lam: "C \<subseteq> A ==> (\<lambda>x\<in>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))";
+ ==> {f`x. x:C} = (if C \<subseteq> 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)
@@ -261,10 +261,10 @@
(*For this lemma and the next, the right-hand side could equivalently
be written \<Union>x\<in>C. {f`x} *)
lemma image_function:
- "[| function(f); C <= domain(f) |] ==> f``C = {f`x. x:C}";
+ "[| function(f); C \<subseteq> 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}"
+lemma image_fun: "[| f \<in> Pi(A,B); C \<subseteq> A |] ==> f``C = {f`x. x:C}"
apply (simp add: Pi_iff)
apply (blast intro: image_function)
done
@@ -280,17 +280,17 @@
subsection{*Properties of @{term "restrict(f,A)"}*}
-lemma restrict_subset: "restrict(f,A) <= f"
+lemma restrict_subset: "restrict(f,A) \<subseteq> f"
by (unfold restrict_def, blast)
lemma function_restrictI:
"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) \<in> 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)"
+lemma restrict: "restrict(f,A) ` a = (if a \<in> A then f`a else 0)"
by (simp add: apply_def restrict_def, blast)
lemma restrict_empty [simp]: "restrict(f,0) = 0"
@@ -300,38 +300,38 @@
by (simp add: restrict_def)
lemma restrict_restrict [simp]:
- "restrict(restrict(r,A),B) = restrict(r, A Int B)"
+ "restrict(restrict(r,A),B) = restrict(r, A \<inter> B)"
by (unfold restrict_def, blast)
-lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) Int C"
+lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) \<inter> C"
apply (unfold restrict_def)
apply (auto simp add: domain_def)
done
-lemma restrict_idem: "f <= Sigma(A,B) ==> restrict(f,A) = f"
+lemma restrict_idem: "f \<subseteq> Sigma(A,B) ==> restrict(f,A) = f"
by (simp add: restrict_def, blast)
(*converse probably holds too*)
lemma domain_restrict_idem:
- "[| domain(r) <= A; relation(r) |] ==> restrict(r,A) = r"
+ "[| domain(r) \<subseteq> A; relation(r) |] ==> restrict(r,A) = r"
by (simp add: restrict_def relation_def, blast)
-lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A Int C"
+lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \<inter> C"
apply (unfold restrict_def lam_def)
apply (rule equalityI)
apply (auto simp add: domain_iff)
done
-lemma restrict_if [simp]: "restrict(f,A) ` a = (if a : A then f`a else 0)"
+lemma restrict_if [simp]: "restrict(f,A) ` a = (if a \<in> A then f`a else 0)"
by (simp add: restrict apply_0)
lemma restrict_lam_eq:
- "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))"
+ "A<=C ==> restrict(\<lambda>x\<in>C. b(x), A) = (\<lambda>x\<in>A. b(x))"
by (unfold restrict_def lam_def, auto)
lemma fun_cons_restrict_eq:
- "f : cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"
+ "f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"
apply (rule equalityI)
prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD])
apply (auto dest!: Pi_memberD simp add: restrict_def lam_def)
@@ -343,21 +343,21 @@
(** The Union of a set of COMPATIBLE functions is a function **)
lemma function_Union:
- "[| ALL x:S. function(x);
- ALL x:S. ALL y:S. x<=y | y<=x |]
- ==> function(Union(S))"
+ "[| \<forall>x\<in>S. function(x);
+ \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x |]
+ ==> function(\<Union>(S))"
by (unfold function_def, blast)
lemma fun_Union:
- "[| ALL f:S. EX C D. f:C->D;
- ALL f:S. ALL y:S. f<=y | y<=f |] ==>
- Union(S) : domain(Union(S)) -> range(Union(S))"
+ "[| \<forall>f\<in>S. \<exists>C D. f:C->D;
+ \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f |] ==>
+ \<Union>(S) \<in> domain(\<Union>(S)) -> range(\<Union>(S))"
apply (unfold Pi_def)
apply (blast intro!: rel_Union function_Union)
done
lemma gen_relation_Union [rule_format]:
- "\<forall>f\<in>F. relation(f) \<Longrightarrow> relation(Union(F))"
+ "\<forall>f\<in>F. relation(f) \<Longrightarrow> relation(\<Union>(F))"
by (simp add: relation_def)
@@ -368,48 +368,48 @@
subset_trans [OF _ Un_upper2]
lemma fun_disjoint_Un:
- "[| f: A->B; g: C->D; A Int C = 0 |]
- ==> (f Un g) : (A Un C) -> (B Un D)"
+ "[| f: A->B; g: C->D; A \<inter> C = 0 |]
+ ==> (f \<union> g) \<in> (A \<union> C) -> (B \<union> D)"
(*Prove the product and domain subgoals using distributive laws*)
apply (simp add: Pi_iff extension Un_rls)
apply (unfold function_def, blast)
done
-lemma fun_disjoint_apply1: "a \<notin> domain(g) ==> (f Un g)`a = f`a"
+lemma fun_disjoint_apply1: "a \<notin> domain(g) ==> (f \<union> g)`a = f`a"
by (simp add: apply_def, blast)
-lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f Un g)`c = g`c"
+lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f \<union> g)`c = g`c"
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"
+lemma domain_of_fun: "f \<in> 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 \<in> Pi(A,B); a: A |] ==> f`a \<in> range(f)"
by (erule apply_Pair [THEN rangeI], assumption)
-lemma range_of_fun: "f : Pi(A,B) ==> f : A->range(f)"
+lemma range_of_fun: "f \<in> Pi(A,B) ==> f \<in> A->range(f)"
by (blast intro: Pi_type apply_rangeI)
subsection{*Extensions of Functions*}
lemma fun_extend:
- "[| f: A->B; c~:A |] ==> cons(<c,b>,f) : cons(c,A) -> cons(b,B)"
+ "[| f: A->B; c\<notin>A |] ==> cons(<c,b>,f) \<in> cons(c,A) -> cons(b,B)"
apply (frule singleton_fun [THEN fun_disjoint_Un], blast)
apply (simp add: cons_eq)
done
lemma fun_extend3:
- "[| f: A->B; c~:A; b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B"
+ "[| f: A->B; c\<notin>A; b: B |] ==> cons(<c,b>,f) \<in> cons(c,A) -> B"
by (blast intro: fun_extend [THEN fun_weaken_type])
lemma extend_apply:
- "c ~: domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+ "c \<notin> domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
by (auto simp add: apply_def)
lemma fun_extend_apply [simp]:
- "[| f: A->B; c~:A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+ "[| f: A->B; c\<notin>A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
apply (rule extend_apply)
apply (simp add: Pi_def, blast)
done
@@ -418,11 +418,11 @@
(*For Finite.ML. Inclusion of right into left is easy*)
lemma cons_fun_eq:
- "c ~: A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
+ "c \<notin> A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
apply (rule equalityI)
apply (safe elim!: fun_extend3)
(*Inclusion of left into right*)
-apply (subgoal_tac "restrict (x, A) : A -> B")
+apply (subgoal_tac "restrict (x, A) \<in> A -> B")
prefer 2 apply (blast intro: restrict_type2)
apply (rule UN_I, assumption)
apply (rule apply_funtype [THEN UN_I])
@@ -443,7 +443,7 @@
definition
update :: "[i,i,i] => i" where
- "update(f,a,b) == lam x: cons(a, domain(f)). if(x=a, b, f`x)"
+ "update(f,a,b) == \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
nonterminal updbinds and updbind
@@ -481,7 +481,7 @@
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 \<in> A; y: B(x) |] ==> f(x:=y) \<in> Pi(A, B)"
apply (unfold update_def)
apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type)
done
@@ -493,59 +493,59 @@
(*Not easy to express monotonicity in P, since any "bigger" predicate
would have to be single-valued*)
-lemma Replace_mono: "A<=B ==> Replace(A,P) <= Replace(B,P)"
+lemma Replace_mono: "A<=B ==> Replace(A,P) \<subseteq> 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} \<subseteq> {f(x). x:B}"
by blast
-lemma Pow_mono: "A<=B ==> Pow(A) <= Pow(B)"
+lemma Pow_mono: "A<=B ==> Pow(A) \<subseteq> Pow(B)"
by blast
-lemma Union_mono: "A<=B ==> Union(A) <= Union(B)"
+lemma Union_mono: "A<=B ==> \<Union>(A) \<subseteq> \<Union>(B)"
by blast
lemma UN_mono:
- "[| A<=C; !!x. x:A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) <= (\<Union>x\<in>C. D(x))"
+ "[| A<=C; !!x. x:A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> (\<Union>x\<in>C. D(x))"
by blast
(*Intersection is ANTI-monotonic. There are TWO premises! *)
-lemma Inter_anti_mono: "[| A<=B; A\<noteq>0 |] ==> Inter(B) <= Inter(A)"
+lemma Inter_anti_mono: "[| A<=B; A\<noteq>0 |] ==> \<Inter>(B) \<subseteq> \<Inter>(A)"
by blast
-lemma cons_mono: "C<=D ==> cons(a,C) <= cons(a,D)"
+lemma cons_mono: "C<=D ==> cons(a,C) \<subseteq> cons(a,D)"
by blast
-lemma Un_mono: "[| A<=C; B<=D |] ==> A Un B <= C Un D"
+lemma Un_mono: "[| A<=C; B<=D |] ==> A \<union> B \<subseteq> C \<union> D"
by blast
-lemma Int_mono: "[| A<=C; B<=D |] ==> A Int B <= C Int D"
+lemma Int_mono: "[| A<=C; B<=D |] ==> A \<inter> B \<subseteq> C \<inter> D"
by blast
-lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B <= C-D"
+lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B \<subseteq> C-D"
by blast
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 \<longrightarrow> B(x) \<subseteq> D(x) |] ==> Sigma(A,B) \<subseteq> Sigma(C,D)"
by blast
-lemma sum_mono: "[| A<=C; B<=D |] ==> A+B <= C+D"
+lemma sum_mono: "[| A<=C; B<=D |] ==> A+B \<subseteq> C+D"
by (unfold sum_def, blast)
(*Note that B->A and C->A are typically disjoint!*)
-lemma Pi_mono: "B<=C ==> A->B <= A->C"
+lemma Pi_mono: "B<=C ==> A->B \<subseteq> A->C"
by (blast intro: lam_type elim: Pi_lamE)
-lemma lam_mono: "A<=B ==> Lambda(A,c) <= Lambda(B,c)"
+lemma lam_mono: "A<=B ==> Lambda(A,c) \<subseteq> Lambda(B,c)"
apply (unfold lam_def)
apply (erule RepFun_mono)
done
subsubsection{*Converse, Domain, Range, Field*}
-lemma converse_mono: "r<=s ==> converse(r) <= converse(s)"
+lemma converse_mono: "r<=s ==> converse(r) \<subseteq> converse(s)"
by blast
lemma domain_mono: "r<=s ==> domain(r)<=domain(s)"
@@ -561,28 +561,28 @@
lemma field_mono: "r<=s ==> field(r)<=field(s)"
by blast
-lemma field_rel_subset: "r <= A*A ==> field(r) <= A"
+lemma field_rel_subset: "r \<subseteq> A*A ==> field(r) \<subseteq> A"
by (erule field_mono [THEN subset_trans], blast)
subsubsection{*Images*}
lemma image_pair_mono:
- "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A <= s``B"
+ "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A \<subseteq> s``B"
by blast
lemma vimage_pair_mono:
- "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A <= s-``B"
+ "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A \<subseteq> s-``B"
by blast
-lemma image_mono: "[| r<=s; A<=B |] ==> r``A <= s``B"
+lemma image_mono: "[| r<=s; A<=B |] ==> r``A \<subseteq> s``B"
by blast
-lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A <= s-``B"
+lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A \<subseteq> s-``B"
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) \<longrightarrow> Q(x) |] ==> Collect(A,P) \<subseteq> Collect(B,Q)"
by blast
(*Used in intr_elim.ML and in individual datatype definitions*)
@@ -592,7 +592,7 @@
(* Useful with simp; contributed by Clemens Ballarin. *)
lemma bex_image_simp:
- "[| f : Pi(X, Y); A \<subseteq> X |] ==> (EX x : f``A. P(x)) <-> (EX x:A. P(f`x))"
+ "[| f \<in> Pi(X, Y); A \<subseteq> X |] ==> (\<exists>x\<in>f``A. P(x)) <-> (\<exists>x\<in>A. P(f`x))"
apply safe
apply rule
prefer 2 apply assumption
@@ -601,7 +601,7 @@
done
lemma ball_image_simp:
- "[| f : Pi(X, Y); A \<subseteq> X |] ==> (ALL x : f``A. P(x)) <-> (ALL x:A. P(f`x))"
+ "[| f \<in> Pi(X, Y); A \<subseteq> X |] ==> (\<forall>x\<in>f``A. P(x)) <-> (\<forall>x\<in>A. P(f`x))"
apply safe
apply (blast intro: apply_Pair)
apply (drule bspec) apply assumption
--- a/src/ZF/pair.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/pair.thy Tue Mar 06 15:15:49 2012 +0000
@@ -17,14 +17,14 @@
ML {* val ZF_ss = @{simpset} *}
-simproc_setup defined_Bex ("EX x:A. P(x) & Q(x)") = {*
+simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) & Q(x)") = {*
let
val unfold_bex_tac = unfold_tac @{thms Bex_def};
fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac;
in fn _ => fn ss => Quantifier1.rearrange_bex (prove_bex_tac ss) ss end
*}
-simproc_setup defined_Ball ("ALL x:A. P(x) --> Q(x)") = {*
+simproc_setup defined_Ball ("\<forall>x\<in>A. P(x) \<longrightarrow> Q(x)") = {*
let
val unfold_ball_tac = unfold_tac @{thms Ball_def};
fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac;
@@ -48,7 +48,7 @@
lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1]
lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2]
-lemma Pair_not_0: "<a,b> ~= 0"
+lemma Pair_not_0: "<a,b> \<noteq> 0"
apply (unfold Pair_def)
apply (blast elim: equalityE)
done
@@ -79,7 +79,7 @@
lemma Sigma_iff [simp]: "<a,b>: Sigma(A,B) <-> a:A & b:B(a)"
by (simp add: Sigma_def)
-lemma SigmaI [TC,intro!]: "[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"
+lemma SigmaI [TC,intro!]: "[| a:A; b:B(a) |] ==> <a,b> \<in> Sigma(A,B)"
by simp
lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1]
@@ -93,7 +93,7 @@
by (unfold Sigma_def, blast)
lemma SigmaE2 [elim!]:
- "[| <a,b> : Sigma(A,B);
+ "[| <a,b> \<in> Sigma(A,B);
[| a:A; b:B(a) |] ==> P
|] ==> P"
by (unfold Sigma_def, blast)
@@ -125,10 +125,10 @@
lemma snd_conv [simp]: "snd(<a,b>) = 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) \<in> 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) \<in> B(fst(p))"
by auto
lemma Pair_fst_snd_eq: "a: Sigma(A,B) ==> <fst(a),snd(a)> = a"
@@ -144,13 +144,13 @@
lemma split_type [TC]:
"[| p:Sigma(A,B);
!!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>)
- |] ==> split(%x y. c(x,y), p) : C(p)"
+ |] ==> split(%x y. c(x,y), p) \<in> C(p)"
apply (erule SigmaE, auto)
done
lemma expand_split:
"u: A*B ==>
- R(split(c,u)) <-> (ALL x:A. ALL y:B. u = <x,y> --> R(c(x,y)))"
+ R(split(c,u)) <-> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x,y> \<longrightarrow> R(c(x,y)))"
apply (simp add: split_def)
apply auto
done
--- a/src/ZF/upair.thy Sun Mar 04 23:20:43 2012 +0100
+++ b/src/ZF/upair.thy Tue Mar 06 15:15:49 2012 +0000
@@ -4,7 +4,7 @@
Observe the order of dependence:
Upair is defined in terms of Replace
- Un is defined in terms of Upair and Union (similarly for Int)
+ \<union> is defined in terms of Upair and \<Union>(similarly for Int)
cons is defined in terms of Upair and Un
Ordered pairs and descriptions are defined using cons ("set notation")
*)
@@ -17,117 +17,117 @@
setup TypeCheck.setup
lemma atomize_ball [symmetric, rulify]:
- "(!!x. x:A ==> P(x)) == Trueprop (ALL x:A. P(x))"
+ "(!!x. x:A ==> P(x)) == Trueprop (\<forall>x\<in>A. P(x))"
by (simp add: Ball_def atomize_all atomize_imp)
subsection{*Unordered Pairs: constant @{term Upair}*}
-lemma Upair_iff [simp]: "c : Upair(a,b) <-> (c=a | c=b)"
+lemma Upair_iff [simp]: "c \<in> Upair(a,b) <-> (c=a | c=b)"
by (unfold Upair_def, blast)
-lemma UpairI1: "a : Upair(a,b)"
+lemma UpairI1: "a \<in> Upair(a,b)"
by simp
-lemma UpairI2: "b : Upair(a,b)"
+lemma UpairI2: "b \<in> Upair(a,b)"
by simp
-lemma UpairE: "[| a : Upair(b,c); a=b ==> P; a=c ==> P |] ==> P"
+lemma UpairE: "[| a \<in> Upair(b,c); a=b ==> P; a=c ==> P |] ==> P"
by (simp, blast)
subsection{*Rules for Binary Union, Defined via @{term Upair}*}
-lemma Un_iff [simp]: "c : A Un B <-> (c:A | c:B)"
+lemma Un_iff [simp]: "c \<in> A \<union> B <-> (c:A | c:B)"
apply (simp add: Un_def)
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
done
-lemma UnI1: "c : A ==> c : A Un B"
+lemma UnI1: "c \<in> A ==> c \<in> A \<union> B"
by simp
-lemma UnI2: "c : B ==> c : A Un B"
+lemma UnI2: "c \<in> B ==> c \<in> A \<union> B"
by simp
declare UnI1 [elim?] UnI2 [elim?]
-lemma UnE [elim!]: "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"
+lemma UnE [elim!]: "[| c \<in> A \<union> B; c:A ==> P; c:B ==> P |] ==> P"
by (simp, blast)
(*Stronger version of the rule above*)
-lemma UnE': "[| c : A Un B; c:A ==> P; [| c:B; c~:A |] ==> P |] ==> P"
+lemma UnE': "[| c \<in> A \<union> B; c:A ==> P; [| c:B; c\<notin>A |] ==> P |] ==> P"
by (simp, blast)
(*Classical introduction rule: no commitment to A vs B*)
-lemma UnCI [intro!]: "(c ~: B ==> c : A) ==> c : A Un B"
+lemma UnCI [intro!]: "(c \<notin> B ==> c \<in> A) ==> c \<in> A \<union> B"
by (simp, blast)
subsection{*Rules for Binary Intersection, Defined via @{term Upair}*}
-lemma Int_iff [simp]: "c : A Int B <-> (c:A & c:B)"
+lemma Int_iff [simp]: "c \<in> A \<inter> B <-> (c:A & c:B)"
apply (unfold Int_def)
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
done
-lemma IntI [intro!]: "[| c : A; c : B |] ==> c : A Int B"
+lemma IntI [intro!]: "[| c \<in> A; c \<in> B |] ==> c \<in> A \<inter> B"
by simp
-lemma IntD1: "c : A Int B ==> c : A"
+lemma IntD1: "c \<in> A \<inter> B ==> c \<in> A"
by simp
-lemma IntD2: "c : A Int B ==> c : B"
+lemma IntD2: "c \<in> A \<inter> B ==> c \<in> B"
by simp
-lemma IntE [elim!]: "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"
+lemma IntE [elim!]: "[| c \<in> A \<inter> 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 \<in> A-B <-> (c:A & c\<notin>B)"
by (unfold Diff_def, blast)
-lemma DiffI [intro!]: "[| c : A; c ~: B |] ==> c : A - B"
+lemma DiffI [intro!]: "[| c \<in> A; c \<notin> B |] ==> c \<in> A - B"
by simp
-lemma DiffD1: "c : A - B ==> c : A"
+lemma DiffD1: "c \<in> A - B ==> c \<in> A"
by simp
-lemma DiffD2: "c : A - B ==> c ~: B"
+lemma DiffD2: "c \<in> A - B ==> c \<notin> B"
by simp
-lemma DiffE [elim!]: "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"
+lemma DiffE [elim!]: "[| c \<in> A - B; [| c:A; c\<notin>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 \<in> cons(b,A) <-> (a=b | a:A)"
apply (unfold cons_def)
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
done
(*risky as a typechecking rule, but solves otherwise unconstrained goals of
-the form x : ?A*)
-lemma consI1 [simp,TC]: "a : cons(a,B)"
+the form x \<in> ?A*)
+lemma consI1 [simp,TC]: "a \<in> cons(a,B)"
by simp
-lemma consI2: "a : B ==> a : cons(b,B)"
+lemma consI2: "a \<in> B ==> a \<in> cons(b,B)"
by simp
-lemma consE [elim!]: "[| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P"
+lemma consE [elim!]: "[| a \<in> 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 \<in> cons(b,A); a=b ==> P; [| a:A; a\<noteq>b |] ==> P |] ==> P"
by (simp, blast)
(*Classical introduction rule*)
-lemma consCI [intro!]: "(a~:B ==> a=b) ==> a: cons(b,B)"
+lemma consCI [intro!]: "(a\<notin>B ==> a=b) ==> a: cons(b,B)"
by (simp, blast)
-lemma cons_not_0 [simp]: "cons(a,B) ~= 0"
+lemma cons_not_0 [simp]: "cons(a,B) \<noteq> 0"
by (blast elim: equalityE)
lemmas cons_neq_0 = cons_not_0 [THEN notE]
@@ -137,10 +137,10 @@
subsection{*Singletons*}
-lemma singleton_iff: "a : {b} <-> a=b"
+lemma singleton_iff: "a \<in> {b} <-> a=b"
by simp
-lemma singletonI [intro!]: "a : {a}"
+lemma singletonI [intro!]: "a \<in> {a}"
by (rule consI1)
lemmas singletonE = singleton_iff [THEN iffD1, elim_format, elim!]
@@ -150,7 +150,7 @@
lemma the_equality [intro]:
"[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
-apply (unfold the_def)
+apply (unfold the_def)
apply (fast dest: subst)
done
@@ -164,8 +164,8 @@
apply (blast+)
done
-(*the_cong is no longer necessary: if (ALL y.P(y)<->Q(y)) then
- (THE x.P(x)) rewrites to (THE x. Q(x)) *)
+(*No congruence rule is necessary: if @{term"\<forall>y.P(y)<->Q(y)"} then
+ @{term "THE x.P(x)"} rewrites to @{term "THE x.Q(x)"} *)
(*If it's "undefined", it's zero!*)
lemma the_0: "~ (EX! x. P(x)) ==> (THE x. P(x))=0"
@@ -203,7 +203,7 @@
(*Never use with case splitting, or if P is known to be true or false*)
lemma if_cong:
- "[| P<->Q; Q ==> a=c; ~Q ==> b=d |]
+ "[| P<->Q; Q ==> a=c; ~Q ==> b=d |]
==> (if P then a else b) = (if Q then c else d)"
by (simp add: if_def cong add: conj_cong)
@@ -221,7 +221,7 @@
by (unfold if_def, blast)
lemma split_if [split]:
- "P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))"
+ "P(if Q then x else y) <-> ((Q \<longrightarrow> P(x)) & (~Q \<longrightarrow> P(y)))"
by (case_tac Q, simp_all)
(** Rewrite rules for boolean case-splitting: faster than split_if [split]
@@ -230,8 +230,8 @@
lemmas split_if_eq1 = split_if [of "%x. x = b"] for b
lemmas split_if_eq2 = split_if [of "%x. a = x"] for x
-lemmas split_if_mem1 = split_if [of "%x. x : b"] for b
-lemmas split_if_mem2 = split_if [of "%x. a : x"] for x
+lemmas split_if_mem1 = split_if [of "%x. x \<in> b"] for b
+lemmas split_if_mem2 = split_if [of "%x. a \<in> x"] for x
lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
@@ -267,39 +267,39 @@
(*mem_irrefl should NOT be added to default databases:
it would be tried on most goals, making proofs slower!*)
-lemma mem_not_refl: "a ~: a"
+lemma mem_not_refl: "a \<notin> a"
apply (rule notI)
apply (erule mem_irrefl)
done
(*Good for proving inequalities by rewriting*)
-lemma mem_imp_not_eq: "a:A ==> a ~= A"
+lemma mem_imp_not_eq: "a:A ==> a \<noteq> A"
by (blast elim!: mem_irrefl)
-lemma eq_imp_not_mem: "a=A ==> a ~: A"
+lemma eq_imp_not_mem: "a=A ==> a \<notin> A"
by (blast intro: elim: mem_irrefl)
subsection{*Rules for Successor*}
-lemma succ_iff: "i : succ(j) <-> i=j | i:j"
+lemma succ_iff: "i \<in> succ(j) <-> i=j | i:j"
by (unfold succ_def, blast)
-lemma succI1 [simp]: "i : succ(i)"
+lemma succI1 [simp]: "i \<in> succ(i)"
by (simp add: succ_iff)
-lemma succI2: "i : j ==> i : succ(j)"
+lemma succI2: "i \<in> j ==> i \<in> succ(j)"
by (simp add: succ_iff)
-lemma succE [elim!]:
- "[| i : succ(j); i=j ==> P; i:j ==> P |] ==> P"
-apply (simp add: succ_iff, blast)
+lemma succE [elim!]:
+ "[| i \<in> 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\<notin>j ==> i=j) ==> i: succ(j)"
by (simp add: succ_iff, blast)
-lemma succ_not_0 [simp]: "succ(n) ~= 0"
+lemma succ_not_0 [simp]: "succ(n) \<noteq> 0"
by (blast elim!: equalityE)
lemmas succ_neq_0 = succ_not_0 [THEN notE, elim!]
@@ -307,10 +307,10 @@
declare succ_not_0 [THEN not_sym, simp]
declare sym [THEN succ_neq_0, elim!]
-(* succ(c) <= B ==> c : B *)
+(* @{term"succ(c) \<subseteq> B ==> c \<in> B"} *)
lemmas succ_subsetD = succI1 [THEN [2] subsetD]
-(* succ(b) ~= b *)
+(* @{term"succ(b) \<noteq> b"} *)
lemmas succ_neq_self = succI1 [THEN mem_imp_not_eq, THEN not_sym]
lemma succ_inject_iff [simp]: "succ(m) = succ(n) <-> m=n"
@@ -322,83 +322,83 @@
subsection{*Miniscoping of the Bounded Universal Quantifier*}
lemma ball_simps1:
- "(ALL x:A. P(x) & Q) <-> (ALL x:A. P(x)) & (A=0 | Q)"
- "(ALL x:A. P(x) | Q) <-> ((ALL x:A. P(x)) | Q)"
- "(ALL x:A. P(x) --> Q) <-> ((EX x:A. P(x)) --> Q)"
- "(~(ALL x:A. P(x))) <-> (EX x:A. ~P(x))"
- "(ALL x:0.P(x)) <-> True"
- "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i. P(x))"
- "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B. P(x))"
- "(ALL x:RepFun(A,f). P(x)) <-> (ALL y:A. P(f(y)))"
- "(ALL x:Union(A).P(x)) <-> (ALL y:A. ALL x:y. P(x))"
+ "(\<forall>x\<in>A. P(x) & Q) <-> (\<forall>x\<in>A. P(x)) & (A=0 | Q)"
+ "(\<forall>x\<in>A. P(x) | Q) <-> ((\<forall>x\<in>A. P(x)) | Q)"
+ "(\<forall>x\<in>A. P(x) \<longrightarrow> Q) <-> ((\<exists>x\<in>A. P(x)) \<longrightarrow> Q)"
+ "(~(\<forall>x\<in>A. P(x))) <-> (\<exists>x\<in>A. ~P(x))"
+ "(\<forall>x\<in>0.P(x)) <-> True"
+ "(\<forall>x\<in>succ(i).P(x)) <-> P(i) & (\<forall>x\<in>i. P(x))"
+ "(\<forall>x\<in>cons(a,B).P(x)) <-> P(a) & (\<forall>x\<in>B. P(x))"
+ "(\<forall>x\<in>RepFun(A,f). P(x)) <-> (\<forall>y\<in>A. P(f(y)))"
+ "(\<forall>x\<in>\<Union>(A).P(x)) <-> (\<forall>y\<in>A. \<forall>x\<in>y. P(x))"
by blast+
lemma ball_simps2:
- "(ALL x:A. P & Q(x)) <-> (A=0 | P) & (ALL x:A. Q(x))"
- "(ALL x:A. P | Q(x)) <-> (P | (ALL x:A. Q(x)))"
- "(ALL x:A. P --> Q(x)) <-> (P --> (ALL x:A. Q(x)))"
+ "(\<forall>x\<in>A. P & Q(x)) <-> (A=0 | P) & (\<forall>x\<in>A. Q(x))"
+ "(\<forall>x\<in>A. P | Q(x)) <-> (P | (\<forall>x\<in>A. Q(x)))"
+ "(\<forall>x\<in>A. P \<longrightarrow> Q(x)) <-> (P \<longrightarrow> (\<forall>x\<in>A. Q(x)))"
by blast+
lemma ball_simps3:
- "(ALL x:Collect(A,Q).P(x)) <-> (ALL x:A. Q(x) --> P(x))"
+ "(\<forall>x\<in>Collect(A,Q).P(x)) <-> (\<forall>x\<in>A. Q(x) \<longrightarrow> P(x))"
by blast+
lemmas ball_simps [simp] = ball_simps1 ball_simps2 ball_simps3
lemma ball_conj_distrib:
- "(ALL x:A. P(x) & Q(x)) <-> ((ALL x:A. P(x)) & (ALL x:A. Q(x)))"
+ "(\<forall>x\<in>A. P(x) & Q(x)) <-> ((\<forall>x\<in>A. P(x)) & (\<forall>x\<in>A. Q(x)))"
by blast
subsection{*Miniscoping of the Bounded Existential Quantifier*}
lemma bex_simps1:
- "(EX x:A. P(x) & Q) <-> ((EX x:A. P(x)) & Q)"
- "(EX x:A. P(x) | Q) <-> (EX x:A. P(x)) | (A~=0 & Q)"
- "(EX x:A. P(x) --> Q) <-> ((ALL x:A. P(x)) --> (A~=0 & Q))"
- "(EX x:0.P(x)) <-> False"
- "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i. P(x))"
- "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B. P(x))"
- "(EX x:RepFun(A,f). P(x)) <-> (EX y:A. P(f(y)))"
- "(EX x:Union(A).P(x)) <-> (EX y:A. EX x:y. P(x))"
- "(~(EX x:A. P(x))) <-> (ALL x:A. ~P(x))"
+ "(\<exists>x\<in>A. P(x) & Q) <-> ((\<exists>x\<in>A. P(x)) & Q)"
+ "(\<exists>x\<in>A. P(x) | Q) <-> (\<exists>x\<in>A. P(x)) | (A\<noteq>0 & Q)"
+ "(\<exists>x\<in>A. P(x) \<longrightarrow> Q) <-> ((\<forall>x\<in>A. P(x)) \<longrightarrow> (A\<noteq>0 & Q))"
+ "(\<exists>x\<in>0.P(x)) <-> False"
+ "(\<exists>x\<in>succ(i).P(x)) <-> P(i) | (\<exists>x\<in>i. P(x))"
+ "(\<exists>x\<in>cons(a,B).P(x)) <-> P(a) | (\<exists>x\<in>B. P(x))"
+ "(\<exists>x\<in>RepFun(A,f). P(x)) <-> (\<exists>y\<in>A. P(f(y)))"
+ "(\<exists>x\<in>\<Union>(A).P(x)) <-> (\<exists>y\<in>A. \<exists>x\<in>y. P(x))"
+ "(~(\<exists>x\<in>A. P(x))) <-> (\<forall>x\<in>A. ~P(x))"
by blast+
lemma bex_simps2:
- "(EX x:A. P & Q(x)) <-> (P & (EX x:A. Q(x)))"
- "(EX x:A. P | Q(x)) <-> (A~=0 & P) | (EX x:A. Q(x))"
- "(EX x:A. P --> Q(x)) <-> ((A=0 | P) --> (EX x:A. Q(x)))"
+ "(\<exists>x\<in>A. P & Q(x)) <-> (P & (\<exists>x\<in>A. Q(x)))"
+ "(\<exists>x\<in>A. P | Q(x)) <-> (A\<noteq>0 & P) | (\<exists>x\<in>A. Q(x))"
+ "(\<exists>x\<in>A. P \<longrightarrow> Q(x)) <-> ((A=0 | P) \<longrightarrow> (\<exists>x\<in>A. Q(x)))"
by blast+
lemma bex_simps3:
- "(EX x:Collect(A,Q).P(x)) <-> (EX x:A. Q(x) & P(x))"
+ "(\<exists>x\<in>Collect(A,Q).P(x)) <-> (\<exists>x\<in>A. Q(x) & P(x))"
by blast
lemmas bex_simps [simp] = bex_simps1 bex_simps2 bex_simps3
lemma bex_disj_distrib:
- "(EX x:A. P(x) | Q(x)) <-> ((EX x:A. P(x)) | (EX x:A. Q(x)))"
+ "(\<exists>x\<in>A. P(x) | Q(x)) <-> ((\<exists>x\<in>A. P(x)) | (\<exists>x\<in>A. Q(x)))"
by blast
(** One-point rule for bounded quantifiers: see HOL/Set.ML **)
-lemma bex_triv_one_point1 [simp]: "(EX x:A. x=a) <-> (a:A)"
+lemma bex_triv_one_point1 [simp]: "(\<exists>x\<in>A. x=a) <-> (a:A)"
by blast
-lemma bex_triv_one_point2 [simp]: "(EX x:A. a=x) <-> (a:A)"
+lemma bex_triv_one_point2 [simp]: "(\<exists>x\<in>A. a=x) <-> (a:A)"
by blast
-lemma bex_one_point1 [simp]: "(EX x:A. x=a & P(x)) <-> (a:A & P(a))"
+lemma bex_one_point1 [simp]: "(\<exists>x\<in>A. x=a & P(x)) <-> (a:A & P(a))"
by blast
-lemma bex_one_point2 [simp]: "(EX x:A. a=x & P(x)) <-> (a:A & P(a))"
+lemma bex_one_point2 [simp]: "(\<exists>x\<in>A. a=x & P(x)) <-> (a:A & P(a))"
by blast
-lemma ball_one_point1 [simp]: "(ALL x:A. x=a --> P(x)) <-> (a:A --> P(a))"
+lemma ball_one_point1 [simp]: "(\<forall>x\<in>A. x=a \<longrightarrow> P(x)) <-> (a:A \<longrightarrow> P(a))"
by blast
-lemma ball_one_point2 [simp]: "(ALL x:A. a=x --> P(x)) <-> (a:A --> P(a))"
+lemma ball_one_point2 [simp]: "(\<forall>x\<in>A. a=x \<longrightarrow> P(x)) <-> (a:A \<longrightarrow> P(a))"
by blast
@@ -418,21 +418,21 @@
subsection{*Miniscoping of Unions*}
lemma UN_simps1:
- "(UN x:C. cons(a, B(x))) = (if C=0 then 0 else cons(a, UN x:C. B(x)))"
- "(UN x:C. A(x) Un B') = (if C=0 then 0 else (UN x:C. A(x)) Un B')"
- "(UN x:C. A' Un B(x)) = (if C=0 then 0 else A' Un (UN x:C. B(x)))"
- "(UN x:C. A(x) Int B') = ((UN x:C. A(x)) Int B')"
- "(UN x:C. A' Int B(x)) = (A' Int (UN x:C. B(x)))"
- "(UN x:C. A(x) - B') = ((UN x:C. A(x)) - B')"
- "(UN x:C. A' - B(x)) = (if C=0 then 0 else A' - (INT x:C. B(x)))"
-apply (simp_all add: Inter_def)
+ "(\<Union>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Union>x\<in>C. B(x)))"
+ "(\<Union>x\<in>C. A(x) \<union> B') = (if C=0 then 0 else (\<Union>x\<in>C. A(x)) \<union> B')"
+ "(\<Union>x\<in>C. A' \<union> B(x)) = (if C=0 then 0 else A' \<union> (\<Union>x\<in>C. B(x)))"
+ "(\<Union>x\<in>C. A(x) \<inter> B') = ((\<Union>x\<in>C. A(x)) \<inter> B')"
+ "(\<Union>x\<in>C. A' \<inter> B(x)) = (A' \<inter> (\<Union>x\<in>C. B(x)))"
+ "(\<Union>x\<in>C. A(x) - B') = ((\<Union>x\<in>C. A(x)) - B')"
+ "(\<Union>x\<in>C. A' - B(x)) = (if C=0 then 0 else A' - (\<Inter>x\<in>C. B(x)))"
+apply (simp_all add: Inter_def)
apply (blast intro!: equalityI )+
done
lemma UN_simps2:
- "(UN x: Union(A). B(x)) = (UN y:A. UN x:y. B(x))"
- "(UN z: (UN x:A. B(x)). C(z)) = (UN x:A. UN z: B(x). C(z))"
- "(UN x: RepFun(A,f). B(x)) = (UN a:A. B(f(a)))"
+ "(\<Union>x\<in>\<Union>(A). B(x)) = (\<Union>y\<in>A. \<Union>x\<in>y. B(x))"
+ "(\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z\<in>B(x). C(z))"
+ "(\<Union>x\<in>RepFun(A,f). B(x)) = (\<Union>a\<in>A. B(f(a)))"
by blast+
lemmas UN_simps [simp] = UN_simps1 UN_simps2
@@ -440,26 +440,26 @@
text{*Opposite of miniscoping: pull the operator out*}
lemma UN_extend_simps1:
- "(UN x:C. A(x)) Un B = (if C=0 then B else (UN x:C. A(x) Un B))"
- "((UN x:C. A(x)) Int B) = (UN x:C. A(x) Int B)"
- "((UN x:C. A(x)) - B) = (UN x:C. A(x) - B)"
-apply simp_all
+ "(\<Union>x\<in>C. A(x)) \<union> B = (if C=0 then B else (\<Union>x\<in>C. A(x) \<union> B))"
+ "((\<Union>x\<in>C. A(x)) \<inter> B) = (\<Union>x\<in>C. A(x) \<inter> B)"
+ "((\<Union>x\<in>C. A(x)) - B) = (\<Union>x\<in>C. A(x) - B)"
+apply simp_all
apply blast+
done
lemma UN_extend_simps2:
- "cons(a, UN x:C. B(x)) = (if C=0 then {a} else (UN x:C. cons(a, B(x))))"
- "A Un (UN x:C. B(x)) = (if C=0 then A else (UN x:C. A Un B(x)))"
- "(A Int (UN x:C. B(x))) = (UN x:C. A Int B(x))"
- "A - (INT x:C. B(x)) = (if C=0 then A else (UN x:C. A - B(x)))"
- "(UN y:A. UN x:y. B(x)) = (UN x: Union(A). B(x))"
- "(UN a:A. B(f(a))) = (UN x: RepFun(A,f). B(x))"
-apply (simp_all add: Inter_def)
+ "cons(a, \<Union>x\<in>C. B(x)) = (if C=0 then {a} else (\<Union>x\<in>C. cons(a, B(x))))"
+ "A \<union> (\<Union>x\<in>C. B(x)) = (if C=0 then A else (\<Union>x\<in>C. A \<union> B(x)))"
+ "(A \<inter> (\<Union>x\<in>C. B(x))) = (\<Union>x\<in>C. A \<inter> B(x))"
+ "A - (\<Inter>x\<in>C. B(x)) = (if C=0 then A else (\<Union>x\<in>C. A - B(x)))"
+ "(\<Union>y\<in>A. \<Union>x\<in>y. B(x)) = (\<Union>x\<in>\<Union>(A). B(x))"
+ "(\<Union>a\<in>A. B(f(a))) = (\<Union>x\<in>RepFun(A,f). B(x))"
+apply (simp_all add: Inter_def)
apply (blast intro!: equalityI)+
done
lemma UN_UN_extend:
- "(UN x:A. UN z: B(x). C(z)) = (UN z: (UN x:A. B(x)). C(z))"
+ "(\<Union>x\<in>A. \<Union>z\<in>B(x). C(z)) = (\<Union>z\<in>(\<Union>x\<in>A. B(x)). C(z))"
by blast
lemmas UN_extend_simps = UN_extend_simps1 UN_extend_simps2 UN_UN_extend
@@ -468,17 +468,17 @@
subsection{*Miniscoping of Intersections*}
lemma INT_simps1:
- "(INT x:C. A(x) Int B) = (INT x:C. A(x)) Int B"
- "(INT x:C. A(x) - B) = (INT x:C. A(x)) - B"
- "(INT x:C. A(x) Un B) = (if C=0 then 0 else (INT x:C. A(x)) Un B)"
+ "(\<Inter>x\<in>C. A(x) \<inter> B) = (\<Inter>x\<in>C. A(x)) \<inter> B"
+ "(\<Inter>x\<in>C. A(x) - B) = (\<Inter>x\<in>C. A(x)) - B"
+ "(\<Inter>x\<in>C. A(x) \<union> B) = (if C=0 then 0 else (\<Inter>x\<in>C. A(x)) \<union> B)"
by (simp_all add: Inter_def, blast+)
lemma INT_simps2:
- "(INT x:C. A Int B(x)) = A Int (INT x:C. B(x))"
- "(INT x:C. A - B(x)) = (if C=0 then 0 else A - (UN x:C. B(x)))"
- "(INT x:C. cons(a, B(x))) = (if C=0 then 0 else cons(a, INT x:C. B(x)))"
- "(INT x:C. A Un B(x)) = (if C=0 then 0 else A Un (INT x:C. B(x)))"
-apply (simp_all add: Inter_def)
+ "(\<Inter>x\<in>C. A \<inter> B(x)) = A \<inter> (\<Inter>x\<in>C. B(x))"
+ "(\<Inter>x\<in>C. A - B(x)) = (if C=0 then 0 else A - (\<Union>x\<in>C. B(x)))"
+ "(\<Inter>x\<in>C. cons(a, B(x))) = (if C=0 then 0 else cons(a, \<Inter>x\<in>C. B(x)))"
+ "(\<Inter>x\<in>C. A \<union> B(x)) = (if C=0 then 0 else A \<union> (\<Inter>x\<in>C. B(x)))"
+apply (simp_all add: Inter_def)
apply (blast intro!: equalityI)+
done
@@ -488,18 +488,18 @@
lemma INT_extend_simps1:
- "(INT x:C. A(x)) Int B = (INT x:C. A(x) Int B)"
- "(INT x:C. A(x)) - B = (INT x:C. A(x) - B)"
- "(INT x:C. A(x)) Un B = (if C=0 then B else (INT x:C. A(x) Un B))"
+ "(\<Inter>x\<in>C. A(x)) \<inter> B = (\<Inter>x\<in>C. A(x) \<inter> B)"
+ "(\<Inter>x\<in>C. A(x)) - B = (\<Inter>x\<in>C. A(x) - B)"
+ "(\<Inter>x\<in>C. A(x)) \<union> B = (if C=0 then B else (\<Inter>x\<in>C. A(x) \<union> B))"
apply (simp_all add: Inter_def, blast+)
done
lemma INT_extend_simps2:
- "A Int (INT x:C. B(x)) = (INT x:C. A Int B(x))"
- "A - (UN x:C. B(x)) = (if C=0 then A else (INT x:C. A - B(x)))"
- "cons(a, INT x:C. B(x)) = (if C=0 then {a} else (INT x:C. cons(a, B(x))))"
- "A Un (INT x:C. B(x)) = (if C=0 then A else (INT x:C. A Un B(x)))"
-apply (simp_all add: Inter_def)
+ "A \<inter> (\<Inter>x\<in>C. B(x)) = (\<Inter>x\<in>C. A \<inter> B(x))"
+ "A - (\<Union>x\<in>C. B(x)) = (if C=0 then A else (\<Inter>x\<in>C. A - B(x)))"
+ "cons(a, \<Inter>x\<in>C. B(x)) = (if C=0 then {a} else (\<Inter>x\<in>C. cons(a, B(x))))"
+ "A \<union> (\<Inter>x\<in>C. B(x)) = (if C=0 then A else (\<Inter>x\<in>C. A \<union> B(x)))"
+apply (simp_all add: Inter_def)
apply (blast intro!: equalityI)+
done
@@ -512,15 +512,15 @@
(*** Miniscoping: pushing in big Unions, Intersections, quantifiers, etc. ***)
lemma misc_simps [simp]:
- "0 Un A = A"
- "A Un 0 = A"
- "0 Int A = 0"
- "A Int 0 = 0"
+ "0 \<union> A = A"
+ "A \<union> 0 = A"
+ "0 \<inter> A = 0"
+ "A \<inter> 0 = 0"
"0 - A = 0"
"A - 0 = A"
- "Union(0) = 0"
- "Union(cons(b,A)) = b Un Union(A)"
- "Inter({b}) = b"
+ "\<Union>(0) = 0"
+ "\<Union>(cons(b,A)) = b \<union> \<Union>(A)"
+ "\<Inter>({b}) = b"
by blast+
end