--- a/src/ZF/AC.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,10 +9,10 @@
text\<open>This definition comes from Halmos (1960), page 59.\<close>
axiomatization where
- AC: "[| a \<in> A; !!x. x \<in> A ==> (\<exists>y. y \<in> B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
+ AC: "\<lbrakk>a \<in> A; \<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> B(x))\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> Pi(A,B)"
(*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)"
+lemma AC_Pi: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> B(x))\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> Pi(A,B)"
apply (case_tac "A=0")
apply (simp add: Pi_empty1)
(*The non-trivial case*)
@@ -20,7 +20,7 @@
done
(*Using dtac, this has the advantage of DELETING the universal quantifier*)
-lemma AC_ball_Pi: "\<forall>x \<in> A. \<exists>y. y \<in> B(x) ==> \<exists>y. y \<in> Pi(A,B)"
+lemma AC_ball_Pi: "\<forall>x \<in> A. \<exists>y. y \<in> B(x) \<Longrightarrow> \<exists>y. y \<in> Pi(A,B)"
apply (rule AC_Pi)
apply (erule bspec, assumption)
done
@@ -31,15 +31,15 @@
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"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> x)\<rbrakk> \<Longrightarrow> \<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)
done
-lemma non_empty_family: "[| 0 \<notin> A; x \<in> A |] ==> \<exists>y. y \<in> x"
+lemma non_empty_family: "\<lbrakk>0 \<notin> A; x \<in> A\<rbrakk> \<Longrightarrow> \<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 \<Longrightarrow> \<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)
done
@@ -51,7 +51,7 @@
apply (erule_tac [2] fun_weaken_type, blast+)
done
-lemma AC_Pi0: "0 \<notin> A ==> \<exists>f. f \<in> (\<Prod>x \<in> A. x)"
+lemma AC_Pi0: "0 \<notin> A \<Longrightarrow> \<exists>f. f \<in> (\<Prod>x \<in> A. x)"
apply (rule AC_Pi)
apply (simp_all add: non_empty_family)
done
--- a/src/ZF/AC/AC15_WO6.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC15_WO6.thy Tue Sep 27 16:51:35 2022 +0100
@@ -4,18 +4,18 @@
The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.
We need the following:
-WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6
+WO1 \<Longrightarrow> AC10(n) \<Longrightarrow> AC11 \<Longrightarrow> AC12 \<Longrightarrow> AC15 \<Longrightarrow> WO6
In order to add the formulations AC13 and AC14 we need:
-AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15
+AC10(succ(n)) \<Longrightarrow> AC13(n) \<Longrightarrow> AC14 \<Longrightarrow> AC15
or
-AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m\<le>n)
+AC1 \<Longrightarrow> AC13(1); AC13(m) \<Longrightarrow> AC13(n) \<Longrightarrow> AC14 \<Longrightarrow> AC15 (m\<le>n)
So we don't have to prove all implications of both cases.
-Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as
+Moreover we don't need to prove AC13(1) \<Longrightarrow> AC1 and AC11 \<Longrightarrow> AC14 as
Rubin & Rubin do.
*)
@@ -31,7 +31,7 @@
(* - ex_fun_AC13_AC15 *)
(* ********************************************************************** *)
-lemma lepoll_Sigma: "A\<noteq>0 ==> B \<lesssim> A*B"
+lemma lepoll_Sigma: "A\<noteq>0 \<Longrightarrow> B \<lesssim> A*B"
apply (unfold lepoll_def)
apply (erule not_emptyE)
apply (rule_tac x = "\<lambda>z \<in> B. <x,z>" in exI)
@@ -39,24 +39,24 @@
done
lemma cons_times_nat_not_Finite:
- "0\<notin>A ==> \<forall>B \<in> {cons(0,x*nat). x \<in> A}. ~Finite(B)"
+ "0\<notin>A \<Longrightarrow> \<forall>B \<in> {cons(0,x*nat). x \<in> A}. \<not>Finite(B)"
apply clarify
apply (rule nat_not_Finite [THEN notE] )
apply (subgoal_tac "x \<noteq> 0")
apply (blast intro: lepoll_Sigma [THEN lepoll_Finite])+
done
-lemma lemma1: "[| \<Union>(C)=A; a \<in> A |] ==> \<exists>B \<in> C. a \<in> B & B \<subseteq> A"
+lemma lemma1: "\<lbrakk>\<Union>(C)=A; a \<in> A\<rbrakk> \<Longrightarrow> \<exists>B \<in> C. a \<in> B & B \<subseteq> A"
by fast
lemma lemma2:
- "[| pairwise_disjoint(A); B \<in> A; C \<in> A; a \<in> B; a \<in> C |] ==> B=C"
+ "\<lbrakk>pairwise_disjoint(A); B \<in> A; C \<in> A; a \<in> B; a \<in> C\<rbrakk> \<Longrightarrow> B=C"
by (unfold pairwise_disjoint_def, blast)
lemma lemma3:
"\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) &
sets_of_size_between(f`B, 2, n) & \<Union>(f`B)=B
- ==> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &
+ \<Longrightarrow> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &
0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"
apply (unfold sets_of_size_between_def)
apply (rule ballI)
@@ -64,7 +64,7 @@
apply (blast dest: lemma1 intro!: lemma2, blast)
done
-lemma lemma4: "[| A \<lesssim> i; Ord(i) |] ==> {P(a). a \<in> A} \<lesssim> i"
+lemma lemma4: "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> {P(a). a \<in> A} \<lesssim> i"
apply (unfold lepoll_def)
apply (erule exE)
apply (rule_tac x = "\<lambda>x \<in> RepFun(A,P). \<mu> j. \<exists>a\<in>A. x=P(a) & f`a=j"
@@ -81,7 +81,7 @@
done
lemma lemma5_1:
- "[| B \<in> A; 2 \<lesssim> u(B) |] ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<noteq> 0"
+ "\<lbrakk>B \<in> A; 2 \<lesssim> u(B)\<rbrakk> \<Longrightarrow> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<noteq> 0"
apply simp
apply (fast dest: lepoll_Diff_sing
elim: lepoll_trans [THEN succ_lepoll_natE] ssubst
@@ -89,24 +89,24 @@
done
lemma lemma5_2:
- "[| B \<in> A; u(B) \<subseteq> cons(0, B*nat) |]
- ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<subseteq> B"
+ "\<lbrakk>B \<in> A; u(B) \<subseteq> cons(0, B*nat)\<rbrakk>
+ \<Longrightarrow> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<subseteq> B"
apply auto
done
lemma lemma5_3:
- "[| n \<in> nat; B \<in> A; 0 \<in> u(B); u(B) \<lesssim> succ(n) |]
- ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<lesssim> n"
+ "\<lbrakk>n \<in> nat; B \<in> A; 0 \<in> u(B); u(B) \<lesssim> succ(n)\<rbrakk>
+ \<Longrightarrow> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<lesssim> n"
apply simp
apply (fast elim!: Diff_lepoll [THEN lemma4 [OF _ nat_into_Ord]])
done
lemma ex_fun_AC13_AC15:
- "[| \<forall>B \<in> {cons(0, x*nat). x \<in> A}.
+ "\<lbrakk>\<forall>B \<in> {cons(0, x*nat). x \<in> A}.
pairwise_disjoint(f`B) &
sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B;
- n \<in> nat |]
- ==> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"
+ n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"
by (fast del: subsetI notI
dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)
@@ -116,39 +116,39 @@
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* AC10(n) ==> AC11 *)
+(* AC10(n) \<Longrightarrow> AC11 *)
(* ********************************************************************** *)
-theorem AC10_AC11: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC11"
+theorem AC10_AC11: "\<lbrakk>n \<in> nat; 1\<le>n; AC10(n)\<rbrakk> \<Longrightarrow> AC11"
by (unfold AC10_def AC11_def, blast)
(* ********************************************************************** *)
-(* AC11 ==> AC12 *)
+(* AC11 \<Longrightarrow> AC12 *)
(* ********************************************************************** *)
-theorem AC11_AC12: "AC11 ==> AC12"
+theorem AC11_AC12: "AC11 \<Longrightarrow> AC12"
by (unfold AC10_def AC11_def AC11_def AC12_def, blast)
(* ********************************************************************** *)
-(* AC12 ==> AC15 *)
+(* AC12 \<Longrightarrow> AC15 *)
(* ********************************************************************** *)
-theorem AC12_AC15: "AC12 ==> AC15"
+theorem AC12_AC15: "AC12 \<Longrightarrow> AC15"
apply (unfold AC12_def AC15_def)
apply (blast del: ballI
intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15)
done
(* ********************************************************************** *)
-(* AC15 ==> WO6 *)
+(* AC15 \<Longrightarrow> WO6 *)
(* ********************************************************************** *)
-lemma OUN_eq_UN: "Ord(x) ==> (\<Union>a<x. F(a)) = (\<Union>a \<in> x. F(a))"
+lemma OUN_eq_UN: "Ord(x) \<Longrightarrow> (\<Union>a<x. F(a)) = (\<Union>a \<in> x. F(a))"
by (fast intro!: ltI dest!: ltD)
lemma AC15_WO6_aux1:
"\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m
- ==> (\<Union>i<\<mu> x. HH(f,A,x)={A}. HH(f,A,i)) = A"
+ \<Longrightarrow> (\<Union>i<\<mu> x. HH(f,A,x)={A}. HH(f,A,i)) = A"
apply (simp add: Ord_Least [THEN OUN_eq_UN])
apply (rule equalityI)
apply (fast dest!: less_Least_subset_x)
@@ -158,7 +158,7 @@
lemma AC15_WO6_aux2:
"\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m
- ==> \<forall>x < (\<mu> x. HH(f,A,x)={A}). HH(f,A,x) \<lesssim> m"
+ \<Longrightarrow> \<forall>x < (\<mu> x. HH(f,A,x)={A}). HH(f,A,x) \<lesssim> m"
apply (rule oallI)
apply (drule ltD [THEN less_Least_subset_x])
apply (frule HH_subset_imp_eq)
@@ -167,7 +167,7 @@
(*but can't use del: DiffE despite the obvious conflict*)
done
-theorem AC15_WO6: "AC15 ==> WO6"
+theorem AC15_WO6: "AC15 \<Longrightarrow> WO6"
apply (unfold AC15_def WO6_def)
apply (rule allI)
apply (erule_tac x = "Pow (A) -{0}" in allE)
@@ -188,14 +188,14 @@
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* AC10(n) ==> AC13(n-1) if 2\<le>n *)
+(* AC10(n) \<Longrightarrow> AC13(n-1) if 2\<le>n *)
(* *)
(* Because of the change to the formal definition of AC10(n) we prove *)
(* the following obviously equivalent theorem \<in> *)
(* AC10(n) implies AC13(n) for (1\<le>n) *)
(* ********************************************************************** *)
-theorem AC10_AC13: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC13(n)"
+theorem AC10_AC13: "\<lbrakk>n \<in> nat; 1\<le>n; AC10(n)\<rbrakk> \<Longrightarrow> AC13(n)"
apply (unfold AC10_def AC13_def, safe)
apply (erule allE)
apply (erule impE [OF _ cons_times_nat_not_Finite], assumption)
@@ -208,10 +208,10 @@
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* AC1 ==> AC13(1) *)
+(* AC1 \<Longrightarrow> AC13(1) *)
(* ********************************************************************** *)
-lemma AC1_AC13: "AC1 ==> AC13(1)"
+lemma AC1_AC13: "AC1 \<Longrightarrow> AC13(1)"
apply (unfold AC1_def AC13_def)
apply (rule allI)
apply (erule allE)
@@ -223,10 +223,10 @@
done
(* ********************************************************************** *)
-(* AC13(m) ==> AC13(n) for m \<subseteq> n *)
+(* AC13(m) \<Longrightarrow> AC13(n) for m \<subseteq> n *)
(* ********************************************************************** *)
-lemma AC13_mono: "[| m\<le>n; AC13(m) |] ==> AC13(n)"
+lemma AC13_mono: "\<lbrakk>m\<le>n; AC13(m)\<rbrakk> \<Longrightarrow> AC13(n)"
apply (unfold AC13_def)
apply (drule le_imp_lepoll)
apply (fast elim!: lepoll_trans)
@@ -237,17 +237,17 @@
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* AC13(n) ==> AC14 if 1 \<subseteq> n *)
+(* AC13(n) \<Longrightarrow> AC14 if 1 \<subseteq> n *)
(* ********************************************************************** *)
-theorem AC13_AC14: "[| n \<in> nat; 1\<le>n; AC13(n) |] ==> AC14"
+theorem AC13_AC14: "\<lbrakk>n \<in> nat; 1\<le>n; AC13(n)\<rbrakk> \<Longrightarrow> AC14"
by (unfold AC13_def AC14_def, auto)
(* ********************************************************************** *)
-(* AC14 ==> AC15 *)
+(* AC14 \<Longrightarrow> AC15 *)
(* ********************************************************************** *)
-theorem AC14_AC15: "AC14 ==> AC15"
+theorem AC14_AC15: "AC14 \<Longrightarrow> AC15"
by (unfold AC13_def AC14_def AC15_def, fast)
(* ********************************************************************** *)
@@ -255,15 +255,15 @@
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* AC13(1) ==> AC1 *)
+(* AC13(1) \<Longrightarrow> AC1 *)
(* ********************************************************************** *)
-lemma lemma_aux: "[| A\<noteq>0; A \<lesssim> 1 |] ==> \<exists>a. A={a}"
+lemma lemma_aux: "\<lbrakk>A\<noteq>0; A \<lesssim> 1\<rbrakk> \<Longrightarrow> \<exists>a. A={a}"
by (fast elim!: not_emptyE lepoll_1_is_sing)
lemma AC13_AC1_lemma:
"\<forall>B \<in> A. f(B)\<noteq>0 & f(B)<=B & f(B) \<lesssim> 1
- ==> (\<lambda>x \<in> A. THE y. f(x)={y}) \<in> (\<Prod>X \<in> A. X)"
+ \<Longrightarrow> (\<lambda>x \<in> A. THE y. f(x)={y}) \<in> (\<Prod>X \<in> A. X)"
apply (rule lam_type)
apply (drule bspec, assumption)
apply (elim conjE)
@@ -271,16 +271,16 @@
apply (simp add: the_equality)
done
-theorem AC13_AC1: "AC13(1) ==> AC1"
+theorem AC13_AC1: "AC13(1) \<Longrightarrow> AC1"
apply (unfold AC13_def AC1_def)
apply (fast elim!: AC13_AC1_lemma)
done
(* ********************************************************************** *)
-(* AC11 ==> AC14 *)
+(* AC11 \<Longrightarrow> AC14 *)
(* ********************************************************************** *)
-theorem AC11_AC14: "AC11 ==> AC14"
+theorem AC11_AC14: "AC11 \<Longrightarrow> AC14"
apply (unfold AC11_def AC14_def)
apply (fast intro!: AC10_AC13)
done
--- a/src/ZF/AC/AC16_WO4.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC16_WO4.thy Tue Sep 27 16:51:35 2022 +0100
@@ -1,7 +1,7 @@
(* Title: ZF/AC/AC16_WO4.thy
Author: Krzysztof Grabczewski
-The proof of AC16(n, k) ==> WO4(n-k)
+The proof of AC16(n, k) \<Longrightarrow> WO4(n-k)
Tidied (using locales) by LCP
*)
@@ -15,8 +15,8 @@
(* ********************************************************************** *)
lemma lemma1:
- "[| Finite(A); 0<m; m \<in> nat |]
- ==> \<exists>a f. Ord(a) & domain(f) = a &
+ "\<lbrakk>Finite(A); 0<m; m \<in> nat\<rbrakk>
+ \<Longrightarrow> \<exists>a f. Ord(a) & domain(f) = a &
(\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
apply (unfold Finite_def)
apply (erule bexE)
@@ -35,10 +35,10 @@
(* The case of infinite set *)
(* ********************************************************************** *)
-(* well_ord(x,r) ==> well_ord({{y,z}. y \<in> x}, Something(x,z)) **)
+(* well_ord(x,r) \<Longrightarrow> well_ord({{y,z}. y \<in> x}, Something(x,z)) **)
lemmas well_ord_paired = paired_bij [THEN bij_is_inj, THEN well_ord_rvimage]
-lemma lepoll_trans1: "[| A \<lesssim> B; ~ A \<lesssim> C |] ==> ~ B \<lesssim> C"
+lemma lepoll_trans1: "\<lbrakk>A \<lesssim> B; \<not> A \<lesssim> C\<rbrakk> \<Longrightarrow> \<not> B \<lesssim> C"
by (blast intro: lepoll_trans)
(* ********************************************************************** *)
@@ -47,7 +47,7 @@
lemmas lepoll_paired = paired_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]
-lemma lemma2: "\<exists>y R. well_ord(y,R) & x \<inter> y = 0 & ~y \<lesssim> z & ~Finite(y)"
+lemma lemma2: "\<exists>y R. well_ord(y,R) & x \<inter> y = 0 & \<not>y \<lesssim> z & \<not>Finite(y)"
apply (rule_tac x = "{{a,x}. a \<in> nat \<union> Hartog (z) }" in exI)
apply (rule well_ord_Un [OF Ord_nat [THEN well_ord_Memrel]
Ord_Hartog [THEN well_ord_Memrel], THEN exE])
@@ -60,7 +60,7 @@
lepoll_paired [THEN lepoll_Finite])
done
-lemma infinite_Un: "~Finite(B) ==> ~Finite(A \<union> B)"
+lemma infinite_Un: "\<not>Finite(B) \<Longrightarrow> \<not>Finite(A \<union> B)"
by (blast intro: subset_Finite)
(* ********************************************************************** *)
@@ -77,8 +77,8 @@
(*Proof simplified by LCP*)
lemma succ_not_lepoll_lemma:
- "[| ~(\<exists>x \<in> A. f`x=y); f \<in> inj(A, B); y \<in> B |]
- ==> (\<lambda>a \<in> succ(A). if(a=A, y, f`a)) \<in> inj(succ(A), B)"
+ "\<lbrakk>\<not>(\<exists>x \<in> A. f`x=y); f \<in> inj(A, B); y \<in> B\<rbrakk>
+ \<Longrightarrow> (\<lambda>a \<in> succ(A). if(a=A, y, f`a)) \<in> inj(succ(A), B)"
apply (rule_tac d = "%z. if (z=y, A, converse (f) `z) " in lam_injective)
apply (force simp add: inj_is_fun [THEN apply_type])
(*this preliminary simplification prevents looping somehow*)
@@ -86,7 +86,7 @@
apply force
done
-lemma succ_not_lepoll_imp_eqpoll: "[| ~A \<approx> B; A \<lesssim> B |] ==> succ(A) \<lesssim> B"
+lemma succ_not_lepoll_imp_eqpoll: "\<lbrakk>\<not>A \<approx> B; A \<lesssim> B\<rbrakk> \<Longrightarrow> succ(A) \<lesssim> B"
apply (unfold lepoll_def eqpoll_def bij_def surj_def)
apply (fast elim!: succ_not_lepoll_lemma inj_is_fun)
done
@@ -100,16 +100,16 @@
ordermap_bij [THEN exI [THEN eqpoll_def [THEN def_imp_iff, THEN iffD2]]]
lemma cons_cons_subset:
- "[| a \<subseteq> y; b \<in> y-a; u \<in> x |] ==> cons(b, cons(u, a)) \<in> Pow(x \<union> y)"
+ "\<lbrakk>a \<subseteq> y; b \<in> y-a; u \<in> x\<rbrakk> \<Longrightarrow> cons(b, cons(u, a)) \<in> Pow(x \<union> y)"
by fast
lemma cons_cons_eqpoll:
- "[| a \<approx> k; a \<subseteq> y; b \<in> y-a; u \<in> x; x \<inter> y = 0 |]
- ==> cons(b, cons(u, a)) \<approx> succ(succ(k))"
+ "\<lbrakk>a \<approx> k; a \<subseteq> y; b \<in> y-a; u \<in> x; x \<inter> y = 0\<rbrakk>
+ \<Longrightarrow> cons(b, cons(u, a)) \<approx> succ(succ(k))"
by (fast intro!: cons_eqpoll_succ)
lemma set_eq_cons:
- "[| succ(k) \<approx> A; k \<approx> B; B \<subseteq> A; a \<in> A-B; k \<in> nat |] ==> A = cons(a, B)"
+ "\<lbrakk>succ(k) \<approx> A; k \<approx> B; B \<subseteq> A; a \<in> A-B; k \<in> nat\<rbrakk> \<Longrightarrow> A = cons(a, B)"
apply (rule equalityI)
prefer 2 apply fast
apply (rule Diff_eq_0_iff [THEN iffD1])
@@ -121,10 +121,10 @@
OF eqpoll_sym [THEN eqpoll_imp_lepoll] subset_imp_lepoll])
done
-lemma cons_eqE: "[| cons(x,a) = cons(y,a); x \<notin> a |] ==> x = y "
+lemma cons_eqE: "\<lbrakk>cons(x,a) = cons(y,a); x \<notin> a\<rbrakk> \<Longrightarrow> x = y "
by (fast elim!: equalityE)
-lemma eq_imp_Int_eq: "A = B ==> A \<inter> C = B \<inter> C"
+lemma eq_imp_Int_eq: "A = B \<Longrightarrow> A \<inter> C = B \<inter> C"
by blast
(* ********************************************************************** *)
@@ -132,8 +132,8 @@
(* ********************************************************************** *)
lemma eqpoll_sum_imp_Diff_lepoll_lemma [rule_format]:
- "[| k \<in> nat; m \<in> nat |]
- ==> \<forall>A B. A \<approx> k #+ m & k \<lesssim> B & B \<subseteq> A \<longrightarrow> A-B \<lesssim> m"
+ "\<lbrakk>k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> \<forall>A B. A \<approx> k #+ m & k \<lesssim> B & B \<subseteq> A \<longrightarrow> A-B \<lesssim> m"
apply (induct_tac "k")
apply (simp add: add_0)
apply (blast intro: eqpoll_imp_lepoll lepoll_trans
@@ -150,8 +150,8 @@
done
lemma eqpoll_sum_imp_Diff_lepoll:
- "[| A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<lesssim> B; k \<in> nat; m \<in> nat |]
- ==> A-B \<lesssim> m"
+ "\<lbrakk>A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<lesssim> B; k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> A-B \<lesssim> m"
apply (simp only: add_succ [symmetric])
apply (blast intro: eqpoll_sum_imp_Diff_lepoll_lemma)
done
@@ -161,8 +161,8 @@
(* ********************************************************************** *)
lemma eqpoll_sum_imp_Diff_eqpoll_lemma [rule_format]:
- "[| k \<in> nat; m \<in> nat |]
- ==> \<forall>A B. A \<approx> k #+ m & k \<approx> B & B \<subseteq> A \<longrightarrow> A-B \<approx> m"
+ "\<lbrakk>k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> \<forall>A B. A \<approx> k #+ m & k \<approx> B & B \<subseteq> A \<longrightarrow> A-B \<approx> m"
apply (induct_tac "k")
apply (force dest!: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_0_is_0])
apply (intro allI impI)
@@ -177,8 +177,8 @@
done
lemma eqpoll_sum_imp_Diff_eqpoll:
- "[| A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<approx> B; k \<in> nat; m \<in> nat |]
- ==> A-B \<approx> m"
+ "\<lbrakk>A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<approx> B; k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> A-B \<approx> m"
apply (simp only: add_succ [symmetric])
apply (blast intro: eqpoll_sum_imp_Diff_eqpoll_lemma)
done
@@ -192,24 +192,24 @@
by (fast dest!: lepoll_0_is_0 intro!: lepoll_refl)
lemma subsets_lepoll_succ:
- "n \<in> nat ==> {z \<in> Pow(y). z \<lesssim> succ(n)} =
+ "n \<in> nat \<Longrightarrow> {z \<in> Pow(y). z \<lesssim> succ(n)} =
{z \<in> Pow(y). z \<lesssim> n} \<union> {z \<in> Pow(y). z \<approx> succ(n)}"
by (blast intro: leI le_imp_lepoll nat_into_Ord
lepoll_trans eqpoll_imp_lepoll
dest!: lepoll_succ_disj)
lemma Int_empty:
- "n \<in> nat ==> {z \<in> Pow(y). z \<lesssim> n} \<inter> {z \<in> Pow(y). z \<approx> succ(n)} = 0"
+ "n \<in> nat \<Longrightarrow> {z \<in> Pow(y). z \<lesssim> n} \<inter> {z \<in> Pow(y). z \<approx> succ(n)} = 0"
by (blast intro: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]
succ_lepoll_natE)
locale AC16 =
fixes x and y and k and l and m and t_n and R and MM and LL and GG and s
- defines k_def: "k == succ(l)"
- and MM_def: "MM == {v \<in> t_n. succ(k) \<lesssim> v \<inter> y}"
- and LL_def: "LL == {v \<inter> y. v \<in> MM}"
- and GG_def: "GG == \<lambda>v \<in> LL. (THE w. w \<in> MM & v \<subseteq> w) - v"
- and s_def: "s(u) == {v \<in> t_n. u \<in> v & k \<lesssim> v \<inter> y}"
+ defines k_def: "k \<equiv> succ(l)"
+ and MM_def: "MM \<equiv> {v \<in> t_n. succ(k) \<lesssim> v \<inter> y}"
+ and LL_def: "LL \<equiv> {v \<inter> y. v \<in> MM}"
+ and GG_def: "GG \<equiv> \<lambda>v \<in> LL. (THE w. w \<in> MM & v \<subseteq> w) - v"
+ and s_def: "s(u) \<equiv> {v \<in> t_n. u \<in> v & k \<lesssim> v \<inter> y}"
assumes all_ex: "\<forall>z \<in> {z \<in> Pow(x \<union> y) . z \<approx> succ(k)}.
\<exists>! w. w \<in> t_n & z \<subseteq> w "
and disjoint[iff]: "x \<inter> y = 0"
@@ -218,8 +218,8 @@
and lnat[iff]: "l \<in> nat"
and mnat[iff]: "m \<in> nat"
and mpos[iff]: "0<m"
- and Infinite[iff]: "~ Finite(y)"
- and noLepoll: "~ y \<lesssim> {v \<in> Pow(x). v \<approx> m}"
+ and Infinite[iff]: "\<not> Finite(y)"
+ and noLepoll: "\<not> y \<lesssim> {v \<in> Pow(x). v \<approx> m}"
begin
lemma knat [iff]: "k \<in> nat"
@@ -231,7 +231,7 @@
(* where a is certain k-2 element subset of y *)
(* ********************************************************************** *)
-lemma Diff_Finite_eqpoll: "[| l \<approx> a; a \<subseteq> y |] ==> y - a \<approx> y"
+lemma Diff_Finite_eqpoll: "\<lbrakk>l \<approx> a; a \<subseteq> y\<rbrakk> \<Longrightarrow> y - a \<approx> y"
apply (insert WO_R Infinite lnat)
apply (rule eqpoll_trans)
apply (rule Diff_lesspoll_eqpoll_Card)
@@ -252,20 +252,20 @@
by (unfold s_def, blast)
lemma sI:
- "[| w \<in> t_n; cons(b,cons(u,a)) \<subseteq> w; a \<subseteq> y; b \<in> y-a; l \<approx> a |]
- ==> w \<in> s(u)"
+ "\<lbrakk>w \<in> t_n; cons(b,cons(u,a)) \<subseteq> w; a \<subseteq> y; b \<in> y-a; l \<approx> a\<rbrakk>
+ \<Longrightarrow> w \<in> s(u)"
apply (unfold s_def succ_def k_def)
apply (blast intro!: eqpoll_imp_lepoll [THEN cons_lepoll_cong]
intro: subset_imp_lepoll lepoll_trans)
done
-lemma in_s_imp_u_in: "v \<in> s(u) ==> u \<in> v"
+lemma in_s_imp_u_in: "v \<in> s(u) \<Longrightarrow> u \<in> v"
by (unfold s_def, blast)
lemma ex1_superset_a:
- "[| l \<approx> a; a \<subseteq> y; b \<in> y - a; u \<in> x |]
- ==> \<exists>! c. c \<in> s(u) & a \<subseteq> c & b \<in> c"
+ "\<lbrakk>l \<approx> a; a \<subseteq> y; b \<in> y - a; u \<in> x\<rbrakk>
+ \<Longrightarrow> \<exists>! c. c \<in> s(u) & a \<subseteq> c & b \<in> c"
apply (rule all_ex [simplified k_def, THEN ballE])
apply (erule ex1E)
apply (rule_tac a = w in ex1I, blast intro: sI)
@@ -275,18 +275,18 @@
done
lemma the_eq_cons:
- "[| \<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y;
- l \<approx> a; a \<subseteq> y; b \<in> y - a; u \<in> x |]
- ==> (THE c. c \<in> s(u) & a \<subseteq> c & b \<in> c) \<inter> y = cons(b, a)"
+ "\<lbrakk>\<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y;
+ l \<approx> a; a \<subseteq> y; b \<in> y - a; u \<in> x\<rbrakk>
+ \<Longrightarrow> (THE c. c \<in> s(u) & a \<subseteq> c & b \<in> c) \<inter> y = cons(b, a)"
apply (frule ex1_superset_a [THEN theI], assumption+)
apply (rule set_eq_cons)
apply (fast+)
done
lemma y_lepoll_subset_s:
- "[| \<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y;
- l \<approx> a; a \<subseteq> y; u \<in> x |]
- ==> y \<lesssim> {v \<in> s(u). a \<subseteq> v}"
+ "\<lbrakk>\<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y;
+ l \<approx> a; a \<subseteq> y; u \<in> x\<rbrakk>
+ \<Longrightarrow> y \<lesssim> {v \<in> s(u). a \<subseteq> v}"
apply (rule Diff_Finite_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll,
THEN lepoll_trans], fast+)
apply (rule_tac f3 = "\<lambda>b \<in> y-a. THE c. c \<in> s (u) & a \<subseteq> c & b \<in> c"
@@ -307,18 +307,18 @@
(*relies on the disjointness of x, y*)
-lemma x_imp_not_y [dest]: "a \<in> x ==> a \<notin> y"
+lemma x_imp_not_y [dest]: "a \<in> x \<Longrightarrow> a \<notin> y"
by (blast dest: disjoint [THEN equalityD1, THEN subsetD, OF IntI])
lemma w_Int_eq_w_Diff:
- "w \<subseteq> x \<union> y ==> w \<inter> (x - {u}) = w - cons(u, w \<inter> y)"
+ "w \<subseteq> x \<union> y \<Longrightarrow> w \<inter> (x - {u}) = w - cons(u, w \<inter> y)"
by blast
lemma w_Int_eqpoll_m:
- "[| w \<in> {v \<in> s(u). a \<subseteq> v};
+ "\<lbrakk>w \<in> {v \<in> s(u). a \<subseteq> v};
l \<approx> a; u \<in> x;
- \<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y |]
- ==> w \<inter> (x - {u}) \<approx> m"
+ \<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y\<rbrakk>
+ \<Longrightarrow> w \<inter> (x - {u}) \<approx> m"
apply (erule CollectE)
apply (subst w_Int_eq_w_Diff)
apply (fast dest!: s_subset [THEN subsetD]
@@ -336,22 +336,22 @@
(* to {v \<in> Pow(x). v \<approx> n-k} *)
(* ********************************************************************** *)
-lemma eqpoll_m_not_empty: "a \<approx> m ==> a \<noteq> 0"
+lemma eqpoll_m_not_empty: "a \<approx> m \<Longrightarrow> a \<noteq> 0"
apply (insert mpos)
apply (fast elim!: zero_lt_natE dest!: eqpoll_succ_imp_not_empty)
done
lemma cons_cons_in:
- "[| z \<in> xa \<inter> (x - {u}); l \<approx> a; a \<subseteq> y; u \<in> x |]
- ==> \<exists>! w. w \<in> t_n & cons(z, cons(u, a)) \<subseteq> w"
+ "\<lbrakk>z \<in> xa \<inter> (x - {u}); l \<approx> a; a \<subseteq> y; u \<in> x\<rbrakk>
+ \<Longrightarrow> \<exists>! w. w \<in> t_n & cons(z, cons(u, a)) \<subseteq> w"
apply (rule all_ex [THEN bspec])
apply (unfold k_def)
apply (fast intro!: cons_eqpoll_succ elim: eqpoll_sym)
done
lemma subset_s_lepoll_w:
- "[| \<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y; a \<subseteq> y; l \<approx> a; u \<in> x |]
- ==> {v \<in> s(u). a \<subseteq> v} \<lesssim> {v \<in> Pow(x). v \<approx> m}"
+ "\<lbrakk>\<forall>v \<in> s(u). succ(l) \<approx> v \<inter> y; a \<subseteq> y; l \<approx> a; u \<in> x\<rbrakk>
+ \<Longrightarrow> {v \<in> s(u). a \<subseteq> v} \<lesssim> {v \<in> Pow(x). v \<approx> m}"
apply (rule_tac f3 = "\<lambda>w \<in> {v \<in> s (u) . a \<subseteq> v}. w \<inter> (x - {u})"
in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]])
apply (simp add: inj_def)
@@ -373,14 +373,14 @@
(* ********************************************************************** *)
lemma well_ord_subsets_eqpoll_n:
- "n \<in> nat ==> \<exists>S. well_ord({z \<in> Pow(y) . z \<approx> succ(n)}, S)"
+ "n \<in> nat \<Longrightarrow> \<exists>S. well_ord({z \<in> Pow(y) . z \<approx> succ(n)}, S)"
apply (rule WO_R [THEN well_ord_infinite_subsets_eqpoll_X,
THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE])
apply (fast intro: bij_is_inj [THEN well_ord_rvimage])+
done
lemma well_ord_subsets_lepoll_n:
- "n \<in> nat ==> \<exists>R. well_ord({z \<in> Pow(y). z \<lesssim> n}, R)"
+ "n \<in> nat \<Longrightarrow> \<exists>R. well_ord({z \<in> Pow(y). z \<lesssim> n}, R)"
apply (induct_tac "n")
apply (force intro!: well_ord_unit simp add: subsets_lepoll_0_eq_unit)
apply (erule exE)
@@ -410,7 +410,7 @@
(* ********************************************************************** *)
lemma unique_superset_in_MM:
- "v \<in> LL ==> \<exists>! w. w \<in> MM & v \<subseteq> w"
+ "v \<in> LL \<Longrightarrow> \<exists>! w. w \<in> MM & v \<subseteq> w"
apply (unfold MM_def LL_def, safe, fast)
apply (rule lepoll_imp_eqpoll_subset [THEN exE], assumption)
apply (rule_tac x = x in all_ex [THEN ballE])
@@ -426,11 +426,11 @@
(* The union of appropriate values is the whole x *)
(* ********************************************************************** *)
-lemma Int_in_LL: "w \<in> MM ==> w \<inter> y \<in> LL"
+lemma Int_in_LL: "w \<in> MM \<Longrightarrow> w \<inter> y \<in> LL"
by (unfold LL_def, fast)
lemma in_LL_eq_Int:
- "v \<in> LL ==> v = (THE x. x \<in> MM & v \<subseteq> x) \<inter> y"
+ "v \<in> LL \<Longrightarrow> v = (THE x. x \<in> MM & v \<subseteq> x) \<inter> y"
apply (unfold LL_def, clarify)
apply (subst unique_superset_in_MM [THEN the_equality2])
apply (auto simp add: Int_in_LL)
@@ -440,13 +440,13 @@
by (erule unique_superset_in_MM [THEN theI, THEN conjunct1])
lemma the_in_MM_subset:
- "v \<in> LL ==> (THE x. x \<in> MM & v \<subseteq> x) \<subseteq> x \<union> y"
+ "v \<in> LL \<Longrightarrow> (THE x. x \<in> MM & v \<subseteq> x) \<subseteq> x \<union> y"
apply (drule unique_superset1)
apply (unfold MM_def)
apply (fast dest!: unique_superset1 "includes" [THEN subsetD])
done
-lemma GG_subset: "v \<in> LL ==> GG ` v \<subseteq> x"
+lemma GG_subset: "v \<in> LL \<Longrightarrow> GG ` v \<subseteq> x"
apply (unfold GG_def)
apply (frule the_in_MM_subset)
apply (frule in_LL_eq_Int)
@@ -461,7 +461,7 @@
apply (rule Infinite)
done
-lemma ex_subset_eqpoll_n: "n \<in> nat ==> \<exists>z. z \<subseteq> y & n \<approx> z"
+lemma ex_subset_eqpoll_n: "n \<in> nat \<Longrightarrow> \<exists>z. z \<subseteq> y & n \<approx> z"
apply (erule nat_lepoll_imp_ex_eqpoll_n)
apply (rule lepoll_trans [OF nat_lepoll_ordertype])
apply (rule ordertype_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll])
@@ -469,7 +469,7 @@
done
-lemma exists_proper_in_s: "u \<in> x ==> \<exists>v \<in> s(u). succ(k) \<lesssim> v \<inter> y"
+lemma exists_proper_in_s: "u \<in> x \<Longrightarrow> \<exists>v \<in> s(u). succ(k) \<lesssim> v \<inter> y"
apply (rule ccontr)
apply (subgoal_tac "\<forall>v \<in> s (u) . k \<approx> v \<inter> y")
prefer 2 apply (simp add: s_def, blast intro: succ_not_lepoll_imp_eqpoll)
@@ -480,12 +480,12 @@
apply (blast intro: lepoll_trans [OF y_lepoll_subset_s subset_s_lepoll_w])
done
-lemma exists_in_MM: "u \<in> x ==> \<exists>w \<in> MM. u \<in> w"
+lemma exists_in_MM: "u \<in> x \<Longrightarrow> \<exists>w \<in> MM. u \<in> w"
apply (erule exists_proper_in_s [THEN bexE])
apply (unfold MM_def s_def, fast)
done
-lemma exists_in_LL: "u \<in> x ==> \<exists>w \<in> LL. u \<in> GG`w"
+lemma exists_in_LL: "u \<in> x \<Longrightarrow> \<exists>w \<in> LL. u \<in> GG`w"
apply (rule exists_in_MM [THEN bexE], assumption)
apply (rule bexI)
apply (erule_tac [2] Int_in_LL)
@@ -495,7 +495,7 @@
apply (fast elim!: Int_in_LL)+
done
-lemma OUN_eq_x: "well_ord(LL,S) ==>
+lemma OUN_eq_x: "well_ord(LL,S) \<Longrightarrow>
(\<Union>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b)) = x"
apply (rule equalityI)
apply (rule subsetI)
@@ -515,19 +515,19 @@
(* Every element of the family is less than or equipollent to n-k (m) *)
(* ********************************************************************** *)
-lemma in_MM_eqpoll_n: "w \<in> MM ==> w \<approx> succ(k #+ m)"
+lemma in_MM_eqpoll_n: "w \<in> MM \<Longrightarrow> w \<approx> succ(k #+ m)"
apply (unfold MM_def)
apply (fast dest: "includes" [THEN subsetD])
done
-lemma in_LL_eqpoll_n: "w \<in> LL ==> succ(k) \<lesssim> w"
+lemma in_LL_eqpoll_n: "w \<in> LL \<Longrightarrow> succ(k) \<lesssim> w"
by (unfold LL_def MM_def, fast)
-lemma in_LL: "w \<in> LL ==> w \<subseteq> (THE x. x \<in> MM \<and> w \<subseteq> x)"
+lemma in_LL: "w \<in> LL \<Longrightarrow> w \<subseteq> (THE x. x \<in> MM \<and> w \<subseteq> x)"
by (erule subset_trans [OF in_LL_eq_Int [THEN equalityD1] Int_lower1])
lemma all_in_lepoll_m:
- "well_ord(LL,S) ==>
+ "well_ord(LL,S) \<Longrightarrow>
\<forall>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b) \<lesssim> m"
apply (unfold GG_def)
apply (rule oallI)
@@ -558,11 +558,11 @@
(* ********************************************************************** *)
-(* The main theorem AC16(n, k) ==> WO4(n-k) *)
+(* The main theorem AC16(n, k) \<Longrightarrow> WO4(n-k) *)
(* ********************************************************************** *)
theorem AC16_WO4:
- "[| AC_Equiv.AC16(k #+ m, k); 0 < k; 0 < m; k \<in> nat; m \<in> nat |] ==> WO4(m)"
+ "\<lbrakk>AC_Equiv.AC16(k #+ m, k); 0 < k; 0 < m; k \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> WO4(m)"
apply (unfold AC_Equiv.AC16_def WO4_def)
apply (rule allI)
apply (case_tac "Finite (A)")
--- a/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 16:51:35 2022 +0100
@@ -8,7 +8,7 @@
imports AC_Equiv Hartog Cardinal_aux
begin
-lemma cons_Diff_eq: "a\<notin>A ==> cons(a,A)-{a}=A"
+lemma cons_Diff_eq: "a\<notin>A \<Longrightarrow> cons(a,A)-{a}=A"
by fast
lemma nat_1_lepoll_iff: "1\<lesssim>X \<longleftrightarrow> (\<exists>x. x \<in> X)"
@@ -28,7 +28,7 @@
apply (fast intro!: singleton_eqpoll_1)
done
-lemma cons_eqpoll_succ: "[| x\<approx>n; y\<notin>x |] ==> cons(y,x)\<approx>succ(n)"
+lemma cons_eqpoll_succ: "\<lbrakk>x\<approx>n; y\<notin>x\<rbrakk> \<Longrightarrow> cons(y,x)\<approx>succ(n)"
apply (unfold succ_def)
apply (fast elim!: cons_eqpoll_cong mem_irrefl)
done
@@ -60,7 +60,7 @@
done
lemma InfCard_Least_in:
- "[| InfCard(x); y \<subseteq> x; y \<approx> succ(z) |] ==> (\<mu> i. i \<in> y) \<in> y"
+ "\<lbrakk>InfCard(x); y \<subseteq> x; y \<approx> succ(z)\<rbrakk> \<Longrightarrow> (\<mu> i. i \<in> y) \<in> y"
apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll,
THEN succ_lepoll_imp_not_empty, THEN not_emptyE])
apply (fast intro: LeastI
@@ -69,8 +69,8 @@
done
lemma subsets_lepoll_lemma1:
- "[| InfCard(x); n \<in> nat |]
- ==> {y \<in> Pow(x). y\<approx>succ(succ(n))} \<lesssim> x*{y \<in> Pow(x). y\<approx>succ(n)}"
+ "\<lbrakk>InfCard(x); n \<in> nat\<rbrakk>
+ \<Longrightarrow> {y \<in> Pow(x). y\<approx>succ(succ(n))} \<lesssim> x*{y \<in> Pow(x). y\<approx>succ(n)}"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>y \<in> {y \<in> Pow(x) . y\<approx>succ (succ (n))}.
<\<mu> i. i \<in> y, y-{\<mu> i. i \<in> y}>" in exI)
@@ -79,7 +79,7 @@
apply (simp, blast intro: InfCard_Least_in)
done
-lemma set_of_Ord_succ_Union: "(\<forall>y \<in> z. Ord(y)) ==> z \<subseteq> succ(\<Union>(z))"
+lemma set_of_Ord_succ_Union: "(\<forall>y \<in> z. Ord(y)) \<Longrightarrow> z \<subseteq> succ(\<Union>(z))"
apply (rule subsetI)
apply (case_tac "\<forall>y \<in> z. y \<subseteq> x", blast )
apply (simp, erule bexE)
@@ -87,11 +87,11 @@
apply (blast dest: le_imp_subset elim: leE ltE)+
done
-lemma subset_not_mem: "j \<subseteq> i ==> i \<notin> j"
+lemma subset_not_mem: "j \<subseteq> i \<Longrightarrow> i \<notin> j"
by (fast elim!: mem_irrefl)
lemma succ_Union_not_mem:
- "(!!y. y \<in> z ==> Ord(y)) ==> succ(\<Union>(z)) \<notin> z"
+ "(\<And>y. y \<in> z \<Longrightarrow> Ord(y)) \<Longrightarrow> succ(\<Union>(z)) \<notin> z"
apply (rule set_of_Ord_succ_Union [THEN subset_not_mem], blast)
done
@@ -99,14 +99,14 @@
"\<Union>(cons(succ(\<Union>(z)),z)) = succ(\<Union>(z))"
by fast
-lemma Un_Ord_disj: "[| Ord(i); Ord(j) |] ==> i \<union> j = i | i \<union> j = j"
+lemma Un_Ord_disj: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> i \<union> j = i | i \<union> j = j"
by (fast dest!: le_imp_subset elim: Ord_linear_le)
-lemma Union_eq_Un: "x \<in> X ==> \<Union>(X) = x \<union> \<Union>(X-{x})"
+lemma Union_eq_Un: "x \<in> X \<Longrightarrow> \<Union>(X) = x \<union> \<Union>(X-{x})"
by fast
lemma Union_in_lemma [rule_format]:
- "n \<in> nat ==> \<forall>z. (\<forall>y \<in> z. Ord(y)) & z\<approx>n & z\<noteq>0 \<longrightarrow> \<Union>(z) \<in> z"
+ "n \<in> nat \<Longrightarrow> \<forall>z. (\<forall>y \<in> z. Ord(y)) & z\<approx>n & z\<noteq>0 \<longrightarrow> \<Union>(z) \<in> z"
apply (induct_tac "n")
apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
apply (intro allI impI)
@@ -126,11 +126,11 @@
apply (drule Un_Ord_disj, assumption, auto)
done
-lemma Union_in: "[| \<forall>x \<in> z. Ord(x); z\<approx>n; z\<noteq>0; n \<in> nat |] ==> \<Union>(z) \<in> z"
+lemma Union_in: "\<lbrakk>\<forall>x \<in> z. Ord(x); z\<approx>n; z\<noteq>0; n \<in> nat\<rbrakk> \<Longrightarrow> \<Union>(z) \<in> z"
by (blast intro: Union_in_lemma)
lemma succ_Union_in_x:
- "[| InfCard(x); z \<in> Pow(x); z\<approx>n; n \<in> nat |] ==> succ(\<Union>(z)) \<in> x"
+ "\<lbrakk>InfCard(x); z \<in> Pow(x); z\<approx>n; n \<in> nat\<rbrakk> \<Longrightarrow> succ(\<Union>(z)) \<in> x"
apply (rule Limit_has_succ [THEN ltE])
prefer 3 apply assumption
apply (erule InfCard_is_Limit)
@@ -142,8 +142,8 @@
done
lemma succ_lepoll_succ_succ:
- "[| InfCard(x); n \<in> nat |]
- ==> {y \<in> Pow(x). y\<approx>succ(n)} \<lesssim> {y \<in> Pow(x). y\<approx>succ(succ(n))}"
+ "\<lbrakk>InfCard(x); n \<in> nat\<rbrakk>
+ \<Longrightarrow> {y \<in> Pow(x). y\<approx>succ(n)} \<lesssim> {y \<in> Pow(x). y\<approx>succ(succ(n))}"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>z \<in> {y\<in>Pow(x). y\<approx>succ(n)}. cons(succ(\<Union>(z)), z)"
in exI)
@@ -159,7 +159,7 @@
done
lemma subsets_eqpoll_X:
- "[| InfCard(X); n \<in> nat |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)} \<approx> X"
+ "\<lbrakk>InfCard(X); n \<in> nat\<rbrakk> \<Longrightarrow> {Y \<in> Pow(X). Y\<approx>succ(n)} \<approx> X"
apply (induct_tac "n")
apply (rule subsets_eqpoll_1_eqpoll)
apply (rule eqpollI)
@@ -172,16 +172,16 @@
done
lemma image_vimage_eq:
- "[| f \<in> surj(A,B); y \<subseteq> B |] ==> f``(converse(f)``y) = y"
+ "\<lbrakk>f \<in> surj(A,B); y \<subseteq> B\<rbrakk> \<Longrightarrow> f``(converse(f)``y) = y"
apply (unfold surj_def)
apply (fast dest: apply_equality2 elim: apply_iff [THEN iffD2])
done
-lemma vimage_image_eq: "[| f \<in> inj(A,B); y \<subseteq> A |] ==> converse(f)``(f``y) = y"
+lemma vimage_image_eq: "\<lbrakk>f \<in> inj(A,B); y \<subseteq> A\<rbrakk> \<Longrightarrow> converse(f)``(f``y) = y"
by (fast elim!: inj_is_fun [THEN apply_Pair] dest: inj_equality)
lemma subsets_eqpoll:
- "A\<approx>B ==> {Y \<in> Pow(A). Y\<approx>n}\<approx>{Y \<in> Pow(B). Y\<approx>n}"
+ "A\<approx>B \<Longrightarrow> {Y \<in> Pow(A). Y\<approx>n}\<approx>{Y \<in> Pow(B). Y\<approx>n}"
apply (unfold eqpoll_def)
apply (erule exE)
apply (rule_tac x = "\<lambda>X \<in> {Y \<in> Pow (A) . \<exists>f. f \<in> bij (Y, n) }. f``X" in exI)
@@ -196,23 +196,23 @@
apply (fast elim!: bij_is_surj [THEN image_vimage_eq])
done
-lemma WO2_imp_ex_Card: "WO2 ==> \<exists>a. Card(a) & X\<approx>a"
+lemma WO2_imp_ex_Card: "WO2 \<Longrightarrow> \<exists>a. Card(a) & X\<approx>a"
apply (unfold WO2_def)
apply (drule spec [of _ X])
apply (blast intro: Card_cardinal eqpoll_trans
well_ord_Memrel [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym])
done
-lemma lepoll_infinite: "[| X\<lesssim>Y; ~Finite(X) |] ==> ~Finite(Y)"
+lemma lepoll_infinite: "\<lbrakk>X\<lesssim>Y; \<not>Finite(X)\<rbrakk> \<Longrightarrow> \<not>Finite(Y)"
by (blast intro: lepoll_Finite)
-lemma infinite_Card_is_InfCard: "[| ~Finite(X); Card(X) |] ==> InfCard(X)"
+lemma infinite_Card_is_InfCard: "\<lbrakk>\<not>Finite(X); Card(X)\<rbrakk> \<Longrightarrow> InfCard(X)"
apply (unfold InfCard_def)
apply (fast elim!: Card_is_Ord [THEN nat_le_infinite_Ord])
done
-lemma WO2_infinite_subsets_eqpoll_X: "[| WO2; n \<in> nat; ~Finite(X) |]
- ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
+lemma WO2_infinite_subsets_eqpoll_X: "\<lbrakk>WO2; n \<in> nat; \<not>Finite(X)\<rbrakk>
+ \<Longrightarrow> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
apply (drule WO2_imp_ex_Card)
apply (elim allE exE conjE)
apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)
@@ -220,12 +220,12 @@
apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans)
done
-lemma well_ord_imp_ex_Card: "well_ord(X,R) ==> \<exists>a. Card(a) & X\<approx>a"
+lemma well_ord_imp_ex_Card: "well_ord(X,R) \<Longrightarrow> \<exists>a. Card(a) & X\<approx>a"
by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym]
intro!: Card_cardinal)
lemma well_ord_infinite_subsets_eqpoll_X:
- "[| well_ord(X,R); n \<in> nat; ~Finite(X) |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
+ "\<lbrakk>well_ord(X,R); n \<in> nat; \<not>Finite(X)\<rbrakk> \<Longrightarrow> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
apply (drule well_ord_imp_ex_Card)
apply (elim allE exE conjE)
apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)
--- a/src/ZF/AC/AC17_AC1.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC17_AC1.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,26 +15,26 @@
(** AC0 is equivalent to AC1.
AC0 comes from Suppes, AC1 from Rubin & Rubin **)
-lemma AC0_AC1_lemma: "[| f:(\<Prod>X \<in> A. X); D \<subseteq> A |] ==> \<exists>g. g:(\<Prod>X \<in> D. X)"
+lemma AC0_AC1_lemma: "\<lbrakk>f:(\<Prod>X \<in> A. X); D \<subseteq> A\<rbrakk> \<Longrightarrow> \<exists>g. g:(\<Prod>X \<in> D. X)"
by (fast intro!: lam_type apply_type)
-lemma AC0_AC1: "AC0 ==> AC1"
+lemma AC0_AC1: "AC0 \<Longrightarrow> AC1"
apply (unfold AC0_def AC1_def)
apply (blast intro: AC0_AC1_lemma)
done
-lemma AC1_AC0: "AC1 ==> AC0"
+lemma AC1_AC0: "AC1 \<Longrightarrow> AC0"
by (unfold AC0_def AC1_def, blast)
-(**** The proof of AC1 ==> AC17 ****)
+(**** The proof of AC1 \<Longrightarrow> AC17 ****)
-lemma AC1_AC17_lemma: "f \<in> (\<Prod>X \<in> Pow(A) - {0}. X) ==> f \<in> (Pow(A) - {0} -> A)"
+lemma AC1_AC17_lemma: "f \<in> (\<Prod>X \<in> Pow(A) - {0}. X) \<Longrightarrow> f \<in> (Pow(A) - {0} -> A)"
apply (rule Pi_type, assumption)
apply (drule apply_type, assumption, fast)
done
-lemma AC1_AC17: "AC1 ==> AC17"
+lemma AC1_AC17: "AC1 \<Longrightarrow> AC17"
apply (unfold AC1_def AC17_def)
apply (rule allI)
apply (rule ballI)
@@ -48,17 +48,17 @@
done
-(**** The proof of AC17 ==> AC1 ****)
+(**** The proof of AC17 \<Longrightarrow> AC1 ****)
(* *********************************************************************** *)
(* more properties of HH *)
(* *********************************************************************** *)
lemma UN_eq_imp_well_ord:
- "[| x - (\<Union>j \<in> \<mu> i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x}.
+ "\<lbrakk>x - (\<Union>j \<in> \<mu> i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x}.
HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, j)) = 0;
- f \<in> Pow(x)-{0} -> x |]
- ==> \<exists>r. well_ord(x,r)"
+ f \<in> Pow(x)-{0} -> x\<rbrakk>
+ \<Longrightarrow> \<exists>r. well_ord(x,r)"
apply (rule exI)
apply (erule well_ord_rvimage
[OF bij_Least_HH_x [THEN bij_converse_bij, THEN bij_is_inj]
@@ -70,7 +70,7 @@
(* *********************************************************************** *)
lemma not_AC1_imp_ex:
- "~AC1 ==> \<exists>A. \<forall>f \<in> Pow(A)-{0} -> A. \<exists>u \<in> Pow(A)-{0}. f`u \<notin> u"
+ "\<not>AC1 \<Longrightarrow> \<exists>A. \<forall>f \<in> Pow(A)-{0} -> A. \<exists>u \<in> Pow(A)-{0}. f`u \<notin> u"
apply (unfold AC1_def)
apply (erule swap)
apply (rule allI)
@@ -80,11 +80,11 @@
done
lemma AC17_AC1_aux1:
- "[| \<forall>f \<in> Pow(x) - {0} -> x. \<exists>u \<in> Pow(x) - {0}. f`u\<notin>u;
+ "\<lbrakk>\<forall>f \<in> Pow(x) - {0} -> x. \<exists>u \<in> Pow(x) - {0}. f`u\<notin>u;
\<exists>f \<in> Pow(x)-{0}->x.
x - (\<Union>a \<in> (\<mu> i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,i)={x}).
- HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,a)) = 0 |]
- ==> P"
+ HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,a)) = 0\<rbrakk>
+ \<Longrightarrow> P"
apply (erule bexE)
apply (erule UN_eq_imp_well_ord [THEN exE], assumption)
apply (erule ex_choice_fun_Pow [THEN exE])
@@ -96,22 +96,22 @@
done
lemma AC17_AC1_aux2:
- "~ (\<exists>f \<in> Pow(x)-{0}->x. x - F(f) = 0)
- ==> (\<lambda>f \<in> Pow(x)-{0}->x . x - F(f))
+ "\<not> (\<exists>f \<in> Pow(x)-{0}->x. x - F(f) = 0)
+ \<Longrightarrow> (\<lambda>f \<in> Pow(x)-{0}->x . x - F(f))
\<in> (Pow(x) -{0} -> x) -> Pow(x) - {0}"
by (fast intro!: lam_type dest!: Diff_eq_0_iff [THEN iffD1])
lemma AC17_AC1_aux3:
- "[| f`Z \<in> Z; Z \<in> Pow(x)-{0} |]
- ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X})`Z \<in> Pow(Z)-{0}"
+ "\<lbrakk>f`Z \<in> Z; Z \<in> Pow(x)-{0}\<rbrakk>
+ \<Longrightarrow> (\<lambda>X \<in> Pow(x)-{0}. {f`X})`Z \<in> Pow(Z)-{0}"
by auto
lemma AC17_AC1_aux4:
"\<exists>f \<in> F. f`((\<lambda>f \<in> F. Q(f))`f) \<in> (\<lambda>f \<in> F. Q(f))`f
- ==> \<exists>f \<in> F. f`Q(f) \<in> Q(f)"
+ \<Longrightarrow> \<exists>f \<in> F. f`Q(f) \<in> Q(f)"
by simp
-lemma AC17_AC1: "AC17 ==> AC1"
+lemma AC17_AC1: "AC17 \<Longrightarrow> AC1"
apply (unfold AC17_def)
apply (rule classical)
apply (erule not_AC1_imp_ex [THEN exE])
@@ -133,25 +133,25 @@
(* **********************************************************************
- AC1 ==> AC2 ==> AC1
- AC1 ==> AC4 ==> AC3 ==> AC1
- AC4 ==> AC5 ==> AC4
+ AC1 \<Longrightarrow> AC2 \<Longrightarrow> AC1
+ AC1 \<Longrightarrow> AC4 \<Longrightarrow> AC3 \<Longrightarrow> AC1
+ AC4 \<Longrightarrow> AC5 \<Longrightarrow> AC4
AC1 \<longleftrightarrow> AC6
************************************************************************* *)
(* ********************************************************************** *)
-(* AC1 ==> AC2 *)
+(* AC1 \<Longrightarrow> AC2 *)
(* ********************************************************************** *)
lemma AC1_AC2_aux1:
- "[| f:(\<Prod>X \<in> A. X); B \<in> A; 0\<notin>A |] ==> {f`B} \<subseteq> B \<inter> {f`C. C \<in> A}"
+ "\<lbrakk>f:(\<Prod>X \<in> A. X); B \<in> A; 0\<notin>A\<rbrakk> \<Longrightarrow> {f`B} \<subseteq> B \<inter> {f`C. C \<in> A}"
by (fast elim!: apply_type)
lemma AC1_AC2_aux2:
- "[| pairwise_disjoint(A); B \<in> A; C \<in> A; D \<in> B; D \<in> C |] ==> f`B = f`C"
+ "\<lbrakk>pairwise_disjoint(A); B \<in> A; C \<in> A; D \<in> B; D \<in> C\<rbrakk> \<Longrightarrow> f`B = f`C"
by (unfold pairwise_disjoint_def, fast)
-lemma AC1_AC2: "AC1 ==> AC2"
+lemma AC1_AC2: "AC1 \<Longrightarrow> AC2"
apply (unfold AC1_def AC2_def)
apply (rule allI)
apply (rule impI)
@@ -163,14 +163,14 @@
(* ********************************************************************** *)
-(* AC2 ==> AC1 *)
+(* AC2 \<Longrightarrow> AC1 *)
(* ********************************************************************** *)
-lemma AC2_AC1_aux1: "0\<notin>A ==> 0 \<notin> {B*{B}. B \<in> A}"
+lemma AC2_AC1_aux1: "0\<notin>A \<Longrightarrow> 0 \<notin> {B*{B}. B \<in> A}"
by (fast dest!: sym [THEN Sigma_empty_iff [THEN iffD1]])
-lemma AC2_AC1_aux2: "[| X*{X} \<inter> C = {y}; X \<in> A |]
- ==> (THE y. X*{X} \<inter> C = {y}): X*A"
+lemma AC2_AC1_aux2: "\<lbrakk>X*{X} \<inter> C = {y}; X \<in> A\<rbrakk>
+ \<Longrightarrow> (THE y. X*{X} \<inter> C = {y}): X*A"
apply (rule subst_elem [of y])
apply (blast elim!: equalityE)
apply (auto simp add: singleton_eq_iff)
@@ -178,13 +178,13 @@
lemma AC2_AC1_aux3:
"\<forall>D \<in> {E*{E}. E \<in> A}. \<exists>y. D \<inter> C = {y}
- ==> (\<lambda>x \<in> A. fst(THE z. (x*{x} \<inter> C = {z}))) \<in> (\<Prod>X \<in> A. X)"
+ \<Longrightarrow> (\<lambda>x \<in> A. fst(THE z. (x*{x} \<inter> C = {z}))) \<in> (\<Prod>X \<in> A. X)"
apply (rule lam_type)
apply (drule bspec, blast)
apply (blast intro: AC2_AC1_aux2 fst_type)
done
-lemma AC2_AC1: "AC2 ==> AC1"
+lemma AC2_AC1: "AC2 \<Longrightarrow> AC1"
apply (unfold AC1_def AC2_def pairwise_disjoint_def)
apply (intro allI impI)
apply (elim allE impE)
@@ -194,13 +194,13 @@
(* ********************************************************************** *)
-(* AC1 ==> AC4 *)
+(* AC1 \<Longrightarrow> AC4 *)
(* ********************************************************************** *)
lemma empty_notin_images: "0 \<notin> {R``{x}. x \<in> domain(R)}"
by blast
-lemma AC1_AC4: "AC1 ==> AC4"
+lemma AC1_AC4: "AC1 \<Longrightarrow> AC4"
apply (unfold AC1_def AC4_def)
apply (intro allI impI)
apply (drule spec, drule mp [OF _ empty_notin_images])
@@ -209,19 +209,19 @@
(* ********************************************************************** *)
-(* AC4 ==> AC3 *)
+(* AC4 \<Longrightarrow> AC3 *)
(* ********************************************************************** *)
-lemma AC4_AC3_aux1: "f \<in> A->B ==> (\<Union>z \<in> A. {z}*f`z) \<subseteq> A*\<Union>(B)"
+lemma AC4_AC3_aux1: "f \<in> A->B \<Longrightarrow> (\<Union>z \<in> A. {z}*f`z) \<subseteq> A*\<Union>(B)"
by (fast dest!: apply_type)
lemma AC4_AC3_aux2: "domain(\<Union>z \<in> A. {z}*f(z)) = {a \<in> A. f(a)\<noteq>0}"
by blast
-lemma AC4_AC3_aux3: "x \<in> A ==> (\<Union>z \<in> A. {z}*f(z))``{x} = f(x)"
+lemma AC4_AC3_aux3: "x \<in> A \<Longrightarrow> (\<Union>z \<in> A. {z}*f(z))``{x} = f(x)"
by fast
-lemma AC4_AC3: "AC4 ==> AC3"
+lemma AC4_AC3: "AC4 \<Longrightarrow> AC3"
apply (unfold AC3_def AC4_def)
apply (intro allI ballI)
apply (elim allE impE)
@@ -230,25 +230,25 @@
done
(* ********************************************************************** *)
-(* AC3 ==> AC1 *)
+(* AC3 \<Longrightarrow> AC1 *)
(* ********************************************************************** *)
lemma AC3_AC1_lemma:
- "b\<notin>A ==> (\<Prod>x \<in> {a \<in> A. id(A)`a\<noteq>b}. id(A)`x) = (\<Prod>x \<in> A. x)"
+ "b\<notin>A \<Longrightarrow> (\<Prod>x \<in> {a \<in> A. id(A)`a\<noteq>b}. id(A)`x) = (\<Prod>x \<in> A. x)"
apply (simp add: id_def cong add: Pi_cong)
apply (rule_tac b = A in subst_context, fast)
done
-lemma AC3_AC1: "AC3 ==> AC1"
+lemma AC3_AC1: "AC3 \<Longrightarrow> AC1"
apply (unfold AC1_def AC3_def)
apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst])
done
(* ********************************************************************** *)
-(* AC4 ==> AC5 *)
+(* AC4 \<Longrightarrow> AC5 *)
(* ********************************************************************** *)
-lemma AC4_AC5: "AC4 ==> AC5"
+lemma AC4_AC5: "AC4 \<Longrightarrow> AC5"
apply (unfold range_def AC4_def AC5_def)
apply (intro allI ballI)
apply (elim allE impE)
@@ -262,27 +262,27 @@
(* ********************************************************************** *)
-(* AC5 ==> AC4, Rubin & Rubin, p. 11 *)
+(* AC5 \<Longrightarrow> AC4, Rubin & Rubin, p. 11 *)
(* ********************************************************************** *)
-lemma AC5_AC4_aux1: "R \<subseteq> A*B ==> (\<lambda>x \<in> R. fst(x)) \<in> R -> A"
+lemma AC5_AC4_aux1: "R \<subseteq> A*B \<Longrightarrow> (\<lambda>x \<in> R. fst(x)) \<in> R -> A"
by (fast intro!: lam_type fst_type)
-lemma AC5_AC4_aux2: "R \<subseteq> A*B ==> range(\<lambda>x \<in> R. fst(x)) = domain(R)"
+lemma AC5_AC4_aux2: "R \<subseteq> A*B \<Longrightarrow> range(\<lambda>x \<in> R. fst(x)) = domain(R)"
by (unfold lam_def, force)
-lemma AC5_AC4_aux3: "[| \<exists>f \<in> A->C. P(f,domain(f)); A=B |] ==> \<exists>f \<in> B->C. P(f,B)"
+lemma AC5_AC4_aux3: "\<lbrakk>\<exists>f \<in> A->C. P(f,domain(f)); A=B\<rbrakk> \<Longrightarrow> \<exists>f \<in> B->C. P(f,B)"
apply (erule bexE)
apply (frule domain_of_fun, fast)
done
-lemma AC5_AC4_aux4: "[| R \<subseteq> A*B; g \<in> C->R; \<forall>x \<in> C. (\<lambda>z \<in> R. fst(z))` (g`x) = x |]
- ==> (\<lambda>x \<in> C. snd(g`x)): (\<Prod>x \<in> C. R``{x})"
+lemma AC5_AC4_aux4: "\<lbrakk>R \<subseteq> A*B; g \<in> C->R; \<forall>x \<in> C. (\<lambda>z \<in> R. fst(z))` (g`x) = x\<rbrakk>
+ \<Longrightarrow> (\<lambda>x \<in> C. snd(g`x)): (\<Prod>x \<in> C. R``{x})"
apply (rule lam_type)
apply (force dest: apply_type)
done
-lemma AC5_AC4: "AC5 ==> AC4"
+lemma AC5_AC4: "AC5 \<Longrightarrow> AC4"
apply (unfold AC4_def AC5_def, clarify)
apply (elim allE ballE)
apply (drule AC5_AC4_aux3 [OF _ AC5_AC4_aux2], assumption)
--- a/src/ZF/AC/AC18_AC19.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC18_AC19.thy Tue Sep 27 16:51:35 2022 +0100
@@ -1,7 +1,7 @@
(* Title: ZF/AC/AC18_AC19.thy
Author: Krzysztof Grabczewski
-The proof of AC1 ==> AC18 ==> AC19 ==> AC1
+The proof of AC1 \<Longrightarrow> AC18 \<Longrightarrow> AC19 \<Longrightarrow> AC1
*)
theory AC18_AC19
@@ -10,21 +10,21 @@
definition
uu :: "i => i" where
- "uu(a) == {c \<union> {0}. c \<in> a}"
+ "uu(a) \<equiv> {c \<union> {0}. c \<in> a}"
(* ********************************************************************** *)
-(* AC1 ==> AC18 *)
+(* AC1 \<Longrightarrow> AC18 *)
(* ********************************************************************** *)
lemma PROD_subsets:
- "[| f \<in> (\<Prod>b \<in> {P(a). a \<in> A}. b); \<forall>a \<in> A. P(a)<=Q(a) |]
- ==> (\<lambda>a \<in> A. f`P(a)) \<in> (\<Prod>a \<in> A. Q(a))"
+ "\<lbrakk>f \<in> (\<Prod>b \<in> {P(a). a \<in> A}. b); \<forall>a \<in> A. P(a)<=Q(a)\<rbrakk>
+ \<Longrightarrow> (\<lambda>a \<in> A. f`P(a)) \<in> (\<Prod>a \<in> A. Q(a))"
by (rule lam_type, drule apply_type, auto)
lemma lemma_AC18:
- "[| \<forall>A. 0 \<notin> A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X)); A \<noteq> 0 |]
- ==> (\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a, b)) \<subseteq>
+ "\<lbrakk>\<forall>A. 0 \<notin> A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X)); A \<noteq> 0\<rbrakk>
+ \<Longrightarrow> (\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a, b)) \<subseteq>
(\<Union>f \<in> \<Prod>a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a))"
apply (rule subsetI)
apply (erule_tac x = "{{b \<in> B (a) . x \<in> X (a,b) }. a \<in> A}" in allE)
@@ -35,14 +35,14 @@
apply (simp, fast elim!: not_emptyE dest: apply_type [OF _ RepFunI])
done
-lemma AC1_AC18: "AC1 ==> PROP AC18"
+lemma AC1_AC18: "AC1 \<Longrightarrow> PROP AC18"
apply (unfold AC1_def)
apply (rule AC18.intro)
apply (fast elim!: lemma_AC18 apply_type intro!: equalityI INT_I UN_I)
done
(* ********************************************************************** *)
-(* AC18 ==> AC19 *)
+(* AC18 \<Longrightarrow> AC19 *)
(* ********************************************************************** *)
theorem (in AC18) AC19
@@ -52,28 +52,28 @@
done
(* ********************************************************************** *)
-(* AC19 ==> AC1 *)
+(* AC19 \<Longrightarrow> AC1 *)
(* ********************************************************************** *)
lemma RepRep_conj:
- "[| A \<noteq> 0; 0 \<notin> A |] ==> {uu(a). a \<in> A} \<noteq> 0 & 0 \<notin> {uu(a). a \<in> A}"
+ "\<lbrakk>A \<noteq> 0; 0 \<notin> A\<rbrakk> \<Longrightarrow> {uu(a). a \<in> A} \<noteq> 0 & 0 \<notin> {uu(a). a \<in> A}"
apply (unfold uu_def, auto)
apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]])
done
-lemma lemma1_1: "[|c \<in> a; x = c \<union> {0}; x \<notin> a |] ==> x - {0} \<in> a"
+lemma lemma1_1: "\<lbrakk>c \<in> a; x = c \<union> {0}; x \<notin> a\<rbrakk> \<Longrightarrow> x - {0} \<in> a"
apply clarify
apply (rule subst_elem, assumption)
apply (fast elim: notE subst_elem)
done
lemma lemma1_2:
- "[| f`(uu(a)) \<notin> a; f \<in> (\<Prod>B \<in> {uu(a). a \<in> A}. B); a \<in> A |]
- ==> f`(uu(a))-{0} \<in> a"
+ "\<lbrakk>f`(uu(a)) \<notin> a; f \<in> (\<Prod>B \<in> {uu(a). a \<in> A}. B); a \<in> A\<rbrakk>
+ \<Longrightarrow> f`(uu(a))-{0} \<in> a"
apply (unfold uu_def, fast elim!: lemma1_1 dest!: apply_type)
done
-lemma lemma1: "\<exists>f. f \<in> (\<Prod>B \<in> {uu(a). a \<in> A}. B) ==> \<exists>f. f \<in> (\<Prod>B \<in> A. B)"
+lemma lemma1: "\<exists>f. f \<in> (\<Prod>B \<in> {uu(a). a \<in> A}. B) \<Longrightarrow> \<exists>f. f \<in> (\<Prod>B \<in> A. B)"
apply (erule exE)
apply (rule_tac x = "\<lambda>a\<in>A. if (f` (uu(a)) \<in> a, f` (uu(a)), f` (uu(a))-{0})"
in exI)
@@ -81,16 +81,16 @@
apply (simp add: lemma1_2)
done
-lemma lemma2_1: "a\<noteq>0 ==> 0 \<in> (\<Union>b \<in> uu(a). b)"
+lemma lemma2_1: "a\<noteq>0 \<Longrightarrow> 0 \<in> (\<Union>b \<in> uu(a). b)"
by (unfold uu_def, auto)
-lemma lemma2: "[| A\<noteq>0; 0\<notin>A |] ==> (\<Inter>x \<in> {uu(a). a \<in> A}. \<Union>b \<in> x. b) \<noteq> 0"
+lemma lemma2: "\<lbrakk>A\<noteq>0; 0\<notin>A\<rbrakk> \<Longrightarrow> (\<Inter>x \<in> {uu(a). a \<in> A}. \<Union>b \<in> x. b) \<noteq> 0"
apply (erule not_emptyE)
apply (rule_tac a = 0 in not_emptyI)
apply (fast intro!: lemma2_1)
done
-lemma AC19_AC1: "AC19 ==> AC1"
+lemma AC19_AC1: "AC19 \<Longrightarrow> AC1"
apply (unfold AC19_def AC1_def, clarify)
apply (case_tac "A=0", force)
apply (erule_tac x = "{uu (a) . a \<in> A}" in allE)
--- a/src/ZF/AC/AC7_AC9.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC7_AC9.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,16 +10,16 @@
begin
(* ********************************************************************** *)
-(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *)
+(* Lemmas used in the proofs AC7 \<Longrightarrow> AC6 and AC9 \<Longrightarrow> AC1 *)
(* - Sigma_fun_space_not0 *)
(* - Sigma_fun_space_eqpoll *)
(* ********************************************************************** *)
-lemma Sigma_fun_space_not0: "[| 0\<notin>A; B \<in> A |] ==> (nat->\<Union>(A)) * B \<noteq> 0"
+lemma Sigma_fun_space_not0: "\<lbrakk>0\<notin>A; B \<in> A\<rbrakk> \<Longrightarrow> (nat->\<Union>(A)) * B \<noteq> 0"
by (blast dest!: Sigma_empty_iff [THEN iffD1] Union_empty_iff [THEN iffD1])
lemma inj_lemma:
- "C \<in> A ==> (\<lambda>g \<in> (nat->\<Union>(A))*C.
+ "C \<in> A \<Longrightarrow> (\<lambda>g \<in> (nat->\<Union>(A))*C.
(\<lambda>n \<in> nat. if(n=0, snd(g), fst(g)`(n #- 1))))
\<in> inj((nat->\<Union>(A))*C, (nat->\<Union>(A)) ) "
apply (unfold inj_def)
@@ -31,7 +31,7 @@
done
lemma Sigma_fun_space_eqpoll:
- "[| C \<in> A; 0\<notin>A |] ==> (nat->\<Union>(A)) * C \<approx> (nat->\<Union>(A))"
+ "\<lbrakk>C \<in> A; 0\<notin>A\<rbrakk> \<Longrightarrow> (nat->\<Union>(A)) * C \<approx> (nat->\<Union>(A))"
apply (rule eqpollI)
apply (simp add: lepoll_def)
apply (fast intro!: inj_lemma)
@@ -41,34 +41,34 @@
(* ********************************************************************** *)
-(* AC6 ==> AC7 *)
+(* AC6 \<Longrightarrow> AC7 *)
(* ********************************************************************** *)
-lemma AC6_AC7: "AC6 ==> AC7"
+lemma AC6_AC7: "AC6 \<Longrightarrow> AC7"
by (unfold AC6_def AC7_def, blast)
(* ********************************************************************** *)
-(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *)
+(* AC7 \<Longrightarrow> AC6, Rubin & Rubin p. 12, Theorem 2.8 *)
(* The case of the empty family of sets added in order to complete *)
(* the proof. *)
(* ********************************************************************** *)
-lemma lemma1_1: "y \<in> (\<Prod>B \<in> A. Y*B) ==> (\<lambda>B \<in> A. snd(y`B)) \<in> (\<Prod>B \<in> A. B)"
+lemma lemma1_1: "y \<in> (\<Prod>B \<in> A. Y*B) \<Longrightarrow> (\<lambda>B \<in> A. snd(y`B)) \<in> (\<Prod>B \<in> A. B)"
by (fast intro!: lam_type snd_type apply_type)
lemma lemma1_2:
- "y \<in> (\<Prod>B \<in> {Y*C. C \<in> A}. B) ==> (\<lambda>B \<in> A. y`(Y*B)) \<in> (\<Prod>B \<in> A. Y*B)"
+ "y \<in> (\<Prod>B \<in> {Y*C. C \<in> A}. B) \<Longrightarrow> (\<lambda>B \<in> A. y`(Y*B)) \<in> (\<Prod>B \<in> A. Y*B)"
apply (fast intro!: lam_type apply_type)
done
lemma AC7_AC6_lemma1:
- "(\<Prod>B \<in> {(nat->\<Union>(A))*C. C \<in> A}. B) \<noteq> 0 ==> (\<Prod>B \<in> A. B) \<noteq> 0"
+ "(\<Prod>B \<in> {(nat->\<Union>(A))*C. C \<in> A}. B) \<noteq> 0 \<Longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
by (fast intro!: equals0I lemma1_1 lemma1_2)
-lemma AC7_AC6_lemma2: "0 \<notin> A ==> 0 \<notin> {(nat -> \<Union>(A)) * C. C \<in> A}"
+lemma AC7_AC6_lemma2: "0 \<notin> A \<Longrightarrow> 0 \<notin> {(nat -> \<Union>(A)) * C. C \<in> A}"
by (blast dest: Sigma_fun_space_not0)
-lemma AC7_AC6: "AC7 ==> AC6"
+lemma AC7_AC6: "AC7 \<Longrightarrow> AC6"
apply (unfold AC6_def AC7_def)
apply (rule allI)
apply (rule impI)
@@ -82,38 +82,38 @@
(* ********************************************************************** *)
-(* AC1 ==> AC8 *)
+(* AC1 \<Longrightarrow> AC8 *)
(* ********************************************************************** *)
lemma AC1_AC8_lemma1:
"\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2
- ==> 0 \<notin> { bij(fst(B),snd(B)). B \<in> A }"
+ \<Longrightarrow> 0 \<notin> { bij(fst(B),snd(B)). B \<in> A }"
apply (unfold eqpoll_def, auto)
done
lemma AC1_AC8_lemma2:
- "[| f \<in> (\<Prod>X \<in> RepFun(A,p). X); D \<in> A |] ==> (\<lambda>x \<in> A. f`p(x))`D \<in> p(D)"
+ "\<lbrakk>f \<in> (\<Prod>X \<in> RepFun(A,p). X); D \<in> A\<rbrakk> \<Longrightarrow> (\<lambda>x \<in> A. f`p(x))`D \<in> p(D)"
apply (simp, fast elim!: apply_type)
done
-lemma AC1_AC8: "AC1 ==> AC8"
+lemma AC1_AC8: "AC1 \<Longrightarrow> AC8"
apply (unfold AC1_def AC8_def)
apply (fast dest: AC1_AC8_lemma1 AC1_AC8_lemma2)
done
(* ********************************************************************** *)
-(* AC8 ==> AC9 *)
+(* AC8 \<Longrightarrow> AC9 *)
(* - this proof replaces the following two from Rubin & Rubin: *)
-(* AC8 ==> AC1 and AC1 ==> AC9 *)
+(* AC8 \<Longrightarrow> AC1 and AC1 \<Longrightarrow> AC9 *)
(* ********************************************************************** *)
lemma AC8_AC9_lemma:
"\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1 \<approx> B2
- ==> \<forall>B \<in> A*A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2"
+ \<Longrightarrow> \<forall>B \<in> A*A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2"
by fast
-lemma AC8_AC9: "AC8 ==> AC9"
+lemma AC8_AC9: "AC8 \<Longrightarrow> AC9"
apply (unfold AC8_def AC9_def)
apply (intro allI impI)
apply (erule allE)
@@ -122,7 +122,7 @@
(* ********************************************************************** *)
-(* AC9 ==> AC1 *)
+(* AC9 \<Longrightarrow> AC1 *)
(* The idea of this proof comes from "Equivalents of the Axiom of Choice" *)
(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to *)
(* (x * y) \<union> {0} when y is a set of total functions acting from nat to *)
@@ -134,15 +134,15 @@
prod_commute_eqpoll)
lemma nat_lepoll_lemma:
- "[|0 \<notin> A; B \<in> A|] ==> nat \<lesssim> ((nat \<rightarrow> \<Union>(A)) \<times> B) \<times> nat"
+ "\<lbrakk>0 \<notin> A; B \<in> A\<rbrakk> \<Longrightarrow> nat \<lesssim> ((nat \<rightarrow> \<Union>(A)) \<times> B) \<times> nat"
by (blast dest: Sigma_fun_space_not0 intro: snd_lepoll_SigmaI)
lemma AC9_AC1_lemma1:
- "[| 0\<notin>A; A\<noteq>0;
+ "\<lbrakk>0\<notin>A; A\<noteq>0;
C = {((nat->\<Union>(A))*B)*nat. B \<in> A} \<union>
{cons(0,((nat->\<Union>(A))*B)*nat). B \<in> A};
- B1 \<in> C; B2 \<in> C |]
- ==> B1 \<approx> B2"
+ B1 \<in> C; B2 \<in> C\<rbrakk>
+ \<Longrightarrow> B1 \<approx> B2"
by (blast intro!: nat_lepoll_lemma Sigma_fun_space_eqpoll
nat_cons_eqpoll [THEN eqpoll_trans]
prod_eqpoll_cong [OF _ eqpoll_refl]
@@ -152,13 +152,13 @@
"\<forall>B1 \<in> {(F*B)*N. B \<in> A} \<union> {cons(0,(F*B)*N). B \<in> A}.
\<forall>B2 \<in> {(F*B)*N. B \<in> A} \<union> {cons(0,(F*B)*N). B \<in> A}.
f`<B1,B2> \<in> bij(B1, B2)
- ==> (\<lambda>B \<in> A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) \<in> (\<Prod>X \<in> A. X)"
+ \<Longrightarrow> (\<lambda>B \<in> A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) \<in> (\<Prod>X \<in> A. X)"
apply (intro lam_type snd_type fst_type)
apply (rule apply_type [OF _ consI1])
apply (fast intro!: fun_weaken_type bij_is_fun)
done
-lemma AC9_AC1: "AC9 ==> AC1"
+lemma AC9_AC1: "AC9 \<Longrightarrow> AC1"
apply (unfold AC1_def AC9_def)
apply (intro allI impI)
apply (erule allE)
--- a/src/ZF/AC/AC_Equiv.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/AC_Equiv.thy Tue Sep 27 16:51:35 2022 +0100
@@ -18,106 +18,106 @@
(* Well Ordering Theorems *)
definition
- "WO1 == \<forall>A. \<exists>R. well_ord(A,R)"
+ "WO1 \<equiv> \<forall>A. \<exists>R. well_ord(A,R)"
definition
- "WO2 == \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
+ "WO2 \<equiv> \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
definition
- "WO3 == \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
+ "WO3 \<equiv> \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
definition
- "WO4(m) == \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &
+ "WO4(m) \<equiv> \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &
(\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
definition
- "WO5 == \<exists>m \<in> nat. 1\<le>m & WO4(m)"
+ "WO5 \<equiv> \<exists>m \<in> nat. 1\<le>m & WO4(m)"
definition
- "WO6 == \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
+ "WO6 \<equiv> \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
& (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
definition
- "WO7 == \<forall>A. Finite(A) \<longleftrightarrow> (\<forall>R. well_ord(A,R) \<longrightarrow> well_ord(A,converse(R)))"
+ "WO7 \<equiv> \<forall>A. Finite(A) \<longleftrightarrow> (\<forall>R. well_ord(A,R) \<longrightarrow> well_ord(A,converse(R)))"
definition
- "WO8 == \<forall>A. (\<exists>f. f \<in> (\<Prod>X \<in> A. X)) \<longrightarrow> (\<exists>R. well_ord(A,R))"
+ "WO8 \<equiv> \<forall>A. (\<exists>f. f \<in> (\<Prod>X \<in> A. X)) \<longrightarrow> (\<exists>R. well_ord(A,R))"
definition
(* Auxiliary concepts needed below *)
pairwise_disjoint :: "i => o" where
- "pairwise_disjoint(A) == \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 \<inter> A2 \<noteq> 0 \<longrightarrow> A1=A2"
+ "pairwise_disjoint(A) \<equiv> \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 \<inter> A2 \<noteq> 0 \<longrightarrow> A1=A2"
definition
sets_of_size_between :: "[i, i, i] => o" where
- "sets_of_size_between(A,m,n) == \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
+ "sets_of_size_between(A,m,n) \<equiv> \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
(* Axioms of Choice *)
definition
- "AC0 == \<forall>A. \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
+ "AC0 \<equiv> \<forall>A. \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
definition
- "AC1 == \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X))"
+ "AC1 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. f \<in> (\<Prod>X \<in> A. X))"
definition
- "AC2 == \<forall>A. 0\<notin>A & pairwise_disjoint(A)
+ "AC2 \<equiv> \<forall>A. 0\<notin>A & pairwise_disjoint(A)
\<longrightarrow> (\<exists>C. \<forall>B \<in> A. \<exists>y. B \<inter> C = {y})"
definition
- "AC3 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Prod>x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
+ "AC3 \<equiv> \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Prod>x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
definition
- "AC4 == \<forall>R A B. (R \<subseteq> A*B \<longrightarrow> (\<exists>f. f \<in> (\<Prod>x \<in> domain(R). R``{x})))"
+ "AC4 \<equiv> \<forall>R A B. (R \<subseteq> A*B \<longrightarrow> (\<exists>f. f \<in> (\<Prod>x \<in> domain(R). R``{x})))"
definition
- "AC5 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g \<in> range(f)->A. \<forall>x \<in> domain(g). f`(g`x) = x"
+ "AC5 \<equiv> \<forall>A B. \<forall>f \<in> A->B. \<exists>g \<in> range(f)->A. \<forall>x \<in> domain(g). f`(g`x) = x"
definition
- "AC6 == \<forall>A. 0\<notin>A \<longrightarrow> (\<Prod>B \<in> A. B)\<noteq>0"
+ "AC6 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<Prod>B \<in> A. B)\<noteq>0"
definition
- "AC7 == \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
+ "AC7 \<equiv> \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
definition
- "AC8 == \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)
+ "AC8 \<equiv> \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)
\<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
definition
- "AC9 == \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow>
+ "AC9 \<equiv> \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow>
(\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
definition
- "AC10(n) == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) \<longrightarrow>
+ "AC10(n) \<equiv> \<forall>A. (\<forall>B \<in> A. \<not>Finite(B)) \<longrightarrow>
(\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &
sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B))"
definition
- "AC11 == \<exists>n \<in> nat. 1\<le>n & AC10(n)"
+ "AC11 \<equiv> \<exists>n \<in> nat. 1\<le>n & AC10(n)"
definition
- "AC12 == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) \<longrightarrow>
+ "AC12 \<equiv> \<forall>A. (\<forall>B \<in> A. \<not>Finite(B)) \<longrightarrow>
(\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &
sets_of_size_between(f`B, 2, succ(n)) & \<Union>(f`B)=B)))"
definition
- "AC13(m) == \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
+ "AC13(m) \<equiv> \<forall>A. 0\<notin>A \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
definition
- "AC14 == \<exists>m \<in> nat. 1\<le>m & AC13(m)"
+ "AC14 \<equiv> \<exists>m \<in> nat. 1\<le>m & AC13(m)"
definition
- "AC15 == \<forall>A. 0\<notin>A \<longrightarrow>
+ "AC15 \<equiv> \<forall>A. 0\<notin>A \<longrightarrow>
(\<exists>m \<in> nat. 1\<le>m & (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m))"
definition
- "AC16(n, k) ==
- \<forall>A. ~Finite(A) \<longrightarrow>
+ "AC16(n, k) \<equiv>
+ \<forall>A. \<not>Finite(A) \<longrightarrow>
(\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &
(\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
definition
- "AC17 == \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.
+ "AC17 \<equiv> \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.
\<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
locale AC18 =
@@ -127,7 +127,7 @@
\<comment> \<open>AC18 cannot be expressed within the object-logic\<close>
definition
- "AC19 == \<forall>A. A\<noteq>0 & 0\<notin>A \<longrightarrow> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =
+ "AC19 \<equiv> \<forall>A. A\<noteq>0 & 0\<notin>A \<longrightarrow> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =
(\<Union>f \<in> (\<Prod>B \<in> A. B). \<Inter>a \<in> A. f`a))"
@@ -148,7 +148,7 @@
(* used only in Hartog.ML *)
lemma ordertype_Int:
- "well_ord(A,r) ==> ordertype(A, r \<inter> A*A) = ordertype(A,r)"
+ "well_ord(A,r) \<Longrightarrow> ordertype(A, r \<inter> A*A) = ordertype(A,r)"
apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst])
apply (erule id_bij [THEN bij_ordertype_vimage])
done
@@ -159,29 +159,29 @@
done
lemma inj_strengthen_type:
- "[| f \<in> inj(A, B); !!a. a \<in> A ==> f`a \<in> C |] ==> f \<in> inj(A,C)"
+ "\<lbrakk>f \<in> inj(A, B); \<And>a. a \<in> A \<Longrightarrow> f`a \<in> C\<rbrakk> \<Longrightarrow> f \<in> inj(A,C)"
by (unfold inj_def, blast intro: Pi_type)
(* ********************************************************************** *)
(* Another elimination rule for \<exists>! *)
(* ********************************************************************** *)
-lemma ex1_two_eq: "[| \<exists>! x. P(x); P(x); P(y) |] ==> x=y"
+lemma ex1_two_eq: "\<lbrakk>\<exists>! x. P(x); P(x); P(y)\<rbrakk> \<Longrightarrow> x=y"
by blast
(* ********************************************************************** *)
-(* Lemmas used in the proofs like WO? ==> AC? *)
+(* Lemmas used in the proofs like WO? \<Longrightarrow> AC? *)
(* ********************************************************************** *)
lemma first_in_B:
- "[| well_ord(\<Union>(A),r); 0 \<notin> A; B \<in> A |] ==> (THE b. first(b,B,r)) \<in> B"
+ "\<lbrakk>well_ord(\<Union>(A),r); 0 \<notin> A; B \<in> A\<rbrakk> \<Longrightarrow> (THE b. first(b,B,r)) \<in> B"
by (blast dest!: well_ord_imp_ex1_first
[THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]])
-lemma ex_choice_fun: "[| well_ord(\<Union>(A), R); 0 \<notin> A |] ==> \<exists>f. f \<in> (\<Prod>X \<in> A. X)"
+lemma ex_choice_fun: "\<lbrakk>well_ord(\<Union>(A), R); 0 \<notin> A\<rbrakk> \<Longrightarrow> \<exists>f. f \<in> (\<Prod>X \<in> A. X)"
by (fast elim!: first_in_B intro!: lam_type)
-lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
+lemma ex_choice_fun_Pow: "well_ord(A, R) \<Longrightarrow> \<exists>f. f \<in> (\<Prod>X \<in> Pow(A)-{0}. X)"
by (fast elim!: well_ord_subset [THEN ex_choice_fun])
@@ -193,7 +193,7 @@
(*Using AC we could trivially prove, for all u, domain(u) \<lesssim> u*)
lemma lepoll_m_imp_domain_lepoll_m:
- "[| m \<in> nat; u \<lesssim> m |] ==> domain(u) \<lesssim> m"
+ "\<lbrakk>m \<in> nat; u \<lesssim> m\<rbrakk> \<Longrightarrow> domain(u) \<lesssim> m"
apply (unfold lepoll_def)
apply (erule exE)
apply (rule_tac x = "\<lambda>x \<in> domain(u). \<mu> i. \<exists>y. <x,y> \<in> u & f`<x,y> = i"
@@ -208,7 +208,7 @@
done
lemma rel_domain_ex1:
- "[| succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat |] ==> function(r)"
+ "\<lbrakk>succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat\<rbrakk> \<Longrightarrow> function(r)"
apply (unfold function_def, safe)
apply (rule ccontr)
apply (fast elim!: lepoll_trans [THEN succ_lepoll_natE]
@@ -217,8 +217,8 @@
done
lemma rel_is_fun:
- "[| succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat;
- r \<subseteq> A*B; A=domain(r) |] ==> r \<in> A->B"
+ "\<lbrakk>succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat;
+ r \<subseteq> A*B; A=domain(r)\<rbrakk> \<Longrightarrow> r \<in> A->B"
by (simp add: Pi_iff rel_domain_ex1)
end
--- a/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 16:51:35 2022 +0100
@@ -6,7 +6,7 @@
theory Cardinal_aux imports AC_Equiv begin
-lemma Diff_lepoll: "[| A \<lesssim> succ(m); B \<subseteq> A; B\<noteq>0 |] ==> A-B \<lesssim> m"
+lemma Diff_lepoll: "\<lbrakk>A \<lesssim> succ(m); B \<subseteq> A; B\<noteq>0\<rbrakk> \<Longrightarrow> A-B \<lesssim> m"
apply (rule not_emptyE, assumption)
apply (blast intro: lepoll_trans [OF subset_imp_lepoll Diff_sing_lepoll])
done
@@ -20,14 +20,14 @@
(* j=|A| *)
lemma lepoll_imp_ex_le_eqpoll:
- "[| A \<lesssim> i; Ord(i) |] ==> \<exists>j. j \<le> i & A \<approx> j"
+ "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j \<le> i & A \<approx> j"
by (blast intro!: lepoll_cardinal_le well_ord_Memrel
well_ord_cardinal_eqpoll [THEN eqpoll_sym]
dest: lepoll_well_ord)
(* j=|A| *)
lemma lesspoll_imp_ex_lt_eqpoll:
- "[| A \<prec> i; Ord(i) |] ==> \<exists>j. j<i & A \<approx> j"
+ "\<lbrakk>A \<prec> i; Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. j<i & A \<approx> j"
by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE)
lemma Un_eqpoll_Inf_Ord:
@@ -70,7 +70,7 @@
(*Finally we reach this result. Surely there's a simpler proof?*)
lemma Un_lepoll_Inf_Ord:
- "[| A \<lesssim> i; B \<lesssim> i; ~Finite(i); Ord(i) |] ==> A \<union> B \<lesssim> i"
+ "\<lbrakk>A \<lesssim> i; B \<lesssim> i; \<not>Finite(i); Ord(i)\<rbrakk> \<Longrightarrow> A \<union> B \<lesssim> i"
apply (rule_tac A1 = i and C1 = i in ex_eqpoll_disjoint [THEN exE])
apply (erule conjE)
apply (drule lepoll_trans)
@@ -79,7 +79,7 @@
apply (blast intro: eqpoll_refl Un_eqpoll_Inf_Ord eqpoll_imp_lepoll)
done
-lemma Least_in_Ord: "[| P(i); i \<in> j; Ord(j) |] ==> (\<mu> i. P(i)) \<in> j"
+lemma Least_in_Ord: "\<lbrakk>P(i); i \<in> j; Ord(j)\<rbrakk> \<Longrightarrow> (\<mu> i. P(i)) \<in> j"
apply (erule Least_le [THEN leE])
apply (erule Ord_in_Ord, assumption)
apply (erule ltE)
@@ -88,8 +88,8 @@
done
lemma Diff_first_lepoll:
- "[| well_ord(x,r); y \<subseteq> x; y \<lesssim> succ(n); n \<in> nat |]
- ==> y - {THE b. first(b,y,r)} \<lesssim> n"
+ "\<lbrakk>well_ord(x,r); y \<subseteq> x; y \<lesssim> succ(n); n \<in> nat\<rbrakk>
+ \<Longrightarrow> y - {THE b. first(b,y,r)} \<lesssim> n"
apply (case_tac "y=0", simp add: empty_lepollI)
apply (fast intro!: Diff_sing_lepoll the_first_in)
done
@@ -98,7 +98,7 @@
"(\<Union>x \<in> X. P(x)) \<subseteq> (\<Union>x \<in> X. P(x)-Q(x)) \<union> (\<Union>x \<in> X. Q(x))"
by blast
-lemma UN_sing_lepoll: "Ord(a) ==> (\<Union>x \<in> a. {P(x)}) \<lesssim> a"
+lemma UN_sing_lepoll: "Ord(a) \<Longrightarrow> (\<Union>x \<in> a. {P(x)}) \<lesssim> a"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>z \<in> (\<Union>x \<in> a. {P (x) }) . (\<mu> i. P (i) =z) " in exI)
apply (rule_tac d = "%z. P (z) " in lam_injective)
@@ -107,8 +107,8 @@
done
lemma UN_fun_lepoll_lemma [rule_format]:
- "[| well_ord(T, R); ~Finite(a); Ord(a); n \<in> nat |]
- ==> \<forall>f. (\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T) \<longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
+ "\<lbrakk>well_ord(T, R); \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<forall>f. (\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T) \<longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
apply (induct_tac "n")
apply (rule allI)
apply (rule impI)
@@ -126,21 +126,21 @@
done
lemma UN_fun_lepoll:
- "[| \<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T; well_ord(T, R);
- ~Finite(a); Ord(a); n \<in> nat |] ==> (\<Union>b \<in> a. f`b) \<lesssim> a"
+ "\<lbrakk>\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T; well_ord(T, R);
+ \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk> \<Longrightarrow> (\<Union>b \<in> a. f`b) \<lesssim> a"
by (blast intro: UN_fun_lepoll_lemma)
lemma UN_lepoll:
- "[| \<forall>b \<in> a. F(b) \<lesssim> n & F(b) \<subseteq> T; well_ord(T, R);
- ~Finite(a); Ord(a); n \<in> nat |]
- ==> (\<Union>b \<in> a. F(b)) \<lesssim> a"
+ "\<lbrakk>\<forall>b \<in> a. F(b) \<lesssim> n & F(b) \<subseteq> T; well_ord(T, R);
+ \<not>Finite(a); Ord(a); n \<in> nat\<rbrakk>
+ \<Longrightarrow> (\<Union>b \<in> a. F(b)) \<lesssim> a"
apply (rule rev_mp)
apply (rule_tac f="\<lambda>b \<in> a. F (b)" in UN_fun_lepoll)
apply auto
done
lemma UN_eq_UN_Diffs:
- "Ord(a) ==> (\<Union>b \<in> a. F(b)) = (\<Union>b \<in> a. F(b) - (\<Union>c \<in> b. F(c)))"
+ "Ord(a) \<Longrightarrow> (\<Union>b \<in> a. F(b)) = (\<Union>b \<in> a. F(b) - (\<Union>c \<in> b. F(c)))"
apply (rule equalityI)
prefer 2 apply fast
apply (rule subsetI)
@@ -153,7 +153,7 @@
done
lemma lepoll_imp_eqpoll_subset:
- "a \<lesssim> X ==> \<exists>Y. Y \<subseteq> X & a \<approx> Y"
+ "a \<lesssim> X \<Longrightarrow> \<exists>Y. Y \<subseteq> X & a \<approx> Y"
apply (unfold lepoll_def eqpoll_def, clarify)
apply (blast intro: restrict_bij
dest: inj_is_fun [THEN fun_is_rel, THEN image_subset])
@@ -164,7 +164,7 @@
(* ********************************************************************** *)
lemma Diff_lesspoll_eqpoll_Card_lemma:
- "[| A\<approx>a; ~Finite(a); Card(a); B \<prec> a; A-B \<prec> a |] ==> P"
+ "\<lbrakk>A\<approx>a; \<not>Finite(a); Card(a); B \<prec> a; A-B \<prec> a\<rbrakk> \<Longrightarrow> P"
apply (elim lesspoll_imp_ex_lt_eqpoll [THEN exE] Card_is_Ord conjE)
apply (frule_tac j=xa in Un_upper1_le [OF lt_Ord lt_Ord], assumption)
apply (frule_tac j=xa in Un_upper2_le [OF lt_Ord lt_Ord], assumption)
@@ -186,7 +186,7 @@
done
lemma Diff_lesspoll_eqpoll_Card:
- "[| A \<approx> a; ~Finite(a); Card(a); B \<prec> a |] ==> A - B \<approx> a"
+ "\<lbrakk>A \<approx> a; \<not>Finite(a); Card(a); B \<prec> a\<rbrakk> \<Longrightarrow> A - B \<approx> a"
apply (rule ccontr)
apply (rule Diff_lesspoll_eqpoll_Card_lemma, (assumption+))
apply (blast intro: lesspoll_def [THEN def_imp_iff, THEN iffD2]
--- a/src/ZF/AC/DC.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/DC.thy Tue Sep 27 16:51:35 2022 +0100
@@ -8,7 +8,7 @@
imports AC_Equiv Hartog Cardinal_aux
begin
-lemma RepFun_lepoll: "Ord(a) ==> {P(b). b \<in> a} \<lesssim> a"
+lemma RepFun_lepoll: "Ord(a) \<Longrightarrow> {P(b). b \<in> a} \<lesssim> a"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>z \<in> RepFun (a,P) . \<mu> i. z=P (i) " in exI)
apply (rule_tac d="%z. P (z)" in lam_injective)
@@ -18,7 +18,7 @@
done
text\<open>Trivial in the presence of AC, but here we need a wellordering of X\<close>
-lemma image_Ord_lepoll: "[| f \<in> X->Y; Ord(X) |] ==> f``X \<lesssim> X"
+lemma image_Ord_lepoll: "\<lbrakk>f \<in> X->Y; Ord(X)\<rbrakk> \<Longrightarrow> f``X \<lesssim> X"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>x \<in> f``X. \<mu> y. f`y = x" in exI)
apply (rule_tac d = "%z. f`z" in lam_injective)
@@ -29,70 +29,70 @@
done
lemma range_subset_domain:
- "[| R \<subseteq> X*X; !!g. g \<in> X ==> \<exists>u. <g,u> \<in> R |]
- ==> range(R) \<subseteq> domain(R)"
+ "\<lbrakk>R \<subseteq> X*X; \<And>g. g \<in> X \<Longrightarrow> \<exists>u. <g,u> \<in> R\<rbrakk>
+ \<Longrightarrow> range(R) \<subseteq> domain(R)"
by blast
-lemma cons_fun_type: "g \<in> n->X ==> cons(<n,x>, g) \<in> succ(n) -> cons(x, X)"
+lemma cons_fun_type: "g \<in> n->X \<Longrightarrow> cons(<n,x>, g) \<in> succ(n) -> cons(x, X)"
apply (unfold succ_def)
apply (fast intro!: fun_extend elim!: mem_irrefl)
done
lemma cons_fun_type2:
- "[| g \<in> n->X; x \<in> X |] ==> cons(<n,x>, g) \<in> succ(n) -> X"
+ "\<lbrakk>g \<in> n->X; x \<in> X\<rbrakk> \<Longrightarrow> cons(<n,x>, g) \<in> succ(n) -> X"
by (erule cons_absorb [THEN subst], erule cons_fun_type)
-lemma cons_image_n: "n \<in> nat ==> cons(<n,x>, g)``n = g``n"
+lemma cons_image_n: "n \<in> nat \<Longrightarrow> cons(<n,x>, g)``n = g``n"
by (fast elim!: mem_irrefl)
-lemma cons_val_n: "g \<in> n->X ==> cons(<n,x>, g)`n = x"
+lemma cons_val_n: "g \<in> n->X \<Longrightarrow> cons(<n,x>, g)`n = x"
by (fast intro!: apply_equality elim!: cons_fun_type)
-lemma cons_image_k: "k \<in> n ==> cons(<n,x>, g)``k = g``k"
+lemma cons_image_k: "k \<in> n \<Longrightarrow> cons(<n,x>, g)``k = g``k"
by (fast elim: mem_asym)
-lemma cons_val_k: "[| k \<in> n; g \<in> n->X |] ==> cons(<n,x>, g)`k = g`k"
+lemma cons_val_k: "\<lbrakk>k \<in> n; g \<in> n->X\<rbrakk> \<Longrightarrow> cons(<n,x>, g)`k = g`k"
by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair)
-lemma domain_cons_eq_succ: "domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)"
+lemma domain_cons_eq_succ: "domain(f)=x \<Longrightarrow> domain(cons(<x,y>, f)) = succ(x)"
by (simp add: domain_cons succ_def)
-lemma restrict_cons_eq: "g \<in> n->X ==> restrict(cons(<n,x>, g), n) = g"
+lemma restrict_cons_eq: "g \<in> n->X \<Longrightarrow> restrict(cons(<n,x>, g), n) = g"
apply (simp add: restrict_def Pi_iff)
apply (blast intro: elim: mem_irrefl)
done
-lemma succ_in_succ: "[| Ord(k); i \<in> k |] ==> succ(i) \<in> succ(k)"
+lemma succ_in_succ: "\<lbrakk>Ord(k); i \<in> k\<rbrakk> \<Longrightarrow> succ(i) \<in> succ(k)"
apply (rule Ord_linear [of "succ(i)" "succ(k)", THEN disjE])
apply (fast elim: Ord_in_Ord mem_irrefl mem_asym)+
done
lemma restrict_eq_imp_val_eq:
- "[|restrict(f, domain(g)) = g; x \<in> domain(g)|]
- ==> f`x = g`x"
+ "\<lbrakk>restrict(f, domain(g)) = g; x \<in> domain(g)\<rbrakk>
+ \<Longrightarrow> f`x = g`x"
by (erule subst, simp add: restrict)
-lemma domain_eq_imp_fun_type: "[| domain(f)=A; f \<in> B->C |] ==> f \<in> A->C"
+lemma domain_eq_imp_fun_type: "\<lbrakk>domain(f)=A; f \<in> B->C\<rbrakk> \<Longrightarrow> f \<in> A->C"
by (frule domain_of_fun, fast)
-lemma ex_in_domain: "[| R \<subseteq> A * B; R \<noteq> 0 |] ==> \<exists>x. x \<in> domain(R)"
+lemma ex_in_domain: "\<lbrakk>R \<subseteq> A * B; R \<noteq> 0\<rbrakk> \<Longrightarrow> \<exists>x. x \<in> domain(R)"
by (fast elim!: not_emptyE)
definition
DC :: "i => o" where
- "DC(a) == \<forall>X R. R \<subseteq> Pow(X)*X &
+ "DC(a) \<equiv> \<forall>X R. R \<subseteq> Pow(X)*X &
(\<forall>Y \<in> Pow(X). Y \<prec> a \<longrightarrow> (\<exists>x \<in> X. <Y,x> \<in> R))
\<longrightarrow> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
definition
DC0 :: o where
- "DC0 == \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R)
+ "DC0 \<equiv> \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R)
\<longrightarrow> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)"
definition
ff :: "[i, i, i, i] => i" where
- "ff(b, X, Q, R) ==
+ "ff(b, X, Q, R) \<equiv>
transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
@@ -101,13 +101,13 @@
assumes all_ex: "\<forall>Y \<in> Pow(X). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R)"
- defines XX_def: "XX == (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})"
- and RR_def: "RR == {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))
+ defines XX_def: "XX \<equiv> (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})"
+ and RR_def: "RR \<equiv> {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))
& restrict(z2, domain(z1)) = z1}"
begin
(* ********************************************************************** *)
-(* DC ==> DC(omega) *)
+(* DC \<Longrightarrow> DC(omega) *)
(* *)
(* The scheme of the proof: *)
(* *)
@@ -169,8 +169,8 @@
done
lemma lemma2:
- "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; n \<in> nat |]
- ==> \<exists>k \<in> nat. f`succ(n) \<in> k -> X & n \<in> k
+ "\<lbrakk>\<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<exists>k \<in> nat. f`succ(n) \<in> k -> X & n \<in> k
& <f`succ(n)``n, f`succ(n)`n> \<in> R"
apply (induct_tac "n")
apply (drule apply_type [OF _ nat_1I])
@@ -198,8 +198,8 @@
done
lemma lemma3_1:
- "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat |]
- ==> {f`succ(x)`x. x \<in> m} = {f`succ(m)`x. x \<in> m}"
+ "\<lbrakk>\<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat\<rbrakk>
+ \<Longrightarrow> {f`succ(x)`x. x \<in> m} = {f`succ(m)`x. x \<in> m}"
apply (subgoal_tac "\<forall>x \<in> m. f`succ (m) `x = f`succ (x) `x")
apply simp
apply (induct_tac "m", blast)
@@ -221,8 +221,8 @@
done
lemma lemma3:
- "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat |]
- ==> (\<lambda>x \<in> nat. f`succ(x)`x) `` m = f`succ(m)``m"
+ "\<lbrakk>\<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat\<rbrakk>
+ \<Longrightarrow> (\<lambda>x \<in> nat. f`succ(x)`x) `` m = f`succ(m)``m"
apply (erule natE, simp)
apply (subst image_lam)
apply (fast elim!: OrdmemD [OF nat_succI Ord_nat])
@@ -234,7 +234,7 @@
end
-theorem DC0_imp_DC_nat: "DC0 ==> DC(nat)"
+theorem DC0_imp_DC_nat: "DC0 \<Longrightarrow> DC(nat)"
apply (unfold DC_def DC0_def, clarify)
apply (elim allE)
apply (erule impE)
@@ -256,7 +256,7 @@
(* ************************************************************************
- DC(omega) ==> DC
+ DC(omega) \<Longrightarrow> DC
The scheme of the proof:
@@ -269,7 +269,7 @@
RR = {<z1,z2>:Fin(XX)*XX.
(domain(z2)=succ(\<Union>f \<in> z1. domain(f)) &
(\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |
- (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) &
+ (\<not> (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) &
(\<forall>f \<in> z1. restrict(g, domain(f)) = f)) &
z2={<0,x>})}
@@ -287,7 +287,7 @@
************************************************************************* *)
lemma singleton_in_funs:
- "x \<in> X ==> {<0,x>} \<in>
+ "x \<in> X \<Longrightarrow> {<0,x>} \<in>
(\<Union>n \<in> nat. {f \<in> succ(n)->X. \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
apply (rule nat_0I [THEN UN_I])
apply (force simp add: singleton_0 [symmetric]
@@ -297,16 +297,16 @@
locale imp_DC0 =
fixes XX and RR and x and R and f and allRR
- defines XX_def: "XX == (\<Union>n \<in> nat.
+ defines XX_def: "XX \<equiv> (\<Union>n \<in> nat.
{f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
and RR_def:
- "RR == {<z1,z2>:Fin(XX)*XX.
+ "RR \<equiv> {<z1,z2>:Fin(XX)*XX.
(domain(z2)=succ(\<Union>f \<in> z1. domain(f))
& (\<forall>f \<in> z1. restrict(z2, domain(f)) = f))
- | (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))
+ | (\<not> (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))
& (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
and allRR_def:
- "allRR == \<forall>b<nat.
+ "allRR \<equiv> \<forall>b<nat.
<f``b, f`b> \<in>
{<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
& (\<Union>f \<in> z1. domain(f)) = b
@@ -314,8 +314,8 @@
begin
lemma lemma4:
- "[| range(R) \<subseteq> domain(R); x \<in> domain(R) |]
- ==> RR \<subseteq> Pow(XX)*XX &
+ "\<lbrakk>range(R) \<subseteq> domain(R); x \<in> domain(R)\<rbrakk>
+ \<Longrightarrow> RR \<subseteq> Pow(XX)*XX &
(\<forall>Y \<in> Pow(XX). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> XX. <Y,x>:RR))"
apply (rule conjI)
apply (force dest!: FinD [THEN PowI] simp add: RR_def)
@@ -332,34 +332,34 @@
done
lemma UN_image_succ_eq:
- "[| f \<in> nat->X; n \<in> nat |]
- ==> (\<Union>x \<in> f``succ(n). P(x)) = P(f`n) \<union> (\<Union>x \<in> f``n. P(x))"
+ "\<lbrakk>f \<in> nat->X; n \<in> nat\<rbrakk>
+ \<Longrightarrow> (\<Union>x \<in> f``succ(n). P(x)) = P(f`n) \<union> (\<Union>x \<in> f``n. P(x))"
by (simp add: image_fun OrdmemD)
lemma UN_image_succ_eq_succ:
- "[| (\<Union>x \<in> f``n. P(x)) = y; P(f`n) = succ(y);
- f \<in> nat -> X; n \<in> nat |] ==> (\<Union>x \<in> f``succ(n). P(x)) = succ(y)"
+ "\<lbrakk>(\<Union>x \<in> f``n. P(x)) = y; P(f`n) = succ(y);
+ f \<in> nat -> X; n \<in> nat\<rbrakk> \<Longrightarrow> (\<Union>x \<in> f``succ(n). P(x)) = succ(y)"
by (simp add: UN_image_succ_eq, blast)
lemma apply_domain_type:
- "[| h \<in> succ(n) -> D; n \<in> nat; domain(h)=succ(y) |] ==> h`y \<in> D"
+ "\<lbrakk>h \<in> succ(n) -> D; n \<in> nat; domain(h)=succ(y)\<rbrakk> \<Longrightarrow> h`y \<in> D"
by (fast elim: apply_type dest!: trans [OF sym domain_of_fun])
lemma image_fun_succ:
- "[| h \<in> nat -> X; n \<in> nat |] ==> h``succ(n) = cons(h`n, h``n)"
+ "\<lbrakk>h \<in> nat -> X; n \<in> nat\<rbrakk> \<Longrightarrow> h``succ(n) = cons(h`n, h``n)"
by (simp add: image_fun OrdmemD)
lemma f_n_type:
- "[| domain(f`n) = succ(k); f \<in> nat -> XX; n \<in> nat |]
- ==> f`n \<in> succ(k) -> domain(R)"
+ "\<lbrakk>domain(f`n) = succ(k); f \<in> nat -> XX; n \<in> nat\<rbrakk>
+ \<Longrightarrow> f`n \<in> succ(k) -> domain(R)"
apply (unfold XX_def)
apply (drule apply_type, assumption)
apply (fast elim: domain_eq_imp_fun_type)
done
lemma f_n_pairs_in_R [rule_format]:
- "[| h \<in> nat -> XX; domain(h`n) = succ(k); n \<in> nat |]
- ==> \<forall>i \<in> k. <h`n`i, h`n`succ(i)> \<in> R"
+ "\<lbrakk>h \<in> nat -> XX; domain(h`n) = succ(k); n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<forall>i \<in> k. <h`n`i, h`n`succ(i)> \<in> R"
apply (unfold XX_def)
apply (drule apply_type, assumption)
apply (elim UN_E CollectE)
@@ -367,8 +367,8 @@
done
lemma restrict_cons_eq_restrict:
- "[| restrict(h, domain(u))=u; h \<in> n->X; domain(u) \<subseteq> n |]
- ==> restrict(cons(<n, y>, h), domain(u)) = u"
+ "\<lbrakk>restrict(h, domain(u))=u; h \<in> n->X; domain(u) \<subseteq> n\<rbrakk>
+ \<Longrightarrow> restrict(cons(<n, y>, h), domain(u)) = u"
apply (unfold restrict_def)
apply (simp add: restrict_def Pi_iff)
apply (erule sym [THEN trans, symmetric])
@@ -376,11 +376,11 @@
done
lemma all_in_image_restrict_eq:
- "[| \<forall>x \<in> f``n. restrict(f`n, domain(x))=x;
+ "\<lbrakk>\<forall>x \<in> f``n. restrict(f`n, domain(x))=x;
f \<in> nat -> XX;
n \<in> nat; domain(f`n) = succ(n);
- (\<Union>x \<in> f``n. domain(x)) \<subseteq> n |]
- ==> \<forall>x \<in> f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x)) = x"
+ (\<Union>x \<in> f``n. domain(x)) \<subseteq> n\<rbrakk>
+ \<Longrightarrow> \<forall>x \<in> f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x)) = x"
apply (rule ballI)
apply (simp add: image_fun_succ)
apply (drule f_n_type, assumption+)
@@ -390,16 +390,16 @@
done
lemma simplify_recursion:
- "[| \<forall>b<nat. <f``b, f`b> \<in> RR;
- f \<in> nat -> XX; range(R) \<subseteq> domain(R); x \<in> domain(R)|]
- ==> allRR"
+ "\<lbrakk>\<forall>b<nat. <f``b, f`b> \<in> RR;
+ f \<in> nat -> XX; range(R) \<subseteq> domain(R); x \<in> domain(R)\<rbrakk>
+ \<Longrightarrow> allRR"
apply (unfold RR_def allRR_def)
apply (rule oallI, drule ltD)
apply (erule nat_induct)
apply (drule_tac x=0 in ospec, blast intro: Limit_has_0)
apply (force simp add: singleton_fun [THEN domain_of_fun] singleton_in_funs)
(*induction step*) (** LEVEL 5 **)
-(*prevent simplification of ~\<exists> to \<forall>~ *)
+(*prevent simplification of \<not>\<exists> to \<forall>\<not> *)
apply (simp only: separation split)
apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI)
apply (elim conjE disjE)
@@ -435,8 +435,8 @@
lemma lemma2:
- "[| allRR; f \<in> nat->XX; range(R) \<subseteq> domain(R); x \<in> domain(R); n \<in> nat |]
- ==> f`n \<in> succ(n) -> domain(R) & (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)"
+ "\<lbrakk>allRR; f \<in> nat->XX; range(R) \<subseteq> domain(R); x \<in> domain(R); n \<in> nat\<rbrakk>
+ \<Longrightarrow> f`n \<in> succ(n) -> domain(R) & (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)"
apply (unfold allRR_def)
apply (drule ospec)
apply (erule ltI [OF _ Ord_nat])
@@ -448,8 +448,8 @@
done
lemma lemma3:
- "[| allRR; f \<in> nat->XX; n\<in>nat; range(R) \<subseteq> domain(R); x \<in> domain(R) |]
- ==> f`n`n = f`succ(n)`n"
+ "\<lbrakk>allRR; f \<in> nat->XX; n\<in>nat; range(R) \<subseteq> domain(R); x \<in> domain(R)\<rbrakk>
+ \<Longrightarrow> f`n`n = f`succ(n)`n"
apply (frule lemma2 [THEN conjunct1, THEN domain_of_fun], assumption+)
apply (unfold allRR_def)
apply (drule ospec)
@@ -462,7 +462,7 @@
end
-theorem DC_nat_imp_DC0: "DC(nat) ==> DC0"
+theorem DC_nat_imp_DC0: "DC(nat) \<Longrightarrow> DC0"
apply (unfold DC_def DC0_def)
apply (intro allI impI)
apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+
@@ -479,13 +479,13 @@
done
(* ********************************************************************** *)
-(* \<forall>K. Card(K) \<longrightarrow> DC(K) ==> WO3 *)
+(* \<forall>K. Card(K) \<longrightarrow> DC(K) \<Longrightarrow> WO3 *)
(* ********************************************************************** *)
lemma fun_Ord_inj:
- "[| f \<in> a->X; Ord(a);
- !!b c. [| b<c; c \<in> a |] ==> f`b\<noteq>f`c |]
- ==> f \<in> inj(a, X)"
+ "\<lbrakk>f \<in> a->X; Ord(a);
+ \<And>b c. \<lbrakk>b<c; c \<in> a\<rbrakk> \<Longrightarrow> f`b\<noteq>f`c\<rbrakk>
+ \<Longrightarrow> f \<in> inj(a, X)"
apply (unfold inj_def, simp)
apply (intro ballI impI)
apply (rule_tac j=x in Ord_in_Ord [THEN Ord_linear_lt], assumption+)
@@ -493,17 +493,17 @@
apply (atomize, blast dest: not_sym)
done
-lemma value_in_image: "[| f \<in> X->Y; A \<subseteq> X; a \<in> A |] ==> f`a \<in> f``A"
+lemma value_in_image: "\<lbrakk>f \<in> X->Y; A \<subseteq> X; a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> f``A"
by (fast elim!: image_fun [THEN ssubst])
-lemma lesspoll_lemma: "[| ~ A \<prec> B; C \<prec> B |] ==> A - C \<noteq> 0"
+lemma lesspoll_lemma: "\<lbrakk>\<not> A \<prec> B; C \<prec> B\<rbrakk> \<Longrightarrow> 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
elim!: eqpollE lepoll_trans)
done
-theorem DC_WO3: "(\<forall>K. Card(K) \<longrightarrow> DC(K)) ==> WO3"
+theorem DC_WO3: "(\<forall>K. Card(K) \<longrightarrow> DC(K)) \<Longrightarrow> WO3"
apply (unfold DC_def WO3_def)
apply (rule allI)
apply (case_tac "A \<prec> Hartog (A)")
@@ -527,17 +527,17 @@
done
(* ********************************************************************** *)
-(* WO1 ==> \<forall>K. Card(K) \<longrightarrow> DC(K) *)
+(* WO1 \<Longrightarrow> \<forall>K. Card(K) \<longrightarrow> DC(K) *)
(* ********************************************************************** *)
lemma images_eq:
- "[| \<forall>x \<in> A. f`x=g`x; f \<in> Df->Cf; g \<in> Dg->Cg; A \<subseteq> Df; A \<subseteq> Dg |]
- ==> f``A = g``A"
+ "\<lbrakk>\<forall>x \<in> A. f`x=g`x; f \<in> Df->Cf; g \<in> Dg->Cg; A \<subseteq> Df; A \<subseteq> Dg\<rbrakk>
+ \<Longrightarrow> f``A = g``A"
apply (simp (no_asm_simp) add: image_fun)
done
lemma lam_images_eq:
- "[| Ord(a); b \<in> a |] ==> (\<lambda>x \<in> a. h(x))``b = (\<lambda>x \<in> b. h(x))``b"
+ "\<lbrakk>Ord(a); b \<in> a\<rbrakk> \<Longrightarrow> (\<lambda>x \<in> a. h(x))``b = (\<lambda>x \<in> b. h(x))``b"
apply (rule images_eq)
apply (rule ballI)
apply (drule OrdmemD [THEN subsetD], assumption+)
@@ -549,16 +549,16 @@
by (fast intro!: lam_type RepFunI)
lemma lemmaX:
- "[| \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R);
- b \<in> K; Z \<in> Pow(X); Z \<prec> K |]
- ==> {x \<in> X. <Z,x> \<in> R} \<noteq> 0"
+ "\<lbrakk>\<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R);
+ b \<in> K; Z \<in> Pow(X); Z \<prec> K\<rbrakk>
+ \<Longrightarrow> {x \<in> X. <Z,x> \<in> R} \<noteq> 0"
by blast
lemma WO1_DC_lemma:
- "[| Card(K); well_ord(X,Q);
- \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R); b \<in> K |]
- ==> ff(b, X, Q, R) \<in> {x \<in> X. <(\<lambda>c \<in> b. ff(c, X, Q, R))``b, x> \<in> R}"
+ "\<lbrakk>Card(K); well_ord(X,Q);
+ \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R); b \<in> K\<rbrakk>
+ \<Longrightarrow> ff(b, X, Q, R) \<in> {x \<in> X. <(\<lambda>c \<in> b. ff(c, X, Q, R))``b, x> \<in> R}"
apply (rule_tac P = "b \<in> K" in impE, (erule_tac [2] asm_rl)+)
apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct],
assumption+)
@@ -571,7 +571,7 @@
apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll)
done
-theorem WO1_DC_Card: "WO1 ==> \<forall>K. Card(K) \<longrightarrow> DC(K)"
+theorem WO1_DC_Card: "WO1 \<Longrightarrow> \<forall>K. Card(K) \<longrightarrow> DC(K)"
apply (unfold DC_def WO1_def)
apply (rule allI impI)+
apply (erule allE exE conjE)+
--- a/src/ZF/AC/HH.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/HH.thy Tue Sep 27 16:51:35 2022 +0100
@@ -2,9 +2,9 @@
Author: Krzysztof Grabczewski
Some properties of the recursive definition of HH used in the proofs of
- AC17 ==> AC1
- AC1 ==> WO2
- AC15 ==> WO6
+ AC17 \<Longrightarrow> AC1
+ AC1 \<Longrightarrow> WO2
+ AC15 \<Longrightarrow> WO6
*)
theory HH
@@ -13,7 +13,7 @@
definition
HH :: "[i, i, i] => i" where
- "HH(f,x,a) == transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
+ "HH(f,x,a) \<equiv> transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
in if f`z \<in> Pow(z)-{0} then f`z else {x})"
subsection\<open>Lemmas useful in each of the three proofs\<close>
@@ -29,26 +29,26 @@
done
lemma subset_imp_Diff_eq:
- "B \<subseteq> A ==> X-(\<Union>a \<in> A. P(a)) = X-(\<Union>a \<in> A-B. P(a))-(\<Union>b \<in> B. P(b))"
+ "B \<subseteq> A \<Longrightarrow> X-(\<Union>a \<in> A. P(a)) = X-(\<Union>a \<in> A-B. P(a))-(\<Union>b \<in> B. P(b))"
by fast
-lemma Ord_DiffE: "[| c \<in> a-b; b<a |] ==> c=b | b<c & c<a"
+lemma Ord_DiffE: "\<lbrakk>c \<in> a-b; b<a\<rbrakk> \<Longrightarrow> c=b | b<c & c<a"
apply (erule ltE)
apply (drule Ord_linear [of _ c])
apply (fast elim: Ord_in_Ord)
apply (fast intro!: ltI intro: Ord_in_Ord)
done
-lemma Diff_UN_eq_self: "(!!y. y\<in>A ==> P(y) = {x}) ==> x - (\<Union>y \<in> A. P(y)) = x"
+lemma Diff_UN_eq_self: "(\<And>y. y\<in>A \<Longrightarrow> P(y) = {x}) \<Longrightarrow> x - (\<Union>y \<in> A. P(y)) = x"
by (simp, fast elim!: mem_irrefl)
lemma HH_eq: "x - (\<Union>b \<in> a. HH(f,x,b)) = x - (\<Union>b \<in> a1. HH(f,x,b))
- ==> HH(f,x,a) = HH(f,x,a1)"
+ \<Longrightarrow> HH(f,x,a) = HH(f,x,a1)"
apply (subst HH_def_satisfies_eq [of _ _ a1])
apply (rule HH_def_satisfies_eq [THEN trans], simp)
done
-lemma HH_is_x_gt_too: "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}"
+lemma HH_is_x_gt_too: "\<lbrakk>HH(f,x,b)={x}; b<a\<rbrakk> \<Longrightarrow> HH(f,x,a)={x}"
apply (rule_tac P = "b<a" in impE)
prefer 2 apply assumption+
apply (erule lt_Ord2 [THEN trans_induct])
@@ -64,7 +64,7 @@
done
lemma HH_subset_x_lt_too:
- "[| HH(f,x,a) \<in> Pow(x)-{0}; b<a |] ==> HH(f,x,b) \<in> Pow(x)-{0}"
+ "\<lbrakk>HH(f,x,a) \<in> Pow(x)-{0}; b<a\<rbrakk> \<Longrightarrow> HH(f,x,b) \<in> Pow(x)-{0}"
apply (rule HH_values [THEN disjE], assumption)
apply (drule HH_is_x_gt_too, assumption)
apply (drule subst, assumption)
@@ -72,7 +72,7 @@
done
lemma HH_subset_x_imp_subset_Diff_UN:
- "HH(f,x,a) \<in> Pow(x)-{0} ==> HH(f,x,a) \<in> Pow(x - (\<Union>b \<in> a. HH(f,x,b)))-{0}"
+ "HH(f,x,a) \<in> Pow(x)-{0} \<Longrightarrow> HH(f,x,a) \<in> Pow(x - (\<Union>b \<in> a. HH(f,x,b)))-{0}"
apply (drule HH_def_satisfies_eq [THEN subst])
apply (rule HH_def_satisfies_eq [THEN ssubst])
apply (simp add: Let_def Diff_subset [THEN PowI])
@@ -81,7 +81,7 @@
done
lemma HH_eq_arg_lt:
- "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v) \<in> Pow(x)-{0}; v \<in> w |] ==> P"
+ "\<lbrakk>HH(f,x,v)=HH(f,x,w); HH(f,x,v) \<in> Pow(x)-{0}; v \<in> w\<rbrakk> \<Longrightarrow> P"
apply (frule_tac P = "%y. y \<in> Pow (x) -{0}" in subst, assumption)
apply (drule_tac a = w in HH_subset_x_imp_subset_Diff_UN)
apply (drule subst_elem, assumption)
@@ -89,7 +89,7 @@
done
lemma HH_eq_imp_arg_eq:
- "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w) \<in> Pow(x)-{0}; Ord(v); Ord(w) |] ==> v=w"
+ "\<lbrakk>HH(f,x,v)=HH(f,x,w); HH(f,x,w) \<in> Pow(x)-{0}; Ord(v); Ord(w)\<rbrakk> \<Longrightarrow> v=w"
apply (rule_tac j = w in Ord_linear_lt)
apply (simp_all (no_asm_simp))
apply (drule subst_elem, assumption)
@@ -98,7 +98,7 @@
done
lemma HH_subset_x_imp_lepoll:
- "[| HH(f, x, i) \<in> Pow(x)-{0}; Ord(i) |] ==> i \<lesssim> Pow(x)-{0}"
+ "\<lbrakk>HH(f, x, i) \<in> Pow(x)-{0}; Ord(i)\<rbrakk> \<Longrightarrow> i \<lesssim> Pow(x)-{0}"
apply (unfold lepoll_def inj_def)
apply (rule_tac x = "\<lambda>j \<in> i. HH (f, x, j) " in exI)
apply (simp (no_asm_simp))
@@ -120,13 +120,13 @@
by (fast intro!: Ord_Hartog HH_Hartog_is_x LeastI)
lemma less_Least_subset_x:
- "a \<in> (\<mu> i. HH(f,x,i)={x}) ==> HH(f,x,a) \<in> Pow(x)-{0}"
+ "a \<in> (\<mu> i. HH(f,x,i)={x}) \<Longrightarrow> HH(f,x,a) \<in> Pow(x)-{0}"
apply (rule HH_values [THEN disjE], assumption)
apply (rule less_LeastE)
apply (erule_tac [2] ltI [OF _ Ord_Least], assumption)
done
-subsection\<open>Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1\<close>
+subsection\<open>Lemmas used in the proofs of AC1 \<Longrightarrow> WO2 and AC17 \<Longrightarrow> AC1\<close>
lemma lam_Least_HH_inj_Pow:
"(\<lambda>a \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,a))
@@ -138,30 +138,30 @@
lemma lam_Least_HH_inj:
"\<forall>a \<in> (\<mu> i. HH(f,x,i)={x}). \<exists>z \<in> x. HH(f,x,a) = {z}
- ==> (\<lambda>a \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,a))
+ \<Longrightarrow> (\<lambda>a \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,a))
\<in> inj(\<mu> i. HH(f,x,i)={x}, {{y}. y \<in> x})"
by (rule lam_Least_HH_inj_Pow [THEN inj_strengthen_type], simp)
lemma lam_surj_sing:
- "[| x - (\<Union>a \<in> A. F(a)) = 0; \<forall>a \<in> A. \<exists>z \<in> x. F(a) = {z} |]
- ==> (\<lambda>a \<in> A. F(a)) \<in> surj(A, {{y}. y \<in> x})"
+ "\<lbrakk>x - (\<Union>a \<in> A. F(a)) = 0; \<forall>a \<in> A. \<exists>z \<in> x. F(a) = {z}\<rbrakk>
+ \<Longrightarrow> (\<lambda>a \<in> A. F(a)) \<in> surj(A, {{y}. y \<in> x})"
apply (simp add: surj_def lam_type Diff_eq_0_iff)
apply (blast elim: equalityE)
done
-lemma not_emptyI2: "y \<in> Pow(x)-{0} ==> x \<noteq> 0"
+lemma not_emptyI2: "y \<in> Pow(x)-{0} \<Longrightarrow> x \<noteq> 0"
by auto
lemma f_subset_imp_HH_subset:
"f`(x - (\<Union>j \<in> i. HH(f,x,j))) \<in> Pow(x - (\<Union>j \<in> i. HH(f,x,j)))-{0}
- ==> HH(f, x, i) \<in> Pow(x) - {0}"
+ \<Longrightarrow> HH(f, x, i) \<in> Pow(x) - {0}"
apply (rule HH_def_satisfies_eq [THEN ssubst])
apply (simp add: Let_def Diff_subset [THEN PowI] not_emptyI2 [THEN if_P], fast)
done
lemma f_subsets_imp_UN_HH_eq_x:
"\<forall>z \<in> Pow(x)-{0}. f`z \<in> Pow(z)-{0}
- ==> x - (\<Union>j \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,j)) = 0"
+ \<Longrightarrow> x - (\<Union>j \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,j)) = 0"
apply (case_tac "P \<in> {0}" for P, fast)
apply (drule Diff_subset [THEN PowI, THEN DiffI])
apply (drule bspec, assumption)
@@ -176,14 +176,14 @@
done
lemma HH_subset_imp_eq:
- "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\<Union>j \<in> i. HH(f,x,j)))"
+ "HH(f,x,i): Pow(x)-{0} \<Longrightarrow> HH(f,x,i)=f`(x - (\<Union>j \<in> i. HH(f,x,j)))"
apply (rule HH_values2 [THEN disjE], assumption)
apply (fast elim!: equalityE mem_irrefl dest!: singleton_subsetD)
done
lemma f_sing_imp_HH_sing:
- "[| f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x};
- a \<in> (\<mu> i. HH(f,x,i)={x}) |] ==> \<exists>z \<in> x. HH(f,x,a) = {z}"
+ "\<lbrakk>f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x};
+ a \<in> (\<mu> i. HH(f,x,i)={x})\<rbrakk> \<Longrightarrow> \<exists>z \<in> x. HH(f,x,a) = {z}"
apply (drule less_Least_subset_x)
apply (frule HH_subset_imp_eq)
apply (drule apply_type)
@@ -192,9 +192,9 @@
done
lemma f_sing_lam_bij:
- "[| x - (\<Union>j \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,j)) = 0;
- f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x} |]
- ==> (\<lambda>a \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,a))
+ "\<lbrakk>x - (\<Union>j \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,j)) = 0;
+ f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x}\<rbrakk>
+ \<Longrightarrow> (\<lambda>a \<in> (\<mu> i. HH(f,x,i)={x}). HH(f,x,a))
\<in> bij(\<mu> i. HH(f,x,i)={x}, {{y}. y \<in> x})"
apply (unfold bij_def)
apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing)
@@ -202,7 +202,7 @@
lemma lam_singI:
"f \<in> (\<Prod>X \<in> Pow(x)-{0}. F(X))
- ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X}) \<in> (\<Prod>X \<in> Pow(x)-{0}. {{z}. z \<in> F(X)})"
+ \<Longrightarrow> (\<lambda>X \<in> Pow(x)-{0}. {f`X}) \<in> (\<Prod>X \<in> Pow(x)-{0}. {{z}. z \<in> F(X)})"
by (fast del: DiffI DiffE
intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type)
@@ -213,7 +213,7 @@
lam_sing_bij [THEN bij_converse_bij]]
-subsection\<open>The proof of AC1 ==> WO2\<close>
+subsection\<open>The proof of AC1 \<Longrightarrow> WO2\<close>
(*Establishing the existence of a bijection, namely
converse
@@ -225,7 +225,7 @@
lemma bijection:
"f \<in> (\<Prod>X \<in> Pow(x) - {0}. X)
- ==> \<exists>g. g \<in> bij(x, \<mu> i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x})"
+ \<Longrightarrow> \<exists>g. g \<in> bij(x, \<mu> i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x})"
apply (rule exI)
apply (rule bij_Least_HH_x [THEN bij_converse_bij])
apply (rule f_subsets_imp_UN_HH_eq_x)
@@ -234,7 +234,7 @@
apply (fast intro: Pi_weaken_type)
done
-lemma AC1_WO2: "AC1 ==> WO2"
+lemma AC1_WO2: "AC1 \<Longrightarrow> WO2"
apply (unfold AC1_def WO2_def eqpoll_def)
apply (intro allI)
apply (drule_tac x = "Pow(A) - {0}" in spec)
--- a/src/ZF/AC/Hartog.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/Hartog.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,16 +10,16 @@
definition
Hartog :: "i => i" where
- "Hartog(X) == \<mu> i. ~ i \<lesssim> X"
+ "Hartog(X) \<equiv> \<mu> i. \<not> i \<lesssim> X"
-lemma Ords_in_set: "\<forall>a. Ord(a) \<longrightarrow> a \<in> X ==> P"
+lemma Ords_in_set: "\<forall>a. Ord(a) \<longrightarrow> a \<in> X \<Longrightarrow> P"
apply (rule_tac X = "{y \<in> X. Ord (y) }" in ON_class [elim_format])
apply fast
done
lemma Ord_lepoll_imp_ex_well_ord:
- "[| Ord(a); a \<lesssim> X |]
- ==> \<exists>Y. Y \<subseteq> X & (\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)"
+ "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk>
+ \<Longrightarrow> \<exists>Y. Y \<subseteq> X & (\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)"
apply (unfold lepoll_def)
apply (erule exE)
apply (intro exI conjI)
@@ -36,14 +36,14 @@
done
lemma Ord_lepoll_imp_eq_ordertype:
- "[| Ord(a); a \<lesssim> X |] ==> \<exists>Y. Y \<subseteq> X & (\<exists>R. R \<subseteq> X*X & ordertype(Y,R)=a)"
+ "\<lbrakk>Ord(a); a \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X & (\<exists>R. R \<subseteq> X*X & ordertype(Y,R)=a)"
apply (drule Ord_lepoll_imp_ex_well_ord, assumption, clarify)
apply (intro exI conjI)
apply (erule_tac [3] ordertype_Int, auto)
done
lemma Ords_lepoll_set_lemma:
- "(\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X) ==>
+ "(\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X) \<Longrightarrow>
\<forall>a. Ord(a) \<longrightarrow>
a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=<Y,R> & ordertype(Y,R)=b}"
apply (intro allI impI)
@@ -51,15 +51,15 @@
apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym)
done
-lemma Ords_lepoll_set: "\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X ==> P"
+lemma Ords_lepoll_set: "\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X \<Longrightarrow> P"
by (erule Ords_lepoll_set_lemma [THEN Ords_in_set])
-lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) & ~a \<lesssim> X"
+lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) & \<not>a \<lesssim> X"
apply (rule ccontr)
apply (best intro: Ords_lepoll_set)
done
-lemma not_Hartog_lepoll_self: "~ Hartog(A) \<lesssim> A"
+lemma not_Hartog_lepoll_self: "\<not> Hartog(A) \<lesssim> A"
apply (unfold Hartog_def)
apply (rule ex_Ord_not_lepoll [THEN exE])
apply (rule LeastI, auto)
@@ -70,10 +70,10 @@
lemma Ord_Hartog: "Ord(Hartog(A))"
by (unfold Hartog_def, rule Ord_Least)
-lemma less_HartogE1: "[| i < Hartog(A); ~ i \<lesssim> A |] ==> P"
+lemma less_HartogE1: "\<lbrakk>i < Hartog(A); \<not> i \<lesssim> A\<rbrakk> \<Longrightarrow> P"
by (unfold Hartog_def, fast elim: less_LeastE)
-lemma less_HartogE: "[| i < Hartog(A); i \<approx> Hartog(A) |] ==> P"
+lemma less_HartogE: "\<lbrakk>i < Hartog(A); i \<approx> Hartog(A)\<rbrakk> \<Longrightarrow> P"
by (blast intro: less_HartogE1 eqpoll_sym eqpoll_imp_lepoll
lepoll_trans [THEN Hartog_lepoll_selfE])
--- a/src/ZF/AC/WO1_AC.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/WO1_AC.thy Tue Sep 27 16:51:35 2022 +0100
@@ -1,7 +1,7 @@
(* Title: ZF/AC/WO1_AC.thy
Author: Krzysztof Grabczewski
-The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
+The proofs of WO1 \<Longrightarrow> AC1 and WO1 \<Longrightarrow> AC10(n) for n >= 1
The latter proof is referred to as clear by the Rubins.
However it seems to be quite complicated.
@@ -29,17 +29,17 @@
begin
(* ********************************************************************** *)
-(* WO1 ==> AC1 *)
+(* WO1 \<Longrightarrow> AC1 *)
(* ********************************************************************** *)
-theorem WO1_AC1: "WO1 ==> AC1"
+theorem WO1_AC1: "WO1 \<Longrightarrow> AC1"
by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun)
(* ********************************************************************** *)
-(* WO1 ==> AC10(n) (n >= 1) *)
+(* WO1 \<Longrightarrow> AC10(n) (n >= 1) *)
(* ********************************************************************** *)
-lemma lemma1: "[| WO1; \<forall>B \<in> A. \<exists>C \<in> D(B). P(C,B) |] ==> \<exists>f. \<forall>B \<in> A. P(f`B,B)"
+lemma lemma1: "\<lbrakk>WO1; \<forall>B \<in> A. \<exists>C \<in> D(B). P(C,B)\<rbrakk> \<Longrightarrow> \<exists>f. \<forall>B \<in> A. P(f`B,B)"
apply (unfold WO1_def)
apply (erule_tac x = "\<Union>({{C \<in> D (B) . P (C,B) }. B \<in> A}) " in allE)
apply (erule exE, drule ex_choice_fun, fast)
@@ -48,7 +48,7 @@
apply (simp, blast dest!: apply_type [OF _ RepFunI])
done
-lemma lemma2_1: "[| ~Finite(B); WO1 |] ==> |B| + |B| \<approx> B"
+lemma lemma2_1: "\<lbrakk>\<not>Finite(B); WO1\<rbrakk> \<Longrightarrow> |B| + |B| \<approx> B"
apply (unfold WO1_def)
apply (rule eqpoll_trans)
prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll)
@@ -61,19 +61,19 @@
done
lemma lemma2_2:
- "f \<in> bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \<in> D} \<in> Pow(Pow(B))"
+ "f \<in> bij(D+D, B) \<Longrightarrow> {{f`Inl(i), f`Inr(i)}. i \<in> D} \<in> Pow(Pow(B))"
by (fast elim!: bij_is_fun [THEN apply_type])
lemma lemma2_3:
- "f \<in> bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \<in> D})"
+ "f \<in> bij(D+D, B) \<Longrightarrow> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \<in> D})"
apply (unfold pairwise_disjoint_def)
apply (blast dest: bij_is_inj [THEN inj_apply_equality])
done
lemma lemma2_4:
- "[| f \<in> bij(D+D, B); 1\<le>n |]
- ==> sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \<in> D}, 2, succ(n))"
+ "\<lbrakk>f \<in> bij(D+D, B); 1\<le>n\<rbrakk>
+ \<Longrightarrow> sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \<in> D}, 2, succ(n))"
apply (simp (no_asm_simp) add: sets_of_size_between_def succ_def)
apply (blast intro!: cons_lepoll_cong
intro: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll]
@@ -82,14 +82,14 @@
done
lemma lemma2_5:
- "f \<in> bij(D+D, B) ==> \<Union>({{f`Inl(i), f`Inr(i)}. i \<in> D})=B"
+ "f \<in> bij(D+D, B) \<Longrightarrow> \<Union>({{f`Inl(i), f`Inr(i)}. i \<in> D})=B"
apply (unfold bij_def surj_def)
apply (fast elim!: inj_is_fun [THEN apply_type])
done
lemma lemma2:
- "[| WO1; ~Finite(B); 1\<le>n |]
- ==> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) &
+ "\<lbrakk>WO1; \<not>Finite(B); 1\<le>n\<rbrakk>
+ \<Longrightarrow> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) &
sets_of_size_between(C, 2, succ(n)) &
\<Union>(C)=B"
apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]],
@@ -97,7 +97,7 @@
apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5)
done
-theorem WO1_AC10: "[| WO1; 1\<le>n |] ==> AC10(n)"
+theorem WO1_AC10: "\<lbrakk>WO1; 1\<le>n\<rbrakk> \<Longrightarrow> AC10(n)"
apply (unfold AC10_def)
apply (fast intro!: lemma1 elim!: lemma2)
done
--- a/src/ZF/AC/WO1_WO7.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/WO1_WO7.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,8 +13,8 @@
begin
definition
- "LEMMA ==
- \<forall>X. ~Finite(X) \<longrightarrow> (\<exists>R. well_ord(X,R) & ~well_ord(X,converse(R)))"
+ "LEMMA \<equiv>
+ \<forall>X. \<not>Finite(X) \<longrightarrow> (\<exists>R. well_ord(X,R) & \<not>well_ord(X,converse(R)))"
(* ********************************************************************** *)
(* It is easy to see that WO7 is equivalent to (**) *)
@@ -29,7 +29,7 @@
(* It is also easy to show that LEMMA implies WO1. *)
(* ********************************************************************** *)
-lemma LEMMA_imp_WO1: "LEMMA ==> WO1"
+lemma LEMMA_imp_WO1: "LEMMA \<Longrightarrow> WO1"
apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def)
apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord])
done
@@ -48,7 +48,7 @@
(* ********************************************************************** *)
lemma converse_Memrel_not_wf_on:
- "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))"
+ "\<lbrakk>Ord(a); \<not>Finite(a)\<rbrakk> \<Longrightarrow> \<not>wf[a](converse(Memrel(a)))"
apply (unfold wf_on_def wf_def)
apply (drule nat_le_infinite_Ord [THEN le_imp_subset], assumption)
apply (rule notI)
@@ -56,7 +56,7 @@
done
lemma converse_Memrel_not_well_ord:
- "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))"
+ "\<lbrakk>Ord(a); \<not>Finite(a)\<rbrakk> \<Longrightarrow> \<not>well_ord(a,converse(Memrel(a)))"
apply (unfold well_ord_def)
apply (blast dest: converse_Memrel_not_wf_on)
done
@@ -69,15 +69,15 @@
Memrel_type [THEN subset_Int_iff [THEN iffD1]] trans)
lemma well_ord_converse_Memrel:
- "[| well_ord(A,r); well_ord(A,converse(r)) |]
- ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))"
+ "\<lbrakk>well_ord(A,r); well_ord(A,converse(r))\<rbrakk>
+ \<Longrightarrow> well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))"
apply (subst well_ord_rvimage_ordertype [symmetric], assumption)
apply (rule rvimage_converse [THEN subst])
apply (blast intro: ordertype_ord_iso ord_iso_sym ord_iso_is_bij
bij_is_inj well_ord_rvimage)
done
-lemma WO1_imp_LEMMA: "WO1 ==> LEMMA"
+lemma WO1_imp_LEMMA: "WO1 \<Longrightarrow> LEMMA"
apply (unfold WO1_def LEMMA_def, clarify)
apply (blast dest: well_ord_converse_Memrel
Ord_ordertype [THEN converse_Memrel_not_well_ord]
@@ -96,12 +96,12 @@
(* The proof of WO8 \<longleftrightarrow> WO1 (Rubin & Rubin p. 6) *)
(* ********************************************************************** *)
-lemma WO1_WO8: "WO1 ==> WO8"
+lemma WO1_WO8: "WO1 \<Longrightarrow> WO8"
by (unfold WO1_def WO8_def, fast)
-(* The implication "WO8 ==> WO1": a faithful image of Rubin & Rubin's proof*)
-lemma WO8_WO1: "WO8 ==> WO1"
+(* The implication "WO8 \<Longrightarrow> WO1": a faithful image of Rubin & Rubin's proof*)
+lemma WO8_WO1: "WO8 \<Longrightarrow> WO1"
apply (unfold WO1_def WO8_def)
apply (rule allI)
apply (erule_tac x = "{{x}. x \<in> A}" in allE)
--- a/src/ZF/AC/WO2_AC16.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/WO2_AC16.thy Tue Sep 27 16:51:35 2022 +0100
@@ -1,7 +1,7 @@
(* Title: ZF/AC/WO2_AC16.thy
Author: Krzysztof Grabczewski
-The proof of WO2 ==> AC16(k #+ m, k)
+The proof of WO2 \<Longrightarrow> AC16(k #+ m, k)
The main part of the proof is the inductive reasoning concerning
properties of constructed family T_gamma.
@@ -15,15 +15,15 @@
theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin
-(**** A recursive definition used in the proof of WO2 ==> AC16 ****)
+(**** A recursive definition used in the proof of WO2 \<Longrightarrow> AC16 ****)
definition
recfunAC16 :: "[i,i,i,i] => i" where
- "recfunAC16(f,h,i,a) ==
+ "recfunAC16(f,h,i,a) \<equiv>
transrec2(i, 0,
%g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
else r \<union> {f`(\<mu> i. h`g \<subseteq> f`i &
- (\<forall>b<a. (h`b \<subseteq> f`i \<longrightarrow> (\<forall>t \<in> r. ~ h`b \<subseteq> t))))})"
+ (\<forall>b<a. (h`b \<subseteq> f`i \<longrightarrow> (\<forall>t \<in> r. \<not> h`b \<subseteq> t))))})"
(* ********************************************************************** *)
(* Basic properties of recfunAC16 *)
@@ -38,12 +38,12 @@
else recfunAC16(f,h,i,a) \<union>
{f ` (\<mu> j. h ` i \<subseteq> f ` j &
(\<forall>b<a. (h`b \<subseteq> f`j
- \<longrightarrow> (\<forall>t \<in> recfunAC16(f,h,i,a). ~ h`b \<subseteq> t))))})"
+ \<longrightarrow> (\<forall>t \<in> recfunAC16(f,h,i,a). \<not> h`b \<subseteq> t))))})"
apply (simp add: recfunAC16_def)
done
lemma recfunAC16_Limit: "Limit(i)
- ==> recfunAC16(f,h,i,a) = (\<Union>j<i. recfunAC16(f,h,j,a))"
+ \<Longrightarrow> recfunAC16(f,h,i,a) = (\<Union>j<i. recfunAC16(f,h,j,a))"
by (simp add: recfunAC16_def transrec2_Limit)
(* ********************************************************************** *)
@@ -51,8 +51,8 @@
(* ********************************************************************** *)
lemma transrec2_mono_lemma [rule_format]:
- "[| !!g r. r \<subseteq> B(g,r); Ord(i) |]
- ==> j<i \<longrightarrow> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
+ "\<lbrakk>\<And>g r. r \<subseteq> B(g,r); Ord(i)\<rbrakk>
+ \<Longrightarrow> j<i \<longrightarrow> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
apply (erule trans_induct)
apply (rule Ord_cases, assumption+, fast)
apply (simp (no_asm_simp))
@@ -63,8 +63,8 @@
done
lemma transrec2_mono:
- "[| !!g r. r \<subseteq> B(g,r); j\<le>i |]
- ==> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
+ "\<lbrakk>\<And>g r. r \<subseteq> B(g,r); j\<le>i\<rbrakk>
+ \<Longrightarrow> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
apply (erule leE)
apply (rule transrec2_mono_lemma)
apply (auto intro: lt_Ord2 )
@@ -75,7 +75,7 @@
(* ********************************************************************** *)
lemma recfunAC16_mono:
- "i\<le>j ==> recfunAC16(f, g, i, a) \<subseteq> recfunAC16(f, g, j, a)"
+ "i\<le>j \<Longrightarrow> recfunAC16(f, g, i, a) \<subseteq> recfunAC16(f, g, j, a)"
apply (unfold recfunAC16_def)
apply (rule transrec2_mono, auto)
done
@@ -86,31 +86,31 @@
lemma lemma3_1:
- "[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+ "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
\<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a;
- V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |]
- ==> V = W"
+ V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk>
+ \<Longrightarrow> V = W"
apply (erule asm_rl allE impE)+
apply (drule subsetD, assumption, blast)
done
lemma lemma3:
- "[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+ "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
\<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); i<x; j<x; z<a;
- V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |]
- ==> V = W"
+ V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk>
+ \<Longrightarrow> V = W"
apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+)
apply (erule lemma3_1 [symmetric], assumption+)
apply (erule lemma3_1, assumption+)
done
lemma lemma4:
- "[| \<forall>y<x. F(y) \<subseteq> X &
+ "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X &
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));
- x < a |]
- ==> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) \<longrightarrow>
+ x < a\<rbrakk>
+ \<Longrightarrow> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(z) \<subseteq> Y)"
apply (intro oallI impI)
apply (drule ospec, assumption, clarify)
@@ -118,11 +118,11 @@
done
lemma lemma5:
- "[| \<forall>y<x. F(y) \<subseteq> X &
+ "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X &
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));
- x < a; Limit(x); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j) |]
- ==> (\<Union>x<x. F(x)) \<subseteq> X &
+ x < a; Limit(x); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk>
+ \<Longrightarrow> (\<Union>x<x. F(x)) \<subseteq> X &
(\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x)
\<longrightarrow> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) & h(xa) \<subseteq> Y))"
apply (rule conjI)
@@ -146,7 +146,7 @@
(*
First quite complicated proof of the fact used in the recursive construction
- of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage
+ of the family T_gamma (WO2 \<Longrightarrow> AC16(k #+ m, k)) - the fact that at any stage
gamma the set (s - \<Union>(...) - k_gamma) is equipollent to s
(Rubin & Rubin page 15).
*)
@@ -157,7 +157,7 @@
lemma dbl_Diff_eqpoll_Card:
- "[| A\<approx>a; Card(a); ~Finite(a); B\<prec>a; C\<prec>a |] ==> A - B - C\<approx>a"
+ "\<lbrakk>A\<approx>a; Card(a); \<not>Finite(a); B\<prec>a; C\<prec>a\<rbrakk> \<Longrightarrow> A - B - C\<approx>a"
by (blast intro: Diff_lesspoll_eqpoll_Card)
(* ********************************************************************** *)
@@ -166,7 +166,7 @@
lemma Finite_lesspoll_infinite_Ord:
- "[| Finite(X); ~Finite(a); Ord(a) |] ==> X\<prec>a"
+ "\<lbrakk>Finite(X); \<not>Finite(a); Ord(a)\<rbrakk> \<Longrightarrow> X\<prec>a"
apply (unfold lesspoll_def)
apply (rule conjI)
apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption)
@@ -177,9 +177,9 @@
done
lemma Union_lesspoll:
- "[| \<forall>x \<in> X. x \<lesssim> n & x \<subseteq> T; well_ord(T, R); X \<lesssim> b;
- b<a; ~Finite(a); Card(a); n \<in> nat |]
- ==> \<Union>(X)\<prec>a"
+ "\<lbrakk>\<forall>x \<in> X. x \<lesssim> n & x \<subseteq> T; well_ord(T, R); X \<lesssim> b;
+ b<a; \<not>Finite(a); Card(a); n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<Union>(X)\<prec>a"
apply (case_tac "Finite (X)")
apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord
lepoll_nat_imp_Finite Finite_Union)
@@ -203,20 +203,20 @@
lemma Un_sing_eq_cons: "A \<union> {a} = cons(a, A)"
by fast
-lemma Un_lepoll_succ: "A \<lesssim> B ==> A \<union> {a} \<lesssim> succ(B)"
+lemma Un_lepoll_succ: "A \<lesssim> B \<Longrightarrow> A \<union> {a} \<lesssim> succ(B)"
apply (simp add: Un_sing_eq_cons succ_def)
apply (blast elim!: mem_irrefl intro: cons_lepoll_cong)
done
-lemma Diff_UN_succ_empty: "Ord(a) ==> F(a) - (\<Union>b<succ(a). F(b)) = 0"
+lemma Diff_UN_succ_empty: "Ord(a) \<Longrightarrow> F(a) - (\<Union>b<succ(a). F(b)) = 0"
by (fast intro!: le_refl)
-lemma Diff_UN_succ_subset: "Ord(a) ==> F(a) \<union> X - (\<Union>b<succ(a). F(b)) \<subseteq> X"
+lemma Diff_UN_succ_subset: "Ord(a) \<Longrightarrow> F(a) \<union> X - (\<Union>b<succ(a). F(b)) \<subseteq> X"
by blast
lemma recfunAC16_Diff_lepoll_1:
"Ord(x)
- ==> recfunAC16(f, g, x, a) - (\<Union>i<x. recfunAC16(f, g, i, a)) \<lesssim> 1"
+ \<Longrightarrow> recfunAC16(f, g, x, a) - (\<Union>i<x. recfunAC16(f, g, i, a)) \<lesssim> 1"
apply (erule Ord_cases)
apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll])
(*Limit case*)
@@ -231,14 +231,14 @@
done
lemma in_Least_Diff:
- "[| z \<in> F(x); Ord(x) |]
- ==> z \<in> F(\<mu> i. z \<in> F(i)) - (\<Union>j<(\<mu> i. z \<in> F(i)). F(j))"
+ "\<lbrakk>z \<in> F(x); Ord(x)\<rbrakk>
+ \<Longrightarrow> z \<in> F(\<mu> i. z \<in> F(i)) - (\<Union>j<(\<mu> i. z \<in> F(i)). F(j))"
by (fast elim: less_LeastE elim!: LeastI)
lemma Least_eq_imp_ex:
- "[| (\<mu> i. w \<in> F(i)) = (\<mu> i. z \<in> F(i));
- w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i)) |]
- ==> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))"
+ "\<lbrakk>(\<mu> i. w \<in> F(i)) = (\<mu> i. z \<in> F(i));
+ w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i))\<rbrakk>
+ \<Longrightarrow> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))"
apply (elim OUN_E)
apply (drule in_Least_Diff, erule lt_Ord)
apply (frule in_Least_Diff, erule lt_Ord)
@@ -247,13 +247,13 @@
done
-lemma two_in_lepoll_1: "[| A \<lesssim> 1; a \<in> A; b \<in> A |] ==> a=b"
+lemma two_in_lepoll_1: "\<lbrakk>A \<lesssim> 1; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a=b"
by (fast dest!: lepoll_1_is_sing)
lemma UN_lepoll_index:
- "[| \<forall>i<a. F(i)-(\<Union>j<i. F(j)) \<lesssim> 1; Limit(a) |]
- ==> (\<Union>x<a. F(x)) \<lesssim> a"
+ "\<lbrakk>\<forall>i<a. F(i)-(\<Union>j<i. F(j)) \<lesssim> 1; Limit(a)\<rbrakk>
+ \<Longrightarrow> (\<Union>x<a. F(x)) \<lesssim> a"
apply (rule lepoll_def [THEN def_imp_iff [THEN iffD2]])
apply (rule_tac x = "\<lambda>z \<in> (\<Union>x<a. F (x)). \<mu> i. z \<in> F (i) " in exI)
apply (unfold inj_def)
@@ -271,7 +271,7 @@
done
-lemma recfunAC16_lepoll_index: "Ord(y) ==> recfunAC16(f, h, y, a) \<lesssim> y"
+lemma recfunAC16_lepoll_index: "Ord(y) \<Longrightarrow> recfunAC16(f, h, y, a) \<lesssim> y"
apply (erule trans_induct3)
(*0 case*)
apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl)
@@ -286,9 +286,9 @@
lemma Union_recfunAC16_lesspoll:
- "[| recfunAC16(f,g,y,a) \<subseteq> {X \<in> Pow(A). X\<approx>n};
- A\<approx>a; y<a; ~Finite(a); Card(a); n \<in> nat |]
- ==> \<Union>(recfunAC16(f,g,y,a))\<prec>a"
+ "\<lbrakk>recfunAC16(f,g,y,a) \<subseteq> {X \<in> Pow(A). X\<approx>n};
+ A\<approx>a; y<a; \<not>Finite(a); Card(a); n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<Union>(recfunAC16(f,g,y,a))\<prec>a"
apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE])
apply (rule_tac T=A in Union_lesspoll, simp_all)
apply (blast intro!: eqpoll_imp_lepoll)
@@ -298,11 +298,11 @@
done
lemma dbl_Diff_eqpoll:
- "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
- Card(a); ~ Finite(a); A\<approx>a;
+ "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
+ Card(a); \<not> Finite(a); A\<approx>a;
k \<in> nat; y<a;
- h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}) |]
- ==> A - \<Union>(recfunAC16(f, h, y, a)) - h`y\<approx>a"
+ h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)})\<rbrakk>
+ \<Longrightarrow> A - \<Union>(recfunAC16(f, h, y, a)) - h`y\<approx>a"
apply (rule dbl_Diff_eqpoll_Card, simp_all)
apply (simp add: Union_recfunAC16_lesspoll)
apply (rule Finite_lesspoll_infinite_Ord)
@@ -318,9 +318,9 @@
OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum]
-lemma Un_in_Collect: "[| x \<in> Pow(A - B - h`i); x\<approx>m;
- h \<in> bij(a, {x \<in> Pow(A) . x\<approx>k}); i<a; k \<in> nat; m \<in> nat |]
- ==> h ` i \<union> x \<in> {x \<in> Pow(A) . x\<approx>k #+ m}"
+lemma Un_in_Collect: "\<lbrakk>x \<in> Pow(A - B - h`i); x\<approx>m;
+ h \<in> bij(a, {x \<in> Pow(A) . x\<approx>k}); i<a; k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> h ` i \<union> x \<in> {x \<in> Pow(A) . x\<approx>k #+ m}"
by (blast intro: disj_Un_eqpoll_nat_sum
dest: ltD bij_is_fun [THEN apply_type])
@@ -330,14 +330,14 @@
(* ********************************************************************** *)
lemma lemma6:
- "[| \<forall>y<succ(j). F(y)<=X & (\<forall>x<a. x<y | P(x,y) \<longrightarrow> Q(x,y)); succ(j)<a |]
- ==> F(j)<=X & (\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j))"
+ "\<lbrakk>\<forall>y<succ(j). F(y)<=X & (\<forall>x<a. x<y | P(x,y) \<longrightarrow> Q(x,y)); succ(j)<a\<rbrakk>
+ \<Longrightarrow> F(j)<=X & (\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j))"
by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl])
lemma lemma7:
- "[| \<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j); succ(j)<a |]
- ==> P(j,j) \<longrightarrow> (\<forall>x<a. x\<le>j | P(x,j) \<longrightarrow> Q(x,j))"
+ "\<lbrakk>\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j); succ(j)<a\<rbrakk>
+ \<Longrightarrow> P(j,j) \<longrightarrow> (\<forall>x<a. x\<le>j | P(x,j) \<longrightarrow> Q(x,j))"
by (fast elim!: leE)
(* ********************************************************************** *)
@@ -346,7 +346,7 @@
(* ********************************************************************** *)
lemma ex_subset_eqpoll:
- "[| A\<approx>a; ~ Finite(a); Ord(a); m \<in> nat |] ==> \<exists>X \<in> Pow(A). X\<approx>m"
+ "\<lbrakk>A\<approx>a; \<not> Finite(a); Ord(a); m \<in> nat\<rbrakk> \<Longrightarrow> \<exists>X \<in> Pow(A). X\<approx>m"
apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE])
apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll])
apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat)
@@ -354,12 +354,12 @@
apply (fast elim!: eqpoll_sym)
done
-lemma subset_Un_disjoint: "[| A \<subseteq> B \<union> C; A \<inter> C = 0 |] ==> A \<subseteq> B"
+lemma subset_Un_disjoint: "\<lbrakk>A \<subseteq> B \<union> C; A \<inter> C = 0\<rbrakk> \<Longrightarrow> A \<subseteq> B"
by blast
lemma Int_empty:
- "[| X \<in> Pow(A - \<Union>(B) -C); T \<in> B; F \<subseteq> T |] ==> F \<inter> X = 0"
+ "\<lbrakk>X \<in> Pow(A - \<Union>(B) -C); T \<in> B; F \<subseteq> T\<rbrakk> \<Longrightarrow> F \<inter> X = 0"
by blast
@@ -369,7 +369,7 @@
lemma subset_imp_eq_lemma:
- "m \<in> nat ==> \<forall>A B. A \<subseteq> B & m \<lesssim> A & B \<lesssim> m \<longrightarrow> A=B"
+ "m \<in> nat \<Longrightarrow> \<forall>A B. A \<subseteq> B & m \<lesssim> A & B \<lesssim> m \<longrightarrow> A=B"
apply (induct_tac "m")
apply (fast dest!: lepoll_0_is_0)
apply (intro allI impI)
@@ -385,13 +385,13 @@
done
-lemma subset_imp_eq: "[| A \<subseteq> B; m \<lesssim> A; B \<lesssim> m; m \<in> nat |] ==> A=B"
+lemma subset_imp_eq: "\<lbrakk>A \<subseteq> B; m \<lesssim> A; B \<lesssim> m; m \<in> nat\<rbrakk> \<Longrightarrow> A=B"
by (blast dest!: subset_imp_eq_lemma)
lemma bij_imp_arg_eq:
- "[| f \<in> bij(a, {Y \<in> X. Y\<approx>succ(k)}); k \<in> nat; f`b \<subseteq> f`y; b<a; y<a |]
- ==> b=y"
+ "\<lbrakk>f \<in> bij(a, {Y \<in> X. Y\<approx>succ(k)}); k \<in> nat; f`b \<subseteq> f`y; b<a; y<a\<rbrakk>
+ \<Longrightarrow> b=y"
apply (drule subset_imp_eq)
apply (erule_tac [3] nat_succI)
apply (unfold bij_def inj_def)
@@ -401,14 +401,14 @@
lemma ex_next_set:
- "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
- Card(a); ~ Finite(a); A\<approx>a;
+ "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
+ Card(a); \<not> Finite(a); A\<approx>a;
k \<in> nat; m \<in> nat; y<a;
h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});
- ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]
- ==> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X &
+ \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk>
+ \<Longrightarrow> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X &
(\<forall>b<a. h`b \<subseteq> X \<longrightarrow>
- (\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
+ (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))"
apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE],
assumption+)
apply (erule Card_is_Ord, assumption)
@@ -427,15 +427,15 @@
(* ********************************************************************** *)
lemma ex_next_Ord:
- "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
- Card(a); ~ Finite(a); A\<approx>a;
+ "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};
+ Card(a); \<not> Finite(a); A\<approx>a;
k \<in> nat; m \<in> nat; y<a;
h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});
f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});
- ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]
- ==> \<exists>c<a. h`y \<subseteq> f`c &
+ \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk>
+ \<Longrightarrow> \<exists>c<a. h`y \<subseteq> f`c &
(\<forall>b<a. h`b \<subseteq> f`c \<longrightarrow>
- (\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
+ (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))"
apply (drule ex_next_set, assumption+)
apply (erule bexE)
apply (rule_tac x="converse(f)`X" in oexI)
@@ -450,10 +450,10 @@
(* ********************************************************************** *)
lemma lemma8:
- "[| \<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa))
+ "\<lbrakk>\<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa))
\<longrightarrow> (\<exists>! Y. Y \<in> F(j) & P(x, Y)); F(j) \<subseteq> X;
- L \<in> X; P(j, L) & (\<forall>x<a. P(x, L) \<longrightarrow> (\<forall>xa \<in> F(j). ~P(x, xa))) |]
- ==> F(j) \<union> {L} \<subseteq> X &
+ L \<in> X; P(j, L) & (\<forall>x<a. P(x, L) \<longrightarrow> (\<forall>xa \<in> F(j). \<not>P(x, xa)))\<rbrakk>
+ \<Longrightarrow> F(j) \<union> {L} \<subseteq> X &
(\<forall>x<a. x\<le>j | (\<exists>xa \<in> (F(j) \<union> {L}). P(x, xa)) \<longrightarrow>
(\<exists>! Y. Y \<in> (F(j) \<union> {L}) & P(x, Y)))"
apply (rule conjI)
@@ -468,10 +468,10 @@
(* ********************************************************************** *)
lemma main_induct:
- "[| b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)});
+ "\<lbrakk>b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)});
h \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k)});
- ~Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat |]
- ==> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} &
+ \<not>Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} &
(\<forall>x<a. x < b | (\<exists>Y \<in> recfunAC16(f, h, b, a). h ` x \<subseteq> Y) \<longrightarrow>
(\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) & h ` x \<subseteq> Y))"
apply (erule lt_induct)
@@ -507,11 +507,11 @@
(* ********************************************************************** *)
lemma lemma_simp_induct:
- "[| \<forall>b. b<a \<longrightarrow> F(b) \<subseteq> S & (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y))
+ "\<lbrakk>\<forall>b. b<a \<longrightarrow> F(b) \<subseteq> S & (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y))
\<longrightarrow> (\<exists>! Y. Y \<in> F(b) & f`x \<subseteq> Y));
f \<in> a->f``(a); Limit(a);
- \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j) |]
- ==> (\<Union>j<a. F(j)) \<subseteq> S &
+ \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk>
+ \<Longrightarrow> (\<Union>j<a. F(j)) \<subseteq> S &
(\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) & x \<subseteq> Y)"
apply (rule conjI)
apply (rule subsetI)
@@ -546,7 +546,7 @@
(* The target theorem *)
(* ********************************************************************** *)
-theorem WO2_AC16: "[| WO2; 0<m; k \<in> nat; m \<in> nat |] ==> AC16(k #+ m,k)"
+theorem WO2_AC16: "\<lbrakk>WO2; 0<m; k \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> AC16(k #+ m,k)"
apply (unfold AC16_def)
apply (rule allI)
apply (rule impI)
--- a/src/ZF/AC/WO6_WO1.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/AC/WO6_WO1.thy Tue Sep 27 16:51:35 2022 +0100
@@ -2,11 +2,11 @@
Author: Krzysztof Grabczewski
Proofs needed to state that formulations WO1,...,WO6 are all equivalent.
-The only hard one is WO6 ==> WO1.
+The only hard one is WO6 \<Longrightarrow> WO1.
-Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear"
+Every proof (except WO6 \<Longrightarrow> WO1 and WO1 \<Longrightarrow> WO2) are described as "clear"
by Rubin & Rubin (page 2).
-They refer reader to a book by Gödel to see the proof WO1 ==> WO2.
+They refer reader to a book by Gödel to see the proof WO1 \<Longrightarrow> WO2.
Fortunately order types made this proof also very easy.
*)
@@ -17,18 +17,18 @@
(* Auxiliary definitions used in proof *)
definition
NN :: "i => i" where
- "NN(y) == {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a &
+ "NN(y) \<equiv> {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a &
(\<Union>b<a. f`b) = y & (\<forall>b<a. f`b \<lesssim> m)}"
definition
uu :: "[i, i, i, i] => i" where
- "uu(f, beta, gamma, delta) == (f`beta * f`gamma) \<inter> f`delta"
+ "uu(f, beta, gamma, delta) \<equiv> (f`beta * f`gamma) \<inter> f`delta"
(** Definitions for case 1 **)
definition
vv1 :: "[i, i, i] => i" where
- "vv1(f,m,b) ==
+ "vv1(f,m,b) \<equiv>
let g = \<mu> g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 &
domain(uu(f,b,g,d)) \<lesssim> m));
d = \<mu> d. domain(uu(f,b,g,d)) \<noteq> 0 &
@@ -37,35 +37,35 @@
definition
ww1 :: "[i, i, i] => i" where
- "ww1(f,m,b) == f`b - vv1(f,m,b)"
+ "ww1(f,m,b) \<equiv> f`b - vv1(f,m,b)"
definition
gg1 :: "[i, i, i] => i" where
- "gg1(f,a,m) == \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
+ "gg1(f,a,m) \<equiv> \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
(** Definitions for case 2 **)
definition
vv2 :: "[i, i, i, i] => i" where
- "vv2(f,b,g,s) ==
+ "vv2(f,b,g,s) \<equiv>
if f`g \<noteq> 0 then {uu(f, b, g, \<mu> d. uu(f,b,g,d) \<noteq> 0)`s} else 0"
definition
ww2 :: "[i, i, i, i] => i" where
- "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)"
+ "ww2(f,b,g,s) \<equiv> f`g - vv2(f,b,g,s)"
definition
gg2 :: "[i, i, i, i] => i" where
- "gg2(f,a,b,s) ==
+ "gg2(f,a,b,s) \<equiv>
\<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
-lemma WO2_WO3: "WO2 ==> WO3"
+lemma WO2_WO3: "WO2 \<Longrightarrow> WO3"
by (unfold WO2_def WO3_def, fast)
(* ********************************************************************** *)
-lemma WO3_WO1: "WO3 ==> WO1"
+lemma WO3_WO1: "WO3 \<Longrightarrow> WO1"
apply (unfold eqpoll_def WO1_def WO3_def)
apply (intro allI)
apply (drule_tac x=A in spec)
@@ -75,25 +75,25 @@
(* ********************************************************************** *)
-lemma WO1_WO2: "WO1 ==> WO2"
+lemma WO1_WO2: "WO1 \<Longrightarrow> WO2"
apply (unfold eqpoll_def WO1_def WO2_def)
apply (blast intro!: Ord_ordertype ordermap_bij)
done
(* ********************************************************************** *)
-lemma lam_sets: "f \<in> A->B ==> (\<lambda>x \<in> A. {f`x}): A -> {{b}. b \<in> B}"
+lemma lam_sets: "f \<in> A->B \<Longrightarrow> (\<lambda>x \<in> A. {f`x}): A -> {{b}. b \<in> B}"
by (fast intro!: lam_type apply_type)
-lemma surj_imp_eq': "f \<in> surj(A,B) ==> (\<Union>a \<in> A. {f`a}) = B"
+lemma surj_imp_eq': "f \<in> surj(A,B) \<Longrightarrow> (\<Union>a \<in> A. {f`a}) = B"
apply (unfold surj_def)
apply (fast elim!: apply_type)
done
-lemma surj_imp_eq: "[| f \<in> surj(A,B); Ord(A) |] ==> (\<Union>a<A. {f`a}) = B"
+lemma surj_imp_eq: "\<lbrakk>f \<in> surj(A,B); Ord(A)\<rbrakk> \<Longrightarrow> (\<Union>a<A. {f`a}) = B"
by (fast dest!: surj_imp_eq' intro!: ltI elim!: ltE)
-lemma WO1_WO4: "WO1 ==> WO4(1)"
+lemma WO1_WO4: "WO1 \<Longrightarrow> WO4(1)"
apply (unfold WO1_def WO4_def)
apply (rule allI)
apply (erule_tac x = A in allE)
@@ -108,32 +108,32 @@
(* ********************************************************************** *)
-lemma WO4_mono: "[| m\<le>n; WO4(m) |] ==> WO4(n)"
+lemma WO4_mono: "\<lbrakk>m\<le>n; WO4(m)\<rbrakk> \<Longrightarrow> WO4(n)"
apply (unfold WO4_def)
apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll])
done
(* ********************************************************************** *)
-lemma WO4_WO5: "[| m \<in> nat; 1\<le>m; WO4(m) |] ==> WO5"
+lemma WO4_WO5: "\<lbrakk>m \<in> nat; 1\<le>m; WO4(m)\<rbrakk> \<Longrightarrow> WO5"
by (unfold WO4_def WO5_def, blast)
(* ********************************************************************** *)
-lemma WO5_WO6: "WO5 ==> WO6"
+lemma WO5_WO6: "WO5 \<Longrightarrow> WO6"
by (unfold WO4_def WO5_def WO6_def, blast)
(* **********************************************************************
- The proof of "WO6 ==> WO1". Simplified by L C Paulson.
+ The proof of "WO6 \<Longrightarrow> WO1". Simplified by L C Paulson.
From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,
pages 2-5
************************************************************************* *)
lemma lt_oadd_odiff_disj:
- "[| k < i++j; Ord(i); Ord(j) |]
- ==> k < i | (~ k<i & k = i ++ (k--i) & (k--i)<j)"
+ "\<lbrakk>k < i++j; Ord(i); Ord(j)\<rbrakk>
+ \<Longrightarrow> k < i | (\<not> k<i & k = i ++ (k--i) & (k--i)<j)"
apply (rule_tac i = k and j = i in Ord_linear2)
prefer 4
apply (drule odiff_lt_mono2, assumption)
@@ -154,7 +154,7 @@
by (unfold uu_def, blast)
lemma quant_domain_uu_lepoll_m:
- "\<forall>b<a. f`b \<lesssim> m ==> \<forall>b<a. \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<lesssim> m"
+ "\<forall>b<a. f`b \<lesssim> m \<Longrightarrow> \<forall>b<a. \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<lesssim> m"
by (blast intro: domain_uu_subset [THEN subset_imp_lepoll] lepoll_trans)
lemma uu_subset1: "uu(f,b,g,d) \<subseteq> f`b * f`g"
@@ -163,7 +163,7 @@
lemma uu_subset2: "uu(f,b,g,d) \<subseteq> f`d"
by (unfold uu_def, blast)
-lemma uu_lepoll_m: "[| \<forall>b<a. f`b \<lesssim> m; d<a |] ==> uu(f,b,g,d) \<lesssim> m"
+lemma uu_lepoll_m: "\<lbrakk>\<forall>b<a. f`b \<lesssim> m; d<a\<rbrakk> \<Longrightarrow> uu(f,b,g,d) \<lesssim> m"
by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans)
(* ********************************************************************** *)
@@ -171,7 +171,7 @@
(* ********************************************************************** *)
lemma cases:
"\<forall>b<a. \<forall>g<a. \<forall>d<a. u(f,b,g,d) \<lesssim> m
- ==> (\<forall>b<a. f`b \<noteq> 0 \<longrightarrow>
+ \<Longrightarrow> (\<forall>b<a. f`b \<noteq> 0 \<longrightarrow>
(\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 & u(f,b,g,d) \<prec> m))
| (\<exists>b<a. f`b \<noteq> 0 & (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 \<longrightarrow>
u(f,b,g,d) \<approx> m))"
@@ -182,7 +182,7 @@
(* ********************************************************************** *)
(* Lemmas used in both cases *)
(* ********************************************************************** *)
-lemma UN_oadd: "Ord(a) ==> (\<Union>b<a++a. C(b)) = (\<Union>b<a. C(b) \<union> C(a++b))"
+lemma UN_oadd: "Ord(a) \<Longrightarrow> (\<Union>b<a++a. C(b)) = (\<Union>b<a. C(b) \<union> C(a++b))"
by (blast intro: ltI lt_oadd1 oadd_lt_mono2 dest!: lt_oadd_disj)
@@ -197,7 +197,7 @@
(* Case 1: Union of images is the whole "y" *)
(* ********************************************************************** *)
lemma UN_gg1_eq:
- "[| Ord(a); m \<in> nat |] ==> (\<Union>b<a++a. gg1(f,a,m)`b) = (\<Union>b<a. f`b)"
+ "\<lbrakk>Ord(a); m \<in> nat\<rbrakk> \<Longrightarrow> (\<Union>b<a++a. gg1(f,a,m)`b) = (\<Union>b<a. f`b)"
by (simp add: gg1_def UN_oadd lt_oadd1 oadd_le_self [THEN le_imp_not_lt]
lt_Ord odiff_oadd_inverse ltD vv1_subset [THEN Diff_partition]
ww1_def)
@@ -209,9 +209,9 @@
(* every value of defined function is less than or equipollent to m *)
(* ********************************************************************** *)
lemma nested_LeastI:
- "[| P(a, b); Ord(a); Ord(b);
- Least_a = (\<mu> a. \<exists>x. Ord(x) & P(a, x)) |]
- ==> P(Least_a, \<mu> b. P(Least_a, b))"
+ "\<lbrakk>P(a, b); Ord(a); Ord(b);
+ Least_a = (\<mu> a. \<exists>x. Ord(x) & P(a, x))\<rbrakk>
+ \<Longrightarrow> P(Least_a, \<mu> b. P(Least_a, b))"
apply (erule ssubst)
apply (rule_tac Q = "%z. P (z, \<mu> b. P (z, b))" in LeastI2)
apply (fast elim!: LeastI)+
@@ -222,12 +222,12 @@
domain(uu(f,b,g,d)) \<lesssim> m"] for f b m
lemma gg1_lepoll_m:
- "[| Ord(a); m \<in> nat;
+ "\<lbrakk>Ord(a); m \<in> nat;
\<forall>b<a. f`b \<noteq>0 \<longrightarrow>
(\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0 &
domain(uu(f,b,g,d)) \<lesssim> m);
- \<forall>b<a. f`b \<lesssim> succ(m); b<a++a |]
- ==> gg1(f,a,m)`b \<lesssim> m"
+ \<forall>b<a. f`b \<lesssim> succ(m); b<a++a\<rbrakk>
+ \<Longrightarrow> gg1(f,a,m)`b \<lesssim> m"
apply (simp add: gg1_def empty_lepollI)
apply (safe dest!: lt_oadd_odiff_disj)
(*Case b<a \<in> show vv1(f,m,b) \<lesssim> m *)
@@ -253,34 +253,34 @@
(* ********************************************************************** *)
lemma ex_d_uu_not_empty:
- "[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0;
- y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]
- ==> \<exists>d<a. uu(f,b,g,d) \<noteq> 0"
+ "\<lbrakk>b<a; g<a; f`b\<noteq>0; f`g\<noteq>0;
+ y*y \<subseteq> y; (\<Union>b<a. f`b)=y\<rbrakk>
+ \<Longrightarrow> \<exists>d<a. uu(f,b,g,d) \<noteq> 0"
by (unfold uu_def, blast)
lemma uu_not_empty:
- "[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]
- ==> uu(f,b,g,\<mu> d. (uu(f,b,g,d) \<noteq> 0)) \<noteq> 0"
+ "\<lbrakk>b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y\<rbrakk>
+ \<Longrightarrow> uu(f,b,g,\<mu> d. (uu(f,b,g,d) \<noteq> 0)) \<noteq> 0"
apply (drule ex_d_uu_not_empty, assumption+)
apply (fast elim!: LeastI lt_Ord)
done
-lemma not_empty_rel_imp_domain: "[| r \<subseteq> A*B; r\<noteq>0 |] ==> domain(r)\<noteq>0"
+lemma not_empty_rel_imp_domain: "\<lbrakk>r \<subseteq> A*B; r\<noteq>0\<rbrakk> \<Longrightarrow> domain(r)\<noteq>0"
by blast
lemma Least_uu_not_empty_lt_a:
- "[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]
- ==> (\<mu> d. uu(f,b,g,d) \<noteq> 0) < a"
+ "\<lbrakk>b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y\<rbrakk>
+ \<Longrightarrow> (\<mu> d. uu(f,b,g,d) \<noteq> 0) < a"
apply (erule ex_d_uu_not_empty [THEN oexE], assumption+)
apply (blast intro: Least_le [THEN lt_trans1] lt_Ord)
done
-lemma subset_Diff_sing: "[| B \<subseteq> A; a\<notin>B |] ==> B \<subseteq> A-{a}"
+lemma subset_Diff_sing: "\<lbrakk>B \<subseteq> A; a\<notin>B\<rbrakk> \<Longrightarrow> B \<subseteq> A-{a}"
by blast
(*Could this be proved more directly?*)
lemma supset_lepoll_imp_eq:
- "[| A \<lesssim> m; m \<lesssim> B; B \<subseteq> A; m \<in> nat |] ==> A=B"
+ "\<lbrakk>A \<lesssim> m; m \<lesssim> B; B \<subseteq> A; m \<in> nat\<rbrakk> \<Longrightarrow> A=B"
apply (erule natE)
apply (fast dest!: lepoll_0_is_0 intro!: equalityI)
apply (safe intro!: equalityI)
@@ -293,12 +293,12 @@
done
lemma uu_Least_is_fun:
- "[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>
+ "\<lbrakk>\<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>
domain(uu(f, b, g, d)) \<approx> succ(m);
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
(\<Union>b<a. f`b)=y; b<a; g<a; d<a;
- f`b\<noteq>0; f`g\<noteq>0; m \<in> nat; s \<in> f`b |]
- ==> uu(f, b, g, \<mu> d. uu(f,b,g,d)\<noteq>0) \<in> f`b -> f`g"
+ f`b\<noteq>0; f`g\<noteq>0; m \<in> nat; s \<in> f`b\<rbrakk>
+ \<Longrightarrow> uu(f, b, g, \<mu> d. uu(f,b,g,d)\<noteq>0) \<in> f`b -> f`g"
apply (drule_tac x2=g in ospec [THEN ospec, THEN mp])
apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty])
apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+)
@@ -312,11 +312,11 @@
done
lemma vv2_subset:
- "[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>
+ "\<lbrakk>\<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>
domain(uu(f, b, g, d)) \<approx> succ(m);
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
- (\<Union>b<a. f`b)=y; b<a; g<a; m \<in> nat; s \<in> f`b |]
- ==> vv2(f,b,g,s) \<subseteq> f`g"
+ (\<Union>b<a. f`b)=y; b<a; g<a; m \<in> nat; s \<in> f`b\<rbrakk>
+ \<Longrightarrow> vv2(f,b,g,s) \<subseteq> f`g"
apply (simp add: vv2_def)
apply (blast intro: uu_Least_is_fun [THEN apply_type])
done
@@ -325,11 +325,11 @@
(* Case 2: Union of images is the whole "y" *)
(* ********************************************************************** *)
lemma UN_gg2_eq:
- "[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>
+ "\<lbrakk>\<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>
domain(uu(f,b,g,d)) \<approx> succ(m);
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
- (\<Union>b<a. f`b)=y; Ord(a); m \<in> nat; s \<in> f`b; b<a |]
- ==> (\<Union>g<a++a. gg2(f,a,b,s) ` g) = y"
+ (\<Union>b<a. f`b)=y; Ord(a); m \<in> nat; s \<in> f`b; b<a\<rbrakk>
+ \<Longrightarrow> (\<Union>g<a++a. gg2(f,a,b,s) ` g) = y"
apply (unfold gg2_def)
apply (drule sym)
apply (simp add: ltD UN_oadd oadd_le_self [THEN le_imp_not_lt]
@@ -345,7 +345,7 @@
(* every value of defined function is less than or equipollent to m *)
(* ********************************************************************** *)
-lemma vv2_lepoll: "[| m \<in> nat; m\<noteq>0 |] ==> vv2(f,b,g,s) \<lesssim> m"
+lemma vv2_lepoll: "\<lbrakk>m \<in> nat; m\<noteq>0\<rbrakk> \<Longrightarrow> vv2(f,b,g,s) \<lesssim> m"
apply (unfold vv2_def)
apply (simp add: empty_lepollI)
apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0]
@@ -355,8 +355,8 @@
done
lemma ww2_lepoll:
- "[| \<forall>b<a. f`b \<lesssim> succ(m); g<a; m \<in> nat; vv2(f,b,g,d) \<subseteq> f`g |]
- ==> ww2(f,b,g,d) \<lesssim> m"
+ "\<lbrakk>\<forall>b<a. f`b \<lesssim> succ(m); g<a; m \<in> nat; vv2(f,b,g,d) \<subseteq> f`g\<rbrakk>
+ \<Longrightarrow> ww2(f,b,g,d) \<lesssim> m"
apply (unfold ww2_def)
apply (case_tac "f`g = 0")
apply (simp add: empty_lepollI)
@@ -366,11 +366,11 @@
done
lemma gg2_lepoll_m:
- "[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>
+ "\<lbrakk>\<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>
domain(uu(f,b,g,d)) \<approx> succ(m);
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
- (\<Union>b<a. f`b)=y; b<a; s \<in> f`b; m \<in> nat; m\<noteq> 0; g<a++a |]
- ==> gg2(f,a,b,s) ` g \<lesssim> m"
+ (\<Union>b<a. f`b)=y; b<a; s \<in> f`b; m \<in> nat; m\<noteq> 0; g<a++a\<rbrakk>
+ \<Longrightarrow> gg2(f,a,b,s) ` g \<lesssim> m"
apply (simp add: gg2_def empty_lepollI)
apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj)
apply (simp add: vv2_lepoll)
@@ -380,7 +380,7 @@
(* ********************************************************************** *)
(* lemma ii *)
(* ********************************************************************** *)
-lemma lemma_ii: "[| succ(m) \<in> NN(y); y*y \<subseteq> y; m \<in> nat; m\<noteq>0 |] ==> m \<in> NN(y)"
+lemma lemma_ii: "\<lbrakk>succ(m) \<in> NN(y); y*y \<subseteq> y; m \<in> nat; m\<noteq>0\<rbrakk> \<Longrightarrow> m \<in> NN(y)"
apply (unfold NN_def)
apply (elim CollectE exE conjE)
apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption)
@@ -413,16 +413,16 @@
by (fast intro: rec_succ [THEN ssubst])
lemma le_subsets:
- "[| \<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat |]
- ==> f(n)<=f(m)"
+ "\<lbrakk>\<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat\<rbrakk>
+ \<Longrightarrow> f(n)<=f(m)"
apply (erule_tac P = "n\<le>m" in rev_mp)
apply (rule_tac P = "%z. n\<le>z \<longrightarrow> f (n) \<subseteq> f (z) " in nat_induct)
apply (auto simp add: le_iff)
done
lemma le_imp_rec_subset:
- "[| n\<le>m; m \<in> nat |]
- ==> rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(m, x, %k r. r \<union> r*r)"
+ "\<lbrakk>n\<le>m; m \<in> nat\<rbrakk>
+ \<Longrightarrow> rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(m, x, %k r. r \<union> r*r)"
apply (rule z_n_subset_z_succ_n [THEN le_subsets])
apply (blast intro: lt_nat_in_nat)+
done
@@ -444,34 +444,34 @@
(* y can be well-ordered" *)
(* In fact we have to prove *)
-(* * WO6 ==> NN(y) \<noteq> 0 *)
+(* * WO6 \<Longrightarrow> NN(y) \<noteq> 0 *)
(* * reverse induction which lets us infer that 1 \<in> NN(y) *)
-(* * 1 \<in> NN(y) ==> y can be well-ordered *)
+(* * 1 \<in> NN(y) \<Longrightarrow> y can be well-ordered *)
(* ********************************************************************** *)
(* ********************************************************************** *)
-(* WO6 ==> NN(y) \<noteq> 0 *)
+(* WO6 \<Longrightarrow> NN(y) \<noteq> 0 *)
(* ********************************************************************** *)
-lemma WO6_imp_NN_not_empty: "WO6 ==> NN(y) \<noteq> 0"
+lemma WO6_imp_NN_not_empty: "WO6 \<Longrightarrow> NN(y) \<noteq> 0"
by (unfold WO6_def NN_def, clarify, blast)
(* ********************************************************************** *)
-(* 1 \<in> NN(y) ==> y can be well-ordered *)
+(* 1 \<in> NN(y) \<Longrightarrow> y can be well-ordered *)
(* ********************************************************************** *)
lemma lemma1:
- "[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |] ==> \<exists>c<a. f`c = {x}"
+ "\<lbrakk>(\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a)\<rbrakk> \<Longrightarrow> \<exists>c<a. f`c = {x}"
by (fast elim!: lepoll_1_is_sing)
lemma lemma2:
- "[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |]
- ==> f` (\<mu> i. f`i = {x}) = {x}"
+ "\<lbrakk>(\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a)\<rbrakk>
+ \<Longrightarrow> f` (\<mu> i. f`i = {x}) = {x}"
apply (drule lemma1, assumption+)
apply (fast elim!: lt_Ord intro: LeastI)
done
-lemma NN_imp_ex_inj: "1 \<in> NN(y) ==> \<exists>a f. Ord(a) & f \<in> inj(y, a)"
+lemma NN_imp_ex_inj: "1 \<in> NN(y) \<Longrightarrow> \<exists>a f. Ord(a) & f \<in> inj(y, a)"
apply (unfold NN_def)
apply (elim CollectE exE conjE)
apply (rule_tac x = a in exI)
@@ -483,7 +483,7 @@
apply (rule lemma2 [THEN ssubst], assumption+, blast)
done
-lemma y_well_ord: "[| y*y \<subseteq> y; 1 \<in> NN(y) |] ==> \<exists>r. well_ord(y, r)"
+lemma y_well_ord: "\<lbrakk>y*y \<subseteq> y; 1 \<in> NN(y)\<rbrakk> \<Longrightarrow> \<exists>r. well_ord(y, r)"
apply (drule NN_imp_ex_inj)
apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel])
done
@@ -493,35 +493,35 @@
(* ********************************************************************** *)
lemma rev_induct_lemma [rule_format]:
- "[| n \<in> nat; !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]
- ==> n\<noteq>0 \<longrightarrow> P(n) \<longrightarrow> P(1)"
+ "\<lbrakk>n \<in> nat; \<And>m. \<lbrakk>m \<in> nat; m\<noteq>0; P(succ(m))\<rbrakk> \<Longrightarrow> P(m)\<rbrakk>
+ \<Longrightarrow> n\<noteq>0 \<longrightarrow> P(n) \<longrightarrow> P(1)"
by (erule nat_induct, blast+)
lemma rev_induct:
- "[| n \<in> nat; P(n); n\<noteq>0;
- !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]
- ==> P(1)"
+ "\<lbrakk>n \<in> nat; P(n); n\<noteq>0;
+ \<And>m. \<lbrakk>m \<in> nat; m\<noteq>0; P(succ(m))\<rbrakk> \<Longrightarrow> P(m)\<rbrakk>
+ \<Longrightarrow> P(1)"
by (rule rev_induct_lemma, blast+)
-lemma NN_into_nat: "n \<in> NN(y) ==> n \<in> nat"
+lemma NN_into_nat: "n \<in> NN(y) \<Longrightarrow> n \<in> nat"
by (simp add: NN_def)
-lemma lemma3: "[| n \<in> NN(y); y*y \<subseteq> y; n\<noteq>0 |] ==> 1 \<in> NN(y)"
+lemma lemma3: "\<lbrakk>n \<in> NN(y); y*y \<subseteq> y; n\<noteq>0\<rbrakk> \<Longrightarrow> 1 \<in> NN(y)"
apply (rule rev_induct [OF NN_into_nat], assumption+)
apply (rule lemma_ii, assumption+)
done
(* ********************************************************************** *)
-(* Main theorem "WO6 ==> WO1" *)
+(* Main theorem "WO6 \<Longrightarrow> WO1" *)
(* ********************************************************************** *)
(* another helpful lemma *)
-lemma NN_y_0: "0 \<in> NN(y) ==> y=0"
+lemma NN_y_0: "0 \<in> NN(y) \<Longrightarrow> y=0"
apply (unfold NN_def)
apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst)
done
-lemma WO6_imp_WO1: "WO6 ==> WO1"
+lemma WO6_imp_WO1: "WO6 \<Longrightarrow> WO1"
apply (unfold WO1_def)
apply (rule allI)
apply (case_tac "A=0")
--- a/src/ZF/Arith.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Arith.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,11 +17,11 @@
definition
pred :: "i=>i" (*inverse of succ*) where
- "pred(y) == nat_case(0, %x. x, y)"
+ "pred(y) \<equiv> nat_case(0, %x. x, y)"
definition
natify :: "i=>i" (*coerces non-nats to nats*) where
- "natify == Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a))
+ "natify \<equiv> Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a))
else 0)"
consts
@@ -44,44 +44,44 @@
definition
add :: "[i,i]=>i" (infixl \<open>#+\<close> 65) where
- "m #+ n == raw_add (natify(m), natify(n))"
+ "m #+ n \<equiv> raw_add (natify(m), natify(n))"
definition
diff :: "[i,i]=>i" (infixl \<open>#-\<close> 65) where
- "m #- n == raw_diff (natify(m), natify(n))"
+ "m #- n \<equiv> raw_diff (natify(m), natify(n))"
definition
mult :: "[i,i]=>i" (infixl \<open>#*\<close> 70) where
- "m #* n == raw_mult (natify(m), natify(n))"
+ "m #* n \<equiv> raw_mult (natify(m), natify(n))"
definition
raw_div :: "[i,i]=>i" where
- "raw_div (m, n) ==
+ "raw_div (m, n) \<equiv>
transrec(m, %j f. if j<n | n=0 then 0 else succ(f`(j#-n)))"
definition
raw_mod :: "[i,i]=>i" where
- "raw_mod (m, n) ==
+ "raw_mod (m, n) \<equiv>
transrec(m, %j f. if j<n | n=0 then j else f`(j#-n))"
definition
div :: "[i,i]=>i" (infixl \<open>div\<close> 70) where
- "m div n == raw_div (natify(m), natify(n))"
+ "m div n \<equiv> raw_div (natify(m), natify(n))"
definition
mod :: "[i,i]=>i" (infixl \<open>mod\<close> 70) where
- "m mod n == raw_mod (natify(m), natify(n))"
+ "m mod n \<equiv> raw_mod (natify(m), natify(n))"
declare rec_type [simp]
nat_0_le [simp]
-lemma zero_lt_lemma: "[| 0<k; k \<in> nat |] ==> \<exists>j\<in>nat. k = succ(j)"
+lemma zero_lt_lemma: "\<lbrakk>0<k; k \<in> nat\<rbrakk> \<Longrightarrow> \<exists>j\<in>nat. k = succ(j)"
apply (erule rev_mp)
apply (induct_tac "k", auto)
done
-(* @{term"[| 0 < k; k \<in> nat; !!j. [| j \<in> nat; k = succ(j) |] ==> Q |] ==> Q"} *)
+(* @{term"\<lbrakk>0 < k; k \<in> nat; \<And>j. \<lbrakk>j \<in> nat; k = succ(j)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"} *)
lemmas zero_lt_natE = zero_lt_lemma [THEN bexE]
@@ -96,7 +96,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 \<noteq> succ(z) ==> natify(x) = 0"
+lemma natify_non_succ: "\<forall>z. x \<noteq> succ(z) \<Longrightarrow> natify(x) = 0"
by (rule natify_def [THEN def_Vrecursor, THEN trans], auto)
lemma natify_in_nat [iff,TC]: "natify(x) \<in> nat"
@@ -105,12 +105,12 @@
apply (auto simp add: natify_succ natify_non_succ)
done
-lemma natify_ident [simp]: "n \<in> nat ==> natify(n) = n"
+lemma natify_ident [simp]: "n \<in> nat \<Longrightarrow> natify(n) = n"
apply (induct_tac "n")
apply (auto simp add: natify_succ)
done
-lemma natify_eqE: "[|natify(x) = y; x \<in> nat|] ==> x=y"
+lemma natify_eqE: "\<lbrakk>natify(x) = y; x \<in> nat\<rbrakk> \<Longrightarrow> x=y"
by auto
@@ -165,7 +165,7 @@
(** Addition **)
-lemma raw_add_type: "[| m\<in>nat; n\<in>nat |] ==> raw_add (m, n) \<in> nat"
+lemma raw_add_type: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> raw_add (m, n) \<in> nat"
by (induct_tac "m", auto)
lemma add_type [iff,TC]: "m #+ n \<in> nat"
@@ -173,7 +173,7 @@
(** Multiplication **)
-lemma raw_mult_type: "[| m\<in>nat; n\<in>nat |] ==> raw_mult (m, n) \<in> nat"
+lemma raw_mult_type: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> raw_mult (m, n) \<in> nat"
apply (induct_tac "m")
apply (simp_all add: raw_add_type)
done
@@ -184,7 +184,7 @@
(** Difference **)
-lemma raw_diff_type: "[| m\<in>nat; n\<in>nat |] ==> raw_diff (m, n) \<in> nat"
+lemma raw_diff_type: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> raw_diff (m, n) \<in> nat"
by (induct_tac "n", auto)
lemma diff_type [iff,TC]: "m #- n \<in> nat"
@@ -208,7 +208,7 @@
lemma diff_0 [simp]: "m #- 0 = natify(m)"
by (simp add: diff_def)
-lemma diff_le_self: "m\<in>nat ==> (m #- n) \<le> m"
+lemma diff_le_self: "m\<in>nat \<Longrightarrow> (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)
@@ -225,7 +225,7 @@
lemma add_succ [simp]: "succ(m) #+ n = succ(m #+ n)"
by (simp add: natify_succ add_def)
-lemma add_0: "m \<in> nat ==> 0 #+ m = m"
+lemma add_0: "m \<in> nat \<Longrightarrow> 0 #+ m = m"
by simp
(*Associative law for addition*)
@@ -250,7 +250,7 @@
apply (auto simp add: natify_succ)
done
-lemma add_0_right: "m \<in> nat ==> m #+ 0 = m"
+lemma add_0_right: "m \<in> nat \<Longrightarrow> m #+ 0 = m"
by auto
(*Commutative law for addition*)
@@ -272,40 +272,40 @@
(*Cancellation law on the left*)
lemma raw_add_left_cancel:
- "[| raw_add(k, m) = raw_add(k, n); k\<in>nat |] ==> m=n"
+ "\<lbrakk>raw_add(k, m) = raw_add(k, n); k\<in>nat\<rbrakk> \<Longrightarrow> m=n"
apply (erule rev_mp)
apply (induct_tac "k", auto)
done
-lemma add_left_cancel_natify: "k #+ m = k #+ n ==> natify(m) = natify(n)"
+lemma add_left_cancel_natify: "k #+ m = k #+ n \<Longrightarrow> natify(m) = natify(n)"
apply (unfold add_def)
apply (drule raw_add_left_cancel, auto)
done
lemma add_left_cancel:
- "[| i = j; i #+ m = j #+ n; m\<in>nat; n\<in>nat |] ==> m = n"
+ "\<lbrakk>i = j; i #+ m = j #+ n; m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> m = n"
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)"
+lemma add_le_elim1_natify: "k#+m \<le> k#+n \<Longrightarrow> 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: "\<lbrakk>k#+m \<le> k#+n; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m \<le> n"
by (drule add_le_elim1_natify, auto)
-lemma add_lt_elim1_natify: "k#+m < k#+n ==> natify(m) < natify(n)"
+lemma add_lt_elim1_natify: "k#+m < k#+n \<Longrightarrow> natify(m) < natify(n)"
apply (rule_tac P = "natify(k) #+m < natify(k) #+n" in rev_mp)
apply (rule_tac [2] n = "natify(k) " in nat_induct)
apply auto
done
-lemma add_lt_elim1: "[| k#+m < k#+n; m \<in> nat; n \<in> nat |] ==> m < n"
+lemma add_lt_elim1: "\<lbrakk>k#+m < k#+n; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m < n"
by (drule add_lt_elim1_natify, auto)
-lemma zero_less_add: "[| n \<in> nat; m \<in> nat |] ==> 0 < m #+ n \<longleftrightarrow> (0<m | 0<n)"
+lemma zero_less_add: "\<lbrakk>n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> 0 < m #+ n \<longleftrightarrow> (0<m | 0<n)"
by (induct_tac "n", auto)
@@ -314,20 +314,20 @@
(*strict, in 1st argument; proof is by rule induction on 'less than'.
Still need j\<in>nat, for consider j = omega. Then we can have i<omega,
which is the same as i\<in>nat, but natify(j)=0, so the conclusion fails.*)
-lemma add_lt_mono1: "[| i<j; j\<in>nat |] ==> i#+k < j#+k"
+lemma add_lt_mono1: "\<lbrakk>i<j; j\<in>nat\<rbrakk> \<Longrightarrow> i#+k < j#+k"
apply (frule lt_nat_in_nat, assumption)
apply (erule succ_lt_induct)
apply (simp_all add: leI)
done
text\<open>strict, in second argument\<close>
-lemma add_lt_mono2: "[| i<j; j\<in>nat |] ==> k#+i < k#+j"
+lemma add_lt_mono2: "\<lbrakk>i<j; j\<in>nat\<rbrakk> \<Longrightarrow> k#+i < k#+j"
by (simp add: add_commute [of k] add_lt_mono1)
text\<open>A [clumsy] way of lifting < monotonicity to \<open>\<le>\<close> monotonicity\<close>
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))"
+ assumes lt_mono: "\<And>i j. \<lbrakk>i<j; j:k\<rbrakk> \<Longrightarrow> f(i) < f(j)"
+ and ford: "\<And>i. i:k \<Longrightarrow> Ord(f(i))"
and leij: "i \<le> j"
and jink: "j:k"
shows "f(i) \<le> f(j)"
@@ -336,29 +336,29 @@
done
text\<open>\<open>\<le>\<close> monotonicity, 1st argument\<close>
-lemma add_le_mono1: "[| i \<le> j; j\<in>nat |] ==> i#+k \<le> j#+k"
+lemma add_le_mono1: "\<lbrakk>i \<le> j; j\<in>nat\<rbrakk> \<Longrightarrow> 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\<open>\<open>\<le>\<close> monotonicity, both arguments\<close>
-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: "\<lbrakk>i \<le> j; k \<le> l; j\<in>nat; l\<in>nat\<rbrakk> \<Longrightarrow> 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
text\<open>Combinations of less-than and less-than-or-equals\<close>
-lemma add_lt_le_mono: "[| i<j; k\<le>l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
+lemma add_lt_le_mono: "\<lbrakk>i<j; k\<le>l; j\<in>nat; l\<in>nat\<rbrakk> \<Longrightarrow> i#+k < j#+l"
apply (rule add_lt_mono1 [THEN lt_trans2], assumption+)
apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+)
done
-lemma add_le_lt_mono: "[| i\<le>j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
+lemma add_le_lt_mono: "\<lbrakk>i\<le>j; k<l; j\<in>nat; l\<in>nat\<rbrakk> \<Longrightarrow> i#+k < j#+l"
by (subst add_commute, subst add_commute, erule add_lt_le_mono, assumption+)
text\<open>Less-than: in other words, strict in both arguments\<close>
-lemma add_lt_mono: "[| i<j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
+lemma add_lt_mono: "\<lbrakk>i<j; k<l; j\<in>nat; l\<in>nat\<rbrakk> \<Longrightarrow> i#+k < j#+l"
apply (rule add_lt_le_mono)
apply (auto intro: leI)
done
@@ -393,21 +393,21 @@
lemma pred_0 [simp]: "pred(0) = 0"
by (simp add: pred_def)
-lemma eq_succ_imp_eq_m1: "[|i = succ(j); i\<in>nat|] ==> j = i #- 1 & j \<in>nat"
+lemma eq_succ_imp_eq_m1: "\<lbrakk>i = succ(j); i\<in>nat\<rbrakk> \<Longrightarrow> j = i #- 1 & j \<in>nat"
by simp
lemma pred_Un_distrib:
- "[|i\<in>nat; j\<in>nat|] ==> pred(i \<union> j) = pred(i) \<union> pred(j)"
+ "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> 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
lemma pred_type [TC,simp]:
- "i \<in> nat ==> pred(i) \<in> nat"
+ "i \<in> nat \<Longrightarrow> pred(i) \<in> nat"
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)"
+lemma nat_diff_pred: "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> i #- succ(j) = pred(i #- j)"
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
@@ -418,13 +418,13 @@
done
lemma nat_diff_Un_distrib:
- "[|i\<in>nat; j\<in>nat; k\<in>nat|] ==> (i \<union> j) #- k = (i#-k) \<union> (j#-k)"
+ "\<lbrakk>i\<in>nat; j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> (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 \<union> j) #- k = (i#-k) \<union> (j#-k)"
+ "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> (i \<union> j) #- k = (i#-k) \<union> (j#-k)"
by (insert nat_diff_Un_distrib [of i j "natify(k)"], simp)
text\<open>We actually prove \<^term>\<open>i #- j #- k = i #- (j #+ k)\<close>\<close>
@@ -451,10 +451,10 @@
by (simp add: diff_cancel)
(*To tidy up the result of a simproc. Only the RHS will be simplified.*)
-lemma eq_cong2: "u = u' ==> (t==u) == (t==u')"
+lemma eq_cong2: "u = u' \<Longrightarrow> (t\<equiv>u) \<equiv> (t\<equiv>u')"
by auto
-lemma iff_cong2: "u \<longleftrightarrow> u' ==> (t==u) == (t==u')"
+lemma iff_cong2: "u \<longleftrightarrow> u' \<Longrightarrow> (t\<equiv>u) \<equiv> (t\<equiv>u')"
by auto
@@ -488,10 +488,10 @@
lemma mult_1_right_natify [simp]: "n #* 1 = natify(n)"
by auto
-lemma mult_1: "n \<in> nat ==> 1 #* n = n"
+lemma mult_1: "n \<in> nat \<Longrightarrow> 1 #* n = n"
by simp
-lemma mult_1_right: "n \<in> nat ==> n #* 1 = n"
+lemma mult_1_right: "n \<in> nat \<Longrightarrow> n #* 1 = n"
by simp
(*Commutative law for multiplication*)
@@ -536,12 +536,12 @@
lemma lt_succ_eq_0_disj:
- "[| m\<in>nat; n\<in>nat |]
- ==> (m < succ(n)) \<longleftrightarrow> (m = 0 | (\<exists>j\<in>nat. m = succ(j) & j < n))"
+ "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk>
+ \<Longrightarrow> (m < succ(n)) \<longleftrightarrow> (m = 0 | (\<exists>j\<in>nat. m = succ(j) & j < n))"
by (induct_tac "m", auto)
lemma less_diff_conv [rule_format]:
- "[| j\<in>nat; k\<in>nat |] ==> \<forall>i\<in>nat. (i < j #- k) \<longleftrightarrow> (i #+ k < j)"
+ "\<lbrakk>j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> \<forall>i\<in>nat. (i < j #- k) \<longleftrightarrow> (i #+ k < j)"
by (erule_tac m = k in diff_induct, auto)
lemmas nat_typechecks = rec_type nat_0I nat_1I nat_succI Ord_nat
--- a/src/ZF/ArithSimp.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ArithSimp.thy Tue Sep 27 16:51:35 2022 +0100
@@ -25,19 +25,19 @@
(*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 \<noteq> 0 = natify(m).*)
-lemma add_diff_inverse: "[| n \<le> m; m:nat |] ==> n #+ (m#-n) = m"
+lemma add_diff_inverse: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> (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: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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)
@@ -45,7 +45,7 @@
done
lemma zero_less_diff [simp]:
- "[| m: nat; n: nat |] ==> 0 < (n #- m) \<longleftrightarrow> m<n"
+ "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> 0 < (n #- m) \<longleftrightarrow> m<n"
apply (rule_tac m = m and n = n in diff_induct)
apply (simp_all (no_asm_simp))
done
@@ -67,7 +67,7 @@
subsection\<open>Remainder\<close>
(*We need m:nat even with natify*)
-lemma div_termination: "[| 0<n; n \<le> m; m:nat |] ==> m #- n < m"
+lemma div_termination: "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> m #- n < m"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (erule rev_mp)
apply (erule rev_mp)
@@ -81,7 +81,7 @@
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) \<in> nat"
+lemma raw_mod_type: "\<lbrakk>m:nat; n:nat\<rbrakk> \<Longrightarrow> 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])
@@ -109,25 +109,25 @@
apply (simp (no_asm_simp))
done (*NOT for adding to default simpset*)
-lemma raw_mod_less: "m<n ==> raw_mod (m,n) = m"
+lemma raw_mod_less: "m<n \<Longrightarrow> raw_mod (m,n) = m"
apply (rule raw_mod_def [THEN def_transrec, THEN trans])
apply (simp (no_asm_simp) add: div_termination [THEN ltD])
done
-lemma mod_less [simp]: "[| m<n; n \<in> nat |] ==> m mod n = m"
+lemma mod_less [simp]: "\<lbrakk>m<n; n \<in> nat\<rbrakk> \<Longrightarrow> 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)"
+ "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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)
@@ -137,7 +137,7 @@
subsection\<open>Division\<close>
-lemma raw_div_type: "[| m:nat; n:nat |] ==> raw_div (m, n) \<in> nat"
+lemma raw_div_type: "\<lbrakk>m:nat; n:nat\<rbrakk> \<Longrightarrow> 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])
@@ -149,17 +149,17 @@
apply (simp (no_asm) add: div_def raw_div_type)
done
-lemma raw_div_less: "m<n ==> raw_div (m,n) = 0"
+lemma raw_div_less: "m<n \<Longrightarrow> raw_div (m,n) = 0"
apply (rule raw_div_def [THEN def_transrec, THEN trans])
apply (simp (no_asm_simp) add: div_termination [THEN ltD])
done
-lemma div_less [simp]: "[| m<n; n \<in> nat |] ==> m div n = 0"
+lemma div_less [simp]: "\<lbrakk>m<n; n \<in> nat\<rbrakk> \<Longrightarrow> 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))"
+lemma raw_div_geq: "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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)
@@ -168,7 +168,7 @@
done
lemma div_geq [simp]:
- "[| 0<n; n \<le> m; m:nat |] ==> m div n = succ ((m#-n) div n)"
+ "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> 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
@@ -177,7 +177,7 @@
(*A key result*)
-lemma mod_div_lemma: "[| m: nat; n: nat |] ==> (m div n)#*n #+ m mod n = m"
+lemma mod_div_lemma: "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m div n)#*n #+ m mod n = m"
apply (case_tac "n=0")
apply (simp add: DIVISION_BY_ZERO_MOD)
apply (simp add: nat_into_Ord [THEN Ord_0_lt_iff])
@@ -195,7 +195,7 @@
apply (subst mod_div_lemma, auto)
done
-lemma mod_div_equality: "m: nat ==> (m div n)#*n #+ m mod n = m"
+lemma mod_div_equality: "m: nat \<Longrightarrow> (m div n)#*n #+ m mod n = m"
apply (simp (no_asm_simp) add: mod_div_equality_natify)
done
@@ -205,8 +205,8 @@
text\<open>(mainly for mutilated chess board)\<close>
lemma mod_succ_lemma:
- "[| 0<n; m:nat; n:nat |]
- ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"
+ "\<lbrakk>0<n; m:nat; n:nat\<rbrakk>
+ \<Longrightarrow> 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\<open>case succ(x) < n\<close>
@@ -221,7 +221,7 @@
done
lemma mod_succ:
- "n:nat ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"
+ "n:nat \<Longrightarrow> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"
apply (case_tac "n=0")
apply (simp (no_asm_simp) add: natify_succ DIVISION_BY_ZERO_MOD)
apply (subgoal_tac "natify (succ (m)) mod n = (if succ (natify (m) mod n) = n then 0 else succ (natify (m) mod n))")
@@ -231,7 +231,7 @@
apply (auto simp del: natify_succ simp add: nat_into_Ord [THEN Ord_0_lt_iff])
done
-lemma mod_less_divisor: "[| 0<n; n:nat |] ==> m mod n < n"
+lemma mod_less_divisor: "\<lbrakk>0<n; n:nat\<rbrakk> \<Longrightarrow> 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)
@@ -242,7 +242,7 @@
lemma mod_1_eq [simp]: "m mod 1 = 0"
by (cut_tac n = 1 in mod_less_divisor, auto)
-lemma mod2_cases: "b<2 ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"
+lemma mod2_cases: "b<2 \<Longrightarrow> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)"
apply (subgoal_tac "k mod 2: 2")
prefer 2 apply (simp add: mod_less_divisor [THEN ltD])
apply (drule ltD, auto)
@@ -266,17 +266,17 @@
subsection\<open>Additional theorems about \<open>\<le>\<close>\<close>
-lemma add_le_self: "m:nat ==> m \<le> (m #+ n)"
+lemma add_le_self: "m:nat \<Longrightarrow> 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 \<Longrightarrow> 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)"
+lemma mult_le_mono1: "\<lbrakk>i \<le> j; j:nat\<rbrakk> \<Longrightarrow> (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)
@@ -284,14 +284,14 @@
done
(* @{text"\<le>"} monotonicity, BOTH arguments*)
-lemma mult_le_mono: "[| i \<le> j; k \<le> l; j:nat; l:nat |] ==> i#*k \<le> j#*l"
+lemma mult_le_mono: "\<lbrakk>i \<le> j; k \<le> l; j:nat; l:nat\<rbrakk> \<Longrightarrow> 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
(*strict, in 1st argument; proof is by induction on k>0.
I can't see how to relax the typing conditions.*)
-lemma mult_lt_mono2: "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j"
+lemma mult_lt_mono2: "\<lbrakk>i<j; 0<k; j:nat; k:nat\<rbrakk> \<Longrightarrow> k#*i < k#*j"
apply (erule zero_lt_natE)
apply (frule_tac [2] lt_nat_in_nat)
apply (simp_all (no_asm_simp))
@@ -299,7 +299,7 @@
apply (simp_all (no_asm_simp) add: add_lt_mono)
done
-lemma mult_lt_mono1: "[| i<j; 0<k; j:nat; k:nat |] ==> i#*k < j#*k"
+lemma mult_lt_mono1: "\<lbrakk>i<j; 0<k; j:nat; k:nat\<rbrakk> \<Longrightarrow> i#*k < j#*k"
apply (simp (no_asm_simp) add: mult_lt_mono2 mult_commute [of _ k])
done
@@ -326,7 +326,7 @@
done
-lemma mult_is_zero: "[|m: nat; n: nat|] ==> (m #* n = 0) \<longleftrightarrow> (m = 0 | n = 0)"
+lemma mult_is_zero: "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m #* n = 0) \<longleftrightarrow> (m = 0 | n = 0)"
apply auto
apply (erule natE)
apply (erule_tac [2] natE, auto)
@@ -342,7 +342,7 @@
subsection\<open>Cancellation Laws for Common Factors in Comparisons\<close>
lemma mult_less_cancel_lemma:
- "[| k: nat; m: nat; n: nat |] ==> (m#*k < n#*k) \<longleftrightarrow> (0<k & m<n)"
+ "\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k < n#*k) \<longleftrightarrow> (0<k & m<n)"
apply (safe intro!: mult_lt_mono1)
apply (erule natE, auto)
apply (rule not_le_iff_lt [THEN iffD1])
@@ -371,14 +371,14 @@
apply auto
done
-lemma mult_le_cancel_le1: "k \<in> nat ==> k #* m \<le> k \<longleftrightarrow> (0 < k \<longrightarrow> natify(m) \<le> 1)"
+lemma mult_le_cancel_le1: "k \<in> nat \<Longrightarrow> 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 \<longleftrightarrow> (m \<le> n & n \<le> m)"
+lemma Ord_eq_iff_le: "\<lbrakk>Ord(m); Ord(n)\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m \<le> n & n \<le> m)"
by (blast intro: le_anti_sym)
lemma mult_cancel2_lemma:
- "[| k: nat; m: nat; n: nat |] ==> (m#*k = n#*k) \<longleftrightarrow> (m=n | k=0)"
+ "\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k = n#*k) \<longleftrightarrow> (m=n | k=0)"
apply (simp (no_asm_simp) add: Ord_eq_iff_le [of "m#*k"] Ord_eq_iff_le [of m])
apply (auto simp add: Ord_0_lt_iff)
done
@@ -398,7 +398,7 @@
(** Cancellation law for division **)
lemma div_cancel_raw:
- "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n"
+ "\<lbrakk>0<n; 0<k; k:nat; m:nat; n:nat\<rbrakk> \<Longrightarrow> (k#*m) div (k#*n) = m div n"
apply (erule_tac i = m in complete_induct)
apply (case_tac "x<n")
apply (simp add: div_less zero_lt_mult_iff mult_lt_mono2)
@@ -407,7 +407,7 @@
done
lemma div_cancel:
- "[| 0 < natify(n); 0 < natify(k) |] ==> (k#*m) div (k#*n) = m div n"
+ "\<lbrakk>0 < natify(n); 0 < natify(k)\<rbrakk> \<Longrightarrow> (k#*m) div (k#*n) = m div n"
apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)"
in div_cancel_raw)
apply auto
@@ -417,7 +417,7 @@
subsection\<open>More Lemmas about Remainder\<close>
lemma mult_mod_distrib_raw:
- "[| k:nat; m:nat; n:nat |] ==> (k#*m) mod (k#*n) = k #* (m mod n)"
+ "\<lbrakk>k:nat; m:nat; n:nat\<rbrakk> \<Longrightarrow> (k#*m) mod (k#*n) = k #* (m mod n)"
apply (case_tac "k=0")
apply (simp add: DIVISION_BY_ZERO_MOD)
apply (case_tac "n=0")
@@ -440,7 +440,7 @@
apply (simp (no_asm) add: mult_commute mod_mult_distrib2)
done
-lemma mod_add_self2_raw: "n \<in> nat ==> (m #+ n) mod n = m mod n"
+lemma mod_add_self2_raw: "n \<in> nat \<Longrightarrow> (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)
@@ -455,7 +455,7 @@
apply (simp (no_asm_simp) add: add_commute mod_add_self2)
done
-lemma mod_mult_self1_raw: "k \<in> nat ==> (m #+ k#*n) mod n = m mod n"
+lemma mod_mult_self1_raw: "k \<in> nat \<Longrightarrow> (m #+ k#*n) mod n = m mod n"
apply (erule nat_induct)
apply (simp_all (no_asm_simp) add: add_left_commute [of _ n])
done
@@ -470,7 +470,7 @@
done
(*Lemma for gcd*)
-lemma mult_eq_self_implies_10: "m = m#*n ==> natify(n)=1 | m=0"
+lemma mult_eq_self_implies_10: "m = m#*n \<Longrightarrow> natify(n)=1 | m=0"
apply (subgoal_tac "m: nat")
prefer 2
apply (erule ssubst)
@@ -486,7 +486,7 @@
done
lemma less_imp_succ_add [rule_format]:
- "[| m<n; n: nat |] ==> \<exists>k\<in>nat. n = succ(m#+k)"
+ "\<lbrakk>m<n; n: nat\<rbrakk> \<Longrightarrow> \<exists>k\<in>nat. n = succ(m#+k)"
apply (frule lt_nat_in_nat, assumption)
apply (erule rev_mp)
apply (induct_tac "n")
@@ -495,7 +495,7 @@
done
lemma less_iff_succ_add:
- "[| m: nat; n: nat |] ==> (m<n) \<longleftrightarrow> (\<exists>k\<in>nat. n = succ(m#+k))"
+ "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m<n) \<longleftrightarrow> (\<exists>k\<in>nat. n = succ(m#+k))"
by (auto intro: less_imp_succ_add)
lemma add_lt_elim2:
@@ -510,7 +510,7 @@
subsubsection\<open>More Lemmas About Difference\<close>
lemma diff_is_0_lemma:
- "[| m: nat; n: nat |] ==> m #- n = 0 \<longleftrightarrow> m \<le> n"
+ "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> m #- n = 0 \<longleftrightarrow> m \<le> n"
apply (rule_tac m = m and n = n in diff_induct, simp_all)
done
@@ -518,11 +518,11 @@
by (simp add: diff_is_0_lemma [symmetric])
lemma nat_lt_imp_diff_eq_0:
- "[| a:nat; b:nat; a<b |] ==> a #- b = 0"
+ "\<lbrakk>a:nat; b:nat; a<b\<rbrakk> \<Longrightarrow> a #- b = 0"
by (simp add: diff_is_0_iff le_iff)
lemma raw_nat_diff_split:
- "[| a:nat; b:nat |] ==>
+ "\<lbrakk>a:nat; b:nat\<rbrakk> \<Longrightarrow>
(P(a #- b)) \<longleftrightarrow> ((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)
@@ -540,7 +540,7 @@
text\<open>Difference and less-than\<close>
-lemma diff_lt_imp_lt: "[|(k#-i) < (k#-j); i\<in>nat; j\<in>nat; k\<in>nat|] ==> j<i"
+lemma diff_lt_imp_lt: "\<lbrakk>(k#-i) < (k#-j); i\<in>nat; j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> j<i"
apply (erule rev_mp)
apply (simp split: nat_diff_split, auto)
apply (blast intro: add_le_self lt_trans1)
@@ -549,7 +549,7 @@
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: "\<lbrakk>j<i; i\<le>k; k\<in>nat\<rbrakk> \<Longrightarrow> (k#-i) < (k#-j)"
apply (frule le_in_nat, assumption)
apply (frule lt_nat_in_nat, assumption)
apply (simp split: nat_diff_split, auto)
@@ -561,7 +561,7 @@
done
-lemma diff_lt_iff_lt: "[|i\<le>k; j\<in>nat; k\<in>nat|] ==> (k#-i) < (k#-j) \<longleftrightarrow> j<i"
+lemma diff_lt_iff_lt: "\<lbrakk>i\<le>k; j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> (k#-i) < (k#-j) \<longleftrightarrow> j<i"
apply (frule le_in_nat, assumption)
apply (blast intro: lt_imp_diff_lt diff_lt_imp_lt)
done
--- a/src/ZF/Bin.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Bin.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,7 +10,7 @@
the numerical interpretation.
The representation expects that (m mod 2) is 0 or 1, even if m is negative;
-For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1
+For instance, \<not>5 div 2 = \<not>3 and \<not>5 mod 2 = 1; thus \<not>5 = (\<not>3)*2 + 1
*)
section\<open>Arithmetic on Binary Integers\<close>
@@ -88,7 +88,7 @@
definition
bin_add :: "[i,i]=>i" where
- "bin_add(v,w) == bin_adder(v)`w"
+ "bin_add(v,w) \<equiv> bin_adder(v)`w"
primrec
@@ -144,26 +144,26 @@
(** Type checking **)
-lemma integ_of_type [TC]: "w \<in> bin ==> integ_of(w) \<in> int"
+lemma integ_of_type [TC]: "w \<in> bin \<Longrightarrow> integ_of(w) \<in> int"
apply (induct_tac "w")
apply (simp_all add: bool_into_nat)
done
-lemma NCons_type [TC]: "[| w \<in> bin; b \<in> bool |] ==> NCons(w,b) \<in> bin"
+lemma NCons_type [TC]: "\<lbrakk>w \<in> bin; b \<in> bool\<rbrakk> \<Longrightarrow> NCons(w,b) \<in> bin"
by (induct_tac "w", auto)
-lemma bin_succ_type [TC]: "w \<in> bin ==> bin_succ(w) \<in> bin"
+lemma bin_succ_type [TC]: "w \<in> bin \<Longrightarrow> bin_succ(w) \<in> bin"
by (induct_tac "w", auto)
-lemma bin_pred_type [TC]: "w \<in> bin ==> bin_pred(w) \<in> bin"
+lemma bin_pred_type [TC]: "w \<in> bin \<Longrightarrow> bin_pred(w) \<in> bin"
by (induct_tac "w", auto)
-lemma bin_minus_type [TC]: "w \<in> bin ==> bin_minus(w) \<in> bin"
+lemma bin_minus_type [TC]: "w \<in> bin \<Longrightarrow> bin_minus(w) \<in> bin"
by (induct_tac "w", auto)
(*This proof is complicated by the mutual recursion*)
lemma bin_add_type [rule_format]:
- "v \<in> bin ==> \<forall>w\<in>bin. bin_add(v,w) \<in> bin"
+ "v \<in> bin \<Longrightarrow> \<forall>w\<in>bin. bin_add(v,w) \<in> bin"
apply (unfold bin_add_def)
apply (induct_tac "v")
apply (rule_tac [3] ballI)
@@ -174,7 +174,7 @@
declare bin_add_type [TC]
-lemma bin_mult_type [TC]: "[| v \<in> bin; w \<in> bin |] ==> bin_mult(v,w) \<in> bin"
+lemma bin_mult_type [TC]: "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk> \<Longrightarrow> bin_mult(v,w) \<in> bin"
by (induct_tac "v", auto)
@@ -183,19 +183,19 @@
(*NCons preserves the integer value of its argument*)
lemma integ_of_NCons [simp]:
- "[| w \<in> bin; b \<in> bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)"
+ "\<lbrakk>w \<in> bin; b \<in> bool\<rbrakk> \<Longrightarrow> integ_of(NCons(w,b)) = integ_of(w BIT b)"
apply (erule bin.cases)
apply (auto elim!: boolE)
done
lemma integ_of_succ [simp]:
- "w \<in> bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)"
+ "w \<in> bin \<Longrightarrow> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)"
apply (erule bin.induct)
apply (auto simp add: zadd_ac elim!: boolE)
done
lemma integ_of_pred [simp]:
- "w \<in> bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)"
+ "w \<in> bin \<Longrightarrow> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)"
apply (erule bin.induct)
apply (auto simp add: zadd_ac elim!: boolE)
done
@@ -203,7 +203,7 @@
subsubsection\<open>\<^term>\<open>bin_minus\<close>: Unary Negation of Binary Integers\<close>
-lemma integ_of_minus: "w \<in> bin ==> integ_of(bin_minus(w)) = $- integ_of(w)"
+lemma integ_of_minus: "w \<in> bin \<Longrightarrow> integ_of(bin_minus(w)) = $- integ_of(w)"
apply (erule bin.induct)
apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE)
done
@@ -211,18 +211,18 @@
subsubsection\<open>\<^term>\<open>bin_add\<close>: Binary Addition\<close>
-lemma bin_add_Pls [simp]: "w \<in> bin ==> bin_add(Pls,w) = w"
+lemma bin_add_Pls [simp]: "w \<in> bin \<Longrightarrow> bin_add(Pls,w) = w"
by (unfold bin_add_def, simp)
-lemma bin_add_Pls_right: "w \<in> bin ==> bin_add(w,Pls) = w"
+lemma bin_add_Pls_right: "w \<in> bin \<Longrightarrow> bin_add(w,Pls) = w"
apply (unfold bin_add_def)
apply (erule bin.induct, auto)
done
-lemma bin_add_Min [simp]: "w \<in> bin ==> bin_add(Min,w) = bin_pred(w)"
+lemma bin_add_Min [simp]: "w \<in> bin \<Longrightarrow> bin_add(Min,w) = bin_pred(w)"
by (unfold bin_add_def, simp)
-lemma bin_add_Min_right: "w \<in> bin ==> bin_add(w,Min) = bin_pred(w)"
+lemma bin_add_Min_right: "w \<in> bin \<Longrightarrow> bin_add(w,Min) = bin_pred(w)"
apply (unfold bin_add_def)
apply (erule bin.induct, auto)
done
@@ -234,13 +234,13 @@
by (unfold bin_add_def, simp)
lemma bin_add_BIT_BIT [simp]:
- "[| w \<in> bin; y \<in> bool |]
- ==> bin_add(v BIT x, w BIT y) =
+ "\<lbrakk>w \<in> bin; y \<in> bool\<rbrakk>
+ \<Longrightarrow> 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 \<in> bin ==>
+ "v \<in> bin \<Longrightarrow>
\<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)
@@ -250,8 +250,8 @@
(*Subtraction*)
lemma diff_integ_of_eq:
- "[| v \<in> bin; w \<in> bin |]
- ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))"
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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)
done
@@ -260,8 +260,8 @@
subsubsection\<open>\<^term>\<open>bin_mult\<close>: Binary Multiplication\<close>
lemma integ_of_mult:
- "[| v \<in> bin; w \<in> bin |]
- ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)"
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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)
@@ -294,16 +294,16 @@
(** extra rules for bin_add **)
-lemma bin_add_BIT_11: "w \<in> bin ==> bin_add(v BIT 1, w BIT 1) =
+lemma bin_add_BIT_11: "w \<in> bin \<Longrightarrow> bin_add(v BIT 1, w BIT 1) =
NCons(bin_add(v, bin_succ(w)), 0)"
by simp
-lemma bin_add_BIT_10: "w \<in> bin ==> bin_add(v BIT 1, w BIT 0) =
+lemma bin_add_BIT_10: "w \<in> bin \<Longrightarrow> bin_add(v BIT 1, w BIT 0) =
NCons(bin_add(v,w), 1)"
by simp
-lemma bin_add_BIT_0: "[| w \<in> bin; y \<in> bool |]
- ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)"
+lemma bin_add_BIT_0: "\<lbrakk>w \<in> bin; y \<in> bool\<rbrakk>
+ \<Longrightarrow> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)"
by simp
(** extra rules for bin_mult **)
@@ -359,8 +359,8 @@
(** Equals (=) **)
lemma eq_integ_of_eq:
- "[| v \<in> bin; w \<in> bin |]
- ==> ((integ_of(v)) = integ_of(w)) \<longleftrightarrow>
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> ((integ_of(v)) = integ_of(w)) \<longleftrightarrow>
iszero (integ_of (bin_add (v, bin_minus(w))))"
apply (unfold iszero_def)
apply (simp add: zcompare_rls integ_of_add integ_of_minus)
@@ -370,14 +370,14 @@
by (unfold iszero_def, simp)
-lemma nonzero_integ_of_Min: "~ iszero (integ_of(Min))"
+lemma nonzero_integ_of_Min: "\<not> iszero (integ_of(Min))"
apply (unfold iszero_def)
apply (simp add: zminus_equation)
done
lemma iszero_integ_of_BIT:
- "[| w \<in> bin; x \<in> bool |]
- ==> iszero (integ_of (w BIT x)) \<longleftrightarrow> (x=0 & iszero (integ_of(w)))"
+ "\<lbrakk>w \<in> bin; x \<in> bool\<rbrakk>
+ \<Longrightarrow> iszero (integ_of (w BIT x)) \<longleftrightarrow> (x=0 & iszero (integ_of(w)))"
apply (unfold iszero_def, simp)
apply (subgoal_tac "integ_of (w) \<in> int")
apply typecheck
@@ -388,10 +388,10 @@
done
lemma iszero_integ_of_0:
- "w \<in> bin ==> iszero (integ_of (w BIT 0)) \<longleftrightarrow> iszero (integ_of(w))"
+ "w \<in> bin \<Longrightarrow> iszero (integ_of (w BIT 0)) \<longleftrightarrow> iszero (integ_of(w))"
by (simp only: iszero_integ_of_BIT, blast)
-lemma iszero_integ_of_1: "w \<in> bin ==> ~ iszero (integ_of (w BIT 1))"
+lemma iszero_integ_of_1: "w \<in> bin \<Longrightarrow> \<not> iszero (integ_of (w BIT 1))"
by (simp only: iszero_integ_of_BIT, blast)
@@ -399,22 +399,22 @@
(** Less-than (<) **)
lemma less_integ_of_eq_neg:
- "[| v \<in> bin; w \<in> bin |]
- ==> integ_of(v) $< integ_of(w)
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> integ_of(v) $< integ_of(w)
\<longleftrightarrow> znegative (integ_of (bin_add (v, bin_minus(w))))"
apply (unfold zless_def zdiff_def)
apply (simp add: integ_of_minus integ_of_add)
done
-lemma not_neg_integ_of_Pls: "~ znegative (integ_of(Pls))"
+lemma not_neg_integ_of_Pls: "\<not> znegative (integ_of(Pls))"
by simp
lemma neg_integ_of_Min: "znegative (integ_of(Min))"
by simp
lemma neg_integ_of_BIT:
- "[| w \<in> bin; x \<in> bool |]
- ==> znegative (integ_of (w BIT x)) \<longleftrightarrow> znegative (integ_of(w))"
+ "\<lbrakk>w \<in> bin; x \<in> bool\<rbrakk>
+ \<Longrightarrow> znegative (integ_of (w BIT x)) \<longleftrightarrow> znegative (integ_of(w))"
apply simp
apply (subgoal_tac "integ_of (w) \<in> int")
apply typecheck
@@ -430,7 +430,7 @@
(** Less-than-or-equals (<=) **)
lemma le_integ_of_eq_not_less:
- "(integ_of(x) $\<le> (integ_of(w))) \<longleftrightarrow> ~ (integ_of(w) $< (integ_of(x)))"
+ "(integ_of(x) $\<le> (integ_of(w))) \<longleftrightarrow> \<not> (integ_of(w) $< (integ_of(x)))"
by (simp add: not_zless_iff_zle [THEN iff_sym])
@@ -461,7 +461,7 @@
(*For making a minimal simpset, one must include these default simprules
- of thy. Also include simp_thms, or at least (~False)=True*)
+ of thy. Also include simp_thms, or at least (\<not>False)=True*)
lemmas bin_arith_simps =
bin_pred_Pls bin_pred_Min
bin_succ_Pls bin_succ_Min
@@ -485,25 +485,25 @@
(** Simplification of arithmetic when nested to the right **)
lemma add_integ_of_left [simp]:
- "[| v \<in> bin; w \<in> bin |]
- ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)"
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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 \<in> bin; w \<in> bin |]
- ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)"
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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 \<in> bin; w \<in> bin |]
- ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)"
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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 \<in> bin; w \<in> bin |]
- ==> integ_of(v) $+ (c $- integ_of(w)) =
+ "\<lbrakk>v \<in> bin; w \<in> bin\<rbrakk>
+ \<Longrightarrow> 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)
@@ -523,10 +523,10 @@
lemma zdiff_self [simp]: "x $- x = #0"
by (simp add: zdiff_def)
-lemma znegative_iff_zless_0: "k \<in> int ==> znegative(k) \<longleftrightarrow> k $< #0"
+lemma znegative_iff_zless_0: "k \<in> int \<Longrightarrow> znegative(k) \<longleftrightarrow> k $< #0"
by (simp add: zless_def)
-lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k \<in> int|] ==> znegative($-k)"
+lemma zero_zless_imp_znegative_zminus: "\<lbrakk>#0 $< k; k \<in> int\<rbrakk> \<Longrightarrow> znegative($-k)"
by (simp add: zless_def)
lemma zero_zle_int_of [simp]: "#0 $\<le> $# n"
@@ -535,21 +535,21 @@
lemma nat_of_0 [simp]: "nat_of(#0) = 0"
by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of)
-lemma nat_le_int0_lemma: "[| z $\<le> $#0; z \<in> int |] ==> nat_of(z) = 0"
+lemma nat_le_int0_lemma: "\<lbrakk>z $\<le> $#0; z \<in> int\<rbrakk> \<Longrightarrow> nat_of(z) = 0"
by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of)
-lemma nat_le_int0: "z $\<le> $#0 ==> nat_of(z) = 0"
+lemma nat_le_int0: "z $\<le> $#0 \<Longrightarrow> nat_of(z) = 0"
apply (subgoal_tac "nat_of (intify (z)) = 0")
apply (rule_tac [2] nat_le_int0_lemma, auto)
done
-lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 ==> natify(n) = 0"
+lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 \<Longrightarrow> natify(n) = 0"
by (rule not_znegative_imp_zero, auto)
lemma nat_of_zminus_int_of: "nat_of($- $# n) = 0"
by (simp add: nat_of_def int_of_def raw_nat_of zminus image_intrel_int)
-lemma int_of_nat_of: "#0 $\<le> z ==> $# nat_of(z) = intify(z)"
+lemma int_of_nat_of: "#0 $\<le> z \<Longrightarrow> $# nat_of(z) = intify(z)"
apply (rule not_zneg_nat_of_intify)
apply (simp add: znegative_iff_zless_0 not_zless_iff_zle)
done
@@ -559,7 +559,7 @@
lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $\<le> z then intify(z) else #0)"
by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless)
-lemma zless_nat_iff_int_zless: "[| m \<in> nat; z \<in> int |] ==> (m < nat_of(z)) \<longleftrightarrow> ($#m $< z)"
+lemma zless_nat_iff_int_zless: "\<lbrakk>m \<in> nat; z \<in> int\<rbrakk> \<Longrightarrow> (m < nat_of(z)) \<longleftrightarrow> ($#m $< z)"
apply (case_tac "znegative (z) ")
apply (erule_tac [2] not_zneg_nat_of [THEN subst])
apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans]
@@ -570,7 +570,7 @@
(** nat_of and zless **)
(*An alternative condition is @{term"$#0 \<subseteq> w"} *)
-lemma zless_nat_conj_lemma: "$#0 $< z ==> (nat_of(w) < nat_of(z)) \<longleftrightarrow> (w $< z)"
+lemma zless_nat_conj_lemma: "$#0 $< z \<Longrightarrow> (nat_of(w) < nat_of(z)) \<longleftrightarrow> (w $< z)"
apply (rule iff_trans)
apply (rule zless_int_of [THEN iff_sym])
apply (auto simp add: int_of_nat_of_if simp del: zless_int_of)
@@ -589,7 +589,7 @@
Conditional rewrite rules are tried after unconditional ones, so a rule
like eq_nat_number_of will be tried first to eliminate #mm=#nn.]
lemma integ_of_reorient [simp]:
- "True ==> (integ_of(w) = x) \<longleftrightarrow> (x = integ_of(w))"
+ "True \<Longrightarrow> (integ_of(w) = x) \<longleftrightarrow> (x = integ_of(w))"
by auto
*)
@@ -639,7 +639,7 @@
lemma zless_iff_zdiff_zless_0: "(x $< y) \<longleftrightarrow> (x$-y $< #0)"
by (simp add: zcompare_rls)
-lemma eq_iff_zdiff_eq_0: "[| x \<in> int; y \<in> int |] ==> (x = y) \<longleftrightarrow> (x$-y = #0)"
+lemma eq_iff_zdiff_eq_0: "\<lbrakk>x \<in> int; y \<in> int\<rbrakk> \<Longrightarrow> (x = y) \<longleftrightarrow> (x$-y = #0)"
by (simp add: zcompare_rls)
lemma zle_iff_zdiff_zle_0: "(x $\<le> y) \<longleftrightarrow> (x$-y $\<le> #0)"
@@ -709,15 +709,15 @@
schematic_goal "y2 $+ ?x42 = y $+ y2" apply simp oops
-lemma "oo : int ==> l $+ (l $+ #2) $+ oo = oo" apply simp oops
+lemma "oo : int \<Longrightarrow> l $+ (l $+ #2) $+ oo = oo" apply simp oops
lemma "#9$*x $+ y = x$*#23 $+ z" apply simp oops
lemma "y $+ x = x $+ z" apply simp oops
-lemma "x : int ==> x $+ y $+ z = x $+ z" apply simp oops
-lemma "x : int ==> y $+ (z $+ x) = z $+ x" apply simp oops
-lemma "z : int ==> x $+ y $+ z = (z $+ y) $+ (x $+ w)" apply simp oops
-lemma "z : int ==> x$*y $+ z = (z $+ y) $+ (y$*x $+ w)" apply simp oops
+lemma "x : int \<Longrightarrow> x $+ y $+ z = x $+ z" apply simp oops
+lemma "x : int \<Longrightarrow> y $+ (z $+ x) = z $+ x" apply simp oops
+lemma "z : int \<Longrightarrow> x $+ y $+ z = (z $+ y) $+ (x $+ w)" apply simp oops
+lemma "z : int \<Longrightarrow> x$*y $+ z = (z $+ y) $+ (y$*x $+ w)" apply simp oops
lemma "#-3 $* x $+ y $\<le> x $* #2 $+ z" apply simp oops
lemma "y $+ x $\<le> x $+ z" apply simp oops
@@ -728,7 +728,7 @@
lemma "x$*y $+ z $< (z $+ y) $+ (y$*x $+ w)" apply simp oops
lemma "l $+ #2 $+ #2 $+ #2 $+ (l $+ #2) $+ (oo $+ #2) = uu" apply simp oops
-lemma "u : int ==> #2 $* u = u" apply simp oops
+lemma "u : int \<Longrightarrow> #2 $* u = u" apply simp oops
lemma "(i $+ j $+ #12 $+ k) $- #15 = y" apply simp oops
lemma "(i $+ j $+ #12 $+ k) $- #5 = y" apply simp oops
--- a/src/ZF/Bool.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Bool.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,31 +9,31 @@
abbreviation
one (\<open>1\<close>) where
- "1 == succ(0)"
+ "1 \<equiv> succ(0)"
abbreviation
two (\<open>2\<close>) where
- "2 == succ(1)"
+ "2 \<equiv> succ(1)"
text\<open>2 is equal to bool, but is used as a number rather than a type.\<close>
-definition "bool == {0,1}"
+definition "bool \<equiv> {0,1}"
-definition "cond(b,c,d) == if(b=1,c,d)"
+definition "cond(b,c,d) \<equiv> if(b=1,c,d)"
-definition "not(b) == cond(b,0,1)"
+definition "not(b) \<equiv> cond(b,0,1)"
definition
"and" :: "[i,i]=>i" (infixl \<open>and\<close> 70) where
- "a and b == cond(a,b,0)"
+ "a and b \<equiv> cond(a,b,0)"
definition
or :: "[i,i]=>i" (infixl \<open>or\<close> 65) where
- "a or b == cond(a,1,b)"
+ "a or b \<equiv> cond(a,1,b)"
definition
xor :: "[i,i]=>i" (infixl \<open>xor\<close> 65) where
- "a xor b == cond(a,not(b),b)"
+ "a xor b \<equiv> cond(a,not(b),b)"
lemmas bool_defs = bool_def cond_def
@@ -52,11 +52,11 @@
lemma one_not_0: "1\<noteq>0"
by (simp add: bool_defs )
-(** 1=0 ==> R **)
+(** 1=0 \<Longrightarrow> R **)
lemmas one_neq_0 = one_not_0 [THEN notE]
lemma boolE:
- "[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P"
+ "\<lbrakk>c: bool; c=1 \<Longrightarrow> P; c=0 \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp add: bool_defs, blast)
(** cond **)
@@ -69,17 +69,17 @@
lemma cond_0 [simp]: "cond(0,c,d) = d"
by (simp add: bool_defs )
-lemma cond_type [TC]: "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"
+lemma cond_type [TC]: "\<lbrakk>b: bool; c: A(1); d: A(0)\<rbrakk> \<Longrightarrow> cond(b,c,d): A(b)"
by (simp add: bool_defs, blast)
(*For Simp_tac and Blast_tac*)
-lemma cond_simple_type: "[| b: bool; c: A; d: A |] ==> cond(b,c,d): A"
+lemma cond_simple_type: "\<lbrakk>b: bool; c: A; d: A\<rbrakk> \<Longrightarrow> cond(b,c,d): A"
by (simp add: bool_defs )
-lemma def_cond_1: "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"
+lemma def_cond_1: "\<lbrakk>\<And>b. j(b)\<equiv>cond(b,c,d)\<rbrakk> \<Longrightarrow> j(1) = c"
by simp
-lemma def_cond_0: "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"
+lemma def_cond_0: "\<lbrakk>\<And>b. j(b)\<equiv>cond(b,c,d)\<rbrakk> \<Longrightarrow> j(0) = d"
by simp
lemmas not_1 = not_def [THEN def_cond_1, simp]
@@ -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) \<in> bool"
+lemma not_type [TC]: "a:bool \<Longrightarrow> not(a) \<in> bool"
by (simp add: not_def)
-lemma and_type [TC]: "[| a:bool; b:bool |] ==> a and b \<in> bool"
+lemma and_type [TC]: "\<lbrakk>a:bool; b:bool\<rbrakk> \<Longrightarrow> a and b \<in> bool"
by (simp add: and_def)
-lemma or_type [TC]: "[| a:bool; b:bool |] ==> a or b \<in> bool"
+lemma or_type [TC]: "\<lbrakk>a:bool; b:bool\<rbrakk> \<Longrightarrow> a or b \<in> bool"
by (simp add: or_def)
-lemma xor_type [TC]: "[| a:bool; b:bool |] ==> a xor b \<in> bool"
+lemma xor_type [TC]: "\<lbrakk>a:bool; b:bool\<rbrakk> \<Longrightarrow> a xor b \<in> bool"
by (simp add: xor_def)
lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type
@@ -111,49 +111,49 @@
subsection\<open>Laws About 'not'\<close>
-lemma not_not [simp]: "a:bool ==> not(not(a)) = a"
+lemma not_not [simp]: "a:bool \<Longrightarrow> not(not(a)) = a"
by (elim boolE, auto)
-lemma not_and [simp]: "a:bool ==> not(a and b) = not(a) or not(b)"
+lemma not_and [simp]: "a:bool \<Longrightarrow> not(a and b) = not(a) or not(b)"
by (elim boolE, auto)
-lemma not_or [simp]: "a:bool ==> not(a or b) = not(a) and not(b)"
+lemma not_or [simp]: "a:bool \<Longrightarrow> not(a or b) = not(a) and not(b)"
by (elim boolE, auto)
subsection\<open>Laws About 'and'\<close>
-lemma and_absorb [simp]: "a: bool ==> a and a = a"
+lemma and_absorb [simp]: "a: bool \<Longrightarrow> a and a = a"
by (elim boolE, auto)
-lemma and_commute: "[| a: bool; b:bool |] ==> a and b = b and a"
+lemma and_commute: "\<lbrakk>a: bool; b:bool\<rbrakk> \<Longrightarrow> a and b = b and a"
by (elim boolE, auto)
-lemma and_assoc: "a: bool ==> (a and b) and c = a and (b and c)"
+lemma and_assoc: "a: bool \<Longrightarrow> (a and b) and c = a and (b and c)"
by (elim boolE, auto)
-lemma and_or_distrib: "[| a: bool; b:bool; c:bool |] ==>
+lemma and_or_distrib: "\<lbrakk>a: bool; b:bool; c:bool\<rbrakk> \<Longrightarrow>
(a or b) and c = (a and c) or (b and c)"
by (elim boolE, auto)
subsection\<open>Laws About 'or'\<close>
-lemma or_absorb [simp]: "a: bool ==> a or a = a"
+lemma or_absorb [simp]: "a: bool \<Longrightarrow> a or a = a"
by (elim boolE, auto)
-lemma or_commute: "[| a: bool; b:bool |] ==> a or b = b or a"
+lemma or_commute: "\<lbrakk>a: bool; b:bool\<rbrakk> \<Longrightarrow> a or b = b or a"
by (elim boolE, auto)
-lemma or_assoc: "a: bool ==> (a or b) or c = a or (b or c)"
+lemma or_assoc: "a: bool \<Longrightarrow> (a or b) or c = a or (b or c)"
by (elim boolE, auto)
-lemma or_and_distrib: "[| a: bool; b: bool; c: bool |] ==>
+lemma or_and_distrib: "\<lbrakk>a: bool; b: bool; c: bool\<rbrakk> \<Longrightarrow>
(a and b) or c = (a or c) and (b or c)"
by (elim boolE, auto)
definition
bool_of_o :: "o=>i" where
- "bool_of_o(P) == (if P then 1 else 0)"
+ "bool_of_o(P) \<equiv> (if P then 1 else 0)"
lemma [simp]: "bool_of_o(True) = 1"
by (simp add: bool_of_o_def)
@@ -167,7 +167,7 @@
lemma [simp]: "(bool_of_o(P) = 1) \<longleftrightarrow> P"
by (simp add: bool_of_o_def)
-lemma [simp]: "(bool_of_o(P) = 0) \<longleftrightarrow> ~P"
+lemma [simp]: "(bool_of_o(P) = 0) \<longleftrightarrow> \<not>P"
by (simp add: bool_of_o_def)
end
--- a/src/ZF/Cardinal.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Cardinal.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,31 +10,31 @@
definition
(*least ordinal operator*)
Least :: "(i=>o) => i" (binder \<open>\<mu> \<close> 10) where
- "Least(P) == THE i. Ord(i) & P(i) & (\<forall>j. j<i \<longrightarrow> ~P(j))"
+ "Least(P) \<equiv> THE i. Ord(i) & P(i) & (\<forall>j. j<i \<longrightarrow> \<not>P(j))"
definition
eqpoll :: "[i,i] => o" (infixl \<open>\<approx>\<close> 50) where
- "A \<approx> B == \<exists>f. f \<in> bij(A,B)"
+ "A \<approx> B \<equiv> \<exists>f. f \<in> bij(A,B)"
definition
lepoll :: "[i,i] => o" (infixl \<open>\<lesssim>\<close> 50) where
- "A \<lesssim> B == \<exists>f. f \<in> inj(A,B)"
+ "A \<lesssim> B \<equiv> \<exists>f. f \<in> inj(A,B)"
definition
lesspoll :: "[i,i] => o" (infixl \<open>\<prec>\<close> 50) where
- "A \<prec> B == A \<lesssim> B & ~(A \<approx> B)"
+ "A \<prec> B \<equiv> A \<lesssim> B & \<not>(A \<approx> B)"
definition
cardinal :: "i=>i" (\<open>|_|\<close>) where
- "|A| == (\<mu> i. i \<approx> A)"
+ "|A| \<equiv> (\<mu> i. i \<approx> A)"
definition
Finite :: "i=>o" where
- "Finite(A) == \<exists>n\<in>nat. A \<approx> n"
+ "Finite(A) \<equiv> \<exists>n\<in>nat. A \<approx> n"
definition
Card :: "i=>o" where
- "Card(i) == (i = |i|)"
+ "Card(i) \<equiv> (i = |i|)"
subsection\<open>The Schroeder-Bernstein Theorem\<close>
@@ -47,7 +47,7 @@
lemma Banach_last_equation:
"g \<in> Y->X
- ==> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) =
+ \<Longrightarrow> 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" for v
in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst])
@@ -55,7 +55,7 @@
done
lemma decomposition:
- "[| f \<in> X->Y; g \<in> Y->X |] ==>
+ "\<lbrakk>f \<in> X->Y; g \<in> Y->X\<rbrakk> \<Longrightarrow>
\<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"
@@ -67,18 +67,18 @@
done
lemma schroeder_bernstein:
- "[| f \<in> inj(X,Y); g \<in> inj(Y,X) |] ==> \<exists>h. h \<in> bij(X,Y)"
+ "\<lbrakk>f \<in> inj(X,Y); g \<in> inj(Y,X)\<rbrakk> \<Longrightarrow> \<exists>h. h \<in> 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 @{term"restrict(f,XA) \<union> converse(restrict(g,YB))"}
- is forced by the context!! *)
+ is forced by the context\<And>*)
done
(** Equipollence is an equivalence relation **)
-lemma bij_imp_eqpoll: "f \<in> bij(A,B) ==> A \<approx> B"
+lemma bij_imp_eqpoll: "f \<in> bij(A,B) \<Longrightarrow> A \<approx> B"
apply (unfold eqpoll_def)
apply (erule exI)
done
@@ -86,20 +86,20 @@
(*A \<approx> A*)
lemmas eqpoll_refl = id_bij [THEN bij_imp_eqpoll, simp]
-lemma eqpoll_sym: "X \<approx> Y ==> Y \<approx> X"
+lemma eqpoll_sym: "X \<approx> Y \<Longrightarrow> Y \<approx> X"
apply (unfold eqpoll_def)
apply (blast intro: bij_converse_bij)
done
lemma eqpoll_trans [trans]:
- "[| X \<approx> Y; Y \<approx> Z |] ==> X \<approx> Z"
+ "\<lbrakk>X \<approx> Y; Y \<approx> Z\<rbrakk> \<Longrightarrow> X \<approx> Z"
apply (unfold eqpoll_def)
apply (blast intro: comp_bij)
done
(** Le-pollence is a partial ordering **)
-lemma subset_imp_lepoll: "X<=Y ==> X \<lesssim> Y"
+lemma subset_imp_lepoll: "X<=Y \<Longrightarrow> X \<lesssim> Y"
apply (unfold lepoll_def)
apply (rule exI)
apply (erule id_subset_inj)
@@ -109,35 +109,35 @@
lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll]
-lemma eqpoll_imp_lepoll: "X \<approx> Y ==> X \<lesssim> Y"
+lemma eqpoll_imp_lepoll: "X \<approx> Y \<Longrightarrow> X \<lesssim> Y"
by (unfold eqpoll_def bij_def lepoll_def, blast)
-lemma lepoll_trans [trans]: "[| X \<lesssim> Y; Y \<lesssim> Z |] ==> X \<lesssim> Z"
+lemma lepoll_trans [trans]: "\<lbrakk>X \<lesssim> Y; Y \<lesssim> Z\<rbrakk> \<Longrightarrow> X \<lesssim> Z"
apply (unfold lepoll_def)
apply (blast intro: comp_inj)
done
-lemma eq_lepoll_trans [trans]: "[| X \<approx> Y; Y \<lesssim> Z |] ==> X \<lesssim> Z"
+lemma eq_lepoll_trans [trans]: "\<lbrakk>X \<approx> Y; Y \<lesssim> Z\<rbrakk> \<Longrightarrow> X \<lesssim> Z"
by (blast intro: eqpoll_imp_lepoll lepoll_trans)
-lemma lepoll_eq_trans [trans]: "[| X \<lesssim> Y; Y \<approx> Z |] ==> X \<lesssim> Z"
+lemma lepoll_eq_trans [trans]: "\<lbrakk>X \<lesssim> Y; Y \<approx> Z\<rbrakk> \<Longrightarrow> X \<lesssim> Z"
by (blast intro: eqpoll_imp_lepoll lepoll_trans)
(*Asymmetry law*)
-lemma eqpollI: "[| X \<lesssim> Y; Y \<lesssim> X |] ==> X \<approx> Y"
+lemma eqpollI: "\<lbrakk>X \<lesssim> Y; Y \<lesssim> X\<rbrakk> \<Longrightarrow> X \<approx> Y"
apply (unfold lepoll_def eqpoll_def)
apply (elim exE)
apply (rule schroeder_bernstein, assumption+)
done
lemma eqpollE:
- "[| X \<approx> Y; [| X \<lesssim> Y; Y \<lesssim> X |] ==> P |] ==> P"
+ "\<lbrakk>X \<approx> Y; \<lbrakk>X \<lesssim> Y; Y \<lesssim> X\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (blast intro: eqpoll_imp_lepoll eqpoll_sym)
lemma eqpoll_iff: "X \<approx> Y \<longleftrightarrow> X \<lesssim> Y & Y \<lesssim> X"
by (blast intro: eqpollI elim!: eqpollE)
-lemma lepoll_0_is_0: "A \<lesssim> 0 ==> A = 0"
+lemma lepoll_0_is_0: "A \<lesssim> 0 \<Longrightarrow> A = 0"
apply (unfold lepoll_def inj_def)
apply (blast dest: apply_type)
done
@@ -149,20 +149,20 @@
by (blast intro: lepoll_0_is_0 lepoll_refl)
lemma Un_lepoll_Un:
- "[| A \<lesssim> B; C \<lesssim> D; B \<inter> D = 0 |] ==> A \<union> C \<lesssim> B \<union> D"
+ "\<lbrakk>A \<lesssim> B; C \<lesssim> D; B \<inter> D = 0\<rbrakk> \<Longrightarrow> A \<union> C \<lesssim> B \<union> D"
apply (unfold lepoll_def)
apply (blast intro: inj_disjoint_Un)
done
-(*A \<approx> 0 ==> A=0*)
+(*A \<approx> 0 \<Longrightarrow> A=0*)
lemmas eqpoll_0_is_0 = eqpoll_imp_lepoll [THEN lepoll_0_is_0]
lemma eqpoll_0_iff: "A \<approx> 0 \<longleftrightarrow> A=0"
by (blast intro: eqpoll_0_is_0 eqpoll_refl)
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"
+ "\<lbrakk>A \<approx> B; C \<approx> D; A \<inter> C = 0; B \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> A \<union> C \<approx> B \<union> D"
apply (unfold eqpoll_def)
apply (blast intro: bij_disjoint_Un)
done
@@ -170,16 +170,16 @@
subsection\<open>lesspoll: contributions by Krzysztof Grabczewski\<close>
-lemma lesspoll_not_refl: "~ (i \<prec> i)"
+lemma lesspoll_not_refl: "\<not> (i \<prec> i)"
by (simp add: lesspoll_def)
-lemma lesspoll_irrefl [elim!]: "i \<prec> i ==> P"
+lemma lesspoll_irrefl [elim!]: "i \<prec> i \<Longrightarrow> P"
by (simp add: lesspoll_def)
-lemma lesspoll_imp_lepoll: "A \<prec> B ==> A \<lesssim> B"
+lemma lesspoll_imp_lepoll: "A \<prec> B \<Longrightarrow> A \<lesssim> B"
by (unfold lesspoll_def, blast)
-lemma lepoll_well_ord: "[| A \<lesssim> B; well_ord(B,r) |] ==> \<exists>s. well_ord(A,s)"
+lemma lepoll_well_ord: "\<lbrakk>A \<lesssim> B; well_ord(B,r)\<rbrakk> \<Longrightarrow> \<exists>s. well_ord(A,s)"
apply (unfold lepoll_def)
apply (blast intro: well_ord_rvimage)
done
@@ -207,36 +207,36 @@
(** Variations on transitivity **)
lemma lesspoll_trans [trans]:
- "[| X \<prec> Y; Y \<prec> Z |] ==> X \<prec> Z"
+ "\<lbrakk>X \<prec> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
done
lemma lesspoll_trans1 [trans]:
- "[| X \<lesssim> Y; Y \<prec> Z |] ==> X \<prec> Z"
+ "\<lbrakk>X \<lesssim> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
done
lemma lesspoll_trans2 [trans]:
- "[| X \<prec> Y; Y \<lesssim> Z |] ==> X \<prec> Z"
+ "\<lbrakk>X \<prec> Y; Y \<lesssim> Z\<rbrakk> \<Longrightarrow> X \<prec> Z"
apply (unfold lesspoll_def)
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
done
lemma eq_lesspoll_trans [trans]:
- "[| X \<approx> Y; Y \<prec> Z |] ==> X \<prec> Z"
+ "\<lbrakk>X \<approx> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z"
by (blast intro: eqpoll_imp_lepoll lesspoll_trans1)
lemma lesspoll_eq_trans [trans]:
- "[| X \<prec> Y; Y \<approx> Z |] ==> X \<prec> Z"
+ "\<lbrakk>X \<prec> Y; Y \<approx> Z\<rbrakk> \<Longrightarrow> X \<prec> Z"
by (blast intro: eqpoll_imp_lepoll lesspoll_trans2)
(** \<mu> -- the least number operator [from HOL/Univ.ML] **)
lemma Least_equality:
- "[| P(i); Ord(i); !!x. x<i ==> ~P(x) |] ==> (\<mu> x. P(x)) = i"
+ "\<lbrakk>P(i); Ord(i); \<And>x. x<i \<Longrightarrow> \<not>P(x)\<rbrakk> \<Longrightarrow> (\<mu> x. P(x)) = i"
apply (unfold Least_def)
apply (rule the_equality, blast)
apply (elim conjE)
@@ -254,7 +254,7 @@
case True thus ?thesis .
next
case False
- hence "\<And>x. x \<in> i \<Longrightarrow> ~P(x)" using step
+ hence "\<And>x. x \<in> i \<Longrightarrow> \<not>P(x)" using step
by blast
hence "(\<mu> a. P(a)) = i" using step
by (blast intro: Least_equality ltD)
@@ -278,7 +278,7 @@
case True thus ?thesis .
next
case False
- hence "\<And>x. x \<in> i \<Longrightarrow> ~ (\<mu> a. P(a)) \<le> i" using step
+ hence "\<And>x. x \<in> i \<Longrightarrow> \<not> (\<mu> a. P(a)) \<le> i" using step
by blast
hence "(\<mu> a. P(a)) = i" using step
by (blast elim: ltE intro: ltI Least_equality lt_trans1)
@@ -291,19 +291,19 @@
qed
(*\<mu> really is the smallest*)
-lemma less_LeastE: "[| P(i); i < (\<mu> x. P(x)) |] ==> Q"
+lemma less_LeastE: "\<lbrakk>P(i); i < (\<mu> x. P(x))\<rbrakk> \<Longrightarrow> Q"
apply (rule Least_le [THEN [2] lt_trans2, THEN lt_irrefl], assumption+)
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(\<mu> j. P(j))"
+ "\<lbrakk>P(i); Ord(i); \<And>j. P(j) \<Longrightarrow> Q(j)\<rbrakk> \<Longrightarrow> Q(\<mu> j. P(j))"
by (blast intro: LeastI )
(*If there is no such P then \<mu> is vacuously 0*)
lemma Least_0:
- "[| ~ (\<exists>i. Ord(i) & P(i)) |] ==> (\<mu> x. P(x)) = 0"
+ "\<lbrakk>\<not> (\<exists>i. Ord(i) & P(i))\<rbrakk> \<Longrightarrow> (\<mu> x. P(x)) = 0"
apply (unfold Least_def)
apply (rule the_0, blast)
done
@@ -326,12 +326,12 @@
subsection\<open>Basic Properties of Cardinals\<close>
(*Not needed for simplification, but helpful below*)
-lemma Least_cong: "(!!y. P(y) \<longleftrightarrow> Q(y)) ==> (\<mu> x. P(x)) = (\<mu> x. Q(x))"
+lemma Least_cong: "(\<And>y. P(y) \<longleftrightarrow> Q(y)) \<Longrightarrow> (\<mu> x. P(x)) = (\<mu> x. Q(x))"
by simp
-(*Need AC to get @{term"X \<lesssim> Y ==> |X| \<le> |Y|"}; see well_ord_lepoll_imp_cardinal_le
+(*Need AC to get @{term"X \<lesssim> Y \<Longrightarrow> |X| \<le> |Y|"}; see well_ord_lepoll_imp_cardinal_le
Converse also requires AC, but see well_ord_cardinal_eqE*)
-lemma cardinal_cong: "X \<approx> Y ==> |X| = |Y|"
+lemma cardinal_cong: "X \<approx> Y \<Longrightarrow> |X| = |Y|"
apply (unfold eqpoll_def cardinal_def)
apply (rule Least_cong)
apply (blast intro: comp_bij bij_converse_bij)
@@ -345,7 +345,7 @@
by (best intro: LeastI Ord_ordertype ordermap_bij bij_converse_bij bij_imp_eqpoll r)
qed
-(* @{term"Ord(A) ==> |A| \<approx> A"} *)
+(* @{term"Ord(A) \<Longrightarrow> |A| \<approx> A"} *)
lemmas Ord_cardinal_eqpoll = well_ord_Memrel [THEN well_ord_cardinal_eqpoll]
lemma Ord_cardinal_idem: "Ord(A) \<Longrightarrow> ||A|| = |A|"
@@ -362,36 +362,36 @@
qed
lemma well_ord_cardinal_eqpoll_iff:
- "[| well_ord(X,r); well_ord(Y,s) |] ==> |X| = |Y| \<longleftrightarrow> X \<approx> Y"
+ "\<lbrakk>well_ord(X,r); well_ord(Y,s)\<rbrakk> \<Longrightarrow> |X| = |Y| \<longleftrightarrow> X \<approx> Y"
by (blast intro: cardinal_cong well_ord_cardinal_eqE)
(** Observations from Kunen, page 28 **)
-lemma Ord_cardinal_le: "Ord(i) ==> |i| \<le> i"
+lemma Ord_cardinal_le: "Ord(i) \<Longrightarrow> |i| \<le> i"
apply (unfold cardinal_def)
apply (erule eqpoll_refl [THEN Least_le])
done
-lemma Card_cardinal_eq: "Card(K) ==> |K| = K"
+lemma Card_cardinal_eq: "Card(K) \<Longrightarrow> |K| = K"
apply (unfold Card_def)
apply (erule sym)
done
-(* Could replace the @{term"~(j \<approx> i)"} by @{term"~(i \<preceq> j)"}. *)
-lemma CardI: "[| Ord(i); !!j. j<i ==> ~(j \<approx> i) |] ==> Card(i)"
+(* Could replace the @{term"\<not>(j \<approx> i)"} by @{term"\<not>(i \<preceq> j)"}. *)
+lemma CardI: "\<lbrakk>Ord(i); \<And>j. j<i \<Longrightarrow> \<not>(j \<approx> i)\<rbrakk> \<Longrightarrow> Card(i)"
apply (unfold Card_def cardinal_def)
apply (subst Least_equality)
apply (blast intro: eqpoll_refl)+
done
-lemma Card_is_Ord: "Card(i) ==> Ord(i)"
+lemma Card_is_Ord: "Card(i) \<Longrightarrow> Ord(i)"
apply (unfold Card_def cardinal_def)
apply (erule ssubst)
apply (rule Ord_Least)
done
-lemma Card_cardinal_le: "Card(K) ==> K \<le> |K|"
+lemma Card_cardinal_le: "Card(K) \<Longrightarrow> K \<le> |K|"
apply (simp (no_asm_simp) add: Card_is_Ord Card_cardinal_eq)
done
@@ -401,7 +401,7 @@
done
text\<open>The cardinals are the initial ordinals.\<close>
-lemma Card_iff_initial: "Card(K) \<longleftrightarrow> Ord(K) & (\<forall>j. j<K \<longrightarrow> ~ j \<approx> K)"
+lemma Card_iff_initial: "Card(K) \<longleftrightarrow> Ord(K) & (\<forall>j. j<K \<longrightarrow> \<not> j \<approx> K)"
proof -
{ fix j
assume K: "Card(K)" "j \<approx> K"
@@ -416,7 +416,7 @@
by (blast intro: CardI Card_is_Ord)
qed
-lemma lt_Card_imp_lesspoll: "[| Card(a); i<a |] ==> i \<prec> a"
+lemma lt_Card_imp_lesspoll: "\<lbrakk>Card(a); i<a\<rbrakk> \<Longrightarrow> i \<prec> a"
apply (unfold lesspoll_def)
apply (drule Card_iff_initial [THEN iffD1])
apply (blast intro!: leI [THEN le_imp_lepoll])
@@ -427,7 +427,7 @@
apply (blast elim!: ltE)
done
-lemma Card_Un: "[| Card(K); Card(L) |] ==> Card(K \<union> L)"
+lemma Card_Un: "\<lbrakk>Card(K); Card(L)\<rbrakk> \<Longrightarrow> 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])
@@ -491,19 +491,19 @@
qed
text\<open>Since we have \<^term>\<open>|succ(nat)| \<le> |nat|\<close>, the converse of \<open>cardinal_mono\<close> fails!\<close>
-lemma cardinal_lt_imp_lt: "[| |i| < |j|; Ord(i); Ord(j) |] ==> i < j"
+lemma cardinal_lt_imp_lt: "\<lbrakk>|i| < |j|; Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> i < j"
apply (rule Ord_linear2 [of i j], assumption+)
apply (erule lt_trans2 [THEN lt_irrefl])
apply (erule cardinal_mono)
done
-lemma Card_lt_imp_lt: "[| |i| < K; Ord(i); Card(K) |] ==> i < K"
+lemma Card_lt_imp_lt: "\<lbrakk>|i| < K; Ord(i); Card(K)\<rbrakk> \<Longrightarrow> i < K"
by (simp (no_asm_simp) add: cardinal_lt_imp_lt Card_is_Ord Card_cardinal_eq)
-lemma Card_lt_iff: "[| Ord(i); Card(K) |] ==> (|i| < K) \<longleftrightarrow> (i < K)"
+lemma Card_lt_iff: "\<lbrakk>Ord(i); Card(K)\<rbrakk> \<Longrightarrow> (|i| < K) \<longleftrightarrow> (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|) \<longleftrightarrow> (K \<le> i)"
+lemma Card_le_iff: "\<lbrakk>Ord(i); Card(K)\<rbrakk> \<Longrightarrow> (K \<le> |i|) \<longleftrightarrow> (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*)
@@ -526,21 +526,21 @@
thus ?thesis by simp
qed
-lemma lepoll_cardinal_le: "[| A \<lesssim> i; Ord(i) |] ==> |A| \<le> i"
+lemma lepoll_cardinal_le: "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> |A| \<le> i"
apply (rule le_trans)
apply (erule well_ord_Memrel [THEN well_ord_lepoll_imp_cardinal_le], assumption)
apply (erule Ord_cardinal_le)
done
-lemma lepoll_Ord_imp_eqpoll: "[| A \<lesssim> i; Ord(i) |] ==> |A| \<approx> A"
+lemma lepoll_Ord_imp_eqpoll: "\<lbrakk>A \<lesssim> i; Ord(i)\<rbrakk> \<Longrightarrow> |A| \<approx> A"
by (blast intro: lepoll_cardinal_le well_ord_Memrel well_ord_cardinal_eqpoll dest!: lepoll_well_ord)
-lemma lesspoll_imp_eqpoll: "[| A \<prec> i; Ord(i) |] ==> |A| \<approx> A"
+lemma lesspoll_imp_eqpoll: "\<lbrakk>A \<prec> i; Ord(i)\<rbrakk> \<Longrightarrow> |A| \<approx> A"
apply (unfold lesspoll_def)
apply (blast intro: lepoll_Ord_imp_eqpoll)
done
-lemma cardinal_subset_Ord: "[|A<=i; Ord(i)|] ==> |A| \<subseteq> i"
+lemma cardinal_subset_Ord: "\<lbrakk>A<=i; Ord(i)\<rbrakk> \<Longrightarrow> |A| \<subseteq> i"
apply (drule subset_imp_lepoll [THEN lepoll_cardinal_le])
apply (auto simp add: lt_def)
apply (blast intro: Ord_trans)
@@ -549,7 +549,7 @@
subsection\<open>The finite cardinals\<close>
lemma cons_lepoll_consD:
- "[| cons(u,A) \<lesssim> cons(v,B); u\<notin>A; v\<notin>B |] ==> A \<lesssim> B"
+ "\<lbrakk>cons(u,A) \<lesssim> cons(v,B); u\<notin>A; v\<notin>B\<rbrakk> \<Longrightarrow> A \<lesssim> B"
apply (unfold lepoll_def inj_def, safe)
apply (rule_tac x = "\<lambda>x\<in>A. if f`x=v then f`u else f`x" in exI)
apply (rule CollectI)
@@ -562,13 +562,13 @@
apply blast
done
-lemma cons_eqpoll_consD: "[| cons(u,A) \<approx> cons(v,B); u\<notin>A; v\<notin>B |] ==> A \<approx> B"
+lemma cons_eqpoll_consD: "\<lbrakk>cons(u,A) \<approx> cons(v,B); u\<notin>A; v\<notin>B\<rbrakk> \<Longrightarrow> A \<approx> B"
apply (simp add: eqpoll_iff)
apply (blast intro: cons_lepoll_consD)
done
(*Lemma suggested by Mike Fourman*)
-lemma succ_lepoll_succD: "succ(m) \<lesssim> succ(n) ==> m \<lesssim> n"
+lemma succ_lepoll_succD: "succ(m) \<lesssim> succ(n) \<Longrightarrow> m \<lesssim> n"
apply (unfold succ_def)
apply (erule cons_lepoll_consD)
apply (rule mem_not_refl)+
@@ -576,7 +576,7 @@
lemma nat_lepoll_imp_le:
- "m \<in> nat ==> n \<in> nat \<Longrightarrow> m \<lesssim> n \<Longrightarrow> m \<le> n"
+ "m \<in> nat \<Longrightarrow> n \<in> nat \<Longrightarrow> m \<lesssim> n \<Longrightarrow> m \<le> n"
proof (induct m arbitrary: n rule: nat_induct)
case 0 thus ?case by (blast intro!: nat_0_le)
next
@@ -591,7 +591,7 @@
qed
qed
-lemma nat_eqpoll_iff: "[| m \<in> nat; n \<in> nat |] ==> m \<approx> n \<longleftrightarrow> m = n"
+lemma nat_eqpoll_iff: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m \<approx> n \<longleftrightarrow> m = n"
apply (rule iffI)
apply (blast intro: nat_lepoll_imp_le le_anti_sym elim!: eqpollE)
apply (simp add: eqpoll_refl)
@@ -616,11 +616,11 @@
(*Part of Kunen's Lemma 10.6*)
-lemma succ_lepoll_natE: "[| succ(n) \<lesssim> n; n \<in> nat |] ==> P"
+lemma succ_lepoll_natE: "\<lbrakk>succ(n) \<lesssim> n; n \<in> nat\<rbrakk> \<Longrightarrow> P"
by (rule nat_lepoll_imp_le [THEN lt_irrefl], auto)
lemma nat_lepoll_imp_ex_eqpoll_n:
- "[| n \<in> nat; nat \<lesssim> X |] ==> \<exists>Y. Y \<subseteq> X & n \<approx> Y"
+ "\<lbrakk>n \<in> nat; nat \<lesssim> X\<rbrakk> \<Longrightarrow> \<exists>Y. Y \<subseteq> X & n \<approx> Y"
apply (unfold lepoll_def eqpoll_def)
apply (fast del: subsetI subsetCE
intro!: subset_SIs
@@ -649,22 +649,22 @@
qed
lemma lesspoll_succ_imp_lepoll:
- "[| A \<prec> succ(m); m \<in> nat |] ==> A \<lesssim> m"
+ "\<lbrakk>A \<prec> succ(m); m \<in> nat\<rbrakk> \<Longrightarrow> A \<lesssim> m"
apply (unfold lesspoll_def lepoll_def eqpoll_def bij_def)
apply (auto dest: inj_not_surj_succ)
done
-lemma lesspoll_succ_iff: "m \<in> nat ==> A \<prec> succ(m) \<longleftrightarrow> A \<lesssim> m"
+lemma lesspoll_succ_iff: "m \<in> nat \<Longrightarrow> A \<prec> succ(m) \<longleftrightarrow> A \<lesssim> m"
by (blast intro!: lepoll_imp_lesspoll_succ lesspoll_succ_imp_lepoll)
-lemma lepoll_succ_disj: "[| A \<lesssim> succ(m); m \<in> nat |] ==> A \<lesssim> m | A \<approx> succ(m)"
+lemma lepoll_succ_disj: "\<lbrakk>A \<lesssim> succ(m); m \<in> nat\<rbrakk> \<Longrightarrow> A \<lesssim> m | A \<approx> succ(m)"
apply (rule disjCI)
apply (rule lesspoll_succ_imp_lepoll)
prefer 2 apply assumption
apply (simp (no_asm_simp) add: lesspoll_def)
done
-lemma lesspoll_cardinal_lt: "[| A \<prec> i; Ord(i) |] ==> |A| < i"
+lemma lesspoll_cardinal_lt: "\<lbrakk>A \<prec> i; Ord(i)\<rbrakk> \<Longrightarrow> |A| < i"
apply (unfold lesspoll_def, clarify)
apply (frule lepoll_cardinal_le, assumption)
apply (blast intro: well_ord_Memrel well_ord_cardinal_eqpoll [THEN eqpoll_sym]
@@ -676,7 +676,7 @@
(*This implies Kunen's Lemma 10.6*)
lemma lt_not_lepoll:
- assumes n: "n<i" "n \<in> nat" shows "~ i \<lesssim> n"
+ assumes n: "n<i" "n \<in> nat" shows "\<not> i \<lesssim> n"
proof -
{ assume i: "i \<lesssim> n"
have "succ(n) \<lesssim> i" using n
@@ -700,8 +700,8 @@
thus ?thesis by (simp add: eqpoll_refl)
next
case gt
- hence "~ i \<lesssim> n" using n by (rule lt_not_lepoll)
- hence "~ i \<approx> n" using n by (blast intro: eqpoll_imp_lepoll)
+ hence "\<not> i \<lesssim> n" using n by (rule lt_not_lepoll)
+ hence "\<not> i \<approx> n" using n by (blast intro: eqpoll_imp_lepoll)
moreover have "i \<noteq> n" using \<open>n<i\<close> by auto
ultimately show ?thesis by blast
qed
@@ -710,7 +710,7 @@
proof -
{ fix i
assume i: "i < nat" "i \<approx> nat"
- hence "~ nat \<lesssim> i"
+ hence "\<not> nat \<lesssim> i"
by (simp add: lt_def lt_not_lepoll)
hence False using i
by (simp add: eqpoll_iff)
@@ -721,12 +721,12 @@
qed
(*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 \<Longrightarrow> nat \<le> |i|"
apply (rule Card_nat [THEN Card_cardinal_eq, THEN subst])
apply (erule cardinal_mono)
done
-lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat"
+lemma n_lesspoll_nat: "n \<in> nat \<Longrightarrow> n \<prec> nat"
by (blast intro: Ord_nat Card_nat ltI lt_Card_imp_lesspoll)
@@ -735,7 +735,7 @@
(*Congruence law for cons under equipollence*)
lemma cons_lepoll_cong:
- "[| A \<lesssim> B; b \<notin> B |] ==> cons(a,A) \<lesssim> cons(b,B)"
+ "\<lbrakk>A \<lesssim> B; b \<notin> B\<rbrakk> \<Longrightarrow> cons(a,A) \<lesssim> cons(b,B)"
apply (unfold lepoll_def, safe)
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 \<in> B then converse (f) `z else a" in lam_injective)
@@ -745,15 +745,15 @@
done
lemma cons_eqpoll_cong:
- "[| A \<approx> B; a \<notin> A; b \<notin> B |] ==> cons(a,A) \<approx> cons(b,B)"
+ "\<lbrakk>A \<approx> B; a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> cons(a,A) \<approx> cons(b,B)"
by (simp add: eqpoll_iff cons_lepoll_cong)
lemma cons_lepoll_cons_iff:
- "[| a \<notin> A; b \<notin> B |] ==> cons(a,A) \<lesssim> cons(b,B) \<longleftrightarrow> A \<lesssim> B"
+ "\<lbrakk>a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> cons(a,A) \<lesssim> cons(b,B) \<longleftrightarrow> A \<lesssim> B"
by (blast intro: cons_lepoll_cong cons_lepoll_consD)
lemma cons_eqpoll_cons_iff:
- "[| a \<notin> A; b \<notin> B |] ==> cons(a,A) \<approx> cons(b,B) \<longleftrightarrow> A \<approx> B"
+ "\<lbrakk>a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> cons(a,A) \<approx> cons(b,B) \<longleftrightarrow> A \<approx> B"
by (blast intro: cons_eqpoll_cong cons_eqpoll_consD)
lemma singleton_eqpoll_1: "{a} \<approx> 1"
@@ -766,7 +766,7 @@
apply (simp (no_asm) add: nat_into_Card [THEN Card_cardinal_eq])
done
-lemma not_0_is_lepoll_1: "A \<noteq> 0 ==> 1 \<lesssim> A"
+lemma not_0_is_lepoll_1: "A \<noteq> 0 \<Longrightarrow> 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)
@@ -774,26 +774,26 @@
done
(*Congruence law for succ under equipollence*)
-lemma succ_eqpoll_cong: "A \<approx> B ==> succ(A) \<approx> succ(B)"
+lemma succ_eqpoll_cong: "A \<approx> B \<Longrightarrow> succ(A) \<approx> succ(B)"
apply (unfold succ_def)
apply (simp add: cons_eqpoll_cong mem_not_refl)
done
(*Congruence law for + under equipollence*)
-lemma sum_eqpoll_cong: "[| A \<approx> C; B \<approx> D |] ==> A+B \<approx> C+D"
+lemma sum_eqpoll_cong: "\<lbrakk>A \<approx> C; B \<approx> D\<rbrakk> \<Longrightarrow> A+B \<approx> C+D"
apply (unfold eqpoll_def)
apply (blast intro!: sum_bij)
done
(*Congruence law for * under equipollence*)
lemma prod_eqpoll_cong:
- "[| A \<approx> C; B \<approx> D |] ==> A*B \<approx> C*D"
+ "\<lbrakk>A \<approx> C; B \<approx> D\<rbrakk> \<Longrightarrow> A*B \<approx> C*D"
apply (unfold eqpoll_def)
apply (blast intro!: prod_bij)
done
lemma inj_disjoint_eqpoll:
- "[| f \<in> inj(A,B); A \<inter> B = 0 |] ==> A \<union> (B - range(f)) \<approx> B"
+ "\<lbrakk>f \<in> inj(A,B); A \<inter> B = 0\<rbrakk> \<Longrightarrow> A \<union> (B - range(f)) \<approx> B"
apply (unfold eqpoll_def)
apply (rule exI)
apply (rule_tac c = "%x. if x \<in> A then f`x else x"
@@ -814,7 +814,7 @@
text\<open>If \<^term>\<open>A\<close> has at most \<^term>\<open>n+1\<close> elements and \<^term>\<open>a \<in> A\<close>
then \<^term>\<open>A-{a}\<close> has at most \<^term>\<open>n\<close>.\<close>
lemma Diff_sing_lepoll:
- "[| a \<in> A; A \<lesssim> succ(n) |] ==> A - {a} \<lesssim> n"
+ "\<lbrakk>a \<in> A; A \<lesssim> succ(n)\<rbrakk> \<Longrightarrow> A - {a} \<lesssim> n"
apply (unfold succ_def)
apply (rule cons_lepoll_consD)
apply (rule_tac [3] mem_not_refl)
@@ -834,12 +834,12 @@
by (blast intro: cons_lepoll_consD mem_irrefl)
qed
-lemma Diff_sing_eqpoll: "[| a \<in> A; A \<approx> succ(n) |] ==> A - {a} \<approx> n"
+lemma Diff_sing_eqpoll: "\<lbrakk>a \<in> A; A \<approx> succ(n)\<rbrakk> \<Longrightarrow> A - {a} \<approx> n"
by (blast intro!: eqpollI
elim!: eqpollE
intro: Diff_sing_lepoll lepoll_Diff_sing)
-lemma lepoll_1_is_sing: "[| A \<lesssim> 1; a \<in> A |] ==> A = {a}"
+lemma lepoll_1_is_sing: "\<lbrakk>A \<lesssim> 1; a \<in> A\<rbrakk> \<Longrightarrow> A = {a}"
apply (frule Diff_sing_lepoll, assumption)
apply (drule lepoll_0_is_0)
apply (blast elim: equalityE)
@@ -854,12 +854,12 @@
done
lemma well_ord_Un:
- "[| well_ord(X,R); well_ord(Y,S) |] ==> \<exists>T. well_ord(X \<union> Y, T)"
+ "\<lbrakk>well_ord(X,R); well_ord(Y,S)\<rbrakk> \<Longrightarrow> \<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 \<inter> B = 0 ==> A \<union> B \<approx> A + B"
+lemma disj_Un_eqpoll_sum: "A \<inter> B = 0 \<Longrightarrow> A \<union> B \<approx> A + B"
apply (unfold eqpoll_def)
apply (rule_tac x = "\<lambda>a\<in>A \<union> B. if a \<in> A then Inl (a) else Inr (a)" in exI)
apply (rule_tac d = "%z. case (%x. x, %x. x, z)" in lam_bijective)
@@ -869,7 +869,7 @@
subsection \<open>Finite and infinite sets\<close>
-lemma eqpoll_imp_Finite_iff: "A \<approx> B ==> Finite(A) \<longleftrightarrow> Finite(B)"
+lemma eqpoll_imp_Finite_iff: "A \<approx> B \<Longrightarrow> Finite(A) \<longleftrightarrow> Finite(B)"
apply (unfold Finite_def)
apply (blast intro: eqpoll_trans eqpoll_sym)
done
@@ -879,7 +879,7 @@
apply (blast intro!: eqpoll_refl nat_0I)
done
-lemma Finite_cons: "Finite(x) ==> Finite(cons(y,x))"
+lemma Finite_cons: "Finite(x) \<Longrightarrow> Finite(cons(y,x))"
apply (unfold Finite_def)
apply (case_tac "y \<in> x")
apply (simp add: cons_absorb)
@@ -889,7 +889,7 @@
apply (simp (no_asm_simp) add: succ_def cons_eqpoll_cong mem_not_refl)
done
-lemma Finite_succ: "Finite(x) ==> Finite(succ(x))"
+lemma Finite_succ: "Finite(x) \<Longrightarrow> Finite(succ(x))"
apply (unfold succ_def)
apply (erule Finite_cons)
done
@@ -911,7 +911,7 @@
qed
lemma lesspoll_nat_is_Finite:
- "A \<prec> nat ==> Finite(A)"
+ "A \<prec> nat \<Longrightarrow> Finite(A)"
apply (unfold Finite_def)
apply (blast dest: ltD lesspoll_cardinal_lt
lesspoll_imp_eqpoll [THEN eqpoll_sym])
@@ -936,13 +936,13 @@
lemma Finite_succ_iff [iff]: "Finite(succ(x)) \<longleftrightarrow> Finite(x)"
by (simp add: succ_def)
-lemma Finite_Int: "Finite(A) | Finite(B) ==> Finite(A \<inter> B)"
+lemma Finite_Int: "Finite(A) | Finite(B) \<Longrightarrow> Finite(A \<inter> B)"
by (blast intro: subset_Finite)
lemmas Finite_Diff = Diff_subset [THEN subset_Finite]
lemma nat_le_infinite_Ord:
- "[| Ord(i); ~ Finite(i) |] ==> nat \<le> i"
+ "\<lbrakk>Ord(i); \<not> Finite(i)\<rbrakk> \<Longrightarrow> nat \<le> i"
apply (unfold Finite_def)
apply (erule Ord_nat [THEN [2] Ord_linear2])
prefer 2 apply assumption
@@ -950,19 +950,19 @@
done
lemma Finite_imp_well_ord:
- "Finite(A) ==> \<exists>r. well_ord(A,r)"
+ "Finite(A) \<Longrightarrow> \<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
-lemma succ_lepoll_imp_not_empty: "succ(x) \<lesssim> y ==> y \<noteq> 0"
+lemma succ_lepoll_imp_not_empty: "succ(x) \<lesssim> y \<Longrightarrow> y \<noteq> 0"
by (fast dest!: lepoll_0_is_0)
-lemma eqpoll_succ_imp_not_empty: "x \<approx> succ(n) ==> x \<noteq> 0"
+lemma eqpoll_succ_imp_not_empty: "x \<approx> succ(n) \<Longrightarrow> x \<noteq> 0"
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) \<longrightarrow> A \<in> Fin(X)"
+ "n \<in> nat \<Longrightarrow> \<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])
@@ -978,10 +978,10 @@
apply (simp add: cons_Diff)
done
-lemma Finite_Fin: "[| Finite(A); A \<subseteq> X |] ==> A \<in> Fin(X)"
+lemma Finite_Fin: "\<lbrakk>Finite(A); A \<subseteq> X\<rbrakk> \<Longrightarrow> A \<in> Fin(X)"
by (unfold Finite_def, blast intro: Finite_Fin_lemma)
-lemma Fin_lemma [rule_format]: "n \<in> nat ==> \<forall>A. A \<approx> n \<longrightarrow> A \<in> Fin(A)"
+lemma Fin_lemma [rule_format]: "n \<in> nat \<Longrightarrow> \<forall>A. A \<approx> n \<longrightarrow> A \<in> Fin(A)"
apply (induct_tac n)
apply (simp add: eqpoll_0_iff, clarify)
apply (subgoal_tac "\<exists>u. u \<in> A")
@@ -997,18 +997,18 @@
apply (blast intro: bij_converse_bij [THEN bij_is_fun, THEN apply_type])
done
-lemma Finite_into_Fin: "Finite(A) ==> A \<in> Fin(A)"
+lemma Finite_into_Fin: "Finite(A) \<Longrightarrow> A \<in> Fin(A)"
apply (unfold Finite_def)
apply (blast intro: Fin_lemma)
done
-lemma Fin_into_Finite: "A \<in> Fin(U) ==> Finite(A)"
+lemma Fin_into_Finite: "A \<in> Fin(U) \<Longrightarrow> Finite(A)"
by (fast intro!: Finite_0 Finite_cons elim: Fin_induct)
lemma Finite_Fin_iff: "Finite(A) \<longleftrightarrow> A \<in> Fin(A)"
by (blast intro: Finite_into_Fin Fin_into_Finite)
-lemma Finite_Un: "[| Finite(A); Finite(B) |] ==> Finite(A \<union> B)"
+lemma Finite_Un: "\<lbrakk>Finite(A); Finite(B)\<rbrakk> \<Longrightarrow> Finite(A \<union> B)"
by (blast intro!: Fin_into_Finite Fin_UnI
dest!: Finite_into_Fin
intro: Un_upper1 [THEN Fin_mono, THEN subsetD]
@@ -1018,7 +1018,7 @@
by (blast intro: subset_Finite Finite_Un)
text\<open>The converse must hold too.\<close>
-lemma Finite_Union: "[| \<forall>y\<in>X. Finite(y); Finite(X) |] ==> Finite(\<Union>(X))"
+lemma Finite_Union: "\<lbrakk>\<forall>y\<in>X. Finite(y); Finite(X)\<rbrakk> \<Longrightarrow> Finite(\<Union>(X))"
apply (simp add: Finite_Fin_iff)
apply (rule Fin_UnionI)
apply (erule Fin_induct, simp)
@@ -1027,15 +1027,15 @@
(* 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 \<notin> B; P(B) |] ==> P(cons(x, B)) |]
- ==> P(A)"
+"\<lbrakk>Finite(A); P(0);
+ \<And>x B. \<lbrakk>Finite(B); x \<notin> B; P(B)\<rbrakk> \<Longrightarrow> P(cons(x, B))\<rbrakk>
+ \<Longrightarrow> P(A)"
apply (erule Finite_into_Fin [THEN Fin_induct])
apply (blast intro: Fin_into_Finite)+
done
-(*Sidi Ehmety. The contrapositive says ~Finite(A) ==> ~Finite(A-{a}) *)
-lemma Diff_sing_Finite: "Finite(A - {a}) ==> Finite(A)"
+(*Sidi Ehmety. The contrapositive says \<not>Finite(A) \<Longrightarrow> \<not>Finite(A-{a}) *)
+lemma Diff_sing_Finite: "Finite(A - {a}) \<Longrightarrow> Finite(A)"
apply (unfold Finite_def)
apply (case_tac "a \<in> A")
apply (subgoal_tac [2] "A-{a}=A", auto)
@@ -1046,8 +1046,8 @@
done
(*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) \<longrightarrow> Finite(A)"
+ \<lbrakk>\<not>Finite(A); Finite(B)\<rbrakk> \<Longrightarrow> \<not>Finite(A-B) *)
+lemma Diff_Finite [rule_format]: "Finite(B) \<Longrightarrow> Finite(A-B) \<longrightarrow> Finite(A)"
apply (erule Finite_induct, auto)
apply (case_tac "x \<in> A")
apply (subgoal_tac [2] "A-cons (x, B) = A - B")
@@ -1055,12 +1055,12 @@
apply (drule Diff_sing_Finite, auto)
done
-lemma Finite_RepFun: "Finite(A) ==> Finite(RepFun(A,f))"
+lemma Finite_RepFun: "Finite(A) \<Longrightarrow> Finite(RepFun(A,f))"
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) \<longrightarrow> Finite(A)"
+ "\<lbrakk>Finite(x); \<And>x y. f(x)=f(y) \<Longrightarrow> x=y\<rbrakk>
+ \<Longrightarrow> \<forall>A. x = RepFun(A,f) \<longrightarrow> Finite(A)"
apply (erule Finite_induct)
apply clarify
apply (case_tac "A=0", simp)
@@ -1078,15 +1078,15 @@
text\<open>I don't know why, but if the premise is expressed using meta-connectives
then the simplifier cannot prove it automatically in conditional rewriting.\<close>
lemma Finite_RepFun_iff:
- "(\<forall>x y. f(x)=f(y) \<longrightarrow> x=y) ==> Finite(RepFun(A,f)) \<longleftrightarrow> Finite(A)"
+ "(\<forall>x y. f(x)=f(y) \<longrightarrow> x=y) \<Longrightarrow> Finite(RepFun(A,f)) \<longleftrightarrow> Finite(A)"
by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f])
-lemma Finite_Pow: "Finite(A) ==> Finite(Pow(A))"
+lemma Finite_Pow: "Finite(A) \<Longrightarrow> Finite(Pow(A))"
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)"
+lemma Finite_Pow_imp_Finite: "Finite(Pow(A)) \<Longrightarrow> Finite(A)"
apply (subgoal_tac "Finite({{x} . x \<in> A})")
apply (simp add: Finite_RepFun_iff )
apply (blast intro: subset_Finite)
@@ -1103,7 +1103,7 @@
(*Krzysztof Grabczewski's proof that the converse of a finite, well-ordered
set is well-ordered. Proofs simplified by lcp. *)
-lemma nat_wf_on_converse_Memrel: "n \<in> nat ==> wf[n](converse(Memrel(n)))"
+lemma nat_wf_on_converse_Memrel: "n \<in> nat \<Longrightarrow> wf[n](converse(Memrel(n)))"
proof (induct n rule: nat_induct)
case 0 thus ?case by (blast intro: wf_onI)
next
@@ -1125,15 +1125,15 @@
qed
qed
-lemma nat_well_ord_converse_Memrel: "n \<in> nat ==> well_ord(n,converse(Memrel(n)))"
+lemma nat_well_ord_converse_Memrel: "n \<in> nat \<Longrightarrow> well_ord(n,converse(Memrel(n)))"
apply (frule Ord_nat [THEN Ord_in_Ord, THEN well_ord_Memrel])
apply (simp add: well_ord_def tot_ord_converse nat_wf_on_converse_Memrel)
done
lemma well_ord_converse:
- "[|well_ord(A,r);
- well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) |]
- ==> well_ord(A,converse(r))"
+ "\<lbrakk>well_ord(A,r);
+ well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))\<rbrakk>
+ \<Longrightarrow> well_ord(A,converse(r))"
apply (rule well_ord_Int_iff [THEN iffD1])
apply (frule ordermap_bij [THEN bij_is_inj, THEN well_ord_rvimage], assumption)
apply (simp add: rvimage_converse converse_Int converse_prod
@@ -1153,16 +1153,16 @@
qed
lemma Finite_well_ord_converse:
- "[| Finite(A); well_ord(A,r) |] ==> well_ord(A,converse(r))"
+ "\<lbrakk>Finite(A); well_ord(A,r)\<rbrakk> \<Longrightarrow> well_ord(A,converse(r))"
apply (unfold Finite_def)
apply (rule well_ord_converse, assumption)
apply (blast dest: ordertype_eq_n intro!: nat_well_ord_converse_Memrel)
done
-lemma nat_into_Finite: "n \<in> nat ==> Finite(n)"
+lemma nat_into_Finite: "n \<in> nat \<Longrightarrow> Finite(n)"
by (auto simp add: Finite_def intro: eqpoll_refl)
-lemma nat_not_Finite: "~ Finite(nat)"
+lemma nat_not_Finite: "\<not> Finite(nat)"
proof -
{ fix n
assume n: "n \<in> nat" "nat \<approx> n"
--- a/src/ZF/CardinalArith.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/CardinalArith.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,19 +9,19 @@
definition
InfCard :: "i=>o" where
- "InfCard(i) == Card(i) & nat \<le> i"
+ "InfCard(i) \<equiv> Card(i) & nat \<le> i"
definition
cmult :: "[i,i]=>i" (infixl \<open>\<otimes>\<close> 70) where
- "i \<otimes> j == |i*j|"
+ "i \<otimes> j \<equiv> |i*j|"
definition
cadd :: "[i,i]=>i" (infixl \<open>\<oplus>\<close> 65) where
- "i \<oplus> j == |i+j|"
+ "i \<oplus> j \<equiv> |i+j|"
definition
csquare_rel :: "i=>i" where
- "csquare_rel(K) ==
+ "csquare_rel(K) \<equiv>
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))))"
@@ -30,14 +30,14 @@
jump_cardinal :: "i=>i" where
\<comment> \<open>This definition is more complex than Kunen's but it more easily proved to
be a cardinal\<close>
- "jump_cardinal(K) ==
+ "jump_cardinal(K) \<equiv>
\<Union>X\<in>Pow(K). {z. r \<in> Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
definition
csucc :: "i=>i" where
\<comment> \<open>needed because \<^term>\<open>jump_cardinal(K)\<close> might not be the successor
of \<^term>\<open>K\<close>\<close>
- "csucc(K) == \<mu> L. Card(L) & K<L"
+ "csucc(K) \<equiv> \<mu> L. Card(L) & K<L"
lemma Card_Union [simp,intro,TC]:
@@ -65,14 +65,14 @@
} thus "\<not> j \<approx> \<Union>A" by blast
qed
-lemma Card_UN: "(!!x. x \<in> A ==> Card(K(x))) ==> Card(\<Union>x\<in>A. K(x))"
+lemma Card_UN: "(\<And>x. x \<in> A \<Longrightarrow> Card(K(x))) \<Longrightarrow> Card(\<Union>x\<in>A. K(x))"
by blast
lemma Card_OUN [simp,intro,TC]:
- "(!!x. x \<in> A ==> Card(K(x))) ==> Card(\<Union>x<A. K(x))"
+ "(\<And>x. x \<in> A \<Longrightarrow> Card(K(x))) \<Longrightarrow> Card(\<Union>x<A. K(x))"
by (auto simp add: OUnion_def Card_0)
-lemma in_Card_imp_lesspoll: "[| Card(K); b \<in> K |] ==> b \<prec> K"
+lemma in_Card_imp_lesspoll: "\<lbrakk>Card(K); b \<in> K\<rbrakk> \<Longrightarrow> b \<prec> K"
apply (unfold lesspoll_def)
apply (simp add: Card_iff_initial)
apply (fast intro!: le_imp_lepoll ltI leI)
@@ -130,7 +130,7 @@
apply (rule bij_0_sum)
done
-lemma cadd_0 [simp]: "Card(K) ==> 0 \<oplus> K = K"
+lemma cadd_0 [simp]: "Card(K) \<Longrightarrow> 0 \<oplus> K = K"
apply (unfold cadd_def)
apply (simp add: sum_0_eqpoll [THEN cardinal_cong] Card_cardinal_eq)
done
@@ -160,7 +160,7 @@
subsubsection\<open>Monotonicity of addition\<close>
lemma sum_lepoll_mono:
- "[| A \<lesssim> C; B \<lesssim> D |] ==> A + B \<lesssim> C + D"
+ "\<lbrakk>A \<lesssim> C; B \<lesssim> D\<rbrakk> \<Longrightarrow> A + B \<lesssim> C + D"
apply (unfold lepoll_def)
apply (elim exE)
apply (rule_tac x = "\<lambda>z\<in>A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI)
@@ -170,7 +170,7 @@
done
lemma cadd_le_mono:
- "[| K' \<le> K; L' \<le> L |] ==> (K' \<oplus> L') \<le> (K \<oplus> L)"
+ "\<lbrakk>K' \<le> K; L' \<le> L\<rbrakk> \<Longrightarrow> (K' \<oplus> L') \<le> (K \<oplus> L)"
apply (unfold cadd_def)
apply (safe dest!: le_subset_iff [THEN iffD1])
apply (rule well_ord_lepoll_imp_cardinal_le)
@@ -292,7 +292,7 @@
apply (rule singleton_prod_bij [THEN bij_converse_bij])
done
-lemma cmult_1 [simp]: "Card(K) ==> 1 \<otimes> K = K"
+lemma cmult_1 [simp]: "Card(K) \<Longrightarrow> 1 \<otimes> K = K"
apply (unfold cmult_def succ_def)
apply (simp add: prod_singleton_eqpoll [THEN cardinal_cong] Card_cardinal_eq)
done
@@ -305,7 +305,7 @@
done
(*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
-lemma cmult_square_le: "Card(K) ==> K \<le> K \<otimes> K"
+lemma cmult_square_le: "Card(K) \<Longrightarrow> K \<le> K \<otimes> K"
apply (unfold cmult_def)
apply (rule le_trans)
apply (rule_tac [2] well_ord_lepoll_imp_cardinal_le)
@@ -316,14 +316,14 @@
subsubsection\<open>Multiplication by a non-zero cardinal\<close>
-lemma prod_lepoll_self: "b \<in> B ==> A \<lesssim> A*B"
+lemma prod_lepoll_self: "b \<in> B \<Longrightarrow> A \<lesssim> A*B"
apply (unfold lepoll_def inj_def)
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 \<otimes> L)"
+ "\<lbrakk>Card(K); Ord(L); 0<L\<rbrakk> \<Longrightarrow> K \<le> (K \<otimes> L)"
apply (unfold cmult_def)
apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_cardinal_le])
apply assumption
@@ -334,7 +334,7 @@
subsubsection\<open>Monotonicity of multiplication\<close>
lemma prod_lepoll_mono:
- "[| A \<lesssim> C; B \<lesssim> D |] ==> A * B \<lesssim> C * D"
+ "\<lbrakk>A \<lesssim> C; B \<lesssim> D\<rbrakk> \<Longrightarrow> A * B \<lesssim> C * D"
apply (unfold lepoll_def)
apply (elim exE)
apply (rule_tac x = "lam <w,y>:A*B. <f`w, fa`y>" in exI)
@@ -344,7 +344,7 @@
done
lemma cmult_le_mono:
- "[| K' \<le> K; L' \<le> L |] ==> (K' \<otimes> L') \<le> (K \<otimes> L)"
+ "\<lbrakk>K' \<le> K; L' \<le> L\<rbrakk> \<Longrightarrow> (K' \<otimes> L') \<le> (K \<otimes> L)"
apply (unfold cmult_def)
apply (safe dest!: le_subset_iff [THEN iffD1])
apply (rule well_ord_lepoll_imp_cardinal_le)
@@ -365,7 +365,7 @@
(*Unconditional version requires AC*)
lemma cmult_succ_lemma:
- "[| Ord(m); Ord(n) |] ==> succ(m) \<otimes> n = n \<oplus> (m \<otimes> n)"
+ "\<lbrakk>Ord(m); Ord(n)\<rbrakk> \<Longrightarrow> succ(m) \<otimes> n = n \<oplus> (m \<otimes> n)"
apply (unfold cmult_def cadd_def)
apply (rule prod_succ_eqpoll [THEN cardinal_cong, THEN trans])
apply (rule cardinal_cong [symmetric])
@@ -373,12 +373,12 @@
apply (blast intro: well_ord_rmult well_ord_Memrel)
done
-lemma nat_cmult_eq_mult: "[| m \<in> nat; n \<in> nat |] ==> m \<otimes> n = m#*n"
+lemma nat_cmult_eq_mult: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m \<otimes> n = m#*n"
apply (induct_tac m)
apply (simp_all add: cmult_succ_lemma nat_cadd_eq_add)
done
-lemma cmult_2: "Card(n) ==> 2 \<otimes> n = n \<oplus> n"
+lemma cmult_2: "Card(n) \<Longrightarrow> 2 \<otimes> n = n \<oplus> n"
by (simp add: cmult_succ_lemma Card_is_Ord cadd_commute [of _ 0])
lemma sum_lepoll_prod:
@@ -391,7 +391,7 @@
finally show "B+B \<lesssim> C*B" .
qed
-lemma lepoll_imp_sum_lepoll_prod: "[| A \<lesssim> B; 2 \<lesssim> A |] ==> A+B \<lesssim> A*B"
+lemma lepoll_imp_sum_lepoll_prod: "\<lbrakk>A \<lesssim> B; 2 \<lesssim> A\<rbrakk> \<Longrightarrow> A+B \<lesssim> A*B"
by (blast intro: sum_lepoll_mono sum_lepoll_prod lepoll_trans lepoll_refl)
@@ -401,7 +401,7 @@
\<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 \<in> nat then nat_case(u, %z. z, y) else y. \
If f \<in> inj(nat,A) then range(f) behaves like the natural numbers.*)
-lemma nat_cons_lepoll: "nat \<lesssim> A ==> cons(u,A) \<lesssim> A"
+lemma nat_cons_lepoll: "nat \<lesssim> A \<Longrightarrow> cons(u,A) \<lesssim> A"
apply (unfold lepoll_def)
apply (erule exE)
apply (rule_tac x =
@@ -419,13 +419,13 @@
inj_converse_fun [THEN apply_funtype])
done
-lemma nat_cons_eqpoll: "nat \<lesssim> A ==> cons(u,A) \<approx> A"
+lemma nat_cons_eqpoll: "nat \<lesssim> A \<Longrightarrow> cons(u,A) \<approx> A"
apply (erule nat_cons_lepoll [THEN eqpollI])
apply (rule subset_consI [THEN subset_imp_lepoll])
done
(*Specialized version required below*)
-lemma nat_succ_eqpoll: "nat \<subseteq> A ==> succ(A) \<approx> A"
+lemma nat_succ_eqpoll: "nat \<subseteq> A \<Longrightarrow> succ(A) \<approx> A"
apply (unfold succ_def)
apply (erule subset_imp_lepoll [THEN nat_cons_eqpoll])
done
@@ -435,19 +435,19 @@
apply (blast intro: Card_nat le_refl Card_is_Ord)
done
-lemma InfCard_is_Card: "InfCard(K) ==> Card(K)"
+lemma InfCard_is_Card: "InfCard(K) \<Longrightarrow> Card(K)"
apply (unfold InfCard_def)
apply (erule conjunct1)
done
lemma InfCard_Un:
- "[| InfCard(K); Card(L) |] ==> InfCard(K \<union> L)"
+ "\<lbrakk>InfCard(K); Card(L)\<rbrakk> \<Longrightarrow> InfCard(K \<union> L)"
apply (unfold InfCard_def)
apply (simp add: Card_Un Un_upper1_le [THEN [2] le_trans] Card_is_Ord)
done
(*Kunen's Lemma 10.11*)
-lemma InfCard_is_Limit: "InfCard(K) ==> Limit(K)"
+lemma InfCard_is_Limit: "InfCard(K) \<Longrightarrow> Limit(K)"
apply (unfold InfCard_def)
apply (erule conjE)
apply (frule Card_is_Ord)
@@ -467,7 +467,7 @@
(*A general fact about ordermap*)
lemma ordermap_eqpoll_pred:
- "[| well_ord(A,r); x \<in> A |] ==> ordermap(A,r)`x \<approx> Order.pred(A,x,r)"
+ "\<lbrakk>well_ord(A,r); x \<in> A\<rbrakk> \<Longrightarrow> ordermap(A,r)`x \<approx> Order.pred(A,x,r)"
apply (unfold eqpoll_def)
apply (rule exI)
apply (simp add: ordermap_eq_image well_ord_is_wf)
@@ -491,7 +491,7 @@
subsubsection\<open>Characterising initial segments of the well-ordering\<close>
lemma csquareD:
- "[| <<x,y>, <z,z>> \<in> csquare_rel(K); x<K; y<K; z<K |] ==> x \<le> z & y \<le> z"
+ "\<lbrakk><<x,y>, <z,z>> \<in> csquare_rel(K); x<K; y<K; z<K\<rbrakk> \<Longrightarrow> x \<le> z & y \<le> z"
apply (unfold csquare_rel_def)
apply (erule rev_mp)
apply (elim ltE)
@@ -501,14 +501,14 @@
done
lemma pred_csquare_subset:
- "z<K ==> Order.pred(K*K, <z,z>, csquare_rel(K)) \<subseteq> succ(z)*succ(z)"
+ "z<K \<Longrightarrow> Order.pred(K*K, <z,z>, csquare_rel(K)) \<subseteq> succ(z)*succ(z)"
apply (unfold Order.pred_def)
apply (safe del: SigmaI dest!: csquareD)
apply (unfold lt_def, auto)
done
lemma csquare_ltI:
- "[| x<z; y<z; z<K |] ==> <<x,y>, <z,z>> \<in> csquare_rel(K)"
+ "\<lbrakk>x<z; y<z; z<K\<rbrakk> \<Longrightarrow> <<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)
@@ -518,7 +518,7 @@
(*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>> \<in> csquare_rel(K) | x=z & y=z"
+ "\<lbrakk>x \<le> z; y \<le> z; z<K\<rbrakk> \<Longrightarrow> <<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)
@@ -532,7 +532,7 @@
subsubsection\<open>The cardinality of initial segments\<close>
lemma ordermap_z_lt:
- "[| Limit(K); x<K; y<K; z=succ(x \<union> y) |] ==>
+ "\<lbrakk>Limit(K); x<K; y<K; z=succ(x \<union> y)\<rbrakk> \<Longrightarrow>
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))")
@@ -668,18 +668,18 @@
finally show ?thesis .
qed
-lemma InfCard_square_eqpoll: "InfCard(K) ==> K \<times> K \<approx> K"
+lemma InfCard_square_eqpoll: "InfCard(K) \<Longrightarrow> 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])
done
-lemma Inf_Card_is_InfCard: "[| Card(i); ~ Finite(i) |] ==> InfCard(i)"
+lemma Inf_Card_is_InfCard: "\<lbrakk>Card(i); \<not> Finite(i)\<rbrakk> \<Longrightarrow> InfCard(i)"
by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord])
subsubsection\<open>Toward's Kunen's Corollary 10.13 (1)\<close>
-lemma InfCard_le_cmult_eq: "[| InfCard(K); L \<le> K; 0<L |] ==> K \<otimes> L = K"
+lemma InfCard_le_cmult_eq: "\<lbrakk>InfCard(K); L \<le> K; 0<L\<rbrakk> \<Longrightarrow> K \<otimes> L = K"
apply (rule le_anti_sym)
prefer 2
apply (erule ltE, blast intro: cmult_le_self InfCard_is_Card)
@@ -689,7 +689,7 @@
done
(*Corollary 10.13 (1), for cardinal multiplication*)
-lemma InfCard_cmult_eq: "[| InfCard(K); InfCard(L) |] ==> K \<otimes> L = K \<union> L"
+lemma InfCard_cmult_eq: "\<lbrakk>InfCard(K); InfCard(L)\<rbrakk> \<Longrightarrow> K \<otimes> 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])
@@ -698,13 +698,13 @@
subset_Un_iff2 [THEN iffD1] le_imp_subset)
done
-lemma InfCard_cdouble_eq: "InfCard(K) ==> K \<oplus> K = K"
+lemma InfCard_cdouble_eq: "InfCard(K) \<Longrightarrow> K \<oplus> K = K"
apply (simp add: cmult_2 [symmetric] InfCard_is_Card cmult_commute)
apply (simp add: InfCard_le_cmult_eq InfCard_is_Limit Limit_has_0 Limit_has_succ)
done
(*Corollary 10.13 (1), for cardinal addition*)
-lemma InfCard_le_cadd_eq: "[| InfCard(K); L \<le> K |] ==> K \<oplus> L = K"
+lemma InfCard_le_cadd_eq: "\<lbrakk>InfCard(K); L \<le> K\<rbrakk> \<Longrightarrow> K \<oplus> L = K"
apply (rule le_anti_sym)
prefer 2
apply (erule ltE, blast intro: cadd_le_self InfCard_is_Card)
@@ -713,7 +713,7 @@
apply (simp add: InfCard_cdouble_eq)
done
-lemma InfCard_cadd_eq: "[| InfCard(K); InfCard(L) |] ==> K \<oplus> L = K \<union> L"
+lemma InfCard_cadd_eq: "\<lbrakk>InfCard(K); InfCard(L)\<rbrakk> \<Longrightarrow> K \<oplus> 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])
@@ -723,7 +723,7 @@
(*The other part, Corollary 10.13 (2), refers to the cardinality of the set
of all n-tuples of elements of K. A better version for the Isabelle theory
- might be InfCard(K) ==> |list(K)| = K.
+ might be InfCard(K) \<Longrightarrow> |list(K)| = K.
*)
subsection\<open>For Every Cardinal Number There Exists A Greater One\<close>
@@ -751,7 +751,7 @@
done
(*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*)
-lemma K_lt_jump_cardinal: "Ord(K) ==> K < jump_cardinal(K)"
+lemma K_lt_jump_cardinal: "Ord(K) \<Longrightarrow> K < jump_cardinal(K)"
apply (rule Ord_jump_cardinal [THEN [2] ltI])
apply (rule jump_cardinal_iff [THEN iffD2])
apply (rule_tac x="Memrel(K)" in exI)
@@ -763,9 +763,9 @@
(*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 \<subseteq> K * K; X \<subseteq> K;
- f \<in> bij(ordertype(X,r), jump_cardinal(K)) |]
- ==> jump_cardinal(K) \<in> jump_cardinal(K)"
+ "\<lbrakk>well_ord(X,r); r \<subseteq> K * K; X \<subseteq> K;
+ f \<in> bij(ordertype(X,r), jump_cardinal(K))\<rbrakk>
+ \<Longrightarrow> 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])
@@ -787,7 +787,7 @@
subsection\<open>Basic Properties of Successor Cardinals\<close>
-lemma csucc_basic: "Ord(K) ==> Card(csucc(K)) & K < csucc(K)"
+lemma csucc_basic: "Ord(K) \<Longrightarrow> Card(csucc(K)) & K < csucc(K)"
apply (unfold csucc_def)
apply (rule LeastI)
apply (blast intro: Card_jump_cardinal K_lt_jump_cardinal Ord_jump_cardinal)+
@@ -797,16 +797,16 @@
lemmas lt_csucc = csucc_basic [THEN conjunct2]
-lemma Ord_0_lt_csucc: "Ord(K) ==> 0 < csucc(K)"
+lemma Ord_0_lt_csucc: "Ord(K) \<Longrightarrow> 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: "\<lbrakk>Card(L); K<L\<rbrakk> \<Longrightarrow> 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) \<longleftrightarrow> |i| \<le> K"
+lemma lt_csucc_iff: "\<lbrakk>Ord(i); Card(K)\<rbrakk> \<Longrightarrow> i < csucc(K) \<longleftrightarrow> |i| \<le> K"
apply (rule iffI)
apply (rule_tac [2] Card_lt_imp_lt)
apply (erule_tac [2] lt_trans1)
@@ -818,10 +818,10 @@
done
lemma Card_lt_csucc_iff:
- "[| Card(K'); Card(K) |] ==> K' < csucc(K) \<longleftrightarrow> K' \<le> K"
+ "\<lbrakk>Card(K'); Card(K)\<rbrakk> \<Longrightarrow> K' < csucc(K) \<longleftrightarrow> K' \<le> K"
by (simp add: lt_csucc_iff Card_cardinal_eq Card_is_Ord)
-lemma InfCard_csucc: "InfCard(K) ==> InfCard(csucc(K))"
+lemma InfCard_csucc: "InfCard(K) \<Longrightarrow> InfCard(csucc(K))"
by (simp add: InfCard_def Card_csucc Card_is_Ord
lt_csucc [THEN leI, THEN [2] le_trans])
@@ -832,7 +832,7 @@
assumes FA: "Finite(A)" and a: "a\<notin>A" shows "|cons(a,A)| = succ(|A|)"
proof -
{ fix X
- have "Finite(X) ==> a \<notin> X \<Longrightarrow> cons(a,X) \<lesssim> X \<Longrightarrow> False"
+ have "Finite(X) \<Longrightarrow> a \<notin> X \<Longrightarrow> cons(a,X) \<lesssim> X \<Longrightarrow> False"
proof (induct X rule: Finite_induct)
case 0 thus False by (simp add: lepoll_0_iff)
next
@@ -842,7 +842,7 @@
thus False using cons by auto
qed
}
- hence [simp]: "~ cons(a,A) \<lesssim> A" using a FA by auto
+ hence [simp]: "\<not> cons(a,A) \<lesssim> A" using a FA by auto
have [simp]: "|A| \<approx> A" using Finite_imp_well_ord [OF FA]
by (blast intro: well_ord_cardinal_eqpoll)
have "(\<mu> i. i \<approx> cons(a, A)) = succ(|A|)"
@@ -864,18 +864,18 @@
qed
lemma Finite_imp_succ_cardinal_Diff:
- "[| Finite(A); a \<in> A |] ==> succ(|A-{a}|) = |A|"
+ "\<lbrakk>Finite(A); a \<in> A\<rbrakk> \<Longrightarrow> succ(|A-{a}|) = |A|"
apply (rule_tac b = A in cons_Diff [THEN subst], assumption)
apply (simp add: Finite_imp_cardinal_cons Diff_subset [THEN subset_Finite])
apply (simp add: cons_Diff)
done
-lemma Finite_imp_cardinal_Diff: "[| Finite(A); a \<in> A |] ==> |A-{a}| < |A|"
+lemma Finite_imp_cardinal_Diff: "\<lbrakk>Finite(A); a \<in> A\<rbrakk> \<Longrightarrow> |A-{a}| < |A|"
apply (rule succ_leE)
apply (simp add: Finite_imp_succ_cardinal_Diff)
done
-lemma Finite_cardinal_in_nat [simp]: "Finite(A) ==> |A| \<in> nat"
+lemma Finite_cardinal_in_nat [simp]: "Finite(A) \<Longrightarrow> |A| \<in> nat"
proof (induct rule: Finite_induct)
case 0 thus ?case by (simp add: cardinal_0)
next
@@ -883,13 +883,13 @@
qed
lemma card_Un_Int:
- "[|Finite(A); Finite(B)|] ==> |A| #+ |B| = |A \<union> B| #+ |A \<inter> B|"
+ "\<lbrakk>Finite(A); Finite(B)\<rbrakk> \<Longrightarrow> |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 \<inter> B = 0|] ==> |A \<union> B| = |A| #+ |B|"
+ "\<lbrakk>Finite(A); Finite(B); A \<inter> B = 0\<rbrakk> \<Longrightarrow> |A \<union> B| = |A| #+ |B|"
by (simp add: Finite_Un card_Un_Int)
lemma card_partition:
@@ -924,7 +924,7 @@
finally show ?thesis .
qed
-lemma Ord_subset_natD [rule_format]: "Ord(i) ==> i \<subseteq> nat \<Longrightarrow> i \<in> nat | i=nat"
+lemma Ord_subset_natD [rule_format]: "Ord(i) \<Longrightarrow> i \<subseteq> nat \<Longrightarrow> i \<in> nat | i=nat"
proof (induct i rule: trans_induct3)
case 0 thus ?case by auto
next
@@ -934,7 +934,7 @@
by (blast dest: nat_le_Limit le_imp_subset)
qed
-lemma Ord_nat_subset_into_Card: "[| Ord(i); i \<subseteq> nat |] ==> Card(i)"
+lemma Ord_nat_subset_into_Card: "\<lbrakk>Ord(i); i \<subseteq> nat\<rbrakk> \<Longrightarrow> Card(i)"
by (blast dest: Ord_subset_natD intro: Card_nat nat_into_Card)
end
--- a/src/ZF/Cardinal_AC.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Cardinal_AC.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,7 +19,7 @@
text\<open>The theorem \<^term>\<open>||A|| = |A|\<close>\<close>
lemmas cardinal_idem = cardinal_eqpoll [THEN cardinal_cong, simp]
-lemma cardinal_eqE: "|X| = |Y| ==> X \<approx> Y"
+lemma cardinal_eqE: "|X| = |Y| \<Longrightarrow> X \<approx> Y"
apply (rule AC_well_ord [THEN exE])
apply (rule AC_well_ord [THEN exE])
apply (rule well_ord_cardinal_eqE, assumption+)
@@ -29,11 +29,11 @@
by (blast intro: cardinal_cong cardinal_eqE)
lemma cardinal_disjoint_Un:
- "[| |A|=|B|; |C|=|D|; A \<inter> C = 0; B \<inter> D = 0 |]
- ==> |A \<union> C| = |B \<union> D|"
+ "\<lbrakk>|A|=|B|; |C|=|D|; A \<inter> C = 0; B \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> |A \<union> C| = |B \<union> D|"
by (simp add: cardinal_eqpoll_iff eqpoll_disjoint_Un)
-lemma lepoll_imp_cardinal_le: "A \<lesssim> B ==> |A| \<le> |B|"
+lemma lepoll_imp_cardinal_le: "A \<lesssim> B \<Longrightarrow> |A| \<le> |B|"
apply (rule AC_well_ord [THEN exE])
apply (erule well_ord_lepoll_imp_cardinal_le, assumption)
done
@@ -59,7 +59,7 @@
apply (rule well_ord_cadd_cmult_distrib, assumption+)
done
-lemma InfCard_square_eq: "InfCard(|A|) ==> A*A \<approx> A"
+lemma InfCard_square_eq: "InfCard(|A|) \<Longrightarrow> A*A \<approx> A"
apply (rule AC_well_ord [THEN exE])
apply (erule well_ord_InfCard_square_eq, assumption)
done
@@ -79,7 +79,7 @@
finally show ?thesis .
qed
-lemma le_Card_iff: "Card(K) ==> |A| \<le> K \<longleftrightarrow> A \<lesssim> K"
+lemma le_Card_iff: "Card(K) \<Longrightarrow> |A| \<le> K \<longleftrightarrow> A \<lesssim> K"
apply (erule Card_cardinal_eq [THEN subst], rule iffI,
erule Card_le_imp_lepoll)
apply (erule lepoll_imp_cardinal_le)
@@ -109,7 +109,7 @@
by (force intro: cardinal_lt_imp_lt lesspoll_cardinal_lt)
qed
-lemma cardinal_le_imp_lepoll: " i \<le> |A| ==> i \<lesssim> A"
+lemma cardinal_le_imp_lepoll: " i \<le> |A| \<Longrightarrow> i \<lesssim> A"
by (blast intro: lt_Ord Card_le_imp_lepoll Ord_cardinal_le le_trans)
@@ -141,11 +141,11 @@
text\<open>Kunen's Lemma 10.21\<close>
lemma cardinal_UN_le:
assumes K: "InfCard(K)"
- shows "(!!i. i\<in>K ==> |X(i)| \<le> K) ==> |\<Union>i\<in>K. X(i)| \<le> K"
+ shows "(\<And>i. i\<in>K \<Longrightarrow> |X(i)| \<le> K) \<Longrightarrow> |\<Union>i\<in>K. X(i)| \<le> K"
proof (simp add: K InfCard_is_Card le_Card_iff)
have [intro]: "Ord(K)" by (blast intro: InfCard_is_Card Card_is_Ord K)
- assume "!!i. i\<in>K ==> X(i) \<lesssim> K"
- hence "!!i. i\<in>K ==> \<exists>f. f \<in> inj(X(i), K)" by (simp add: lepoll_def)
+ assume "\<And>i. i\<in>K \<Longrightarrow> X(i) \<lesssim> K"
+ hence "\<And>i. i\<in>K \<Longrightarrow> \<exists>f. f \<in> inj(X(i), K)" by (simp add: lepoll_def)
with AC_Pi obtain f where f: "f \<in> (\<Prod>i\<in>K. inj(X(i), K))"
by force
{ fix z
@@ -173,15 +173,15 @@
text\<open>The same again, using \<^term>\<open>csucc\<close>\<close>
lemma cardinal_UN_lt_csucc:
- "[| InfCard(K); \<And>i. i\<in>K \<Longrightarrow> |X(i)| < csucc(K) |]
- ==> |\<Union>i\<in>K. X(i)| < csucc(K)"
+ "\<lbrakk>InfCard(K); \<And>i. i\<in>K \<Longrightarrow> |X(i)| < csucc(K)\<rbrakk>
+ \<Longrightarrow> |\<Union>i\<in>K. X(i)| < csucc(K)"
by (simp add: Card_lt_csucc_iff cardinal_UN_le InfCard_is_Card Card_cardinal)
text\<open>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).\<close>
lemma cardinal_UN_Ord_lt_csucc:
- "[| InfCard(K); \<And>i. i\<in>K \<Longrightarrow> j(i) < csucc(K) |]
- ==> (\<Union>i\<in>K. j(i)) < csucc(K)"
+ "\<lbrakk>InfCard(K); \<And>i. i\<in>K \<Longrightarrow> j(i) < csucc(K)\<rbrakk>
+ \<Longrightarrow> (\<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)
apply (blast intro!: Ord_UN elim: ltE)
--- a/src/ZF/Coind/Dynamic.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Dynamic.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,28 +11,28 @@
domains "EvalRel" \<subseteq> "ValEnv * Exp * Val"
intros
eval_constI:
- " [| ve \<in> ValEnv; c \<in> Const |] ==>
+ " \<lbrakk>ve \<in> ValEnv; c \<in> Const\<rbrakk> \<Longrightarrow>
<ve,e_const(c),v_const(c)>:EvalRel"
eval_varI:
- " [| ve \<in> ValEnv; x \<in> ExVar; x \<in> ve_dom(ve) |] ==>
+ " \<lbrakk>ve \<in> ValEnv; x \<in> ExVar; x \<in> ve_dom(ve)\<rbrakk> \<Longrightarrow>
<ve,e_var(x),ve_app(ve,x)>:EvalRel"
eval_fnI:
- " [| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp |] ==>
+ " \<lbrakk>ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp\<rbrakk> \<Longrightarrow>
<ve,e_fn(x,e),v_clos(x,e,ve)>:EvalRel "
eval_fixI:
- " [| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val;
- v_clos(x,e,ve_owr(ve,f,cl))=cl |] ==>
+ " \<lbrakk>ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val;
+ v_clos(x,e,ve_owr(ve,f,cl))=cl\<rbrakk> \<Longrightarrow>
<ve,e_fix(f,x,e),cl>:EvalRel "
eval_appI1:
- " [| ve \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; c1 \<in> Const; c2 \<in> Const;
+ " \<lbrakk>ve \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; c1 \<in> Const; c2 \<in> Const;
<ve,e1,v_const(c1)>:EvalRel;
- <ve,e2,v_const(c2)>:EvalRel |] ==>
+ <ve,e2,v_const(c2)>:EvalRel\<rbrakk> \<Longrightarrow>
<ve,e_app(e1,e2),v_const(c_app(c1,c2))>:EvalRel "
eval_appI2:
- " [| ve \<in> ValEnv; vem \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; em \<in> Exp; xm \<in> ExVar; v \<in> Val;
+ " \<lbrakk>ve \<in> ValEnv; vem \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; em \<in> Exp; xm \<in> ExVar; v \<in> Val;
<ve,e1,v_clos(xm,em,vem)>:EvalRel;
<ve,e2,v2>:EvalRel;
- <ve_owr(vem,xm,v2),em,v>:EvalRel |] ==>
+ <ve_owr(vem,xm,v2),em,v>:EvalRel\<rbrakk> \<Longrightarrow>
<ve,e_app(e1,e2),v>:EvalRel "
type_intros c_appI ve_appI ve_empI ve_owrI Exp.intros Val_ValEnv.intros
--- a/src/ZF/Coind/ECR.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/ECR.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,13 +13,13 @@
domains "HasTyRel" \<subseteq> "Val * Ty"
intros
htr_constI [intro!]:
- "[| c \<in> Const; t \<in> Ty; isof(c,t) |] ==> <v_const(c),t> \<in> HasTyRel"
+ "\<lbrakk>c \<in> Const; t \<in> Ty; isof(c,t)\<rbrakk> \<Longrightarrow> <v_const(c),t> \<in> HasTyRel"
htr_closI [intro]:
- "[| x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv;
+ "\<lbrakk>x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv;
<te,e_fn(x,e),t> \<in> ElabRel;
ve_dom(ve) = te_dom(te);
- {<ve_app(ve,y),te_app(te,y)>.y \<in> ve_dom(ve)} \<in> Pow(HasTyRel) |]
- ==> <v_clos(x,e,ve),t> \<in> HasTyRel"
+ {<ve_app(ve,y),te_app(te,y)>.y \<in> ve_dom(ve)} \<in> Pow(HasTyRel)\<rbrakk>
+ \<Longrightarrow> <v_clos(x,e,ve),t> \<in> HasTyRel"
monos Pow_mono
type_intros Val_ValEnv.intros
@@ -27,18 +27,18 @@
definition
hastyenv :: "[i,i] => o" where
- "hastyenv(ve,te) ==
+ "hastyenv(ve,te) \<equiv>
ve_dom(ve) = te_dom(te) &
(\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
(* Specialised co-induction rule *)
lemma htr_closCI [intro]:
- "[| x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv;
+ "\<lbrakk>x \<in> ExVar; e \<in> Exp; t \<in> Ty; ve \<in> ValEnv; te \<in> TyEnv;
<te, e_fn(x, e), t> \<in> ElabRel; ve_dom(ve) = te_dom(te);
{<ve_app(ve,y),te_app(te,y)>.y \<in> ve_dom(ve)} \<in>
- Pow({<v_clos(x,e,ve),t>} \<union> HasTyRel) |]
- ==> <v_clos(x, e, ve),t> \<in> HasTyRel"
+ Pow({<v_clos(x,e,ve),t>} \<union> HasTyRel)\<rbrakk>
+ \<Longrightarrow> <v_clos(x, e, ve),t> \<in> HasTyRel"
apply (rule singletonI [THEN HasTyRel.coinduct], auto)
done
@@ -55,12 +55,12 @@
HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE]
lemma hastyenv_owr:
- "[| ve \<in> ValEnv; te \<in> TyEnv; hastyenv(ve,te); <v,t> \<in> HasTyRel |]
- ==> hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))"
+ "\<lbrakk>ve \<in> ValEnv; te \<in> TyEnv; hastyenv(ve,te); <v,t> \<in> HasTyRel\<rbrakk>
+ \<Longrightarrow> hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))"
by (auto simp add: hastyenv_def ve_app_owr HasTyRel_non_zero)
lemma basic_consistency_lem:
- "[| ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te) |] ==> hastyenv(ve,te)"
+ "\<lbrakk>ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te)\<rbrakk> \<Longrightarrow> hastyenv(ve,te)"
apply (unfold isofenv_def hastyenv_def)
apply (force intro: te_appI ve_domI)
done
@@ -71,20 +71,20 @@
(* ############################################################ *)
lemma consistency_const:
- "[| c \<in> Const; hastyenv(ve,te);<te,e_const(c),t> \<in> ElabRel |]
- ==> <v_const(c), t> \<in> HasTyRel"
+ "\<lbrakk>c \<in> Const; hastyenv(ve,te);<te,e_const(c),t> \<in> ElabRel\<rbrakk>
+ \<Longrightarrow> <v_const(c), t> \<in> HasTyRel"
by blast
lemma consistency_var:
- "[| x \<in> ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t> \<in> ElabRel |] ==>
+ "\<lbrakk>x \<in> ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t> \<in> ElabRel\<rbrakk> \<Longrightarrow>
<ve_app(ve,x),t> \<in> HasTyRel"
by (unfold hastyenv_def, blast)
lemma consistency_fn:
- "[| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; hastyenv(ve,te);
+ "\<lbrakk>ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; hastyenv(ve,te);
<te,e_fn(x,e),t> \<in> ElabRel
- |] ==> <v_clos(x, e, ve), t> \<in> HasTyRel"
+\<rbrakk> \<Longrightarrow> <v_clos(x, e, ve), t> \<in> HasTyRel"
by (unfold hastyenv_def, blast)
declare ElabRel.dom_subset [THEN subsetD, dest]
@@ -94,9 +94,9 @@
declare Val_ValEnv.intros [simp, intro!]
lemma consistency_fix:
- "[| ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val;
+ "\<lbrakk>ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val;
v_clos(x,e,ve_owr(ve,f,cl)) = cl;
- hastyenv(ve,te); <te,e_fix(f,x,e),t> \<in> ElabRel |] ==>
+ hastyenv(ve,te); <te,e_fix(f,x,e),t> \<in> ElabRel\<rbrakk> \<Longrightarrow>
<cl,t> \<in> HasTyRel"
apply (unfold hastyenv_def)
apply (erule elab_fixE, safe)
@@ -111,7 +111,7 @@
lemma consistency_app1:
- "[| ve \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; c1 \<in> Const; c2 \<in> Const;
+ "\<lbrakk>ve \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; c1 \<in> Const; c2 \<in> Const;
<ve,e1,v_const(c1)> \<in> EvalRel;
\<forall>t te.
hastyenv(ve,te) \<longrightarrow> <te,e1,t> \<in> ElabRel \<longrightarrow> <v_const(c1),t> \<in> HasTyRel;
@@ -119,12 +119,12 @@
\<forall>t te.
hastyenv(ve,te) \<longrightarrow> <te,e2,t> \<in> ElabRel \<longrightarrow> <v_const(c2),t> \<in> HasTyRel;
hastyenv(ve, te);
- <te,e_app(e1,e2),t> \<in> ElabRel |]
- ==> <v_const(c_app(c1, c2)),t> \<in> HasTyRel"
+ <te,e_app(e1,e2),t> \<in> ElabRel\<rbrakk>
+ \<Longrightarrow> <v_const(c_app(c1, c2)),t> \<in> HasTyRel"
by (blast intro!: c_appI intro: isof_app)
lemma consistency_app2:
- "[| ve \<in> ValEnv; vem \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; em \<in> Exp; xm \<in> ExVar;
+ "\<lbrakk>ve \<in> ValEnv; vem \<in> ValEnv; e1 \<in> Exp; e2 \<in> Exp; em \<in> Exp; xm \<in> ExVar;
v \<in> Val;
<ve,e1,v_clos(xm,em,vem)> \<in> EvalRel;
\<forall>t te. hastyenv(ve,te) \<longrightarrow>
@@ -134,8 +134,8 @@
<ve_owr(vem,xm,v2),em,v> \<in> EvalRel;
\<forall>t te. hastyenv(ve_owr(vem,xm,v2),te) \<longrightarrow>
<te,em,t> \<in> ElabRel \<longrightarrow> <v,t> \<in> HasTyRel;
- hastyenv(ve,te); <te,e_app(e1,e2),t> \<in> ElabRel |]
- ==> <v,t> \<in> HasTyRel"
+ hastyenv(ve,te); <te,e_app(e1,e2),t> \<in> ElabRel\<rbrakk>
+ \<Longrightarrow> <v,t> \<in> HasTyRel"
apply (erule elab_appE)
apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+)
apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+)
@@ -151,7 +151,7 @@
lemma consistency [rule_format]:
"<ve,e,v> \<in> EvalRel
- ==> (\<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e,t> \<in> ElabRel \<longrightarrow> <v,t> \<in> HasTyRel)"
+ \<Longrightarrow> (\<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e,t> \<in> ElabRel \<longrightarrow> <v,t> \<in> HasTyRel)"
apply (erule EvalRel.induct)
apply (simp_all add: consistency_const consistency_var consistency_fn
consistency_fix consistency_app1)
@@ -159,9 +159,9 @@
done
lemma basic_consistency:
- "[| ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te);
- <ve,e,v_const(c)> \<in> EvalRel; <te,e,t> \<in> ElabRel |]
- ==> isof(c,t)"
+ "\<lbrakk>ve \<in> ValEnv; te \<in> TyEnv; isofenv(ve,te);
+ <ve,e,v_const(c)> \<in> EvalRel; <te,e,t> \<in> ElabRel\<rbrakk>
+ \<Longrightarrow> isof(c,t)"
by (blast dest: consistency intro!: basic_consistency_lem)
end
--- a/src/ZF/Coind/Language.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Language.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,8 +12,8 @@
Const :: i and (* Abstract type of constants *)
c_app :: "[i,i] => i" (* Abstract constructor for fun application*)
where
- constNEE: "c \<in> Const ==> c \<noteq> 0" and
- c_appI: "[| c1 \<in> Const; c2 \<in> Const |] ==> c_app(c1,c2) \<in> Const"
+ constNEE: "c \<in> Const \<Longrightarrow> c \<noteq> 0" and
+ c_appI: "\<lbrakk>c1 \<in> Const; c2 \<in> Const\<rbrakk> \<Longrightarrow> c_app(c1,c2) \<in> Const"
consts
--- a/src/ZF/Coind/Map.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Map.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,25 +9,25 @@
definition
TMap :: "[i,i] => i" where
- "TMap(A,B) == {m \<in> Pow(A*\<Union>(B)).\<forall>a \<in> A. m``{a} \<in> B}"
+ "TMap(A,B) \<equiv> {m \<in> Pow(A*\<Union>(B)).\<forall>a \<in> A. m``{a} \<in> B}"
definition
PMap :: "[i,i] => i" where
- "PMap(A,B) == TMap(A,cons(0,B))"
+ "PMap(A,B) \<equiv> TMap(A,cons(0,B))"
-(* Note: 0 \<in> B ==> TMap(A,B) = PMap(A,B) *)
+(* Note: 0 \<in> B \<Longrightarrow> TMap(A,B) = PMap(A,B) *)
definition
map_emp :: i where
- "map_emp == 0"
+ "map_emp \<equiv> 0"
definition
map_owr :: "[i,i,i]=>i" where
- "map_owr(m,a,b) == \<Sum>x \<in> {a} \<union> domain(m). if x=a then b else m``{x}"
+ "map_owr(m,a,b) \<equiv> \<Sum>x \<in> {a} \<union> domain(m). if x=a then b else m``{x}"
definition
map_app :: "[i,i]=>i" where
- "map_app(m,a) == m``{a}"
+ "map_app(m,a) \<equiv> m``{a}"
lemma "{0,1} \<subseteq> {1} \<union> TMap(I, {0,1})"
by (unfold TMap_def, blast)
@@ -51,13 +51,13 @@
lemma qbeta_if: "Sigma(A,B)``{a} = (if a \<in> A then B(a) else 0)"
by auto
-lemma qbeta: "a \<in> A ==> Sigma(A,B)``{a} = B(a)"
+lemma qbeta: "a \<in> A \<Longrightarrow> Sigma(A,B)``{a} = B(a)"
by fast
-lemma qbeta_emp: "a\<notin>A ==> Sigma(A,B)``{a} = 0"
+lemma qbeta_emp: "a\<notin>A \<Longrightarrow> Sigma(A,B)``{a} = 0"
by fast
-lemma image_Sigma1: "a \<notin> A ==> Sigma(A,B)``{a}=0"
+lemma image_Sigma1: "a \<notin> A \<Longrightarrow> Sigma(A,B)``{a}=0"
by fast
@@ -68,7 +68,7 @@
(* Lemmas *)
lemma MapQU_lemma:
- "A \<subseteq> univ(X) ==> Pow(A * \<Union>(quniv(X))) \<subseteq> quniv(X)"
+ "A \<subseteq> univ(X) \<Longrightarrow> Pow(A * \<Union>(quniv(X))) \<subseteq> quniv(X)"
apply (unfold quniv_def)
apply (rule Pow_mono)
apply (rule subset_trans [OF Sigma_mono product_univ])
@@ -80,7 +80,7 @@
(* Theorems *)
lemma mapQU:
- "[| m \<in> PMap(A,quniv(B)); !!x. x \<in> A ==> x \<in> univ(B) |] ==> m \<in> quniv(B)"
+ "\<lbrakk>m \<in> PMap(A,quniv(B)); \<And>x. x \<in> A \<Longrightarrow> x \<in> univ(B)\<rbrakk> \<Longrightarrow> m \<in> quniv(B)"
apply (unfold PMap_def TMap_def)
apply (blast intro!: MapQU_lemma [THEN subsetD])
done
@@ -89,7 +89,7 @@
(* Monotonicity *)
(* ############################################################ *)
-lemma PMap_mono: "B \<subseteq> C ==> PMap(A,B)<=PMap(A,C)"
+lemma PMap_mono: "B \<subseteq> C \<Longrightarrow> PMap(A,B)<=PMap(A,C)"
by (unfold PMap_def TMap_def, blast)
@@ -105,7 +105,7 @@
(** map_owr **)
lemma pmap_owrI:
- "[| m \<in> PMap(A,B); a \<in> A; b \<in> B |] ==> map_owr(m,a,b):PMap(A,B)"
+ "\<lbrakk>m \<in> PMap(A,B); a \<in> A; b \<in> B\<rbrakk> \<Longrightarrow> map_owr(m,a,b):PMap(A,B)"
apply (unfold map_owr_def PMap_def TMap_def, safe)
apply (simp_all add: if_iff, auto)
(*Remaining subgoal*)
@@ -118,15 +118,15 @@
(** map_app **)
lemma tmap_app_notempty:
- "[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> map_app(m,a) \<noteq>0"
+ "\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a) \<noteq>0"
by (unfold TMap_def map_app_def, blast)
lemma tmap_appI:
- "[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> map_app(m,a):B"
+ "\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a):B"
by (unfold TMap_def map_app_def domain_def, blast)
lemma pmap_appI:
- "[| m \<in> PMap(A,B); a \<in> domain(m) |] ==> map_app(m,a):B"
+ "\<lbrakk>m \<in> PMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a):B"
apply (unfold PMap_def)
apply (frule tmap_app_notempty, assumption)
apply (drule tmap_appI, auto)
@@ -135,11 +135,11 @@
(** domain **)
lemma tmap_domainD:
- "[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> a \<in> A"
+ "\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> a \<in> A"
by (unfold TMap_def, blast)
lemma pmap_domainD:
- "[| m \<in> PMap(A,B); a \<in> domain(m) |] ==> a \<in> A"
+ "\<lbrakk>m \<in> PMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> a \<in> A"
by (unfold PMap_def TMap_def, blast)
@@ -164,7 +164,7 @@
by (unfold map_emp_def, blast)
lemma map_domain_owr:
- "b \<noteq> 0 ==> domain(map_owr(f,a,b)) = {a} \<union> domain(f)"
+ "b \<noteq> 0 \<Longrightarrow> domain(map_owr(f,a,b)) = {a} \<union> domain(f)"
apply (unfold map_owr_def)
apply (auto simp add: domain_Sigma)
done
@@ -178,7 +178,7 @@
lemma map_app_owr1: "map_app(map_owr(f,a,b),a) = b"
by (simp add: map_app_owr)
-lemma map_app_owr2: "c \<noteq> a ==> map_app(map_owr(f,a,b),c)= map_app(f,c)"
+lemma map_app_owr2: "c \<noteq> a \<Longrightarrow> map_app(map_owr(f,a,b),c)= map_app(f,c)"
by (simp add: map_app_owr)
end
--- a/src/ZF/Coind/Static.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Static.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,13 +10,13 @@
***)
axiomatization isof :: "[i,i] => o"
- where isof_app: "[|isof(c1,t_fun(t1,t2)); isof(c2,t1)|] ==> isof(c_app(c1,c2),t2)"
+ where isof_app: "\<lbrakk>isof(c1,t_fun(t1,t2)); isof(c2,t1)\<rbrakk> \<Longrightarrow> isof(c_app(c1,c2),t2)"
(*Its extension to environments*)
definition
isofenv :: "[i,i] => o" where
- "isofenv(ve,te) ==
+ "isofenv(ve,te) \<equiv>
ve_dom(ve) = te_dom(te) &
(\<forall>x \<in> ve_dom(ve).
\<exists>c \<in> Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))"
@@ -30,23 +30,23 @@
domains "ElabRel" \<subseteq> "TyEnv * Exp * Ty"
intros
constI [intro!]:
- "[| te \<in> TyEnv; c \<in> Const; t \<in> Ty; isof(c,t) |] ==>
+ "\<lbrakk>te \<in> TyEnv; c \<in> Const; t \<in> Ty; isof(c,t)\<rbrakk> \<Longrightarrow>
<te,e_const(c),t> \<in> ElabRel"
varI [intro!]:
- "[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==>
+ "\<lbrakk>te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te)\<rbrakk> \<Longrightarrow>
<te,e_var(x),te_app(te,x)> \<in> ElabRel"
fnI [intro!]:
- "[| te \<in> TyEnv; x \<in> ExVar; e \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
- <te_owr(te,x,t1),e,t2> \<in> ElabRel |] ==>
+ "\<lbrakk>te \<in> TyEnv; x \<in> ExVar; e \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
+ <te_owr(te,x,t1),e,t2> \<in> ElabRel\<rbrakk> \<Longrightarrow>
<te,e_fn(x,e),t_fun(t1,t2)> \<in> ElabRel"
fixI [intro!]:
- "[| te \<in> TyEnv; f \<in> ExVar; x \<in> ExVar; t1 \<in> Ty; t2 \<in> Ty;
- <te_owr(te_owr(te,f,t_fun(t1,t2)),x,t1),e,t2> \<in> ElabRel |] ==>
+ "\<lbrakk>te \<in> TyEnv; f \<in> ExVar; x \<in> ExVar; t1 \<in> Ty; t2 \<in> Ty;
+ <te_owr(te_owr(te,f,t_fun(t1,t2)),x,t1),e,t2> \<in> ElabRel\<rbrakk> \<Longrightarrow>
<te,e_fix(f,x,e),t_fun(t1,t2)> \<in> ElabRel"
appI [intro]:
- "[| te \<in> TyEnv; e1 \<in> Exp; e2 \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
+ "\<lbrakk>te \<in> TyEnv; e1 \<in> Exp; e2 \<in> Exp; t1 \<in> Ty; t2 \<in> Ty;
<te,e1,t_fun(t1,t2)> \<in> ElabRel;
- <te,e2,t1> \<in> ElabRel |] ==> <te,e_app(e1,e2),t2> \<in> ElabRel"
+ <te,e2,t1> \<in> ElabRel\<rbrakk> \<Longrightarrow> <te,e_app(e1,e2),t2> \<in> ElabRel"
type_intros te_appI Exp.intros Ty.intros
--- a/src/ZF/Coind/Types.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Types.thy Tue Sep 27 16:51:35 2022 +0100
@@ -43,7 +43,7 @@
by simp
(*redundant??*)
-lemma te_app_owr2: "x \<noteq> y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)"
+lemma te_app_owr2: "x \<noteq> y \<Longrightarrow> te_app(te_owr(te,x,t),y) = te_app(te,y)"
by auto
lemma te_app_owr [simp]:
@@ -51,7 +51,7 @@
by auto
lemma te_appI:
- "[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==> te_app(te,x) \<in> Ty"
+ "\<lbrakk>te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te)\<rbrakk> \<Longrightarrow> te_app(te,x) \<in> Ty"
apply (erule_tac P = "x \<in> te_dom (te) " in rev_mp)
apply (erule TyEnv.induct, auto)
done
--- a/src/ZF/Coind/Values.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Coind/Values.thy Tue Sep 27 16:51:35 2022 +0100
@@ -34,24 +34,24 @@
definition
ve_emp :: i where
- "ve_emp == ve_mk(map_emp)"
+ "ve_emp \<equiv> ve_mk(map_emp)"
(* Elimination rules *)
lemma ValEnvE:
- "[| ve \<in> ValEnv; !!m.[| ve=ve_mk(m); m \<in> PMap(ExVar,Val) |] ==> Q |] ==> Q"
+ "\<lbrakk>ve \<in> ValEnv; \<And>m.\<lbrakk>ve=ve_mk(m); m \<in> PMap(ExVar,Val)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
apply (unfold Part_def Val_def ValEnv_def, clarify)
apply (erule Val_ValEnv.cases)
apply (auto simp add: Val_def Part_def Val_ValEnv.con_defs)
done
lemma ValE:
- "[| v \<in> Val;
- !!c. [| v = v_const(c); c \<in> Const |] ==> Q;
- !!e ve x.
- [| v = v_clos(x,e,ve); x \<in> ExVar; e \<in> Exp; ve \<in> ValEnv |] ==> Q
- |] ==>
+ "\<lbrakk>v \<in> Val;
+ \<And>c. \<lbrakk>v = v_const(c); c \<in> Const\<rbrakk> \<Longrightarrow> Q;
+ \<And>e ve x.
+ \<lbrakk>v = v_clos(x,e,ve); x \<in> ExVar; e \<in> Exp; ve \<in> ValEnv\<rbrakk> \<Longrightarrow> Q
+\<rbrakk> \<Longrightarrow>
Q"
apply (unfold Part_def Val_def ValEnv_def, clarify)
apply (erule Val_ValEnv.cases)
@@ -66,7 +66,7 @@
declare v_closNE [THEN notE, elim!]
-lemma v_constNE [simp]: "c \<in> Const ==> v_const(c) \<noteq> 0"
+lemma v_constNE [simp]: "c \<in> Const \<Longrightarrow> v_const(c) \<noteq> 0"
apply (unfold QPair_def QInl_def QInr_def Val_ValEnv.con_defs)
apply (drule constNEE, auto)
done
@@ -74,28 +74,28 @@
(* Proving that the empty set is not a value *)
-lemma ValNEE: "v \<in> Val ==> v \<noteq> 0"
+lemma ValNEE: "v \<in> Val \<Longrightarrow> v \<noteq> 0"
by (erule ValE, auto)
(* Equalities for value environments *)
lemma ve_dom_owr [simp]:
- "[| ve \<in> ValEnv; v \<noteq>0 |] ==> ve_dom(ve_owr(ve,x,v)) = ve_dom(ve) \<union> {x}"
+ "\<lbrakk>ve \<in> ValEnv; v \<noteq>0\<rbrakk> \<Longrightarrow> ve_dom(ve_owr(ve,x,v)) = ve_dom(ve) \<union> {x}"
apply (erule ValEnvE)
apply (auto simp add: map_domain_owr)
done
lemma ve_app_owr [simp]:
"ve \<in> ValEnv
- ==> ve_app(ve_owr(ve,y,v),x) = (if x=y then v else ve_app(ve,x))"
+ \<Longrightarrow> ve_app(ve_owr(ve,y,v),x) = (if x=y then v else ve_app(ve,x))"
by (erule ValEnvE, simp add: map_app_owr)
(* Introduction rules for operators on value environments *)
-lemma ve_appI: "[| ve \<in> ValEnv; x \<in> ve_dom(ve) |] ==> ve_app(ve,x):Val"
+lemma ve_appI: "\<lbrakk>ve \<in> ValEnv; x \<in> ve_dom(ve)\<rbrakk> \<Longrightarrow> ve_app(ve,x):Val"
by (erule ValEnvE, simp add: pmap_appI)
-lemma ve_domI: "[| ve \<in> ValEnv; x \<in> ve_dom(ve) |] ==> x \<in> ExVar"
+lemma ve_domI: "\<lbrakk>ve \<in> ValEnv; x \<in> ve_dom(ve)\<rbrakk> \<Longrightarrow> x \<in> ExVar"
apply (erule ValEnvE, simp)
apply (blast dest: pmap_domainD)
done
@@ -107,7 +107,7 @@
done
lemma ve_owrI:
- "[|ve \<in> ValEnv; x \<in> ExVar; v \<in> Val |] ==> ve_owr(ve,x,v):ValEnv"
+ "\<lbrakk>ve \<in> ValEnv; x \<in> ExVar; v \<in> Val\<rbrakk> \<Longrightarrow> ve_owr(ve,x,v):ValEnv"
apply (erule ValEnvE, simp)
apply (blast intro: pmap_owrI Val_ValEnv.intros)
done
--- a/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,17 +17,17 @@
domains "rlist(A,r)" \<subseteq> "list(A) * list(A)"
intros
shorterI:
- "[| length(l') < length(l); l' \<in> list(A); l \<in> list(A) |]
- ==> <l', l> \<in> rlist(A,r)"
+ "\<lbrakk>length(l') < length(l); l' \<in> list(A); l \<in> list(A)\<rbrakk>
+ \<Longrightarrow> <l', l> \<in> rlist(A,r)"
sameI:
- "[| <l',l> \<in> rlist(A,r); a \<in> A |]
- ==> <Cons(a,l'), Cons(a,l)> \<in> rlist(A,r)"
+ "\<lbrakk><l',l> \<in> rlist(A,r); a \<in> A\<rbrakk>
+ \<Longrightarrow> <Cons(a,l'), Cons(a,l)> \<in> rlist(A,r)"
diffI:
- "[| length(l') = length(l); <a',a> \<in> r;
- l' \<in> list(A); l \<in> list(A); a' \<in> A; a \<in> A |]
- ==> <Cons(a',l'), Cons(a,l)> \<in> rlist(A,r)"
+ "\<lbrakk>length(l') = length(l); <a',a> \<in> r;
+ l' \<in> list(A); l \<in> list(A); a' \<in> A; a \<in> A\<rbrakk>
+ \<Longrightarrow> <Cons(a',l'), Cons(a,l)> \<in> rlist(A,r)"
type_intros list.intros
@@ -40,7 +40,7 @@
subsubsection\<open>Linearity\<close>
lemma rlist_Nil_Cons [intro]:
- "[|a \<in> A; l \<in> list(A)|] ==> <[], Cons(a,l)> \<in> rlist(A, r)"
+ "\<lbrakk>a \<in> A; l \<in> list(A)\<rbrakk> \<Longrightarrow> <[], Cons(a,l)> \<in> rlist(A, r)"
by (simp add: shorterI)
lemma linear_rlist:
@@ -82,13 +82,13 @@
lemma not_rlist_Nil [simp]: " <l,[]> \<notin> rlist(A,r)"
by (blast intro: elim: rlist_NilE)
-lemma rlist_imp_length_le: "<l',l> \<in> rlist(A,r) ==> length(l') \<le> length(l)"
+lemma rlist_imp_length_le: "<l',l> \<in> rlist(A,r) \<Longrightarrow> length(l') \<le> length(l)"
apply (erule rlist.induct)
apply (simp_all add: leI)
done
lemma wf_on_rlist_n:
- "[| n \<in> nat; wf[A](r) |] ==> wf[{l \<in> list(A). length(l) = n}](rlist(A,r))"
+ "\<lbrakk>n \<in> nat; wf[A](r)\<rbrakk> \<Longrightarrow> wf[{l \<in> list(A). length(l) = n}](rlist(A,r))"
apply (induct_tac n)
apply (rule wf_onI2, simp)
apply (rule wf_onI2, clarify)
@@ -112,7 +112,7 @@
by (blast intro: length_type)
-lemma wf_on_rlist: "wf[A](r) ==> wf[list(A)](rlist(A,r))"
+lemma wf_on_rlist: "wf[A](r) \<Longrightarrow> wf[list(A)](rlist(A,r))"
apply (subst list_eq_UN_length)
apply (rule wf_on_Union)
apply (rule wf_imp_wf_on [OF wf_Memrel [of nat]])
@@ -125,14 +125,14 @@
done
-lemma wf_rlist: "wf(r) ==> wf(rlist(field(r),r))"
+lemma wf_rlist: "wf(r) \<Longrightarrow> wf(rlist(field(r),r))"
apply (simp add: wf_iff_wf_on_field)
apply (rule wf_on_subset_A [OF _ field_rlist])
apply (blast intro: wf_on_rlist)
done
lemma well_ord_rlist:
- "well_ord(A,r) ==> well_ord(list(A), rlist(A,r))"
+ "well_ord(A,r) \<Longrightarrow> well_ord(list(A), rlist(A,r))"
apply (rule well_ordI)
apply (simp add: well_ord_def wf_on_rlist)
apply (simp add: well_ord_def tot_ord_def linear_rlist)
@@ -167,20 +167,20 @@
"enum(f, Forall(p)) = f ` <succ(2), enum(f,p)>"
lemma (in Nat_Times_Nat) fn_type [TC,simp]:
- "[|x \<in> nat; y \<in> nat|] ==> fn`<x,y> \<in> nat"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> fn`<x,y> \<in> nat"
by (blast intro: inj_is_fun [OF fn_inj] apply_funtype)
lemma (in Nat_Times_Nat) fn_iff:
- "[|x \<in> nat; y \<in> nat; u \<in> nat; v \<in> nat|]
- ==> (fn`<x,y> = fn`<u,v>) \<longleftrightarrow> (x=u & y=v)"
+ "\<lbrakk>x \<in> nat; y \<in> nat; u \<in> nat; v \<in> nat\<rbrakk>
+ \<Longrightarrow> (fn`<x,y> = fn`<u,v>) \<longleftrightarrow> (x=u & y=v)"
by (blast dest: inj_apply_equality [OF fn_inj])
lemma (in Nat_Times_Nat) enum_type [TC,simp]:
- "p \<in> formula ==> enum(fn,p) \<in> nat"
+ "p \<in> formula \<Longrightarrow> enum(fn,p) \<in> nat"
by (induct_tac p, simp_all)
lemma (in Nat_Times_Nat) enum_inject [rule_format]:
- "p \<in> formula ==> \<forall>q\<in>formula. enum(fn,p) = enum(fn,q) \<longrightarrow> p=q"
+ "p \<in> formula \<Longrightarrow> \<forall>q\<in>formula. enum(fn,p) = enum(fn,q) \<longrightarrow> p=q"
apply (induct_tac p, simp_all)
apply (rule ballI)
apply (erule formula.cases)
@@ -232,7 +232,7 @@
definition
env_form_r :: "[i,i,i]=>i" where
\<comment> \<open>wellordering on (environment, formula) pairs\<close>
- "env_form_r(f,r,A) ==
+ "env_form_r(f,r,A) \<equiv>
rmult(list(A), rlist(A, r),
formula, measure(formula, enum(f)))"
@@ -240,12 +240,12 @@
env_form_map :: "[i,i,i,i]=>i" where
\<comment> \<open>map from (environment, formula) pairs to ordinals\<close>
"env_form_map(f,r,A,z)
- == ordermap(list(A) * formula, env_form_r(f,r,A)) ` z"
+ \<equiv> ordermap(list(A) * formula, env_form_r(f,r,A)) ` z"
definition
DPow_ord :: "[i,i,i,i,i]=>o" where
\<comment> \<open>predicate that holds if \<^term>\<open>k\<close> is a valid index for \<^term>\<open>X\<close>\<close>
- "DPow_ord(f,r,A,X,k) ==
+ "DPow_ord(f,r,A,X,k) \<equiv>
\<exists>env \<in> list(A). \<exists>p \<in> formula.
arity(p) \<le> succ(length(env)) &
X = {x\<in>A. sats(A, p, Cons(x,env))} &
@@ -254,61 +254,61 @@
definition
DPow_least :: "[i,i,i,i]=>i" where
\<comment> \<open>function yielding the smallest index for \<^term>\<open>X\<close>\<close>
- "DPow_least(f,r,A,X) == \<mu> k. DPow_ord(f,r,A,X,k)"
+ "DPow_least(f,r,A,X) \<equiv> \<mu> k. DPow_ord(f,r,A,X,k)"
definition
DPow_r :: "[i,i,i]=>i" where
\<comment> \<open>a wellordering on \<^term>\<open>DPow(A)\<close>\<close>
- "DPow_r(f,r,A) == measure(DPow(A), DPow_least(f,r,A))"
+ "DPow_r(f,r,A) \<equiv> measure(DPow(A), DPow_least(f,r,A))"
lemma (in Nat_Times_Nat) well_ord_env_form_r:
"well_ord(A,r)
- ==> well_ord(list(A) * formula, env_form_r(fn,r,A))"
+ \<Longrightarrow> well_ord(list(A) * formula, env_form_r(fn,r,A))"
by (simp add: env_form_r_def well_ord_rmult well_ord_rlist well_ord_formula)
lemma (in Nat_Times_Nat) Ord_env_form_map:
- "[|well_ord(A,r); z \<in> list(A) * formula|]
- ==> Ord(env_form_map(fn,r,A,z))"
+ "\<lbrakk>well_ord(A,r); z \<in> list(A) * formula\<rbrakk>
+ \<Longrightarrow> Ord(env_form_map(fn,r,A,z))"
by (simp add: env_form_map_def Ord_ordermap well_ord_env_form_r)
lemma DPow_imp_ex_DPow_ord:
- "X \<in> DPow(A) ==> \<exists>k. DPow_ord(fn,r,A,X,k)"
+ "X \<in> DPow(A) \<Longrightarrow> \<exists>k. DPow_ord(fn,r,A,X,k)"
apply (simp add: DPow_ord_def)
apply (blast dest!: DPowD)
done
lemma (in Nat_Times_Nat) DPow_ord_imp_Ord:
- "[|DPow_ord(fn,r,A,X,k); well_ord(A,r)|] ==> Ord(k)"
+ "\<lbrakk>DPow_ord(fn,r,A,X,k); well_ord(A,r)\<rbrakk> \<Longrightarrow> Ord(k)"
apply (simp add: DPow_ord_def, clarify)
apply (simp add: Ord_env_form_map)
done
lemma (in Nat_Times_Nat) DPow_imp_DPow_least:
- "[|X \<in> DPow(A); well_ord(A,r)|]
- ==> DPow_ord(fn, r, A, X, DPow_least(fn,r,A,X))"
+ "\<lbrakk>X \<in> DPow(A); well_ord(A,r)\<rbrakk>
+ \<Longrightarrow> DPow_ord(fn, r, A, X, DPow_least(fn,r,A,X))"
apply (simp add: DPow_least_def)
apply (blast dest: DPow_imp_ex_DPow_ord intro: DPow_ord_imp_Ord LeastI)
done
lemma (in Nat_Times_Nat) env_form_map_inject:
- "[|env_form_map(fn,r,A,u) = env_form_map(fn,r,A,v); well_ord(A,r);
- u \<in> list(A) * formula; v \<in> list(A) * formula|]
- ==> u=v"
+ "\<lbrakk>env_form_map(fn,r,A,u) = env_form_map(fn,r,A,v); well_ord(A,r);
+ u \<in> list(A) * formula; v \<in> list(A) * formula\<rbrakk>
+ \<Longrightarrow> u=v"
apply (simp add: env_form_map_def)
apply (rule inj_apply_equality [OF bij_is_inj, OF ordermap_bij,
OF well_ord_env_form_r], assumption+)
done
lemma (in Nat_Times_Nat) DPow_ord_unique:
- "[|DPow_ord(fn,r,A,X,k); DPow_ord(fn,r,A,Y,k); well_ord(A,r)|]
- ==> X=Y"
+ "\<lbrakk>DPow_ord(fn,r,A,X,k); DPow_ord(fn,r,A,Y,k); well_ord(A,r)\<rbrakk>
+ \<Longrightarrow> X=Y"
apply (simp add: DPow_ord_def, clarify)
apply (drule env_form_map_inject, auto)
done
lemma (in Nat_Times_Nat) well_ord_DPow_r:
- "well_ord(A,r) ==> well_ord(DPow(A), DPow_r(fn,r,A))"
+ "well_ord(A,r) \<Longrightarrow> well_ord(DPow(A), DPow_r(fn,r,A))"
apply (simp add: DPow_r_def)
apply (rule well_ord_measure)
apply (simp add: DPow_least_def)
@@ -332,7 +332,7 @@
rlimit :: "[i,i=>i]=>i" where
\<comment> \<open>Expresses the wellordering at limit ordinals. The conditional
lets us remove the premise \<^term>\<open>Limit(i)\<close> from some theorems.\<close>
- "rlimit(i,r) ==
+ "rlimit(i,r) \<equiv>
if Limit(i) then
{z: Lset(i) * Lset(i).
\<exists>x' x. z = <x',x> &
@@ -344,10 +344,10 @@
Lset_new :: "i=>i" where
\<comment> \<open>This constant denotes the set of elements introduced at level
\<^term>\<open>succ(i)\<close>\<close>
- "Lset_new(i) == {x \<in> Lset(succ(i)). lrank(x) = i}"
+ "Lset_new(i) \<equiv> {x \<in> Lset(succ(i)). lrank(x) = i}"
lemma Limit_Lset_eq2:
- "Limit(i) ==> Lset(i) = (\<Union>j\<in>i. Lset_new(j))"
+ "Limit(i) \<Longrightarrow> Lset(i) = (\<Union>j\<in>i. Lset_new(j))"
apply (simp add: Limit_Lset_eq)
apply (rule equalityI)
apply safe
@@ -362,14 +362,14 @@
done
lemma wf_on_Lset:
- "wf[Lset(succ(j))](r(succ(j))) ==> wf[Lset_new(j)](rlimit(i,r))"
+ "wf[Lset(succ(j))](r(succ(j))) \<Longrightarrow> wf[Lset_new(j)](rlimit(i,r))"
apply (simp add: wf_on_def Lset_new_def)
apply (erule wf_subset)
apply (simp add: rlimit_def, force)
done
lemma wf_on_rlimit:
- "(\<forall>j<i. wf[Lset(j)](r(j))) ==> wf[Lset(i)](rlimit(i,r))"
+ "(\<forall>j<i. wf[Lset(j)](r(j))) \<Longrightarrow> wf[Lset(i)](rlimit(i,r))"
apply (case_tac "Limit(i)")
prefer 2
apply (simp add: rlimit_def wf_on_any_0)
@@ -382,8 +382,8 @@
done
lemma linear_rlimit:
- "[|Limit(i); \<forall>j<i. linear(Lset(j), r(j)) |]
- ==> linear(Lset(i), rlimit(i,r))"
+ "\<lbrakk>Limit(i); \<forall>j<i. linear(Lset(j), r(j))\<rbrakk>
+ \<Longrightarrow> linear(Lset(i), rlimit(i,r))"
apply (frule Limit_is_Ord)
apply (simp add: Limit_Lset_eq2 Lset_new_def)
apply (simp add: linear_def rlimit_def Ball_def lt_Ord Lset_iff_lrank_lt)
@@ -397,13 +397,13 @@
done
lemma well_ord_rlimit:
- "[|Limit(i); \<forall>j<i. well_ord(Lset(j), r(j)) |]
- ==> well_ord(Lset(i), rlimit(i,r))"
+ "\<lbrakk>Limit(i); \<forall>j<i. well_ord(Lset(j), r(j))\<rbrakk>
+ \<Longrightarrow> well_ord(Lset(i), rlimit(i,r))"
by (blast intro: well_ordI wf_on_rlimit well_ord_is_wf
linear_rlimit well_ord_is_linear)
lemma rlimit_cong:
- "(!!j. j<i ==> r'(j) = r(j)) ==> rlimit(i,r) = rlimit(i,r')"
+ "(\<And>j. j<i \<Longrightarrow> r'(j) = r(j)) \<Longrightarrow> rlimit(i,r) = rlimit(i,r')"
apply (simp add: rlimit_def, clarify)
apply (rule refl iff_refl Collect_cong ex_cong conj_cong)+
apply (simp add: Limit_is_Ord Lset_lrank_lt)
@@ -414,7 +414,7 @@
definition
L_r :: "[i, i] => i" where
- "L_r(f) == %i.
+ "L_r(f) \<equiv> %i.
transrec3(i, 0, \<lambda>x r. DPow_r(f, r, Lset(x)),
\<lambda>x r. rlimit(x, \<lambda>y. r`y))"
@@ -427,25 +427,25 @@
text\<open>The limit case is non-trivial because of the distinction between
object-level and meta-level abstraction.\<close>
-lemma [simp]: "Limit(i) ==> L_r(f,i) = rlimit(i, L_r(f))"
+lemma [simp]: "Limit(i) \<Longrightarrow> L_r(f,i) = rlimit(i, L_r(f))"
by (simp cong: rlimit_cong add: transrec3_Limit L_r_def ltD)
lemma (in Nat_Times_Nat) L_r_type:
- "Ord(i) ==> L_r(fn,i) \<subseteq> Lset(i) * Lset(i)"
+ "Ord(i) \<Longrightarrow> L_r(fn,i) \<subseteq> Lset(i) * Lset(i)"
apply (induct i rule: trans_induct3)
apply (simp_all add: Lset_succ DPow_r_type well_ord_DPow_r rlimit_def
Transset_subset_DPow [OF Transset_Lset], blast)
done
lemma (in Nat_Times_Nat) well_ord_L_r:
- "Ord(i) ==> well_ord(Lset(i), L_r(fn,i))"
+ "Ord(i) \<Longrightarrow> well_ord(Lset(i), L_r(fn,i))"
apply (induct i rule: trans_induct3)
apply (simp_all add: well_ord0 Lset_succ L_r_type well_ord_DPow_r
well_ord_rlimit ltD)
done
lemma well_ord_L_r:
- "Ord(i) ==> \<exists>r. well_ord(Lset(i), r)"
+ "Ord(i) \<Longrightarrow> \<exists>r. well_ord(Lset(i), r)"
apply (insert nat_times_nat_lepoll_nat)
apply (unfold lepoll_def)
apply (blast intro: Nat_Times_Nat.well_ord_L_r Nat_Times_Nat.intro)
@@ -453,7 +453,7 @@
text\<open>Every constructible set is well-ordered! Therefore the Wellordering Theorem and
- the Axiom of Choice hold in \<^term>\<open>L\<close>!!\<close>
+ the Axiom of Choice hold in \<^term>\<open>L\<close>\<And>\<close>
theorem L_implies_AC: assumes x: "L(x)" shows "\<exists>r. well_ord(x,r)"
using Transset_Lset x
apply (simp add: Transset_def L_def)
--- a/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,10 +13,10 @@
subsubsection\<open>The Operator \<^term>\<open>is_formula_rec\<close>\<close>
text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0. It is buried
- within 11 quantifiers!!\<close>
+ within 11 quantifiers\<And>\<close>
(* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o"
- "is_formula_rec(M,MH,p,z) ==
+ "is_formula_rec(M,MH,p,z) \<equiv>
\<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) &
2 1 0
successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
@@ -24,7 +24,7 @@
definition
formula_rec_fm :: "[i, i, i]=>i" where
- "formula_rec_fm(mh,p,z) ==
+ "formula_rec_fm(mh,p,z) \<equiv>
Exists(Exists(Exists(
And(finite_ordinal_fm(2),
And(depth_fm(p#+3,2),
@@ -32,42 +32,42 @@
And(fun_apply_fm(0,p#+3,z#+3), is_transrec_fm(mh,1,0))))))))"
lemma is_formula_rec_type [TC]:
- "[| p \<in> formula; x \<in> nat; z \<in> nat |]
- ==> formula_rec_fm(p,x,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> formula_rec_fm(p,x,z) \<in> formula"
by (simp add: formula_rec_fm_def)
lemma sats_formula_rec_fm:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow>
+ "\<And>a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow>
sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
Cons(a4,Cons(a5,Cons(a6,Cons(a7,
Cons(a8,Cons(a9,Cons(a10,env))))))))))))"
shows
- "[|x \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, formula_rec_fm(p,x,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, formula_rec_fm(p,x,z), env) \<longleftrightarrow>
is_formula_rec(##A, MH, nth(x,env), nth(z,env))"
by (simp add: formula_rec_fm_def sats_is_transrec_fm is_formula_rec_def
MH_iff_sats [THEN iff_sym])
lemma formula_rec_iff_sats:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow>
+ "\<And>a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A; a8\<in>A; a9\<in>A; a10\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow>
sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
Cons(a4,Cons(a5,Cons(a6,Cons(a7,
Cons(a8,Cons(a9,Cons(a10,env))))))))))))"
shows
- "[|nth(i,env) = x; nth(k,env) = z;
- i \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_formula_rec(##A, MH, x, z) \<longleftrightarrow> sats(A, formula_rec_fm(p,i,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(k,env) = z;
+ i \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_formula_rec(##A, MH, x, z) \<longleftrightarrow> sats(A, formula_rec_fm(p,i,k), env)"
by (simp add: sats_formula_rec_fm [OF MH_iff_sats])
theorem formula_rec_reflection:
assumes MH_reflection:
- "!!f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
+ "\<And>f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
\<lambda>i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. is_formula_rec(L, MH(L,x), f(x), h(x)),
\<lambda>i x. is_formula_rec(##Lset(i), MH(##Lset(i),x), f(x), h(x))]"
@@ -79,25 +79,25 @@
subsubsection\<open>The Operator \<^term>\<open>is_satisfies\<close>\<close>
-(* is_satisfies(M,A,p,z) == is_formula_rec (M, satisfies_MH(M,A), p, z) *)
+(* is_satisfies(M,A,p,z) \<equiv> is_formula_rec (M, satisfies_MH(M,A), p, z) *)
definition
satisfies_fm :: "[i,i,i]=>i" where
- "satisfies_fm(x) == formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))"
+ "satisfies_fm(x) \<equiv> formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))"
lemma is_satisfies_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> satisfies_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> satisfies_fm(x,y,z) \<in> formula"
by (simp add: satisfies_fm_def)
lemma sats_satisfies_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_fm(x,y,z), env) \<longleftrightarrow>
is_satisfies(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: satisfies_fm_def is_satisfies_def sats_formula_rec_fm)
lemma satisfies_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_satisfies(##A, x, y, z) \<longleftrightarrow> sats(A, satisfies_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_satisfies(##A, x, y, z) \<longleftrightarrow> sats(A, satisfies_fm(i,j,k), env)"
by (simp)
theorem satisfies_reflection:
@@ -121,56 +121,56 @@
(the comprehension).\<close>
definition
is_DPow_sats :: "[i=>o,i,i,i,i] => o" where
- "is_DPow_sats(M,A,env,p,x) ==
+ "is_DPow_sats(M,A,env,p,x) \<equiv>
\<forall>n1[M]. \<forall>e[M]. \<forall>sp[M].
is_satisfies(M,A,p,sp) \<longrightarrow> is_Cons(M,x,env,e) \<longrightarrow>
fun_apply(M, sp, e, n1) \<longrightarrow> number1(M, n1)"
lemma (in M_satisfies) DPow_sats_abs:
- "[| M(A); env \<in> list(A); p \<in> formula; M(x) |]
- ==> is_DPow_sats(M,A,env,p,x) \<longleftrightarrow> sats(A, p, Cons(x,env))"
+ "\<lbrakk>M(A); env \<in> list(A); p \<in> formula; M(x)\<rbrakk>
+ \<Longrightarrow> is_DPow_sats(M,A,env,p,x) \<longleftrightarrow> sats(A, p, Cons(x,env))"
apply (subgoal_tac "M(env)")
apply (simp add: is_DPow_sats_def satisfies_closed satisfies_abs)
apply (blast dest: transM)
done
lemma (in M_satisfies) Collect_DPow_sats_abs:
- "[| M(A); env \<in> list(A); p \<in> formula |]
- ==> Collect(A, is_DPow_sats(M,A,env,p)) =
+ "\<lbrakk>M(A); env \<in> list(A); p \<in> formula\<rbrakk>
+ \<Longrightarrow> Collect(A, is_DPow_sats(M,A,env,p)) =
{x \<in> A. sats(A, p, Cons(x,env))}"
by (simp add: DPow_sats_abs transM [of _ A])
subsubsection\<open>The Operator \<^term>\<open>is_DPow_sats\<close>, Internalized\<close>
-(* is_DPow_sats(M,A,env,p,x) ==
+(* is_DPow_sats(M,A,env,p,x) \<equiv>
\<forall>n1[M]. \<forall>e[M]. \<forall>sp[M].
is_satisfies(M,A,p,sp) \<longrightarrow> is_Cons(M,x,env,e) \<longrightarrow>
fun_apply(M, sp, e, n1) \<longrightarrow> number1(M, n1) *)
definition
DPow_sats_fm :: "[i,i,i,i]=>i" where
- "DPow_sats_fm(A,env,p,x) ==
+ "DPow_sats_fm(A,env,p,x) \<equiv>
Forall(Forall(Forall(
Implies(satisfies_fm(A#+3,p#+3,0),
Implies(Cons_fm(x#+3,env#+3,1),
Implies(fun_apply_fm(0,1,2), number1_fm(2)))))))"
lemma is_DPow_sats_type [TC]:
- "[| A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> DPow_sats_fm(A,x,y,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> DPow_sats_fm(A,x,y,z) \<in> formula"
by (simp add: DPow_sats_fm_def)
lemma sats_DPow_sats_fm [simp]:
- "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, DPow_sats_fm(u,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, DPow_sats_fm(u,x,y,z), env) \<longleftrightarrow>
is_DPow_sats(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
by (simp add: DPow_sats_fm_def is_DPow_sats_def)
lemma DPow_sats_iff_sats:
- "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
- u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> is_DPow_sats(##A,nu,nx,ny,nz) \<longleftrightarrow>
+ "\<lbrakk>nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
+ u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_DPow_sats(##A,nu,nx,ny,nz) \<longleftrightarrow>
sats(A, DPow_sats_fm(u,x,y,z), env)"
by simp
@@ -187,23 +187,23 @@
locale M_DPow = M_satisfies +
assumes sep:
- "[| M(A); env \<in> list(A); p \<in> formula |]
- ==> separation(M, \<lambda>x. is_DPow_sats(M,A,env,p,x))"
+ "\<lbrakk>M(A); env \<in> list(A); p \<in> formula\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>x. is_DPow_sats(M,A,env,p,x))"
and rep:
"M(A)
- ==> strong_replacement (M,
+ \<Longrightarrow> strong_replacement (M,
\<lambda>ep z. \<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
pair(M,env,p,ep) &
is_Collect(M, A, \<lambda>x. is_DPow_sats(M,A,env,p,x), z))"
lemma (in M_DPow) sep':
- "[| M(A); env \<in> list(A); p \<in> formula |]
- ==> separation(M, \<lambda>x. sats(A, p, Cons(x,env)))"
+ "\<lbrakk>M(A); env \<in> list(A); p \<in> formula\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>x. sats(A, p, Cons(x,env)))"
by (insert sep [of A env p], simp add: DPow_sats_abs)
lemma (in M_DPow) rep':
"M(A)
- ==> strong_replacement (M,
+ \<Longrightarrow> strong_replacement (M,
\<lambda>ep z. \<exists>env\<in>list(A). \<exists>p\<in>formula.
ep = <env,p> & z = {x \<in> A . sats(A, p, Cons(x, env))})"
by (insert rep [of A], simp add: Collect_DPow_sats_abs)
@@ -213,7 +213,7 @@
"univalent (M, A, \<lambda>xy z. \<exists>x\<in>B. \<exists>y\<in>C. xy = \<langle>x,y\<rangle> \<and> z = f(x,y))"
by (simp add: univalent_def, blast)
-lemma (in M_DPow) DPow'_closed: "M(A) ==> M(DPow'(A))"
+lemma (in M_DPow) DPow'_closed: "M(A) \<Longrightarrow> M(DPow'(A))"
apply (simp add: DPow'_eq)
apply (fast intro: rep' sep' univalent_pair_eq)
done
@@ -221,14 +221,14 @@
text\<open>Relativization of the Operator \<^term>\<open>DPow'\<close>\<close>
definition
is_DPow' :: "[i=>o,i,i] => o" where
- "is_DPow'(M,A,Z) ==
+ "is_DPow'(M,A,Z) \<equiv>
\<forall>X[M]. X \<in> Z \<longleftrightarrow>
subset(M,X,A) &
(\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
is_Collect(M, A, is_DPow_sats(M,A,env,p), X))"
lemma (in M_DPow) DPow'_abs:
- "[|M(A); M(Z)|] ==> is_DPow'(M,A,Z) \<longleftrightarrow> Z = DPow'(A)"
+ "\<lbrakk>M(A); M(Z)\<rbrakk> \<Longrightarrow> is_DPow'(M,A,Z) \<longleftrightarrow> Z = DPow'(A)"
apply (rule iffI)
prefer 2 apply (simp add: is_DPow'_def DPow'_def Collect_DPow_sats_abs)
apply (rule M_equalityI)
@@ -242,8 +242,8 @@
subsubsection\<open>The Instance of Separation\<close>
lemma DPow_separation:
- "[| L(A); env \<in> list(A); p \<in> formula |]
- ==> separation(L, \<lambda>x. is_DPow_sats(L,A,env,p,x))"
+ "\<lbrakk>L(A); env \<in> list(A); p \<in> formula\<rbrakk>
+ \<Longrightarrow> separation(L, \<lambda>x. is_DPow_sats(L,A,env,p,x))"
apply (rule gen_separation_multi [OF DPow_sats_reflection, of "{A,env,p}"],
auto intro: transL)
apply (rule_tac env="[A,env,p]" in DPow_LsetI)
@@ -271,7 +271,7 @@
lemma DPow_replacement:
"L(A)
- ==> strong_replacement (L,
+ \<Longrightarrow> strong_replacement (L,
\<lambda>ep z. \<exists>env[L]. \<exists>p[L]. mem_formula(L,p) & mem_list(L,A,env) &
pair(L,env,p,ep) &
is_Collect(L, A, \<lambda>x. is_DPow_sats(L,A,env,p,x), z))"
@@ -309,35 +309,35 @@
enclosed within a single quantifier.\<close>
(* is_Collect :: "[i=>o,i,i=>o,i] => o"
- "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)" *)
+ "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)" *)
definition
Collect_fm :: "[i, i, i]=>i" where
- "Collect_fm(A,is_P,z) ==
+ "Collect_fm(A,is_P,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
And(Member(0,succ(A)), is_P)))"
lemma is_Collect_type [TC]:
- "[| is_P \<in> formula; x \<in> nat; y \<in> nat |]
- ==> Collect_fm(x,is_P,y) \<in> formula"
+ "\<lbrakk>is_P \<in> formula; x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> Collect_fm(x,is_P,y) \<in> formula"
by (simp add: Collect_fm_def)
lemma sats_Collect_fm:
assumes is_P_iff_sats:
- "!!a. a \<in> A ==> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
+ "\<And>a. a \<in> A \<Longrightarrow> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
shows
- "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, Collect_fm(x,p,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Collect_fm(x,p,y), env) \<longleftrightarrow>
is_Collect(##A, nth(x,env), is_P, nth(y,env))"
by (simp add: Collect_fm_def is_Collect_def is_P_iff_sats [THEN iff_sym])
lemma Collect_iff_sats:
assumes is_P_iff_sats:
- "!!a. a \<in> A ==> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
+ "\<And>a. a \<in> A \<Longrightarrow> is_P(a) \<longleftrightarrow> sats(A, p, Cons(a, env))"
shows
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_Collect(##A, x, is_P, y) \<longleftrightarrow> sats(A, Collect_fm(i,p,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Collect(##A, x, is_P, y) \<longleftrightarrow> sats(A, Collect_fm(i,p,j), env)"
by (simp add: sats_Collect_fm [OF is_P_iff_sats])
@@ -345,7 +345,7 @@
which is essential for handling free variable references.\<close>
theorem Collect_reflection:
assumes is_P_reflection:
- "!!h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x)),
+ "\<And>h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x)),
\<lambda>i x. is_P(##Lset(i), f(x), g(x))]"
shows "REFLECTS[\<lambda>x. is_Collect(L, f(x), is_P(L,x), g(x)),
\<lambda>i x. is_Collect(##Lset(i), f(x), is_P(##Lset(i), x), g(x))]"
@@ -360,37 +360,37 @@
and not the usual 1, 0! It is enclosed within two quantifiers.\<close>
(* is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
- "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))" *)
+ "is_Replace(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))" *)
definition
Replace_fm :: "[i, i, i]=>i" where
- "Replace_fm(A,is_P,z) ==
+ "Replace_fm(A,is_P,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,A#+2), is_P))))"
lemma is_Replace_type [TC]:
- "[| is_P \<in> formula; x \<in> nat; y \<in> nat |]
- ==> Replace_fm(x,is_P,y) \<in> formula"
+ "\<lbrakk>is_P \<in> formula; x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> Replace_fm(x,is_P,y) \<in> formula"
by (simp add: Replace_fm_def)
lemma sats_Replace_fm:
assumes is_P_iff_sats:
- "!!a b. [|a \<in> A; b \<in> A|]
- ==> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
+ "\<And>a b. \<lbrakk>a \<in> A; b \<in> A\<rbrakk>
+ \<Longrightarrow> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
shows
- "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, Replace_fm(x,p,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Replace_fm(x,p,y), env) \<longleftrightarrow>
is_Replace(##A, nth(x,env), is_P, nth(y,env))"
by (simp add: Replace_fm_def is_Replace_def is_P_iff_sats [THEN iff_sym])
lemma Replace_iff_sats:
assumes is_P_iff_sats:
- "!!a b. [|a \<in> A; b \<in> A|]
- ==> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
+ "\<And>a b. \<lbrakk>a \<in> A; b \<in> A\<rbrakk>
+ \<Longrightarrow> is_P(a,b) \<longleftrightarrow> sats(A, p, Cons(a,Cons(b,env)))"
shows
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_Replace(##A, x, is_P, y) \<longleftrightarrow> sats(A, Replace_fm(i,p,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Replace(##A, x, is_P, y) \<longleftrightarrow> sats(A, Replace_fm(i,p,j), env)"
by (simp add: sats_Replace_fm [OF is_P_iff_sats])
@@ -398,7 +398,7 @@
which is essential for handling free variable references.\<close>
theorem Replace_reflection:
assumes is_P_reflection:
- "!!h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x), h(x)),
+ "\<And>h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x), h(x)),
\<lambda>i x. is_P(##Lset(i), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. is_Replace(L, f(x), is_P(L,x), g(x)),
\<lambda>i x. is_Replace(##Lset(i), f(x), is_P(##Lset(i), x), g(x))]"
@@ -410,7 +410,7 @@
subsubsection\<open>The Operator \<^term>\<open>is_DPow'\<close>, Internalized\<close>
-(* "is_DPow'(M,A,Z) ==
+(* "is_DPow'(M,A,Z) \<equiv>
\<forall>X[M]. X \<in> Z \<longleftrightarrow>
subset(M,X,A) &
(\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
@@ -418,7 +418,7 @@
definition
DPow'_fm :: "[i,i]=>i" where
- "DPow'_fm(A,Z) ==
+ "DPow'_fm(A,Z) \<equiv>
Forall(
Iff(Member(0,succ(Z)),
And(subset_fm(0,succ(A)),
@@ -429,19 +429,19 @@
DPow_sats_fm(A#+4, 2, 1, 0), 2))))))))"
lemma is_DPow'_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> DPow'_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> DPow'_fm(x,y) \<in> formula"
by (simp add: DPow'_fm_def)
lemma sats_DPow'_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, DPow'_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, DPow'_fm(x,y), env) \<longleftrightarrow>
is_DPow'(##A, nth(x,env), nth(y,env))"
by (simp add: DPow'_fm_def is_DPow'_def sats_subset_fm' sats_Collect_fm)
lemma DPow'_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_DPow'(##A, x, y) \<longleftrightarrow> sats(A, DPow'_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_DPow'(##A, x, y) \<longleftrightarrow> sats(A, DPow'_fm(i,j), env)"
by (simp)
theorem DPow'_reflection:
@@ -457,39 +457,39 @@
definition
transrec_body :: "[i=>o,i,i,i,i] => o" where
- "transrec_body(M,g,x) ==
+ "transrec_body(M,g,x) \<equiv>
\<lambda>y z. \<exists>gy[M]. y \<in> x & fun_apply(M,g,y,gy) & is_DPow'(M,gy,z)"
lemma (in M_DPow) transrec_body_abs:
- "[|M(x); M(g); M(z)|]
- ==> transrec_body(M,g,x,y,z) \<longleftrightarrow> y \<in> x & z = DPow'(g`y)"
+ "\<lbrakk>M(x); M(g); M(z)\<rbrakk>
+ \<Longrightarrow> transrec_body(M,g,x,y,z) \<longleftrightarrow> y \<in> x & z = DPow'(g`y)"
by (simp add: transrec_body_def DPow'_abs transM [of _ x])
locale M_Lset = M_DPow +
assumes strong_rep:
- "[|M(x); M(g)|] ==> strong_replacement(M, \<lambda>y z. transrec_body(M,g,x,y,z))"
+ "\<lbrakk>M(x); M(g)\<rbrakk> \<Longrightarrow> strong_replacement(M, \<lambda>y z. transrec_body(M,g,x,y,z))"
and transrec_rep:
- "M(i) ==> transrec_replacement(M, \<lambda>x f u.
+ "M(i) \<Longrightarrow> transrec_replacement(M, \<lambda>x f u.
\<exists>r[M]. is_Replace(M, x, transrec_body(M,f,x), r) &
big_union(M, r, u), i)"
lemma (in M_Lset) strong_rep':
- "[|M(x); M(g)|]
- ==> strong_replacement(M, \<lambda>y z. y \<in> x & z = DPow'(g`y))"
+ "\<lbrakk>M(x); M(g)\<rbrakk>
+ \<Longrightarrow> strong_replacement(M, \<lambda>y z. y \<in> x & z = DPow'(g`y))"
by (insert strong_rep [of x g], simp add: transrec_body_abs)
lemma (in M_Lset) DPow_apply_closed:
- "[|M(f); M(x); y\<in>x|] ==> M(DPow'(f`y))"
+ "\<lbrakk>M(f); M(x); y\<in>x\<rbrakk> \<Longrightarrow> M(DPow'(f`y))"
by (blast intro: DPow'_closed dest: transM)
lemma (in M_Lset) RepFun_DPow_apply_closed:
- "[|M(f); M(x)|] ==> M({DPow'(f`y). y\<in>x})"
+ "\<lbrakk>M(f); M(x)\<rbrakk> \<Longrightarrow> M({DPow'(f`y). y\<in>x})"
by (blast intro: DPow_apply_closed RepFun_closed2 strong_rep')
lemma (in M_Lset) RepFun_DPow_abs:
- "[|M(x); M(f); M(r) |]
- ==> is_Replace(M, x, \<lambda>y z. transrec_body(M,f,x,y,z), r) \<longleftrightarrow>
+ "\<lbrakk>M(x); M(f); M(r)\<rbrakk>
+ \<Longrightarrow> is_Replace(M, x, \<lambda>y z. transrec_body(M,f,x,y,z), r) \<longleftrightarrow>
r = {DPow'(f`y). y\<in>x}"
apply (simp add: transrec_body_abs RepFun_def)
apply (rule iff_trans)
@@ -498,7 +498,7 @@
done
lemma (in M_Lset) transrec_rep':
- "M(i) ==> transrec_replacement(M, \<lambda>x f u. u = (\<Union>y\<in>x. DPow'(f ` y)), i)"
+ "M(i) \<Longrightarrow> transrec_replacement(M, \<lambda>x f u. u = (\<Union>y\<in>x. DPow'(f ` y)), i)"
apply (insert transrec_rep [of i])
apply (simp add: RepFun_DPow_apply_closed RepFun_DPow_abs
transrec_replacement_def)
@@ -512,18 +512,18 @@
\<comment> \<open>We can use the term language below because \<^term>\<open>is_Lset\<close> will
not have to be internalized: it isn't used in any instance of
separation.\<close>
- "is_Lset(M,a,z) == is_transrec(M, %x f u. u = (\<Union>y\<in>x. DPow'(f`y)), a, z)"
+ "is_Lset(M,a,z) \<equiv> is_transrec(M, %x f u. u = (\<Union>y\<in>x. DPow'(f`y)), a, z)"
lemma (in M_Lset) Lset_abs:
- "[|Ord(i); M(i); M(z)|]
- ==> is_Lset(M,i,z) \<longleftrightarrow> z = Lset(i)"
+ "\<lbrakk>Ord(i); M(i); M(z)\<rbrakk>
+ \<Longrightarrow> is_Lset(M,i,z) \<longleftrightarrow> z = Lset(i)"
apply (simp add: is_Lset_def Lset_eq_transrec_DPow')
apply (rule transrec_abs)
apply (simp_all add: transrec_rep' relation2_def RepFun_DPow_apply_closed)
done
lemma (in M_Lset) Lset_closed:
- "[|Ord(i); M(i)|] ==> M(Lset(i))"
+ "\<lbrakk>Ord(i); M(i)\<rbrakk> \<Longrightarrow> M(Lset(i))"
apply (simp add: Lset_eq_transrec_DPow')
apply (rule transrec_closed [OF transrec_rep'])
apply (simp_all add: relation2_def RepFun_DPow_apply_closed)
@@ -542,7 +542,7 @@
by (intro FOL_reflections function_reflections DPow'_reflection)
lemma strong_rep:
- "[|L(x); L(g)|] ==> strong_replacement(L, \<lambda>y z. transrec_body(L,g,x,y,z))"
+ "\<lbrakk>L(x); L(g)\<rbrakk> \<Longrightarrow> strong_replacement(L, \<lambda>y z. transrec_body(L,g,x,y,z))"
apply (unfold transrec_body_def)
apply (rule strong_replacementI)
apply (rule_tac u="{x,g,B}"
@@ -577,8 +577,8 @@
lemma transrec_rep:
- "[|L(j)|]
- ==> transrec_replacement(L, \<lambda>x f u.
+ "\<lbrakk>L(j)\<rbrakk>
+ \<Longrightarrow> transrec_replacement(L, \<lambda>x f u.
\<exists>r[L]. is_Replace(L, x, transrec_body(L,f,x), r) &
big_union(L, r, u), j)"
apply (rule L.transrec_replacementI, assumption)
@@ -614,7 +614,7 @@
definition
constructible :: "[i=>o,i] => o" where
- "constructible(M,x) ==
+ "constructible(M,x) \<equiv>
\<exists>i[M]. \<exists>Li[M]. ordinal(M,i) & is_Lset(M,i,Li) & x \<in> Li"
theorem V_equals_L_in_L:
--- a/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,25 +11,25 @@
definition
directed :: "i=>o" where
- "directed(A) == A\<noteq>0 & (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)"
+ "directed(A) \<equiv> A\<noteq>0 & (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)"
definition
contin :: "(i=>i) => o" where
- "contin(h) == (\<forall>A. directed(A) \<longrightarrow> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
+ "contin(h) \<equiv> (\<forall>A. directed(A) \<longrightarrow> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
-lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \<in> nat|] ==> h^n (0) \<subseteq> D"
+lemma bnd_mono_iterates_subset: "\<lbrakk>bnd_mono(D, h); n \<in> nat\<rbrakk> \<Longrightarrow> h^n (0) \<subseteq> D"
apply (induct_tac n)
apply (simp_all add: bnd_mono_def, blast)
done
lemma bnd_mono_increasing [rule_format]:
- "[|i \<in> nat; j \<in> nat; bnd_mono(D,h)|] ==> i \<le> j \<longrightarrow> h^i(0) \<subseteq> h^j(0)"
+ "\<lbrakk>i \<in> nat; j \<in> nat; bnd_mono(D,h)\<rbrakk> \<Longrightarrow> i \<le> j \<longrightarrow> h^i(0) \<subseteq> h^j(0)"
apply (rule_tac m=i and n=j in diff_induct, simp_all)
apply (blast del: subsetI
intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h])
done
-lemma directed_iterates: "bnd_mono(D,h) ==> directed({h^n (0). n\<in>nat})"
+lemma directed_iterates: "bnd_mono(D,h) \<Longrightarrow> directed({h^n (0). n\<in>nat})"
apply (simp add: directed_def, clarify)
apply (rename_tac i j)
apply (rule_tac x="i \<union> j" in bexI)
@@ -42,8 +42,8 @@
lemma contin_iterates_eq:
- "[|bnd_mono(D, h); contin(h)|]
- ==> h(\<Union>n\<in>nat. h^n (0)) = (\<Union>n\<in>nat. h^n (0))"
+ "\<lbrakk>bnd_mono(D, h); contin(h)\<rbrakk>
+ \<Longrightarrow> h(\<Union>n\<in>nat. h^n (0)) = (\<Union>n\<in>nat. h^n (0))"
apply (simp add: contin_def directed_iterates)
apply (rule trans)
apply (rule equalityI)
@@ -56,14 +56,14 @@
done
lemma lfp_subset_Union:
- "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) \<subseteq> (\<Union>n\<in>nat. h^n(0))"
+ "\<lbrakk>bnd_mono(D, h); contin(h)\<rbrakk> \<Longrightarrow> lfp(D,h) \<subseteq> (\<Union>n\<in>nat. h^n(0))"
apply (rule lfp_lowerbound)
apply (simp add: contin_iterates_eq)
apply (simp add: contin_def bnd_mono_iterates_subset UN_subset_iff)
done
lemma Union_subset_lfp:
- "bnd_mono(D,h) ==> (\<Union>n\<in>nat. h^n(0)) \<subseteq> lfp(D,h)"
+ "bnd_mono(D,h) \<Longrightarrow> (\<Union>n\<in>nat. h^n(0)) \<subseteq> lfp(D,h)"
apply (simp add: UN_subset_iff)
apply (rule ballI)
apply (induct_tac n, simp_all)
@@ -73,23 +73,23 @@
done
lemma lfp_eq_Union:
- "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) = (\<Union>n\<in>nat. h^n(0))"
+ "\<lbrakk>bnd_mono(D, h); contin(h)\<rbrakk> \<Longrightarrow> lfp(D,h) = (\<Union>n\<in>nat. h^n(0))"
by (blast del: subsetI
intro: lfp_subset_Union Union_subset_lfp)
subsubsection\<open>Some Standard Datatype Constructions Preserve Continuity\<close>
-lemma contin_imp_mono: "[|X\<subseteq>Y; contin(F)|] ==> F(X) \<subseteq> F(Y)"
+lemma contin_imp_mono: "\<lbrakk>X\<subseteq>Y; contin(F)\<rbrakk> \<Longrightarrow> F(X) \<subseteq> F(Y)"
apply (simp add: contin_def)
apply (drule_tac x="{X,Y}" in spec)
apply (simp add: directed_def subset_Un_iff2 Un_commute)
done
-lemma sum_contin: "[|contin(F); contin(G)|] ==> contin(\<lambda>X. F(X) + G(X))"
+lemma sum_contin: "\<lbrakk>contin(F); contin(G)\<rbrakk> \<Longrightarrow> contin(\<lambda>X. F(X) + G(X))"
by (simp add: contin_def, blast)
-lemma prod_contin: "[|contin(F); contin(G)|] ==> contin(\<lambda>X. F(X) * G(X))"
+lemma prod_contin: "\<lbrakk>contin(F); contin(G)\<rbrakk> \<Longrightarrow> contin(\<lambda>X. F(X) * G(X))"
apply (subgoal_tac "\<forall>B C. F(B) \<subseteq> F(B \<union> C)")
prefer 2 apply (simp add: Un_upper1 contin_imp_mono)
apply (subgoal_tac "\<forall>B C. G(C) \<subseteq> G(B \<union> C)")
@@ -115,38 +115,38 @@
definition
iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where
- "iterates_MH(M,isF,v,n,g,z) ==
+ "iterates_MH(M,isF,v,n,g,z) \<equiv>
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
n, z)"
definition
is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where
- "is_iterates(M,isF,v,n,Z) ==
+ "is_iterates(M,isF,v,n,Z) \<equiv>
\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"
definition
iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where
- "iterates_replacement(M,isF,v) ==
+ "iterates_replacement(M,isF,v) \<equiv>
\<forall>n[M]. n\<in>nat \<longrightarrow>
wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))"
lemma (in M_basic) iterates_MH_abs:
- "[| relation1(M,isF,F); M(n); M(g); M(z) |]
- ==> iterates_MH(M,isF,v,n,g,z) \<longleftrightarrow> z = nat_case(v, \<lambda>m. F(g`m), n)"
+ "\<lbrakk>relation1(M,isF,F); M(n); M(g); M(z)\<rbrakk>
+ \<Longrightarrow> iterates_MH(M,isF,v,n,g,z) \<longleftrightarrow> z = nat_case(v, \<lambda>m. F(g`m), n)"
by (simp add: nat_case_abs [of _ "\<lambda>m. F(g ` m)"]
relation1_def iterates_MH_def)
lemma (in M_trancl) iterates_imp_wfrec_replacement:
- "[|relation1(M,isF,F); n \<in> nat; iterates_replacement(M,isF,v)|]
- ==> wfrec_replacement(M, \<lambda>n f z. z = nat_case(v, \<lambda>m. F(f`m), n),
+ "\<lbrakk>relation1(M,isF,F); n \<in> nat; iterates_replacement(M,isF,v)\<rbrakk>
+ \<Longrightarrow> wfrec_replacement(M, \<lambda>n f z. z = nat_case(v, \<lambda>m. F(f`m), n),
Memrel(succ(n)))"
by (simp add: iterates_replacement_def iterates_MH_abs)
theorem (in M_trancl) iterates_abs:
- "[| iterates_replacement(M,isF,v); relation1(M,isF,F);
- n \<in> nat; M(v); M(z); \<forall>x[M]. M(F(x)) |]
- ==> is_iterates(M,isF,v,n,z) \<longleftrightarrow> z = iterates(F,n,v)"
+ "\<lbrakk>iterates_replacement(M,isF,v); relation1(M,isF,F);
+ n \<in> nat; M(v); M(z); \<forall>x[M]. M(F(x))\<rbrakk>
+ \<Longrightarrow> is_iterates(M,isF,v,n,z) \<longleftrightarrow> z = iterates(F,n,v)"
apply (frule iterates_imp_wfrec_replacement, assumption+)
apply (simp add: wf_Memrel trans_Memrel relation_Memrel
is_iterates_def relation2_def iterates_MH_abs
@@ -157,9 +157,9 @@
lemma (in M_trancl) iterates_closed [intro,simp]:
- "[| iterates_replacement(M,isF,v); relation1(M,isF,F);
- n \<in> nat; M(v); \<forall>x[M]. M(F(x)) |]
- ==> M(iterates(F,n,v))"
+ "\<lbrakk>iterates_replacement(M,isF,v); relation1(M,isF,F);
+ n \<in> nat; M(v); \<forall>x[M]. M(F(x))\<rbrakk>
+ \<Longrightarrow> M(iterates(F,n,v))"
apply (frule iterates_imp_wfrec_replacement, assumption+)
apply (simp add: wf_Memrel trans_Memrel relation_Memrel
relation2_def iterates_MH_abs
@@ -210,12 +210,12 @@
definition
is_list_functor :: "[i=>o,i,i,i] => o" where
- "is_list_functor(M,A,X,Z) ==
+ "is_list_functor(M,A,X,Z) \<equiv>
\<exists>n1[M]. \<exists>AX[M].
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
lemma (in M_basic) list_functor_abs [simp]:
- "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) \<longleftrightarrow> (Z = {0} + A*X)"
+ "\<lbrakk>M(A); M(X); M(Z)\<rbrakk> \<Longrightarrow> is_list_functor(M,A,X,Z) \<longleftrightarrow> (Z = {0} + A*X)"
by (simp add: is_list_functor_def singleton_0 nat_into_M)
@@ -263,7 +263,7 @@
definition
is_formula_functor :: "[i=>o,i,i] => o" where
- "is_formula_functor(M,X,Z) ==
+ "is_formula_functor(M,X,Z) \<equiv>
\<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
omega(M,nat') & cartprod(M,nat',nat',natnat) &
is_sum(M,natnat,natnat,natnatsum) &
@@ -271,8 +271,8 @@
is_sum(M,natnatsum,X3,Z)"
lemma (in M_trancl) formula_functor_abs [simp]:
- "[| M(X); M(Z) |]
- ==> is_formula_functor(M,X,Z) \<longleftrightarrow>
+ "\<lbrakk>M(X); M(Z)\<rbrakk>
+ \<Longrightarrow> is_formula_functor(M,X,Z) \<longleftrightarrow>
Z = ((nat*nat) + (nat*nat)) + (X*X + X)"
by (simp add: is_formula_functor_def)
@@ -281,7 +281,7 @@
definition
list_N :: "[i,i] => i" where
- "list_N(A,n) == (\<lambda>X. {0} + A * X)^n (0)"
+ "list_N(A,n) \<equiv> (\<lambda>X. {0} + A * X)^n (0)"
lemma Nil_in_list_N [simp]: "[] \<in> list_N(A,succ(n))"
by (simp add: list_N_def Nil_def)
@@ -299,25 +299,25 @@
by (simp add: list_N_def)
lemma list_N_imp_list:
- "[| l \<in> list_N(A,n); n \<in> nat |] ==> l \<in> list(A)"
+ "\<lbrakk>l \<in> list_N(A,n); n \<in> nat\<rbrakk> \<Longrightarrow> l \<in> list(A)"
by (force simp add: list_eq_Union list_N_def)
lemma list_N_imp_length_lt [rule_format]:
- "n \<in> nat ==> \<forall>l \<in> list_N(A,n). length(l) < n"
+ "n \<in> nat \<Longrightarrow> \<forall>l \<in> list_N(A,n). length(l) < n"
apply (induct_tac n)
apply (auto simp add: list_N_0 list_N_succ
Nil_def [symmetric] Cons_def [symmetric])
done
lemma list_imp_list_N [rule_format]:
- "l \<in> list(A) ==> \<forall>n\<in>nat. length(l) < n \<longrightarrow> l \<in> list_N(A, n)"
+ "l \<in> list(A) \<Longrightarrow> \<forall>n\<in>nat. length(l) < n \<longrightarrow> l \<in> list_N(A, n)"
apply (induct_tac l)
apply (force elim: natE)+
done
lemma list_N_imp_eq_length:
- "[|n \<in> nat; l \<notin> list_N(A, n); l \<in> list_N(A, succ(n))|]
- ==> n = length(l)"
+ "\<lbrakk>n \<in> nat; l \<notin> list_N(A, n); l \<in> list_N(A, succ(n))\<rbrakk>
+ \<Longrightarrow> n = length(l)"
apply (rule le_anti_sym)
prefer 2 apply (simp add: list_N_imp_length_lt)
apply (frule list_N_imp_list, simp)
@@ -328,7 +328,7 @@
text\<open>Express \<^term>\<open>list_rec\<close> without using \<^term>\<open>rank\<close> or \<^term>\<open>Vset\<close>,
neither of which is absolute.\<close>
lemma (in M_trivial) list_rec_eq:
- "l \<in> list(A) ==>
+ "l \<in> list(A) \<Longrightarrow>
list_rec(a,g,l) =
transrec (succ(length(l)),
\<lambda>x h. Lambda (list(A),
@@ -342,19 +342,19 @@
definition
is_list_N :: "[i=>o,i,i,i] => o" where
- "is_list_N(M,A,n,Z) ==
+ "is_list_N(M,A,n,Z) \<equiv>
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_list_functor(M,A), zero, n, Z)"
definition
mem_list :: "[i=>o,i,i] => o" where
- "mem_list(M,A,l) ==
+ "mem_list(M,A,l) \<equiv>
\<exists>n[M]. \<exists>listn[M].
finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn"
definition
is_list :: "[i=>o,i,i] => o" where
- "is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l)"
+ "is_list(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l)"
subsubsection\<open>Towards Absoluteness of \<^term>\<open>formula_rec\<close>\<close>
@@ -365,13 +365,13 @@
"depth(Nand(p,q)) = succ(depth(p) \<union> depth(q))"
"depth(Forall(p)) = succ(depth(p))"
-lemma depth_type [TC]: "p \<in> formula ==> depth(p) \<in> nat"
+lemma depth_type [TC]: "p \<in> formula \<Longrightarrow> depth(p) \<in> nat"
by (induct_tac p, simp_all)
definition
formula_N :: "i => i" where
- "formula_N(n) == (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
+ "formula_N(n) \<equiv> (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
lemma Member_in_formula_N [simp]:
"Member(x,y) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> nat & y \<in> nat"
@@ -400,11 +400,11 @@
by (simp add: formula_N_def)
lemma formula_N_imp_formula:
- "[| p \<in> formula_N(n); n \<in> nat |] ==> p \<in> formula"
+ "\<lbrakk>p \<in> formula_N(n); n \<in> nat\<rbrakk> \<Longrightarrow> p \<in> formula"
by (force simp add: formula_eq_Union formula_N_def)
lemma formula_N_imp_depth_lt [rule_format]:
- "n \<in> nat ==> \<forall>p \<in> formula_N(n). depth(p) < n"
+ "n \<in> nat \<Longrightarrow> \<forall>p \<in> formula_N(n). depth(p) < n"
apply (induct_tac n)
apply (auto simp add: formula_N_0 formula_N_succ
depth_type formula_N_imp_formula Un_least_lt_iff
@@ -413,15 +413,15 @@
done
lemma formula_imp_formula_N [rule_format]:
- "p \<in> formula ==> \<forall>n\<in>nat. depth(p) < n \<longrightarrow> p \<in> formula_N(n)"
+ "p \<in> formula \<Longrightarrow> \<forall>n\<in>nat. depth(p) < n \<longrightarrow> p \<in> formula_N(n)"
apply (induct_tac p)
apply (simp_all add: succ_Un_distrib Un_least_lt_iff)
apply (force elim: natE)+
done
lemma formula_N_imp_eq_depth:
- "[|n \<in> nat; p \<notin> formula_N(n); p \<in> formula_N(succ(n))|]
- ==> n = depth(p)"
+ "\<lbrakk>n \<in> nat; p \<notin> formula_N(n); p \<in> formula_N(succ(n))\<rbrakk>
+ \<Longrightarrow> n = depth(p)"
apply (rule le_anti_sym)
prefer 2 apply (simp add: formula_N_imp_depth_lt)
apply (frule formula_N_imp_formula, simp)
@@ -432,13 +432,13 @@
text\<open>This result and the next are unused.\<close>
lemma formula_N_mono [rule_format]:
- "[| m \<in> nat; n \<in> nat |] ==> m\<le>n \<longrightarrow> formula_N(m) \<subseteq> formula_N(n)"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m\<le>n \<longrightarrow> formula_N(m) \<subseteq> formula_N(n)"
apply (rule_tac m = m and n = n in diff_induct)
apply (simp_all add: formula_N_0 formula_N_succ, blast)
done
lemma formula_N_distrib:
- "[| m \<in> nat; n \<in> nat |] ==> formula_N(m \<union> n) = formula_N(m) \<union> formula_N(n)"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> formula_N(m \<union> n) = formula_N(m) \<union> formula_N(n)"
apply (rule_tac i = m and j = n in Ord_linear_le, auto)
apply (simp_all add: subset_Un_iff [THEN iffD1] subset_Un_iff2 [THEN iffD1]
le_imp_subset formula_N_mono)
@@ -446,26 +446,26 @@
definition
is_formula_N :: "[i=>o,i,i] => o" where
- "is_formula_N(M,n,Z) ==
+ "is_formula_N(M,n,Z) \<equiv>
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_formula_functor(M), zero, n, Z)"
definition
mem_formula :: "[i=>o,i] => o" where
- "mem_formula(M,p) ==
+ "mem_formula(M,p) \<equiv>
\<exists>n[M]. \<exists>formn[M].
finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn"
definition
is_formula :: "[i=>o,i] => o" where
- "is_formula(M,Z) == \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p)"
+ "is_formula(M,Z) \<equiv> \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p)"
locale M_datatypes = M_trancl +
assumes list_replacement1:
- "M(A) ==> iterates_replacement(M, is_list_functor(M,A), 0)"
+ "M(A) \<Longrightarrow> iterates_replacement(M, is_list_functor(M,A), 0)"
and list_replacement2:
- "M(A) ==> strong_replacement(M,
+ "M(A) \<Longrightarrow> strong_replacement(M,
\<lambda>n y. n\<in>nat & is_iterates(M, is_list_functor(M,A), 0, n, y))"
and formula_replacement1:
"iterates_replacement(M, is_formula_functor(M), 0)"
@@ -473,13 +473,13 @@
"strong_replacement(M,
\<lambda>n y. n\<in>nat & is_iterates(M, is_formula_functor(M), 0, n, y))"
and nth_replacement:
- "M(l) ==> iterates_replacement(M, %l t. is_tl(M,l,t), l)"
+ "M(l) \<Longrightarrow> iterates_replacement(M, %l t. is_tl(M,l,t), l)"
subsubsection\<open>Absoluteness of the List Construction\<close>
lemma (in M_datatypes) list_replacement2':
- "M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
+ "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
apply (insert list_replacement2 [of A])
apply (rule strong_replacement_cong [THEN iffD1])
apply (rule conj_cong [OF iff_refl iterates_abs [of "is_list_functor(M,A)"]])
@@ -487,7 +487,7 @@
done
lemma (in M_datatypes) list_closed [intro,simp]:
- "M(A) ==> M(list(A))"
+ "M(A) \<Longrightarrow> M(list(A))"
apply (insert list_replacement1)
by (simp add: RepFun_closed2 list_eq_Union
list_replacement2' relation1_def
@@ -497,29 +497,29 @@
lemmas (in M_datatypes) list_into_M = transM [OF _ list_closed]
lemma (in M_datatypes) list_N_abs [simp]:
- "[|M(A); n\<in>nat; M(Z)|]
- ==> is_list_N(M,A,n,Z) \<longleftrightarrow> Z = list_N(A,n)"
+ "\<lbrakk>M(A); n\<in>nat; M(Z)\<rbrakk>
+ \<Longrightarrow> is_list_N(M,A,n,Z) \<longleftrightarrow> Z = list_N(A,n)"
apply (insert list_replacement1)
apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M
iterates_abs [of "is_list_functor(M,A)" _ "\<lambda>X. {0} + A*X"])
done
lemma (in M_datatypes) list_N_closed [intro,simp]:
- "[|M(A); n\<in>nat|] ==> M(list_N(A,n))"
+ "\<lbrakk>M(A); n\<in>nat\<rbrakk> \<Longrightarrow> M(list_N(A,n))"
apply (insert list_replacement1)
apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M
iterates_closed [of "is_list_functor(M,A)"])
done
lemma (in M_datatypes) mem_list_abs [simp]:
- "M(A) ==> mem_list(M,A,l) \<longleftrightarrow> l \<in> list(A)"
+ "M(A) \<Longrightarrow> mem_list(M,A,l) \<longleftrightarrow> l \<in> list(A)"
apply (insert list_replacement1)
apply (simp add: mem_list_def list_N_def relation1_def list_eq_Union
iterates_closed [of "is_list_functor(M,A)"])
done
lemma (in M_datatypes) list_abs [simp]:
- "[|M(A); M(Z)|] ==> is_list(M,A,Z) \<longleftrightarrow> Z = list(A)"
+ "\<lbrakk>M(A); M(Z)\<rbrakk> \<Longrightarrow> is_list(M,A,Z) \<longleftrightarrow> Z = list(A)"
apply (simp add: is_list_def, safe)
apply (rule M_equalityI, simp_all)
done
@@ -545,8 +545,8 @@
lemmas (in M_datatypes) formula_into_M = transM [OF _ formula_closed]
lemma (in M_datatypes) formula_N_abs [simp]:
- "[|n\<in>nat; M(Z)|]
- ==> is_formula_N(M,n,Z) \<longleftrightarrow> Z = formula_N(n)"
+ "\<lbrakk>n\<in>nat; M(Z)\<rbrakk>
+ \<Longrightarrow> is_formula_N(M,n,Z) \<longleftrightarrow> Z = formula_N(n)"
apply (insert formula_replacement1)
apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M
iterates_abs [of "is_formula_functor(M)" _
@@ -554,7 +554,7 @@
done
lemma (in M_datatypes) formula_N_closed [intro,simp]:
- "n\<in>nat ==> M(formula_N(n))"
+ "n\<in>nat \<Longrightarrow> M(formula_N(n))"
apply (insert formula_replacement1)
apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M
iterates_closed [of "is_formula_functor(M)"])
@@ -568,7 +568,7 @@
done
lemma (in M_datatypes) formula_abs [simp]:
- "[|M(Z)|] ==> is_formula(M,Z) \<longleftrightarrow> Z = formula"
+ "\<lbrakk>M(Z)\<rbrakk> \<Longrightarrow> is_formula(M,Z) \<longleftrightarrow> Z = formula"
apply (simp add: is_formula_def, safe)
apply (rule M_equalityI, simp_all)
done
@@ -589,28 +589,28 @@
definition
is_eclose_n :: "[i=>o,i,i,i] => o" where
- "is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z)"
+ "is_eclose_n(M,A,n,Z) \<equiv> is_iterates(M, big_union(M), A, n, Z)"
definition
mem_eclose :: "[i=>o,i,i] => o" where
- "mem_eclose(M,A,l) ==
+ "mem_eclose(M,A,l) \<equiv>
\<exists>n[M]. \<exists>eclosen[M].
finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen"
definition
is_eclose :: "[i=>o,i,i] => o" where
- "is_eclose(M,A,Z) == \<forall>u[M]. u \<in> Z \<longleftrightarrow> mem_eclose(M,A,u)"
+ "is_eclose(M,A,Z) \<equiv> \<forall>u[M]. u \<in> Z \<longleftrightarrow> mem_eclose(M,A,u)"
locale M_eclose = M_datatypes +
assumes eclose_replacement1:
- "M(A) ==> iterates_replacement(M, big_union(M), A)"
+ "M(A) \<Longrightarrow> iterates_replacement(M, big_union(M), A)"
and eclose_replacement2:
- "M(A) ==> strong_replacement(M,
+ "M(A) \<Longrightarrow> strong_replacement(M,
\<lambda>n y. n\<in>nat & is_iterates(M, big_union(M), A, n, y))"
lemma (in M_eclose) eclose_replacement2':
- "M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = Union^n (A))"
+ "M(A) \<Longrightarrow> strong_replacement(M, \<lambda>n y. n\<in>nat & y = Union^n (A))"
apply (insert eclose_replacement2 [of A])
apply (rule strong_replacement_cong [THEN iffD1])
apply (rule conj_cong [OF iff_refl iterates_abs [of "big_union(M)"]])
@@ -618,28 +618,28 @@
done
lemma (in M_eclose) eclose_closed [intro,simp]:
- "M(A) ==> M(eclose(A))"
+ "M(A) \<Longrightarrow> M(eclose(A))"
apply (insert eclose_replacement1)
by (simp add: RepFun_closed2 eclose_eq_Union
eclose_replacement2' relation1_def
iterates_closed [of "big_union(M)"])
lemma (in M_eclose) is_eclose_n_abs [simp]:
- "[|M(A); n\<in>nat; M(Z)|] ==> is_eclose_n(M,A,n,Z) \<longleftrightarrow> Z = Union^n (A)"
+ "\<lbrakk>M(A); n\<in>nat; M(Z)\<rbrakk> \<Longrightarrow> is_eclose_n(M,A,n,Z) \<longleftrightarrow> Z = Union^n (A)"
apply (insert eclose_replacement1)
apply (simp add: is_eclose_n_def relation1_def nat_into_M
iterates_abs [of "big_union(M)" _ "Union"])
done
lemma (in M_eclose) mem_eclose_abs [simp]:
- "M(A) ==> mem_eclose(M,A,l) \<longleftrightarrow> l \<in> eclose(A)"
+ "M(A) \<Longrightarrow> mem_eclose(M,A,l) \<longleftrightarrow> l \<in> eclose(A)"
apply (insert eclose_replacement1)
apply (simp add: mem_eclose_def relation1_def eclose_eq_Union
iterates_closed [of "big_union(M)"])
done
lemma (in M_eclose) eclose_abs [simp]:
- "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) \<longleftrightarrow> Z = eclose(A)"
+ "\<lbrakk>M(A); M(Z)\<rbrakk> \<Longrightarrow> is_eclose(M,A,Z) \<longleftrightarrow> Z = eclose(A)"
apply (simp add: is_eclose_def, safe)
apply (rule M_equalityI, simp_all)
done
@@ -651,14 +651,14 @@
definition
is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
- "is_transrec(M,MH,a,z) ==
+ "is_transrec(M,MH,a,z) \<equiv>
\<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
is_wfrec(M,MH,mesa,a,z)"
definition
transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
- "transrec_replacement(M,MH,a) ==
+ "transrec_replacement(M,MH,a) \<equiv>
\<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
wfrec_replacement(M,MH,mesa)"
@@ -667,30 +667,30 @@
\<open>trans_wfrec_abs\<close> rather than \<open>trans_wfrec_abs\<close>,
which I haven't even proved yet.\<close>
theorem (in M_eclose) transrec_abs:
- "[|transrec_replacement(M,MH,i); relation2(M,MH,H);
+ "\<lbrakk>transrec_replacement(M,MH,i); relation2(M,MH,H);
Ord(i); M(i); M(z);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> is_transrec(M,MH,i,z) \<longleftrightarrow> z = transrec(i,H)"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> is_transrec(M,MH,i,z) \<longleftrightarrow> z = transrec(i,H)"
by (simp add: trans_wfrec_abs transrec_replacement_def is_transrec_def
transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
theorem (in M_eclose) transrec_closed:
- "[|transrec_replacement(M,MH,i); relation2(M,MH,H);
+ "\<lbrakk>transrec_replacement(M,MH,i); relation2(M,MH,H);
Ord(i); M(i);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> M(transrec(i,H))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> M(transrec(i,H))"
by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def
transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
text\<open>Helps to prove instances of \<^term>\<open>transrec_replacement\<close>\<close>
lemma (in M_eclose) transrec_replacementI:
- "[|M(a);
+ "\<lbrakk>M(a);
strong_replacement (M,
\<lambda>x z. \<exists>y[M]. pair(M, x, y, z) &
- is_wfrec(M,MH,Memrel(eclose({a})),x,y))|]
- ==> transrec_replacement(M,MH,a)"
+ is_wfrec(M,MH,Memrel(eclose({a})),x,y))\<rbrakk>
+ \<Longrightarrow> transrec_replacement(M,MH,a)"
by (simp add: transrec_replacement_def wfrec_replacement_def)
@@ -699,14 +699,14 @@
definition
is_length :: "[i=>o,i,i,i] => o" where
- "is_length(M,A,l,n) ==
+ "is_length(M,A,l,n) \<equiv>
\<exists>sn[M]. \<exists>list_n[M]. \<exists>list_sn[M].
is_list_N(M,A,n,list_n) & l \<notin> list_n &
successor(M,n,sn) & is_list_N(M,A,sn,list_sn) & l \<in> list_sn"
lemma (in M_datatypes) length_abs [simp]:
- "[|M(A); l \<in> list(A); n \<in> nat|] ==> is_length(M,A,l,n) \<longleftrightarrow> n = length(l)"
+ "\<lbrakk>M(A); l \<in> list(A); n \<in> nat\<rbrakk> \<Longrightarrow> is_length(M,A,l,n) \<longleftrightarrow> n = length(l)"
apply (subgoal_tac "M(l) & M(n)")
prefer 2 apply (blast dest: transM)
apply (simp add: is_length_def)
@@ -716,14 +716,14 @@
text\<open>Proof is trivial since \<^term>\<open>length\<close> returns natural numbers.\<close>
lemma (in M_trivial) length_closed [intro,simp]:
- "l \<in> list(A) ==> M(length(l))"
+ "l \<in> list(A) \<Longrightarrow> M(length(l))"
by (simp add: nat_into_M)
subsection \<open>Absoluteness for the List Operator \<^term>\<open>nth\<close>\<close>
lemma nth_eq_hd_iterates_tl [rule_format]:
- "xs \<in> list(A) ==> \<forall>n \<in> nat. nth(n,xs) = hd' (tl'^n (xs))"
+ "xs \<in> list(A) \<Longrightarrow> \<forall>n \<in> nat. nth(n,xs) = hd' (tl'^n (xs))"
apply (induct_tac xs)
apply (simp add: iterates_tl_Nil hd'_Nil, clarify)
apply (erule natE)
@@ -732,14 +732,14 @@
done
lemma (in M_basic) iterates_tl'_closed:
- "[|n \<in> nat; M(x)|] ==> M(tl'^n (x))"
+ "\<lbrakk>n \<in> nat; M(x)\<rbrakk> \<Longrightarrow> M(tl'^n (x))"
apply (induct_tac n, simp)
apply (simp add: tl'_Cons tl'_closed)
done
text\<open>Immediate by type-checking\<close>
lemma (in M_datatypes) nth_closed [intro,simp]:
- "[|xs \<in> list(A); n \<in> nat; M(A)|] ==> M(nth(n,xs))"
+ "\<lbrakk>xs \<in> list(A); n \<in> nat; M(A)\<rbrakk> \<Longrightarrow> M(nth(n,xs))"
apply (case_tac "n < length(xs)")
apply (blast intro: nth_type transM)
apply (simp add: not_lt_iff_le nth_eq_0)
@@ -747,12 +747,12 @@
definition
is_nth :: "[i=>o,i,i,i] => o" where
- "is_nth(M,n,l,Z) ==
+ "is_nth(M,n,l,Z) \<equiv>
\<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)"
lemma (in M_datatypes) nth_abs [simp]:
- "[|M(A); n \<in> nat; l \<in> list(A); M(Z)|]
- ==> is_nth(M,n,l,Z) \<longleftrightarrow> Z = nth(n,l)"
+ "\<lbrakk>M(A); n \<in> nat; l \<in> list(A); M(Z)\<rbrakk>
+ \<Longrightarrow> is_nth(M,n,l,Z) \<longleftrightarrow> Z = nth(n,l)"
apply (subgoal_tac "M(l)")
prefer 2 apply (blast intro: transM)
apply (simp add: is_nth_def nth_eq_hd_iterates_tl nat_into_M
@@ -766,11 +766,11 @@
definition
is_Member :: "[i=>o,i,i,i] => o" where
\<comment> \<open>because \<^term>\<open>Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))\<close>\<close>
- "is_Member(M,x,y,Z) ==
+ "is_Member(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
lemma (in M_trivial) Member_abs [simp]:
- "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) \<longleftrightarrow> (Z = Member(x,y))"
+ "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Member(M,x,y,Z) \<longleftrightarrow> (Z = Member(x,y))"
by (simp add: is_Member_def Member_def)
lemma (in M_trivial) Member_in_M_iff [iff]:
@@ -780,11 +780,11 @@
definition
is_Equal :: "[i=>o,i,i,i] => o" where
\<comment> \<open>because \<^term>\<open>Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))\<close>\<close>
- "is_Equal(M,x,y,Z) ==
+ "is_Equal(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
lemma (in M_trivial) Equal_abs [simp]:
- "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) \<longleftrightarrow> (Z = Equal(x,y))"
+ "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Equal(M,x,y,Z) \<longleftrightarrow> (Z = Equal(x,y))"
by (simp add: is_Equal_def Equal_def)
lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \<longleftrightarrow> M(x) & M(y)"
@@ -793,11 +793,11 @@
definition
is_Nand :: "[i=>o,i,i,i] => o" where
\<comment> \<open>because \<^term>\<open>Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))\<close>\<close>
- "is_Nand(M,x,y,Z) ==
+ "is_Nand(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
lemma (in M_trivial) Nand_abs [simp]:
- "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) \<longleftrightarrow> (Z = Nand(x,y))"
+ "\<lbrakk>M(x); M(y); M(Z)\<rbrakk> \<Longrightarrow> is_Nand(M,x,y,Z) \<longleftrightarrow> (Z = Nand(x,y))"
by (simp add: is_Nand_def Nand_def)
lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \<longleftrightarrow> M(x) & M(y)"
@@ -806,10 +806,10 @@
definition
is_Forall :: "[i=>o,i,i] => o" where
\<comment> \<open>because \<^term>\<open>Forall(x) \<equiv> Inr(Inr(p))\<close>\<close>
- "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)"
+ "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)"
lemma (in M_trivial) Forall_abs [simp]:
- "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) \<longleftrightarrow> (Z = Forall(x))"
+ "\<lbrakk>M(x); M(Z)\<rbrakk> \<Longrightarrow> is_Forall(M,x,Z) \<longleftrightarrow> (Z = Forall(x))"
by (simp add: is_Forall_def Forall_def)
lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) \<longleftrightarrow> M(x)"
@@ -822,7 +822,7 @@
definition
formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where
\<comment> \<open>the instance of \<^term>\<open>formula_case\<close> in \<^term>\<open>formula_rec\<close>\<close>
- "formula_rec_case(a,b,c,d,h) ==
+ "formula_rec_case(a,b,c,d,h) \<equiv>
formula_case (a, b,
\<lambda>u v. c(u, v, h ` succ(depth(u)) ` u,
h ` succ(depth(v)) ` v),
@@ -832,7 +832,7 @@
Express \<^term>\<open>formula_rec\<close> without using \<^term>\<open>rank\<close> or \<^term>\<open>Vset\<close>,
neither of which is absolute.\<close>
lemma (in M_trivial) formula_rec_eq:
- "p \<in> formula ==>
+ "p \<in> formula \<Longrightarrow>
formula_rec(a,b,c,d,p) =
transrec (succ(depth(p)),
\<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h))) ` p"
@@ -855,14 +855,14 @@
definition
is_depth :: "[i=>o,i,i] => o" where
- "is_depth(M,p,n) ==
+ "is_depth(M,p,n) \<equiv>
\<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M].
is_formula_N(M,n,formula_n) & p \<notin> formula_n &
successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn"
lemma (in M_datatypes) depth_abs [simp]:
- "[|p \<in> formula; n \<in> nat|] ==> is_depth(M,p,n) \<longleftrightarrow> n = depth(p)"
+ "\<lbrakk>p \<in> formula; n \<in> nat\<rbrakk> \<Longrightarrow> is_depth(M,p,n) \<longleftrightarrow> n = depth(p)"
apply (subgoal_tac "M(p) & M(n)")
prefer 2 apply (blast dest: transM)
apply (simp add: is_depth_def)
@@ -872,7 +872,7 @@
text\<open>Proof is trivial since \<^term>\<open>depth\<close> returns natural numbers.\<close>
lemma (in M_trivial) depth_closed [intro,simp]:
- "p \<in> formula ==> M(depth(p))"
+ "p \<in> formula \<Longrightarrow> M(depth(p))"
by (simp add: nat_into_M)
@@ -882,7 +882,7 @@
is_formula_case ::
"[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where
\<comment> \<open>no constraint on non-formulas\<close>
- "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) ==
+ "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) \<equiv>
(\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
is_Member(M,x,y,p) \<longrightarrow> is_a(x,y,z)) &
(\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
@@ -892,10 +892,10 @@
(\<forall>x[M]. mem_formula(M,x) \<longrightarrow> is_Forall(M,x,p) \<longrightarrow> is_d(x,z))"
lemma (in M_datatypes) formula_case_abs [simp]:
- "[| Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b);
+ "\<lbrakk>Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b);
Relation2(M,formula,formula,is_c,c); Relation1(M,formula,is_d,d);
- p \<in> formula; M(z) |]
- ==> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) \<longleftrightarrow>
+ p \<in> formula; M(z)\<rbrakk>
+ \<Longrightarrow> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) \<longleftrightarrow>
z = formula_case(a,b,c,d,p)"
apply (simp add: formula_into_M is_formula_case_def)
apply (erule formula.cases)
@@ -903,11 +903,11 @@
done
lemma (in M_datatypes) formula_case_closed [intro,simp]:
- "[|p \<in> formula;
+ "\<lbrakk>p \<in> formula;
\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> M(a(x,y));
\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> M(b(x,y));
\<forall>x[M]. \<forall>y[M]. x\<in>formula \<longrightarrow> y\<in>formula \<longrightarrow> M(c(x,y));
- \<forall>x[M]. x\<in>formula \<longrightarrow> M(d(x))|] ==> M(formula_case(a,b,c,d,p))"
+ \<forall>x[M]. x\<in>formula \<longrightarrow> M(d(x))\<rbrakk> \<Longrightarrow> M(formula_case(a,b,c,d,p))"
by (erule formula.cases, simp_all)
@@ -916,7 +916,7 @@
definition
is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
\<comment> \<open>predicate to relativize the functional \<^term>\<open>formula_rec\<close>\<close>
- "is_formula_rec(M,MH,p,z) ==
+ "is_formula_rec(M,MH,p,z) \<equiv>
\<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) &
successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
@@ -924,14 +924,14 @@
text\<open>Sufficient conditions to relativize the instance of \<^term>\<open>formula_case\<close>
in \<^term>\<open>formula_rec\<close>\<close>
lemma (in M_datatypes) Relation1_formula_rec_case:
- "[|Relation2(M, nat, nat, is_a, a);
+ "\<lbrakk>Relation2(M, nat, nat, is_a, a);
Relation2(M, nat, nat, is_b, b);
Relation2 (M, formula, formula,
is_c, \<lambda>u v. c(u, v, h`succ(depth(u))`u, h`succ(depth(v))`v));
Relation1(M, formula,
is_d, \<lambda>u. d(u, h ` succ(depth(u)) ` u));
- M(h) |]
- ==> Relation1(M, formula,
+ M(h)\<rbrakk>
+ \<Longrightarrow> Relation1(M, formula,
is_formula_case (M, is_a, is_b, is_c, is_d),
formula_rec_case(a, b, c, d, h))"
apply (simp (no_asm) add: formula_rec_case_def Relation1_def)
@@ -945,38 +945,38 @@
locale Formula_Rec = M_eclose +
fixes a and is_a and b and is_b and c and is_c and d and is_d and MH
defines
- "MH(u::i,f,z) ==
+ "MH(u::i,f,z) \<equiv>
\<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
is_lambda
(M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)"
- assumes a_closed: "[|x\<in>nat; y\<in>nat|] ==> M(a(x,y))"
+ assumes a_closed: "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> M(a(x,y))"
and a_rel: "Relation2(M, nat, nat, is_a, a)"
- and b_closed: "[|x\<in>nat; y\<in>nat|] ==> M(b(x,y))"
+ and b_closed: "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> M(b(x,y))"
and b_rel: "Relation2(M, nat, nat, is_b, b)"
- and c_closed: "[|x \<in> formula; y \<in> formula; M(gx); M(gy)|]
- ==> M(c(x, y, gx, gy))"
+ and c_closed: "\<lbrakk>x \<in> formula; y \<in> formula; M(gx); M(gy)\<rbrakk>
+ \<Longrightarrow> M(c(x, y, gx, gy))"
and c_rel:
- "M(f) ==>
+ "M(f) \<Longrightarrow>
Relation2 (M, formula, formula, is_c(f),
\<lambda>u v. c(u, v, f ` succ(depth(u)) ` u, f ` succ(depth(v)) ` v))"
- and d_closed: "[|x \<in> formula; M(gx)|] ==> M(d(x, gx))"
+ and d_closed: "\<lbrakk>x \<in> formula; M(gx)\<rbrakk> \<Longrightarrow> M(d(x, gx))"
and d_rel:
- "M(f) ==>
+ "M(f) \<Longrightarrow>
Relation1(M, formula, is_d(f), \<lambda>u. d(u, f ` succ(depth(u)) ` u))"
- and fr_replace: "n \<in> nat ==> transrec_replacement(M,MH,n)"
+ and fr_replace: "n \<in> nat \<Longrightarrow> transrec_replacement(M,MH,n)"
and fr_lam_replace:
- "M(g) ==>
+ "M(g) \<Longrightarrow>
strong_replacement
(M, \<lambda>x y. x \<in> formula &
y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)"
lemma (in Formula_Rec) formula_rec_case_closed:
- "[|M(g); p \<in> formula|] ==> M(formula_rec_case(a, b, c, d, g, p))"
+ "\<lbrakk>M(g); p \<in> formula\<rbrakk> \<Longrightarrow> M(formula_rec_case(a, b, c, d, g, p))"
by (simp add: formula_rec_case_def a_closed b_closed c_closed d_closed)
lemma (in Formula_Rec) formula_rec_lam_closed:
- "M(g) ==> M(Lambda (formula, formula_rec_case(a,b,c,d,g)))"
+ "M(g) \<Longrightarrow> M(Lambda (formula, formula_rec_case(a,b,c,d,g)))"
by (simp add: lam_closed2 fr_lam_replace formula_rec_case_closed)
lemma (in Formula_Rec) MH_rel2:
@@ -990,20 +990,20 @@
lemma (in Formula_Rec) fr_transrec_closed:
"n \<in> nat
- ==> M(transrec
+ \<Longrightarrow> M(transrec
(n, \<lambda>x h. Lambda(formula, formula_rec_case(a, b, c, d, h))))"
by (simp add: transrec_closed [OF fr_replace MH_rel2]
nat_into_M formula_rec_lam_closed)
text\<open>The main two results: \<^term>\<open>formula_rec\<close> is absolute for \<^term>\<open>M\<close>.\<close>
theorem (in Formula_Rec) formula_rec_closed:
- "p \<in> formula ==> M(formula_rec(a,b,c,d,p))"
+ "p \<in> formula \<Longrightarrow> M(formula_rec(a,b,c,d,p))"
by (simp add: formula_rec_eq fr_transrec_closed
transM [OF _ formula_closed])
theorem (in Formula_Rec) formula_rec_abs:
- "[| p \<in> formula; M(z)|]
- ==> is_formula_rec(M,MH,p,z) \<longleftrightarrow> z = formula_rec(a,b,c,d,p)"
+ "\<lbrakk>p \<in> formula; M(z)\<rbrakk>
+ \<Longrightarrow> is_formula_rec(M,MH,p,z) \<longleftrightarrow> z = formula_rec(a,b,c,d,p)"
by (simp add: is_formula_rec_def formula_rec_eq transM [OF _ formula_closed]
transrec_abs [OF fr_replace MH_rel2] depth_type
fr_transrec_closed formula_rec_lam_closed eq_commute)
--- a/src/ZF/Constructible/Formula.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Formula.thy Tue Sep 27 16:51:35 2022 +0100
@@ -22,46 +22,46 @@
definition
Neg :: "i=>i" where
- "Neg(p) == Nand(p,p)"
+ "Neg(p) \<equiv> Nand(p,p)"
definition
And :: "[i,i]=>i" where
- "And(p,q) == Neg(Nand(p,q))"
+ "And(p,q) \<equiv> Neg(Nand(p,q))"
definition
Or :: "[i,i]=>i" where
- "Or(p,q) == Nand(Neg(p),Neg(q))"
+ "Or(p,q) \<equiv> Nand(Neg(p),Neg(q))"
definition
Implies :: "[i,i]=>i" where
- "Implies(p,q) == Nand(p,Neg(q))"
+ "Implies(p,q) \<equiv> Nand(p,Neg(q))"
definition
Iff :: "[i,i]=>i" where
- "Iff(p,q) == And(Implies(p,q), Implies(q,p))"
+ "Iff(p,q) \<equiv> And(Implies(p,q), Implies(q,p))"
definition
Exists :: "i=>i" where
- "Exists(p) == Neg(Forall(Neg(p)))"
+ "Exists(p) \<equiv> Neg(Forall(Neg(p)))"
-lemma Neg_type [TC]: "p \<in> formula ==> Neg(p) \<in> formula"
+lemma Neg_type [TC]: "p \<in> formula \<Longrightarrow> Neg(p) \<in> formula"
by (simp add: Neg_def)
-lemma And_type [TC]: "[| p \<in> formula; q \<in> formula |] ==> And(p,q) \<in> formula"
+lemma And_type [TC]: "\<lbrakk>p \<in> formula; q \<in> formula\<rbrakk> \<Longrightarrow> And(p,q) \<in> formula"
by (simp add: And_def)
-lemma Or_type [TC]: "[| p \<in> formula; q \<in> formula |] ==> Or(p,q) \<in> formula"
+lemma Or_type [TC]: "\<lbrakk>p \<in> formula; q \<in> formula\<rbrakk> \<Longrightarrow> Or(p,q) \<in> formula"
by (simp add: Or_def)
lemma Implies_type [TC]:
- "[| p \<in> formula; q \<in> formula |] ==> Implies(p,q) \<in> formula"
+ "\<lbrakk>p \<in> formula; q \<in> formula\<rbrakk> \<Longrightarrow> Implies(p,q) \<in> formula"
by (simp add: Implies_def)
lemma Iff_type [TC]:
- "[| p \<in> formula; q \<in> formula |] ==> Iff(p,q) \<in> formula"
+ "\<lbrakk>p \<in> formula; q \<in> formula\<rbrakk> \<Longrightarrow> Iff(p,q) \<in> formula"
by (simp add: Iff_def)
-lemma Exists_type [TC]: "p \<in> formula ==> Exists(p) \<in> formula"
+lemma Exists_type [TC]: "p \<in> formula \<Longrightarrow> Exists(p) \<in> formula"
by (simp add: Exists_def)
@@ -85,7 +85,7 @@
abbreviation
sats :: "[i,i,i] => o" where
- "sats(A,p,env) == satisfies(A,p)`env = 1"
+ "sats(A,p,env) \<equiv> satisfies(A,p)`env = 1"
lemma sats_Member_iff [simp]:
"env \<in> list(A) \<Longrightarrow> sats(A, Member(x,y), env) \<longleftrightarrow> nth(x,env) \<in> nth(y,env)"
@@ -97,12 +97,12 @@
lemma sats_Nand_iff [simp]:
"env \<in> list(A)
- ==> (sats(A, Nand(p,q), env)) \<longleftrightarrow> ~ (sats(A,p,env) & sats(A,q,env))"
+ \<Longrightarrow> (sats(A, Nand(p,q), env)) \<longleftrightarrow> \<not> (sats(A,p,env) & sats(A,q,env))"
by (simp add: Bool.and_def Bool.not_def cond_def)
lemma sats_Forall_iff [simp]:
"env \<in> list(A)
- ==> sats(A, Forall(p), env) \<longleftrightarrow> (\<forall>x\<in>A. sats(A, p, Cons(x,env)))"
+ \<Longrightarrow> sats(A, Forall(p), env) \<longleftrightarrow> (\<forall>x\<in>A. sats(A, p, Cons(x,env)))"
by simp
declare satisfies.simps [simp del]
@@ -111,80 +111,80 @@
lemma sats_Neg_iff [simp]:
"env \<in> list(A)
- ==> sats(A, Neg(p), env) \<longleftrightarrow> ~ sats(A,p,env)"
+ \<Longrightarrow> sats(A, Neg(p), env) \<longleftrightarrow> \<not> sats(A,p,env)"
by (simp add: Neg_def)
lemma sats_And_iff [simp]:
"env \<in> list(A)
- ==> (sats(A, And(p,q), env)) \<longleftrightarrow> sats(A,p,env) & sats(A,q,env)"
+ \<Longrightarrow> (sats(A, And(p,q), env)) \<longleftrightarrow> sats(A,p,env) & sats(A,q,env)"
by (simp add: And_def)
lemma sats_Or_iff [simp]:
"env \<in> list(A)
- ==> (sats(A, Or(p,q), env)) \<longleftrightarrow> sats(A,p,env) | sats(A,q,env)"
+ \<Longrightarrow> (sats(A, Or(p,q), env)) \<longleftrightarrow> sats(A,p,env) | sats(A,q,env)"
by (simp add: Or_def)
lemma sats_Implies_iff [simp]:
"env \<in> list(A)
- ==> (sats(A, Implies(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longrightarrow> sats(A,q,env))"
+ \<Longrightarrow> (sats(A, Implies(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longrightarrow> sats(A,q,env))"
by (simp add: Implies_def, blast)
lemma sats_Iff_iff [simp]:
"env \<in> list(A)
- ==> (sats(A, Iff(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longleftrightarrow> sats(A,q,env))"
+ \<Longrightarrow> (sats(A, Iff(p,q), env)) \<longleftrightarrow> (sats(A,p,env) \<longleftrightarrow> sats(A,q,env))"
by (simp add: Iff_def, blast)
lemma sats_Exists_iff [simp]:
"env \<in> list(A)
- ==> sats(A, Exists(p), env) \<longleftrightarrow> (\<exists>x\<in>A. sats(A, p, Cons(x,env)))"
+ \<Longrightarrow> sats(A, Exists(p), env) \<longleftrightarrow> (\<exists>x\<in>A. sats(A, p, Cons(x,env)))"
by (simp add: Exists_def)
subsubsection\<open>Derived rules to help build up formulas\<close>
lemma mem_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
- ==> (x\<in>y) \<longleftrightarrow> sats(A, Member(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (x\<in>y) \<longleftrightarrow> sats(A, Member(i,j), env)"
by (simp add: satisfies.simps)
lemma equal_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
- ==> (x=y) \<longleftrightarrow> sats(A, Equal(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (x=y) \<longleftrightarrow> sats(A, Equal(i,j), env)"
by (simp add: satisfies.simps)
lemma not_iff_sats:
- "[| P \<longleftrightarrow> sats(A,p,env); env \<in> list(A)|]
- ==> (~P) \<longleftrightarrow> sats(A, Neg(p), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (\<not>P) \<longleftrightarrow> sats(A, Neg(p), env)"
by simp
lemma conj_iff_sats:
- "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
- ==> (P & Q) \<longleftrightarrow> sats(A, And(p,q), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (P & Q) \<longleftrightarrow> sats(A, And(p,q), env)"
by (simp)
lemma disj_iff_sats:
- "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
- ==> (P | Q) \<longleftrightarrow> sats(A, Or(p,q), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (P | Q) \<longleftrightarrow> sats(A, Or(p,q), env)"
by (simp)
lemma iff_iff_sats:
- "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
- ==> (P \<longleftrightarrow> Q) \<longleftrightarrow> sats(A, Iff(p,q), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> sats(A, Iff(p,q), env)"
by (simp)
lemma imp_iff_sats:
- "[| P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)|]
- ==> (P \<longrightarrow> Q) \<longleftrightarrow> sats(A, Implies(p,q), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A,p,env); Q \<longleftrightarrow> sats(A,q,env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (P \<longrightarrow> Q) \<longleftrightarrow> sats(A, Implies(p,q), env)"
by (simp)
lemma ball_iff_sats:
- "[| !!x. x\<in>A ==> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)|]
- ==> (\<forall>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Forall(p), env)"
+ "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (\<forall>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Forall(p), env)"
by (simp)
lemma bex_iff_sats:
- "[| !!x. x\<in>A ==> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)|]
- ==> (\<exists>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Exists(p), env)"
+ "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> P(x) \<longleftrightarrow> sats(A, p, Cons(x, env)); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> (\<exists>x\<in>A. P(x)) \<longleftrightarrow> sats(A, Exists(p), env)"
by (simp)
lemmas FOL_iff_sats =
@@ -206,7 +206,7 @@
"arity(Forall(p)) = Arith.pred(arity(p))"
-lemma arity_type [TC]: "p \<in> formula ==> arity(p) \<in> nat"
+lemma arity_type [TC]: "p \<in> formula \<Longrightarrow> arity(p) \<in> nat"
by (induct_tac p, simp_all)
lemma arity_Neg [simp]: "arity(Neg(p)) = arity(p)"
@@ -229,8 +229,8 @@
lemma arity_sats_iff [rule_format]:
- "[| p \<in> formula; extra \<in> list(A) |]
- ==> \<forall>env \<in> list(A).
+ "\<lbrakk>p \<in> formula; extra \<in> list(A)\<rbrakk>
+ \<Longrightarrow> \<forall>env \<in> list(A).
arity(p) \<le> length(env) \<longrightarrow>
sats(A, p, env @ extra) \<longleftrightarrow> sats(A, p, env)"
apply (induct_tac p)
@@ -239,9 +239,9 @@
done
lemma arity_sats1_iff:
- "[| arity(p) \<le> succ(length(env)); p \<in> formula; x \<in> A; env \<in> list(A);
- extra \<in> list(A) |]
- ==> sats(A, p, Cons(x, env @ extra)) \<longleftrightarrow> sats(A, p, Cons(x, env))"
+ "\<lbrakk>arity(p) \<le> succ(length(env)); p \<in> formula; x \<in> A; env \<in> list(A);
+ extra \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, p, Cons(x, env @ extra)) \<longleftrightarrow> sats(A, p, Cons(x, env))"
apply (insert arity_sats_iff [of p extra A "Cons(x,env)"])
apply simp
done
@@ -251,12 +251,12 @@
definition
incr_var :: "[i,i]=>i" where
- "incr_var(x,nq) == if x<nq then x else succ(x)"
+ "incr_var(x,nq) \<equiv> if x<nq then x else succ(x)"
-lemma incr_var_lt: "x<nq ==> incr_var(x,nq) = x"
+lemma incr_var_lt: "x<nq \<Longrightarrow> incr_var(x,nq) = x"
by (simp add: incr_var_def)
-lemma incr_var_le: "nq\<le>x ==> incr_var(x,nq) = succ(x)"
+lemma incr_var_le: "nq\<le>x \<Longrightarrow> incr_var(x,nq) = succ(x)"
apply (simp add: incr_var_def)
apply (blast dest: lt_trans1)
done
@@ -276,10 +276,10 @@
(\<lambda>nq \<in> nat. Forall (incr_bv(p) ` succ(nq)))"
-lemma [TC]: "x \<in> nat ==> incr_var(x,nq) \<in> nat"
+lemma [TC]: "x \<in> nat \<Longrightarrow> incr_var(x,nq) \<in> nat"
by (simp add: incr_var_def)
-lemma incr_bv_type [TC]: "p \<in> formula ==> incr_bv(p) \<in> nat -> formula"
+lemma incr_bv_type [TC]: "p \<in> formula \<Longrightarrow> incr_bv(p) \<in> nat -> formula"
by (induct_tac p, simp_all)
text\<open>Obviously, \<^term>\<open>DPow\<close> is closed under complements and finite
@@ -287,8 +287,8 @@
parameters to be combined.\<close>
lemma sats_incr_bv_iff [rule_format]:
- "[| p \<in> formula; env \<in> list(A); x \<in> A |]
- ==> \<forall>bvs \<in> list(A).
+ "\<lbrakk>p \<in> formula; env \<in> list(A); x \<in> A\<rbrakk>
+ \<Longrightarrow> \<forall>bvs \<in> list(A).
sats(A, incr_bv(p) ` length(bvs), bvs @ Cons(x,env)) \<longleftrightarrow>
sats(A, p, bvs@env)"
apply (induct_tac p)
@@ -299,8 +299,8 @@
(*the following two lemmas prevent huge case splits in arity_incr_bv_lemma*)
lemma incr_var_lemma:
- "[| x \<in> nat; y \<in> nat; nq \<le> x |]
- ==> succ(x) \<union> incr_var(y,nq) = succ(x \<union> y)"
+ "\<lbrakk>x \<in> nat; y \<in> nat; nq \<le> x\<rbrakk>
+ \<Longrightarrow> succ(x) \<union> incr_var(y,nq) = succ(x \<union> y)"
apply (simp add: incr_var_def Ord_Un_if, auto)
apply (blast intro: leI)
apply (simp add: not_lt_iff_le)
@@ -309,14 +309,14 @@
done
lemma incr_And_lemma:
- "y < x ==> y \<union> succ(x) = succ(x \<union> y)"
+ "y < x \<Longrightarrow> y \<union> succ(x) = succ(x \<union> y)"
apply (simp add: Ord_Un_if lt_Ord lt_Ord2 succ_lt_iff)
apply (blast dest: lt_asym)
done
lemma arity_incr_bv_lemma [rule_format]:
"p \<in> formula
- ==> \<forall>n \<in> nat. arity (incr_bv(p) ` n) =
+ \<Longrightarrow> \<forall>n \<in> nat. arity (incr_bv(p) ` n) =
(if n < arity(p) then succ(arity(p)) else arity(p))"
apply (induct_tac p)
apply (simp_all add: imp_disj not_lt_iff_le Un_least_lt_iff lt_Un_iff le_Un_iff
@@ -340,24 +340,24 @@
definition
incr_bv1 :: "i => i" where
- "incr_bv1(p) == incr_bv(p)`1"
+ "incr_bv1(p) \<equiv> incr_bv(p)`1"
-lemma incr_bv1_type [TC]: "p \<in> formula ==> incr_bv1(p) \<in> formula"
+lemma incr_bv1_type [TC]: "p \<in> formula \<Longrightarrow> incr_bv1(p) \<in> formula"
by (simp add: incr_bv1_def)
(*For renaming all but the bound variable at level 0*)
lemma sats_incr_bv1_iff:
- "[| p \<in> formula; env \<in> list(A); x \<in> A; y \<in> A |]
- ==> sats(A, incr_bv1(p), Cons(x, Cons(y, env))) \<longleftrightarrow>
+ "\<lbrakk>p \<in> formula; env \<in> list(A); x \<in> A; y \<in> A\<rbrakk>
+ \<Longrightarrow> sats(A, incr_bv1(p), Cons(x, Cons(y, env))) \<longleftrightarrow>
sats(A, p, Cons(x,env))"
apply (insert sats_incr_bv_iff [of p env A y "Cons(x,Nil)"])
apply (simp add: incr_bv1_def)
done
lemma formula_add_params1 [rule_format]:
- "[| p \<in> formula; n \<in> nat; x \<in> A |]
- ==> \<forall>bvs \<in> list(A). \<forall>env \<in> list(A).
+ "\<lbrakk>p \<in> formula; n \<in> nat; x \<in> A\<rbrakk>
+ \<Longrightarrow> \<forall>bvs \<in> list(A). \<forall>env \<in> list(A).
length(bvs) = n \<longrightarrow>
sats(A, iterates(incr_bv1, n, p), Cons(x, bvs@env)) \<longleftrightarrow>
sats(A, p, Cons(x,env))"
@@ -369,15 +369,15 @@
lemma arity_incr_bv1_eq:
"p \<in> formula
- ==> arity(incr_bv1(p)) =
+ \<Longrightarrow> arity(incr_bv1(p)) =
(if 1 < arity(p) then succ(arity(p)) else arity(p))"
apply (insert arity_incr_bv_lemma [of p 1])
apply (simp add: incr_bv1_def)
done
lemma arity_iterates_incr_bv1_eq:
- "[| p \<in> formula; n \<in> nat |]
- ==> arity(incr_bv1^n(p)) =
+ "\<lbrakk>p \<in> formula; n \<in> nat\<rbrakk>
+ \<Longrightarrow> arity(incr_bv1^n(p)) =
(if 1 < arity(p) then n #+ arity(p) else arity(p))"
apply (induct_tac n)
apply (simp_all add: arity_incr_bv1_eq)
@@ -392,26 +392,26 @@
text\<open>The definable powerset operation: Kunen's definition VI 1.1, page 165.\<close>
definition
DPow :: "i => i" where
- "DPow(A) == {X \<in> Pow(A).
+ "DPow(A) \<equiv> {X \<in> Pow(A).
\<exists>env \<in> list(A). \<exists>p \<in> formula.
arity(p) \<le> succ(length(env)) &
X = {x\<in>A. sats(A, p, Cons(x,env))}}"
lemma DPowI:
- "[|env \<in> list(A); p \<in> formula; arity(p) \<le> succ(length(env))|]
- ==> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
+ "\<lbrakk>env \<in> list(A); p \<in> formula; arity(p) \<le> succ(length(env))\<rbrakk>
+ \<Longrightarrow> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
by (simp add: DPow_def, blast)
text\<open>With this rule we can specify \<^term>\<open>p\<close> later.\<close>
lemma DPowI2 [rule_format]:
- "[|\<forall>x\<in>A. P(x) \<longleftrightarrow> sats(A, p, Cons(x,env));
- env \<in> list(A); p \<in> formula; arity(p) \<le> succ(length(env))|]
- ==> {x\<in>A. P(x)} \<in> DPow(A)"
+ "\<lbrakk>\<forall>x\<in>A. P(x) \<longleftrightarrow> sats(A, p, Cons(x,env));
+ env \<in> list(A); p \<in> formula; arity(p) \<le> succ(length(env))\<rbrakk>
+ \<Longrightarrow> {x\<in>A. P(x)} \<in> DPow(A)"
by (simp add: DPow_def, blast)
lemma DPowD:
"X \<in> DPow(A)
- ==> X \<subseteq> A &
+ \<Longrightarrow> X \<subseteq> A &
(\<exists>env \<in> list(A).
\<exists>p \<in> formula. arity(p) \<le> succ(length(env)) &
X = {x\<in>A. sats(A, p, Cons(x,env))})"
@@ -420,8 +420,8 @@
lemmas DPow_imp_subset = DPowD [THEN conjunct1]
(*Kunen's Lemma VI 1.2*)
-lemma "[| p \<in> formula; env \<in> list(A); arity(p) \<le> succ(length(env)) |]
- ==> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
+lemma "\<lbrakk>p \<in> formula; env \<in> list(A); arity(p) \<le> succ(length(env))\<rbrakk>
+ \<Longrightarrow> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
by (blast intro: DPowI)
lemma DPow_subset_Pow: "DPow(A) \<subseteq> Pow(A)"
@@ -434,14 +434,14 @@
apply (auto simp add: Un_least_lt_iff)
done
-lemma Compl_in_DPow: "X \<in> DPow(A) ==> (A-X) \<in> DPow(A)"
+lemma Compl_in_DPow: "X \<in> DPow(A) \<Longrightarrow> (A-X) \<in> DPow(A)"
apply (simp add: DPow_def, clarify, auto)
apply (rule bexI)
apply (rule_tac x="Neg(p)" in bexI)
apply auto
done
-lemma Int_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X \<inter> Y \<in> DPow(A)"
+lemma Int_in_DPow: "\<lbrakk>X \<in> DPow(A); Y \<in> DPow(A)\<rbrakk> \<Longrightarrow> X \<inter> Y \<in> DPow(A)"
apply (simp add: DPow_def, auto)
apply (rename_tac envp p envq q)
apply (rule_tac x="envp@envq" in bexI)
@@ -454,13 +454,13 @@
apply (simp add: arity_sats1_iff formula_add_params1, blast)
done
-lemma Un_in_DPow: "[| X \<in> DPow(A); Y \<in> DPow(A) |] ==> X \<union> Y \<in> DPow(A)"
+lemma Un_in_DPow: "\<lbrakk>X \<in> DPow(A); Y \<in> DPow(A)\<rbrakk> \<Longrightarrow> X \<union> Y \<in> DPow(A)"
apply (subgoal_tac "X \<union> Y = A - ((A-X) \<inter> (A-Y))")
apply (simp add: Int_in_DPow Compl_in_DPow)
apply (simp add: DPow_def, blast)
done
-lemma singleton_in_DPow: "a \<in> A ==> {a} \<in> DPow(A)"
+lemma singleton_in_DPow: "a \<in> A \<Longrightarrow> {a} \<in> DPow(A)"
apply (simp add: DPow_def)
apply (rule_tac x="Cons(a,Nil)" in bexI)
apply (rule_tac x="Equal(0,1)" in bexI)
@@ -468,13 +468,13 @@
apply (force simp add: succ_Un_distrib [symmetric])
done
-lemma cons_in_DPow: "[| a \<in> A; X \<in> DPow(A) |] ==> cons(a,X) \<in> DPow(A)"
+lemma cons_in_DPow: "\<lbrakk>a \<in> A; X \<in> DPow(A)\<rbrakk> \<Longrightarrow> cons(a,X) \<in> DPow(A)"
apply (rule cons_eq [THEN subst])
apply (blast intro: singleton_in_DPow Un_in_DPow)
done
(*Part of Lemma 1.3*)
-lemma Fin_into_DPow: "X \<in> Fin(A) ==> X \<in> DPow(A)"
+lemma Fin_into_DPow: "X \<in> Fin(A) \<Longrightarrow> X \<in> DPow(A)"
apply (erule Fin.induct)
apply (rule empty_in_DPow)
apply (blast intro: cons_in_DPow)
@@ -491,10 +491,10 @@
oops
*)
-lemma Finite_Pow_subset_Pow: "Finite(A) ==> Pow(A) \<subseteq> DPow(A)"
+lemma Finite_Pow_subset_Pow: "Finite(A) \<Longrightarrow> Pow(A) \<subseteq> DPow(A)"
by (blast intro: Fin_into_DPow Finite_into_Fin Fin_subset)
-lemma Finite_DPow_eq_Pow: "Finite(A) ==> DPow(A) = Pow(A)"
+lemma Finite_DPow_eq_Pow: "Finite(A) \<Longrightarrow> DPow(A) = Pow(A)"
apply (rule equalityI)
apply (rule DPow_subset_Pow)
apply (erule Finite_Pow_subset_Pow)
@@ -515,18 +515,18 @@
definition
subset_fm :: "[i,i]=>i" where
- "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
+ "subset_fm(x,y) \<equiv> Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
-lemma subset_type [TC]: "[| x \<in> nat; y \<in> nat |] ==> subset_fm(x,y) \<in> formula"
+lemma subset_type [TC]: "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> subset_fm(x,y) \<in> formula"
by (simp add: subset_fm_def)
lemma arity_subset_fm [simp]:
- "[| x \<in> nat; y \<in> nat |] ==> arity(subset_fm(x,y)) = succ(x) \<union> succ(y)"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> arity(subset_fm(x,y)) = succ(x) \<union> succ(y)"
by (simp add: subset_fm_def succ_Un_distrib [symmetric])
lemma sats_subset_fm [simp]:
- "[|x < length(env); y \<in> nat; env \<in> list(A); Transset(A)|]
- ==> sats(A, subset_fm(x,y), env) \<longleftrightarrow> nth(x,env) \<subseteq> nth(y,env)"
+ "\<lbrakk>x < length(env); y \<in> nat; env \<in> list(A); Transset(A)\<rbrakk>
+ \<Longrightarrow> sats(A, subset_fm(x,y), env) \<longleftrightarrow> nth(x,env) \<subseteq> nth(y,env)"
apply (frule lt_length_in_nat, assumption)
apply (simp add: subset_fm_def Transset_def)
apply (blast intro: nth_type)
@@ -536,18 +536,18 @@
definition
transset_fm :: "i=>i" where
- "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
+ "transset_fm(x) \<equiv> Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
-lemma transset_type [TC]: "x \<in> nat ==> transset_fm(x) \<in> formula"
+lemma transset_type [TC]: "x \<in> nat \<Longrightarrow> transset_fm(x) \<in> formula"
by (simp add: transset_fm_def)
lemma arity_transset_fm [simp]:
- "x \<in> nat ==> arity(transset_fm(x)) = succ(x)"
+ "x \<in> nat \<Longrightarrow> arity(transset_fm(x)) = succ(x)"
by (simp add: transset_fm_def succ_Un_distrib [symmetric])
lemma sats_transset_fm [simp]:
- "[|x < length(env); env \<in> list(A); Transset(A)|]
- ==> sats(A, transset_fm(x), env) \<longleftrightarrow> Transset(nth(x,env))"
+ "\<lbrakk>x < length(env); env \<in> list(A); Transset(A)\<rbrakk>
+ \<Longrightarrow> sats(A, transset_fm(x), env) \<longleftrightarrow> Transset(nth(x,env))"
apply (frule lt_nat_in_nat, erule length_type)
apply (simp add: transset_fm_def Transset_def)
apply (blast intro: nth_type)
@@ -557,19 +557,19 @@
definition
ordinal_fm :: "i=>i" where
- "ordinal_fm(x) ==
+ "ordinal_fm(x) \<equiv>
And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
-lemma ordinal_type [TC]: "x \<in> nat ==> ordinal_fm(x) \<in> formula"
+lemma ordinal_type [TC]: "x \<in> nat \<Longrightarrow> ordinal_fm(x) \<in> formula"
by (simp add: ordinal_fm_def)
lemma arity_ordinal_fm [simp]:
- "x \<in> nat ==> arity(ordinal_fm(x)) = succ(x)"
+ "x \<in> nat \<Longrightarrow> arity(ordinal_fm(x)) = succ(x)"
by (simp add: ordinal_fm_def succ_Un_distrib [symmetric])
lemma sats_ordinal_fm:
- "[|x < length(env); env \<in> list(A); Transset(A)|]
- ==> sats(A, ordinal_fm(x), env) \<longleftrightarrow> Ord(nth(x,env))"
+ "\<lbrakk>x < length(env); env \<in> list(A); Transset(A)\<rbrakk>
+ \<Longrightarrow> sats(A, ordinal_fm(x), env) \<longleftrightarrow> Ord(nth(x,env))"
apply (frule lt_nat_in_nat, erule length_type)
apply (simp add: ordinal_fm_def Ord_def Transset_def)
apply (blast intro: nth_type)
@@ -577,7 +577,7 @@
text\<open>The subset consisting of the ordinals is definable. Essential lemma for
\<open>Ord_in_Lset\<close>. This result is the objective of the present subsection.\<close>
-theorem Ords_in_DPow: "Transset(A) ==> {x \<in> A. Ord(x)} \<in> DPow(A)"
+theorem Ords_in_DPow: "Transset(A) \<Longrightarrow> {x \<in> A. Ord(x)} \<in> DPow(A)"
apply (simp add: DPow_def Collect_subset)
apply (rule_tac x=Nil in bexI)
apply (rule_tac x="ordinal_fm(0)" in bexI)
@@ -589,40 +589,40 @@
definition
Lset :: "i=>i" where
- "Lset(i) == transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
+ "Lset(i) \<equiv> transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
definition
L :: "i=>o" where \<comment> \<open>Kunen's definition VI 1.5, page 167\<close>
- "L(x) == \<exists>i. Ord(i) & x \<in> Lset(i)"
+ "L(x) \<equiv> \<exists>i. Ord(i) & x \<in> Lset(i)"
text\<open>NOT SUITABLE FOR REWRITING -- RECURSIVE!\<close>
lemma Lset: "Lset(i) = (\<Union>j\<in>i. DPow(Lset(j)))"
by (subst Lset_def [THEN def_transrec], simp)
-lemma LsetI: "[|y\<in>x; A \<in> DPow(Lset(y))|] ==> A \<in> Lset(x)"
+lemma LsetI: "\<lbrakk>y\<in>x; A \<in> DPow(Lset(y))\<rbrakk> \<Longrightarrow> A \<in> Lset(x)"
by (subst Lset, blast)
-lemma LsetD: "A \<in> Lset(x) ==> \<exists>y\<in>x. A \<in> DPow(Lset(y))"
+lemma LsetD: "A \<in> Lset(x) \<Longrightarrow> \<exists>y\<in>x. A \<in> DPow(Lset(y))"
apply (insert Lset [of x])
apply (blast intro: elim: equalityE)
done
subsubsection\<open>Transitivity\<close>
-lemma elem_subset_in_DPow: "[|X \<in> A; X \<subseteq> A|] ==> X \<in> DPow(A)"
+lemma elem_subset_in_DPow: "\<lbrakk>X \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> X \<in> DPow(A)"
apply (simp add: Transset_def DPow_def)
apply (rule_tac x="[X]" in bexI)
apply (rule_tac x="Member(0,1)" in bexI)
apply (auto simp add: Un_least_lt_iff)
done
-lemma Transset_subset_DPow: "Transset(A) ==> A \<subseteq> DPow(A)"
+lemma Transset_subset_DPow: "Transset(A) \<Longrightarrow> A \<subseteq> DPow(A)"
apply clarify
apply (simp add: Transset_def)
apply (blast intro: elem_subset_in_DPow)
done
-lemma Transset_DPow: "Transset(A) ==> Transset(DPow(A))"
+lemma Transset_DPow: "Transset(A) \<Longrightarrow> Transset(DPow(A))"
apply (simp add: Transset_def)
apply (blast intro: elem_subset_in_DPow dest: DPowD)
done
@@ -634,7 +634,7 @@
apply (blast intro!: Transset_Union_family Transset_Un Transset_DPow)
done
-lemma mem_Lset_imp_subset_Lset: "a \<in> Lset(i) ==> a \<subseteq> Lset(i)"
+lemma mem_Lset_imp_subset_Lset: "a \<in> Lset(i) \<Longrightarrow> a \<subseteq> Lset(i)"
apply (insert Transset_Lset)
apply (simp add: Transset_def)
done
@@ -666,7 +666,7 @@
text\<open>Useful with Reflection to bump up the ordinal\<close>
-lemma subset_Lset_ltD: "[|A \<subseteq> Lset(i); i < j|] ==> A \<subseteq> Lset(j)"
+lemma subset_Lset_ltD: "\<lbrakk>A \<subseteq> Lset(i); i < j\<rbrakk> \<Longrightarrow> A \<subseteq> Lset(j)"
by (blast dest: ltD [THEN Lset_mono_mem])
subsubsection\<open>0, successor and limit equations for Lset\<close>
@@ -707,16 +707,16 @@
subsubsection\<open>Lset applied to Limit ordinals\<close>
lemma Limit_Lset_eq:
- "Limit(i) ==> Lset(i) = (\<Union>y\<in>i. Lset(y))"
+ "Limit(i) \<Longrightarrow> Lset(i) = (\<Union>y\<in>i. Lset(y))"
by (simp add: Lset_Union [symmetric] Limit_Union_eq)
-lemma lt_LsetI: "[| a \<in> Lset(j); j<i |] ==> a \<in> Lset(i)"
+lemma lt_LsetI: "\<lbrakk>a \<in> Lset(j); j<i\<rbrakk> \<Longrightarrow> a \<in> Lset(i)"
by (blast dest: Lset_mono [OF le_imp_subset [OF leI]])
lemma Limit_LsetE:
- "[| a \<in> Lset(i); ~R ==> Limit(i);
- !!x. [| x<i; a \<in> Lset(x) |] ==> R
- |] ==> R"
+ "\<lbrakk>a \<in> Lset(i); \<not>R \<Longrightarrow> Limit(i);
+ \<And>x. \<lbrakk>x<i; a \<in> Lset(x)\<rbrakk> \<Longrightarrow> R
+\<rbrakk> \<Longrightarrow> R"
apply (rule classical)
apply (rule Limit_Lset_eq [THEN equalityD1, THEN subsetD, THEN UN_E])
prefer 2 apply assumption
@@ -726,7 +726,7 @@
subsubsection\<open>Basic closure properties\<close>
-lemma zero_in_Lset: "y \<in> x ==> 0 \<in> Lset(x)"
+lemma zero_in_Lset: "y \<in> x \<Longrightarrow> 0 \<in> Lset(x)"
by (subst Lset, blast intro: empty_in_DPow)
lemma notin_Lset: "x \<notin> Lset(x)"
@@ -738,7 +738,7 @@
subsection\<open>Constructible Ordinals: Kunen's VI 1.9 (b)\<close>
-lemma Ords_of_Lset_eq: "Ord(i) ==> {x\<in>Lset(i). Ord(x)} = i"
+lemma Ords_of_Lset_eq: "Ord(i) \<Longrightarrow> {x\<in>Lset(i). Ord(x)} = i"
apply (erule trans_induct3)
apply (simp_all add: Lset_succ Limit_Lset_eq Limit_Union_eq)
txt\<open>The successor case remains.\<close>
@@ -758,22 +758,22 @@
done
-lemma Ord_subset_Lset: "Ord(i) ==> i \<subseteq> Lset(i)"
+lemma Ord_subset_Lset: "Ord(i) \<Longrightarrow> i \<subseteq> Lset(i)"
by (subst Ords_of_Lset_eq [symmetric], assumption, fast)
-lemma Ord_in_Lset: "Ord(i) ==> i \<in> Lset(succ(i))"
+lemma Ord_in_Lset: "Ord(i) \<Longrightarrow> i \<in> Lset(succ(i))"
apply (simp add: Lset_succ)
apply (subst Ords_of_Lset_eq [symmetric], assumption,
rule Ords_in_DPow [OF Transset_Lset])
done
-lemma Ord_in_L: "Ord(i) ==> L(i)"
+lemma Ord_in_L: "Ord(i) \<Longrightarrow> L(i)"
by (simp add: L_def, blast intro: Ord_in_Lset)
subsubsection\<open>Unions\<close>
lemma Union_in_Lset:
- "X \<in> Lset(i) ==> \<Union>(X) \<in> Lset(succ(i))"
+ "X \<in> Lset(i) \<Longrightarrow> \<Union>(X) \<in> Lset(succ(i))"
apply (insert Transset_Lset)
apply (rule LsetI [OF succI1])
apply (simp add: Transset_def DPow_def)
@@ -785,20 +785,20 @@
apply (simp add: succ_Un_distrib [symmetric], blast)
done
-theorem Union_in_L: "L(X) ==> L(\<Union>(X))"
+theorem Union_in_L: "L(X) \<Longrightarrow> L(\<Union>(X))"
by (simp add: L_def, blast dest: Union_in_Lset)
subsubsection\<open>Finite sets and ordered pairs\<close>
-lemma singleton_in_Lset: "a \<in> Lset(i) ==> {a} \<in> Lset(succ(i))"
+lemma singleton_in_Lset: "a \<in> Lset(i) \<Longrightarrow> {a} \<in> Lset(succ(i))"
by (simp add: Lset_succ singleton_in_DPow)
lemma doubleton_in_Lset:
- "[| a \<in> Lset(i); b \<in> Lset(i) |] ==> {a,b} \<in> Lset(succ(i))"
+ "\<lbrakk>a \<in> Lset(i); b \<in> Lset(i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Lset(succ(i))"
by (simp add: Lset_succ empty_in_DPow cons_in_DPow)
lemma Pair_in_Lset:
- "[| a \<in> Lset(i); b \<in> Lset(i); Ord(i) |] ==> <a,b> \<in> Lset(succ(succ(i)))"
+ "\<lbrakk>a \<in> Lset(i); b \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Lset(succ(succ(i)))"
apply (unfold Pair_def)
apply (blast intro: doubleton_in_Lset)
done
@@ -808,21 +808,21 @@
text\<open>Hard work is finding a single \<^term>\<open>j \<in> i\<close> such that \<^term>\<open>{a,b} \<subseteq> Lset(j)\<close>\<close>
lemma doubleton_in_LLimit:
- "[| a \<in> Lset(i); b \<in> Lset(i); Limit(i) |] ==> {a,b} \<in> Lset(i)"
+ "\<lbrakk>a \<in> Lset(i); b \<in> Lset(i); Limit(i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Lset(i)"
apply (erule Limit_LsetE, assumption)
apply (erule Limit_LsetE, assumption)
apply (blast intro: lt_LsetI [OF doubleton_in_Lset]
Lset_UnI1 Lset_UnI2 Limit_has_succ Un_least_lt)
done
-theorem doubleton_in_L: "[| L(a); L(b) |] ==> L({a, b})"
+theorem doubleton_in_L: "\<lbrakk>L(a); L(b)\<rbrakk> \<Longrightarrow> L({a, b})"
apply (simp add: L_def, clarify)
apply (drule Ord2_imp_greater_Limit, assumption)
apply (blast intro: lt_LsetI doubleton_in_LLimit Limit_is_Ord)
done
lemma Pair_in_LLimit:
- "[| a \<in> Lset(i); b \<in> Lset(i); Limit(i) |] ==> <a,b> \<in> Lset(i)"
+ "\<lbrakk>a \<in> Lset(i); b \<in> Lset(i); Limit(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Lset(i)"
txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
apply (erule Limit_LsetE, assumption)
apply (erule Limit_LsetE, assumption)
@@ -836,18 +836,18 @@
text\<open>The rank function for the constructible universe\<close>
definition
lrank :: "i=>i" where \<comment> \<open>Kunen's definition VI 1.7\<close>
- "lrank(x) == \<mu> i. x \<in> Lset(succ(i))"
+ "lrank(x) \<equiv> \<mu> i. x \<in> Lset(succ(i))"
-lemma L_I: "[|x \<in> Lset(i); Ord(i)|] ==> L(x)"
+lemma L_I: "\<lbrakk>x \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> L(x)"
by (simp add: L_def, blast)
-lemma L_D: "L(x) ==> \<exists>i. Ord(i) & x \<in> Lset(i)"
+lemma L_D: "L(x) \<Longrightarrow> \<exists>i. Ord(i) & x \<in> Lset(i)"
by (simp add: L_def)
lemma Ord_lrank [simp]: "Ord(lrank(a))"
by (simp add: lrank_def)
-lemma Lset_lrank_lt [rule_format]: "Ord(i) ==> x \<in> Lset(i) \<longrightarrow> lrank(x) < i"
+lemma Lset_lrank_lt [rule_format]: "Ord(i) \<Longrightarrow> x \<in> Lset(i) \<longrightarrow> lrank(x) < i"
apply (erule trans_induct3)
apply simp
apply (simp only: lrank_def)
@@ -859,7 +859,7 @@
text\<open>Kunen's VI 1.8. The proof is much harder than the text would
suggest. For a start, it needs the previous lemma, which is proved by
induction.\<close>
-lemma Lset_iff_lrank_lt: "Ord(i) ==> x \<in> Lset(i) \<longleftrightarrow> L(x) & lrank(x) < i"
+lemma Lset_iff_lrank_lt: "Ord(i) \<Longrightarrow> x \<in> Lset(i) \<longleftrightarrow> L(x) & lrank(x) < i"
apply (simp add: L_def, auto)
apply (blast intro: Lset_lrank_lt)
apply (unfold lrank_def)
@@ -872,7 +872,7 @@
by (simp add: Lset_iff_lrank_lt)
text\<open>Kunen's VI 1.9 (a)\<close>
-lemma lrank_of_Ord: "Ord(i) ==> lrank(i) = i"
+lemma lrank_of_Ord: "Ord(i) \<Longrightarrow> lrank(i) = i"
apply (unfold lrank_def)
apply (rule Least_equality)
apply (erule Ord_in_Lset)
@@ -893,7 +893,7 @@
apply auto
done
-lemma lrank_Lset: "Ord(i) ==> lrank(Lset(i)) = i"
+lemma lrank_Lset: "Ord(i) \<Longrightarrow> lrank(Lset(i)) = i"
apply (unfold lrank_def)
apply (rule Least_equality)
apply (rule Lset_in_Lset_succ)
@@ -905,7 +905,7 @@
done
text\<open>Kunen's VI 1.11\<close>
-lemma Lset_subset_Vset: "Ord(i) ==> Lset(i) \<subseteq> Vset(i)"
+lemma Lset_subset_Vset: "Ord(i) \<Longrightarrow> Lset(i) \<subseteq> Vset(i)"
apply (erule trans_induct)
apply (subst Lset)
apply (subst Vset)
@@ -915,7 +915,7 @@
done
text\<open>Kunen's VI 1.12\<close>
-lemma Lset_subset_Vset': "i \<in> nat ==> Lset(i) = Vset(i)"
+lemma Lset_subset_Vset': "i \<in> nat \<Longrightarrow> Lset(i) = Vset(i)"
apply (erule nat_induct)
apply (simp add: Vfrom_0)
apply (simp add: Lset_succ Vset_succ Finite_Vset Finite_DPow_eq_Pow)
@@ -923,24 +923,24 @@
text\<open>Every set of constructible sets is included in some \<^term>\<open>Lset\<close>\<close>
lemma subset_Lset:
- "(\<forall>x\<in>A. L(x)) ==> \<exists>i. Ord(i) & A \<subseteq> Lset(i)"
+ "(\<forall>x\<in>A. L(x)) \<Longrightarrow> \<exists>i. Ord(i) & A \<subseteq> Lset(i)"
by (rule_tac x = "\<Union>x\<in>A. succ(lrank(x))" in exI, force)
lemma subset_LsetE:
- "[|\<forall>x\<in>A. L(x);
- !!i. [|Ord(i); A \<subseteq> Lset(i)|] ==> P|]
- ==> P"
+ "\<lbrakk>\<forall>x\<in>A. L(x);
+ \<And>i. \<lbrakk>Ord(i); A \<subseteq> Lset(i)\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by (blast dest: subset_Lset)
subsubsection\<open>For L to satisfy the Powerset axiom\<close>
lemma LPow_env_typing:
- "[| y \<in> Lset(i); Ord(i); y \<subseteq> X |]
- ==> \<exists>z \<in> Pow(X). y \<in> Lset(succ(lrank(z)))"
+ "\<lbrakk>y \<in> Lset(i); Ord(i); y \<subseteq> X\<rbrakk>
+ \<Longrightarrow> \<exists>z \<in> Pow(X). y \<in> Lset(succ(lrank(z)))"
by (auto intro: L_I iff: Lset_succ_lrank_iff)
lemma LPow_in_Lset:
- "[|X \<in> Lset(i); Ord(i)|] ==> \<exists>j. Ord(j) & {y \<in> Pow(X). L(y)} \<in> Lset(j)"
+ "\<lbrakk>X \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. Ord(j) & {y \<in> Pow(X). L(y)} \<in> Lset(j)"
apply (rule_tac x="succ(\<Union>y \<in> Pow(X). succ(lrank(y)))" in exI)
apply simp
apply (rule LsetI [OF succI1])
@@ -958,34 +958,34 @@
apply (auto intro: L_I iff: Lset_succ_lrank_iff)
done
-theorem LPow_in_L: "L(X) ==> L({y \<in> Pow(X). L(y)})"
+theorem LPow_in_L: "L(X) \<Longrightarrow> L({y \<in> Pow(X). L(y)})"
by (blast intro: L_I dest: L_D LPow_in_Lset)
subsection\<open>Eliminating \<^term>\<open>arity\<close> from the Definition of \<^term>\<open>Lset\<close>\<close>
-lemma nth_zero_eq_0: "n \<in> nat ==> nth(n,[0]) = 0"
+lemma nth_zero_eq_0: "n \<in> nat \<Longrightarrow> nth(n,[0]) = 0"
by (induct_tac n, auto)
lemma sats_app_0_iff [rule_format]:
- "[| p \<in> formula; 0 \<in> A |]
- ==> \<forall>env \<in> list(A). sats(A,p, env@[0]) \<longleftrightarrow> sats(A,p,env)"
+ "\<lbrakk>p \<in> formula; 0 \<in> A\<rbrakk>
+ \<Longrightarrow> \<forall>env \<in> list(A). sats(A,p, env@[0]) \<longleftrightarrow> sats(A,p,env)"
apply (induct_tac p)
apply (simp_all del: app_Cons add: app_Cons [symmetric]
add: nth_zero_eq_0 nth_append not_lt_iff_le nth_eq_0)
done
lemma sats_app_zeroes_iff:
- "[| p \<in> formula; 0 \<in> A; env \<in> list(A); n \<in> nat |]
- ==> sats(A,p,env @ repeat(0,n)) \<longleftrightarrow> sats(A,p,env)"
+ "\<lbrakk>p \<in> formula; 0 \<in> A; env \<in> list(A); n \<in> nat\<rbrakk>
+ \<Longrightarrow> sats(A,p,env @ repeat(0,n)) \<longleftrightarrow> sats(A,p,env)"
apply (induct_tac n, simp)
apply (simp del: repeat.simps
add: repeat_succ_app sats_app_0_iff app_assoc [symmetric])
done
lemma exists_bigger_env:
- "[| p \<in> formula; 0 \<in> A; env \<in> list(A) |]
- ==> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) &
+ "\<lbrakk>p \<in> formula; 0 \<in> A; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> \<exists>env' \<in> list(A). arity(p) \<le> succ(length(env')) &
(\<forall>a\<in>A. sats(A,p,Cons(a,env')) \<longleftrightarrow> sats(A,p,Cons(a,env)))"
apply (rule_tac x="env @ repeat(0,arity(p))" in bexI)
apply (simp del: app_Cons add: app_Cons [symmetric]
@@ -996,7 +996,7 @@
text\<open>A simpler version of \<^term>\<open>DPow\<close>: no arity check!\<close>
definition
DPow' :: "i => i" where
- "DPow'(A) == {X \<in> Pow(A).
+ "DPow'(A) \<equiv> {X \<in> Pow(A).
\<exists>env \<in> list(A). \<exists>p \<in> formula.
X = {x\<in>A. sats(A, p, Cons(x,env))}}"
@@ -1006,12 +1006,12 @@
lemma DPow'_0: "DPow'(0) = {0}"
by (auto simp add: DPow'_def)
-lemma DPow'_subset_DPow: "0 \<in> A ==> DPow'(A) \<subseteq> DPow(A)"
+lemma DPow'_subset_DPow: "0 \<in> A \<Longrightarrow> DPow'(A) \<subseteq> DPow(A)"
apply (auto simp add: DPow'_def DPow_def)
apply (frule exists_bigger_env, assumption+, force)
done
-lemma DPow_eq_DPow': "Transset(A) ==> DPow(A) = DPow'(A)"
+lemma DPow_eq_DPow': "Transset(A) \<Longrightarrow> DPow(A) = DPow'(A)"
apply (drule Transset_0_disj)
apply (erule disjE)
apply (simp add: DPow'_0 Finite_DPow_eq_Pow)
@@ -1032,9 +1032,9 @@
text\<open>With this rule we can specify \<^term>\<open>p\<close> later and don't worry about
arities at all!\<close>
lemma DPow_LsetI [rule_format]:
- "[|\<forall>x\<in>Lset(i). P(x) \<longleftrightarrow> sats(Lset(i), p, Cons(x,env));
- env \<in> list(Lset(i)); p \<in> formula|]
- ==> {x\<in>Lset(i). P(x)} \<in> DPow(Lset(i))"
+ "\<lbrakk>\<forall>x\<in>Lset(i). P(x) \<longleftrightarrow> sats(Lset(i), p, Cons(x,env));
+ env \<in> list(Lset(i)); p \<in> formula\<rbrakk>
+ \<Longrightarrow> {x\<in>Lset(i). P(x)} \<in> DPow(Lset(i))"
by (simp add: DPow_eq_DPow' [OF Transset_Lset] DPow'_def, blast)
end
--- a/src/ZF/Constructible/Internalize.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Internalize.thy Tue Sep 27 16:51:35 2022 +0100
@@ -8,24 +8,24 @@
subsubsection\<open>The Formula \<^term>\<open>is_Inl\<close>, Internalized\<close>
-(* is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *)
+(* is_Inl(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *)
definition
Inl_fm :: "[i,i]=>i" where
- "Inl_fm(a,z) == Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))"
+ "Inl_fm(a,z) \<equiv> Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))"
lemma Inl_type [TC]:
- "[| x \<in> nat; z \<in> nat |] ==> Inl_fm(x,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Inl_fm(x,z) \<in> formula"
by (simp add: Inl_fm_def)
lemma sats_Inl_fm [simp]:
- "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Inl_fm(x,z), env) \<longleftrightarrow> is_Inl(##A, nth(x,env), nth(z,env))"
+ "\<lbrakk>x \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Inl_fm(x,z), env) \<longleftrightarrow> is_Inl(##A, nth(x,env), nth(z,env))"
by (simp add: Inl_fm_def is_Inl_def)
lemma Inl_iff_sats:
- "[| nth(i,env) = x; nth(k,env) = z;
- i \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_Inl(##A, x, z) \<longleftrightarrow> sats(A, Inl_fm(i,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(k,env) = z;
+ i \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Inl(##A, x, z) \<longleftrightarrow> sats(A, Inl_fm(i,k), env)"
by simp
theorem Inl_reflection:
@@ -38,24 +38,24 @@
subsubsection\<open>The Formula \<^term>\<open>is_Inr\<close>, Internalized\<close>
-(* is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *)
+(* is_Inr(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *)
definition
Inr_fm :: "[i,i]=>i" where
- "Inr_fm(a,z) == Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))"
+ "Inr_fm(a,z) \<equiv> Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))"
lemma Inr_type [TC]:
- "[| x \<in> nat; z \<in> nat |] ==> Inr_fm(x,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Inr_fm(x,z) \<in> formula"
by (simp add: Inr_fm_def)
lemma sats_Inr_fm [simp]:
- "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Inr_fm(x,z), env) \<longleftrightarrow> is_Inr(##A, nth(x,env), nth(z,env))"
+ "\<lbrakk>x \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Inr_fm(x,z), env) \<longleftrightarrow> is_Inr(##A, nth(x,env), nth(z,env))"
by (simp add: Inr_fm_def is_Inr_def)
lemma Inr_iff_sats:
- "[| nth(i,env) = x; nth(k,env) = z;
- i \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_Inr(##A, x, z) \<longleftrightarrow> sats(A, Inr_fm(i,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(k,env) = z;
+ i \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Inr(##A, x, z) \<longleftrightarrow> sats(A, Inr_fm(i,k), env)"
by simp
theorem Inr_reflection:
@@ -68,23 +68,23 @@
subsubsection\<open>The Formula \<^term>\<open>is_Nil\<close>, Internalized\<close>
-(* is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *)
+(* is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *)
definition
Nil_fm :: "i=>i" where
- "Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))"
+ "Nil_fm(x) \<equiv> Exists(And(empty_fm(0), Inl_fm(0,succ(x))))"
-lemma Nil_type [TC]: "x \<in> nat ==> Nil_fm(x) \<in> formula"
+lemma Nil_type [TC]: "x \<in> nat \<Longrightarrow> Nil_fm(x) \<in> formula"
by (simp add: Nil_fm_def)
lemma sats_Nil_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, Nil_fm(x), env) \<longleftrightarrow> is_Nil(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Nil_fm(x), env) \<longleftrightarrow> is_Nil(##A, nth(x,env))"
by (simp add: Nil_fm_def is_Nil_def)
lemma Nil_iff_sats:
- "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
- ==> is_Nil(##A, x) \<longleftrightarrow> sats(A, Nil_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Nil(##A, x) \<longleftrightarrow> sats(A, Nil_fm(i), env)"
by simp
theorem Nil_reflection:
@@ -98,26 +98,26 @@
subsubsection\<open>The Formula \<^term>\<open>is_Cons\<close>, Internalized\<close>
-(* "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
+(* "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
definition
Cons_fm :: "[i,i,i]=>i" where
- "Cons_fm(a,l,Z) ==
+ "Cons_fm(a,l,Z) \<equiv>
Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))"
lemma Cons_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Cons_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Cons_fm(x,y,z) \<in> formula"
by (simp add: Cons_fm_def)
lemma sats_Cons_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Cons_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Cons_fm(x,y,z), env) \<longleftrightarrow>
is_Cons(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: Cons_fm_def is_Cons_def)
lemma Cons_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==>is_Cons(##A, x, y, z) \<longleftrightarrow> sats(A, Cons_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow>is_Cons(##A, x, y, z) \<longleftrightarrow> sats(A, Cons_fm(i,j,k), env)"
by simp
theorem Cons_reflection:
@@ -129,24 +129,24 @@
subsubsection\<open>The Formula \<^term>\<open>is_quasilist\<close>, Internalized\<close>
-(* is_quasilist(M,xs) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
+(* is_quasilist(M,xs) \<equiv> is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
definition
quasilist_fm :: "i=>i" where
- "quasilist_fm(x) ==
+ "quasilist_fm(x) \<equiv>
Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))"
-lemma quasilist_type [TC]: "x \<in> nat ==> quasilist_fm(x) \<in> formula"
+lemma quasilist_type [TC]: "x \<in> nat \<Longrightarrow> quasilist_fm(x) \<in> formula"
by (simp add: quasilist_fm_def)
lemma sats_quasilist_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, quasilist_fm(x), env) \<longleftrightarrow> is_quasilist(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, quasilist_fm(x), env) \<longleftrightarrow> is_quasilist(##A, nth(x,env))"
by (simp add: quasilist_fm_def is_quasilist_def)
lemma quasilist_iff_sats:
- "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
- ==> is_quasilist(##A, x) \<longleftrightarrow> sats(A, quasilist_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_quasilist(##A, x) \<longleftrightarrow> sats(A, quasilist_fm(i), env)"
by simp
theorem quasilist_reflection:
@@ -162,30 +162,30 @@
subsubsection\<open>The Formula \<^term>\<open>is_hd\<close>, Internalized\<close>
-(* "is_hd(M,xs,H) ==
+(* "is_hd(M,xs,H) \<equiv>
(is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
- (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
+ (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | H=x) &
(is_quasilist(M,xs) | empty(M,H))" *)
definition
hd_fm :: "[i,i]=>i" where
- "hd_fm(xs,H) ==
+ "hd_fm(xs,H) \<equiv>
And(Implies(Nil_fm(xs), empty_fm(H)),
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(H#+2,1)))),
Or(quasilist_fm(xs), empty_fm(H))))"
lemma hd_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> hd_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> hd_fm(x,y) \<in> formula"
by (simp add: hd_fm_def)
lemma sats_hd_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, hd_fm(x,y), env) \<longleftrightarrow> is_hd(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, hd_fm(x,y), env) \<longleftrightarrow> is_hd(##A, nth(x,env), nth(y,env))"
by (simp add: hd_fm_def is_hd_def)
lemma hd_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_hd(##A, x, y) \<longleftrightarrow> sats(A, hd_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_hd(##A, x, y) \<longleftrightarrow> sats(A, hd_fm(i,j), env)"
by simp
theorem hd_reflection:
@@ -199,30 +199,30 @@
subsubsection\<open>The Formula \<^term>\<open>is_tl\<close>, Internalized\<close>
-(* "is_tl(M,xs,T) ==
+(* "is_tl(M,xs,T) \<equiv>
(is_Nil(M,xs) \<longrightarrow> T=xs) &
- (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
+ (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | T=l) &
(is_quasilist(M,xs) | empty(M,T))" *)
definition
tl_fm :: "[i,i]=>i" where
- "tl_fm(xs,T) ==
+ "tl_fm(xs,T) \<equiv>
And(Implies(Nil_fm(xs), Equal(T,xs)),
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))),
Or(quasilist_fm(xs), empty_fm(T))))"
lemma tl_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> tl_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> tl_fm(x,y) \<in> formula"
by (simp add: tl_fm_def)
lemma sats_tl_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, tl_fm(x,y), env) \<longleftrightarrow> is_tl(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, tl_fm(x,y), env) \<longleftrightarrow> is_tl(##A, nth(x,env), nth(y,env))"
by (simp add: tl_fm_def is_tl_def)
lemma tl_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_tl(##A, x, y) \<longleftrightarrow> sats(A, tl_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_tl(##A, x, y) \<longleftrightarrow> sats(A, tl_fm(i,j), env)"
by simp
theorem tl_reflection:
@@ -237,34 +237,34 @@
subsubsection\<open>The Operator \<^term>\<open>is_bool_of_o\<close>\<close>
(* is_bool_of_o :: "[i=>o, o, i] => o"
- "is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))" *)
+ "is_bool_of_o(M,P,z) \<equiv> (P & number1(M,z)) | (\<not>P & empty(M,z))" *)
text\<open>The formula \<^term>\<open>p\<close> has no free variables.\<close>
definition
bool_of_o_fm :: "[i, i]=>i" where
- "bool_of_o_fm(p,z) ==
+ "bool_of_o_fm(p,z) \<equiv>
Or(And(p,number1_fm(z)),
And(Neg(p),empty_fm(z)))"
lemma is_bool_of_o_type [TC]:
- "[| p \<in> formula; z \<in> nat |] ==> bool_of_o_fm(p,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; z \<in> nat\<rbrakk> \<Longrightarrow> bool_of_o_fm(p,z) \<in> formula"
by (simp add: bool_of_o_fm_def)
lemma sats_bool_of_o_fm:
assumes p_iff_sats: "P \<longleftrightarrow> sats(A, p, env)"
shows
- "[|z \<in> nat; env \<in> list(A)|]
- ==> sats(A, bool_of_o_fm(p,z), env) \<longleftrightarrow>
+ "\<lbrakk>z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, bool_of_o_fm(p,z), env) \<longleftrightarrow>
is_bool_of_o(##A, P, nth(z,env))"
by (simp add: bool_of_o_fm_def is_bool_of_o_def p_iff_sats [THEN iff_sym])
lemma is_bool_of_o_iff_sats:
- "[| P \<longleftrightarrow> sats(A, p, env); nth(k,env) = z; k \<in> nat; env \<in> list(A)|]
- ==> is_bool_of_o(##A, P, z) \<longleftrightarrow> sats(A, bool_of_o_fm(p,k), env)"
+ "\<lbrakk>P \<longleftrightarrow> sats(A, p, env); nth(k,env) = z; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_bool_of_o(##A, P, z) \<longleftrightarrow> sats(A, bool_of_o_fm(p,k), env)"
by (simp add: sats_bool_of_o_fm)
theorem bool_of_o_reflection:
- "REFLECTS [P(L), \<lambda>i. P(##Lset(i))] ==>
+ "REFLECTS [P(L), \<lambda>i. P(##Lset(i))] \<Longrightarrow>
REFLECTS[\<lambda>x. is_bool_of_o(L, P(L,x), f(x)),
\<lambda>i x. is_bool_of_o(##Lset(i), P(##Lset(i),x), f(x))]"
apply (simp (no_asm) only: is_bool_of_o_def)
@@ -280,12 +280,12 @@
\<^term>\<open>p\<close> will be enclosed by three quantifiers.\<close>
(* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o"
- "is_lambda(M, A, is_b, z) ==
+ "is_lambda(M, A, is_b, z) \<equiv>
\<forall>p[M]. p \<in> z \<longleftrightarrow>
(\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))" *)
definition
lambda_fm :: "[i, i, i]=>i" where
- "lambda_fm(p,A,z) ==
+ "lambda_fm(p,A,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(Exists(And(Member(1,A#+3),
And(pair_fm(1,0,2), p))))))"
@@ -294,24 +294,24 @@
the corresponding quantified variables with de Bruijn indices 1, 0.\<close>
lemma is_lambda_type [TC]:
- "[| p \<in> formula; x \<in> nat; y \<in> nat |]
- ==> lambda_fm(p,x,y) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> lambda_fm(p,x,y) \<in> formula"
by (simp add: lambda_fm_def)
lemma sats_lambda_fm:
assumes is_b_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> is_b(a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> is_b(a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))"
shows
- "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, lambda_fm(p,x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, lambda_fm(p,x,y), env) \<longleftrightarrow>
is_lambda(##A, nth(x,env), is_b, nth(y,env))"
by (simp add: lambda_fm_def is_lambda_def is_b_iff_sats [THEN iff_sym])
theorem is_lambda_reflection:
assumes is_b_reflection:
- "!!f g h. REFLECTS[\<lambda>x. is_b(L, f(x), g(x), h(x)),
+ "\<And>f g h. REFLECTS[\<lambda>x. is_b(L, f(x), g(x), h(x)),
\<lambda>i x. is_b(##Lset(i), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. is_lambda(L, A(x), is_b(L,x), f(x)),
\<lambda>i x. is_lambda(##Lset(i), A(x), is_b(##Lset(i),x), f(x))]"
@@ -321,28 +321,28 @@
subsubsection\<open>The Operator \<^term>\<open>is_Member\<close>, Internalized\<close>
-(* "is_Member(M,x,y,Z) ==
+(* "is_Member(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
definition
Member_fm :: "[i,i,i]=>i" where
- "Member_fm(x,y,Z) ==
+ "Member_fm(x,y,Z) \<equiv>
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inl_fm(1,0), Inl_fm(0,Z#+2)))))"
lemma is_Member_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Member_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Member_fm(x,y,z) \<in> formula"
by (simp add: Member_fm_def)
lemma sats_Member_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Member_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Member_fm(x,y,z), env) \<longleftrightarrow>
is_Member(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: Member_fm_def is_Member_def)
lemma Member_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_Member(##A, x, y, z) \<longleftrightarrow> sats(A, Member_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Member(##A, x, y, z) \<longleftrightarrow> sats(A, Member_fm(i,j,k), env)"
by (simp)
theorem Member_reflection:
@@ -354,28 +354,28 @@
subsubsection\<open>The Operator \<^term>\<open>is_Equal\<close>, Internalized\<close>
-(* "is_Equal(M,x,y,Z) ==
+(* "is_Equal(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
definition
Equal_fm :: "[i,i,i]=>i" where
- "Equal_fm(x,y,Z) ==
+ "Equal_fm(x,y,Z) \<equiv>
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inr_fm(1,0), Inl_fm(0,Z#+2)))))"
lemma is_Equal_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Equal_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Equal_fm(x,y,z) \<in> formula"
by (simp add: Equal_fm_def)
lemma sats_Equal_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Equal_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Equal_fm(x,y,z), env) \<longleftrightarrow>
is_Equal(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: Equal_fm_def is_Equal_def)
lemma Equal_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_Equal(##A, x, y, z) \<longleftrightarrow> sats(A, Equal_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Equal(##A, x, y, z) \<longleftrightarrow> sats(A, Equal_fm(i,j,k), env)"
by (simp)
theorem Equal_reflection:
@@ -387,28 +387,28 @@
subsubsection\<open>The Operator \<^term>\<open>is_Nand\<close>, Internalized\<close>
-(* "is_Nand(M,x,y,Z) ==
+(* "is_Nand(M,x,y,Z) \<equiv>
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
definition
Nand_fm :: "[i,i,i]=>i" where
- "Nand_fm(x,y,Z) ==
+ "Nand_fm(x,y,Z) \<equiv>
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inl_fm(1,0), Inr_fm(0,Z#+2)))))"
lemma is_Nand_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Nand_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> Nand_fm(x,y,z) \<in> formula"
by (simp add: Nand_fm_def)
lemma sats_Nand_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, Nand_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Nand_fm(x,y,z), env) \<longleftrightarrow>
is_Nand(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: Nand_fm_def is_Nand_def)
lemma Nand_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_Nand(##A, x, y, z) \<longleftrightarrow> sats(A, Nand_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Nand(##A, x, y, z) \<longleftrightarrow> sats(A, Nand_fm(i,j,k), env)"
by (simp)
theorem Nand_reflection:
@@ -420,26 +420,26 @@
subsubsection\<open>The Operator \<^term>\<open>is_Forall\<close>, Internalized\<close>
-(* "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *)
+(* "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *)
definition
Forall_fm :: "[i,i]=>i" where
- "Forall_fm(x,Z) ==
+ "Forall_fm(x,Z) \<equiv>
Exists(And(Inr_fm(succ(x),0), Inr_fm(0,succ(Z))))"
lemma is_Forall_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> Forall_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> Forall_fm(x,y) \<in> formula"
by (simp add: Forall_fm_def)
lemma sats_Forall_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, Forall_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Forall_fm(x,y), env) \<longleftrightarrow>
is_Forall(##A, nth(x,env), nth(y,env))"
by (simp add: Forall_fm_def is_Forall_def)
lemma Forall_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_Forall(##A, x, y) \<longleftrightarrow> sats(A, Forall_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_Forall(##A, x, y) \<longleftrightarrow> sats(A, Forall_fm(i,j), env)"
by (simp)
theorem Forall_reflection:
@@ -452,28 +452,28 @@
subsubsection\<open>The Operator \<^term>\<open>is_and\<close>, Internalized\<close>
-(* is_and(M,a,b,z) == (number1(M,a) & z=b) |
- (~number1(M,a) & empty(M,z)) *)
+(* is_and(M,a,b,z) \<equiv> (number1(M,a) & z=b) |
+ (\<not>number1(M,a) & empty(M,z)) *)
definition
and_fm :: "[i,i,i]=>i" where
- "and_fm(a,b,z) ==
+ "and_fm(a,b,z) \<equiv>
Or(And(number1_fm(a), Equal(z,b)),
And(Neg(number1_fm(a)),empty_fm(z)))"
lemma is_and_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> and_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> and_fm(x,y,z) \<in> formula"
by (simp add: and_fm_def)
lemma sats_and_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, and_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, and_fm(x,y,z), env) \<longleftrightarrow>
is_and(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: and_fm_def is_and_def)
lemma is_and_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_and(##A, x, y, z) \<longleftrightarrow> sats(A, and_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_and(##A, x, y, z) \<longleftrightarrow> sats(A, and_fm(i,j,k), env)"
by simp
theorem is_and_reflection:
@@ -486,29 +486,29 @@
subsubsection\<open>The Operator \<^term>\<open>is_or\<close>, Internalized\<close>
-(* is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) |
- (~number1(M,a) & z=b) *)
+(* is_or(M,a,b,z) \<equiv> (number1(M,a) & number1(M,z)) |
+ (\<not>number1(M,a) & z=b) *)
definition
or_fm :: "[i,i,i]=>i" where
- "or_fm(a,b,z) ==
+ "or_fm(a,b,z) \<equiv>
Or(And(number1_fm(a), number1_fm(z)),
And(Neg(number1_fm(a)), Equal(z,b)))"
lemma is_or_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> or_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> or_fm(x,y,z) \<in> formula"
by (simp add: or_fm_def)
lemma sats_or_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, or_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, or_fm(x,y,z), env) \<longleftrightarrow>
is_or(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: or_fm_def is_or_def)
lemma is_or_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_or(##A, x, y, z) \<longleftrightarrow> sats(A, or_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_or(##A, x, y, z) \<longleftrightarrow> sats(A, or_fm(i,j,k), env)"
by simp
theorem is_or_reflection:
@@ -522,27 +522,27 @@
subsubsection\<open>The Operator \<^term>\<open>is_not\<close>, Internalized\<close>
-(* is_not(M,a,z) == (number1(M,a) & empty(M,z)) |
- (~number1(M,a) & number1(M,z)) *)
+(* is_not(M,a,z) \<equiv> (number1(M,a) & empty(M,z)) |
+ (\<not>number1(M,a) & number1(M,z)) *)
definition
not_fm :: "[i,i]=>i" where
- "not_fm(a,z) ==
+ "not_fm(a,z) \<equiv>
Or(And(number1_fm(a), empty_fm(z)),
And(Neg(number1_fm(a)), number1_fm(z)))"
lemma is_not_type [TC]:
- "[| x \<in> nat; z \<in> nat |] ==> not_fm(x,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> not_fm(x,z) \<in> formula"
by (simp add: not_fm_def)
lemma sats_is_not_fm [simp]:
- "[| x \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, not_fm(x,z), env) \<longleftrightarrow> is_not(##A, nth(x,env), nth(z,env))"
+ "\<lbrakk>x \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, not_fm(x,z), env) \<longleftrightarrow> is_not(##A, nth(x,env), nth(z,env))"
by (simp add: not_fm_def is_not_def)
lemma is_not_iff_sats:
- "[| nth(i,env) = x; nth(k,env) = z;
- i \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_not(##A, x, z) \<longleftrightarrow> sats(A, not_fm(i,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(k,env) = z;
+ i \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_not(##A, x, z) \<longleftrightarrow> sats(A, not_fm(i,k), env)"
by simp
theorem is_not_reflection:
@@ -579,7 +579,7 @@
(* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
- "M_is_recfun(M,MH,r,a,f) ==
+ "M_is_recfun(M,MH,r,a,f) \<equiv>
\<forall>z[M]. z \<in> f \<longleftrightarrow>
2 1 0
new def (\<exists>x[M]. \<exists>f_r_sx[M]. \<exists>y[M].
@@ -593,7 +593,7 @@
text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0 and z\<close>
definition
is_recfun_fm :: "[i, i, i, i]=>i" where
- "is_recfun_fm(p,r,a,f) ==
+ "is_recfun_fm(p,r,a,f) \<equiv>
Forall(Iff(Member(0,succ(f)),
Exists(Exists(Exists(
And(p,
@@ -605,31 +605,31 @@
And(restriction_fm(f#+7,0,4), Member(2,r#+7)))))))))))))))"
lemma is_recfun_type [TC]:
- "[| p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> is_recfun_fm(p,x,y,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> is_recfun_fm(p,x,y,z) \<in> formula"
by (simp add: is_recfun_fm_def)
lemma sats_is_recfun_fm:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
+ "\<And>a0 a1 a2 a3.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
shows
- "[|x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, is_recfun_fm(p,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_recfun_fm(p,x,y,z), env) \<longleftrightarrow>
M_is_recfun(##A, MH, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: is_recfun_fm_def M_is_recfun_iff MH_iff_sats [THEN iff_sym])
lemma is_recfun_iff_sats:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
+ "\<And>a0 a1 a2 a3.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))"
shows
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> M_is_recfun(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_recfun_fm(p,i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> M_is_recfun(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_recfun_fm(p,i,j,k), env)"
by (simp add: sats_is_recfun_fm [OF MH_iff_sats])
text\<open>The additional variable in the premise, namely \<^term>\<open>f'\<close>, is essential.
@@ -637,7 +637,7 @@
The same thing occurs in \<open>is_wfrec_reflection\<close>.\<close>
theorem is_recfun_reflection:
assumes MH_reflection:
- "!!f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
+ "\<And>f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
\<lambda>i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. M_is_recfun(L, MH(L,x), f(x), g(x), h(x)),
\<lambda>i x. M_is_recfun(##Lset(i), MH(##Lset(i),x), f(x), g(x), h(x))]"
@@ -652,11 +652,11 @@
\<^term>\<open>p\<close> is enclosed by 5 quantifiers.\<close>
(* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
- "is_wfrec(M,MH,r,a,z) ==
+ "is_wfrec(M,MH,r,a,z) \<equiv>
\<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" *)
definition
is_wfrec_fm :: "[i, i, i, i]=>i" where
- "is_wfrec_fm(p,r,a,z) ==
+ "is_wfrec_fm(p,r,a,z) \<equiv>
Exists(And(is_recfun_fm(p, succ(r), succ(a), 0),
Exists(Exists(Exists(Exists(
And(Equal(2,a#+5), And(Equal(1,4), And(Equal(0,z#+5), p)))))))))"
@@ -668,18 +668,18 @@
environments in both calls to MH have the same length.\<close>
lemma is_wfrec_type [TC]:
- "[| p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> is_wfrec_fm(p,x,y,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> is_wfrec_fm(p,x,y,z) \<in> formula"
by (simp add: is_wfrec_fm_def)
lemma sats_is_wfrec_fm:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
+ "\<And>a0 a1 a2 a3 a4.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
shows
- "[|x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
- ==> sats(A, is_wfrec_fm(p,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_wfrec_fm(p,x,y,z), env) \<longleftrightarrow>
is_wfrec(##A, MH, nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=z in lt_length_in_nat, assumption)
apply (frule lt_length_in_nat, assumption)
@@ -689,18 +689,18 @@
lemma is_wfrec_iff_sats:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
+ "\<And>a0 a1 a2 a3 a4.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow> sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))"
shows
- "[|nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
- ==> is_wfrec(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_wfrec_fm(p,i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j < length(env); k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_wfrec(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_wfrec_fm(p,i,j,k), env)"
by (simp add: sats_is_wfrec_fm [OF MH_iff_sats])
theorem is_wfrec_reflection:
assumes MH_reflection:
- "!!f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
+ "\<And>f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
\<lambda>i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. is_wfrec(L, MH(L,x), f(x), g(x), h(x)),
\<lambda>i x. is_wfrec(##Lset(i), MH(##Lset(i),x), f(x), g(x), h(x))]"
@@ -715,28 +715,28 @@
definition
cartprod_fm :: "[i,i,i]=>i" where
-(* "cartprod(M,A,B,z) ==
+(* "cartprod(M,A,B,z) \<equiv>
\<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
- "cartprod_fm(A,B,z) ==
+ "cartprod_fm(A,B,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(A))),
Exists(And(Member(0,succ(succ(succ(B)))),
pair_fm(1,0,2)))))))"
lemma cartprod_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> cartprod_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> cartprod_fm(x,y,z) \<in> formula"
by (simp add: cartprod_fm_def)
lemma sats_cartprod_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, cartprod_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, cartprod_fm(x,y,z), env) \<longleftrightarrow>
cartprod(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: cartprod_fm_def cartprod_def)
lemma cartprod_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> cartprod(##A, x, y, z) \<longleftrightarrow> sats(A, cartprod_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> cartprod(##A, x, y, z) \<longleftrightarrow> sats(A, cartprod_fm(i,j,k), env)"
by (simp)
theorem cartprod_reflection:
@@ -749,14 +749,14 @@
subsubsection\<open>Binary Sums, Internalized\<close>
-(* "is_sum(M,A,B,Z) ==
+(* "is_sum(M,A,B,Z) \<equiv>
\<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
3 2 1 0
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *)
definition
sum_fm :: "[i,i,i]=>i" where
- "sum_fm(A,B,Z) ==
+ "sum_fm(A,B,Z) \<equiv>
Exists(Exists(Exists(Exists(
And(number1_fm(2),
And(cartprod_fm(2,A#+4,3),
@@ -764,19 +764,19 @@
And(cartprod_fm(1,B#+4,0), union_fm(3,0,Z#+4)))))))))"
lemma sum_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> sum_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> sum_fm(x,y,z) \<in> formula"
by (simp add: sum_fm_def)
lemma sats_sum_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, sum_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, sum_fm(x,y,z), env) \<longleftrightarrow>
is_sum(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: sum_fm_def is_sum_def)
lemma sum_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_sum(##A, x, y, z) \<longleftrightarrow> sats(A, sum_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_sum(##A, x, y, z) \<longleftrightarrow> sats(A, sum_fm(i,j,k), env)"
by simp
theorem sum_reflection:
@@ -789,24 +789,24 @@
subsubsection\<open>The Operator \<^term>\<open>quasinat\<close>\<close>
-(* "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *)
+(* "is_quasinat(M,z) \<equiv> empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *)
definition
quasinat_fm :: "i=>i" where
- "quasinat_fm(z) == Or(empty_fm(z), Exists(succ_fm(0,succ(z))))"
+ "quasinat_fm(z) \<equiv> Or(empty_fm(z), Exists(succ_fm(0,succ(z))))"
lemma quasinat_type [TC]:
- "x \<in> nat ==> quasinat_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> quasinat_fm(x) \<in> formula"
by (simp add: quasinat_fm_def)
lemma sats_quasinat_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, quasinat_fm(x), env) \<longleftrightarrow> is_quasinat(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, quasinat_fm(x), env) \<longleftrightarrow> is_quasinat(##A, nth(x,env))"
by (simp add: quasinat_fm_def is_quasinat_def)
lemma quasinat_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> is_quasinat(##A, x) \<longleftrightarrow> sats(A, quasinat_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_quasinat(##A, x) \<longleftrightarrow> sats(A, quasinat_fm(i), env)"
by simp
theorem quasinat_reflection:
@@ -823,43 +823,43 @@
stand for \<^term>\<open>m\<close> and \<^term>\<open>b\<close>, respectively.\<close>
(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
- "is_nat_case(M, a, is_b, k, z) ==
+ "is_nat_case(M, a, is_b, k, z) \<equiv>
(empty(M,k) \<longrightarrow> z=a) &
(\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))" *)
text\<open>The formula \<^term>\<open>is_b\<close> has free variables 1 and 0.\<close>
definition
is_nat_case_fm :: "[i, i, i, i]=>i" where
- "is_nat_case_fm(a,is_b,k,z) ==
+ "is_nat_case_fm(a,is_b,k,z) \<equiv>
And(Implies(empty_fm(k), Equal(z,a)),
And(Forall(Implies(succ_fm(0,succ(k)),
Forall(Implies(Equal(0,succ(succ(z))), is_b)))),
Or(quasinat_fm(k), empty_fm(z))))"
lemma is_nat_case_type [TC]:
- "[| is_b \<in> formula;
- x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> is_nat_case_fm(x,is_b,y,z) \<in> formula"
+ "\<lbrakk>is_b \<in> formula;
+ x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> is_nat_case_fm(x,is_b,y,z) \<in> formula"
by (simp add: is_nat_case_fm_def)
lemma sats_is_nat_case_fm:
assumes is_b_iff_sats:
- "!!a. a \<in> A ==> is_b(a,nth(z, env)) \<longleftrightarrow>
+ "\<And>a. a \<in> A \<Longrightarrow> is_b(a,nth(z, env)) \<longleftrightarrow>
sats(A, p, Cons(nth(z,env), Cons(a, env)))"
shows
- "[|x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
- ==> sats(A, is_nat_case_fm(x,p,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_nat_case_fm(x,p,y,z), env) \<longleftrightarrow>
is_nat_case(##A, nth(x,env), is_b, nth(y,env), nth(z,env))"
apply (frule lt_length_in_nat, assumption)
apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym])
done
lemma is_nat_case_iff_sats:
- "[| (!!a. a \<in> A ==> is_b(a,z) \<longleftrightarrow>
+ "\<lbrakk>(\<And>a. a \<in> A \<Longrightarrow> is_b(a,z) \<longleftrightarrow>
sats(A, p, Cons(z, Cons(a,env))));
nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
- ==> is_nat_case(##A, x, is_b, y, z) \<longleftrightarrow> sats(A, is_nat_case_fm(i,p,j,k), env)"
+ i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_nat_case(##A, x, is_b, y, z) \<longleftrightarrow> sats(A, is_nat_case_fm(i,p,j,k), env)"
by (simp add: sats_is_nat_case_fm [of A is_b])
@@ -868,7 +868,7 @@
argument, we cannot prove reflection for \<^term>\<open>iterates_MH\<close>.\<close>
theorem is_nat_case_reflection:
assumes is_b_reflection:
- "!!h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)),
+ "\<And>h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)),
\<lambda>i x. is_b(##Lset(i), h(x), f(x), g(x))]"
shows "REFLECTS[\<lambda>x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)),
\<lambda>i x. is_nat_case(##Lset(i), f(x), is_b(##Lset(i), x), g(x), h(x))]"
@@ -881,31 +881,31 @@
subsection\<open>The Operator \<^term>\<open>iterates_MH\<close>, Needed for Iteration\<close>
(* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o"
- "iterates_MH(M,isF,v,n,g,z) ==
+ "iterates_MH(M,isF,v,n,g,z) \<equiv>
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
n, z)" *)
definition
iterates_MH_fm :: "[i, i, i, i, i]=>i" where
- "iterates_MH_fm(isF,v,n,g,z) ==
+ "iterates_MH_fm(isF,v,n,g,z) \<equiv>
is_nat_case_fm(v,
Exists(And(fun_apply_fm(succ(succ(succ(g))),2,0),
Forall(Implies(Equal(0,2), isF)))),
n, z)"
lemma iterates_MH_type [TC]:
- "[| p \<in> formula;
- v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> iterates_MH_fm(p,v,x,y,z) \<in> formula"
+ "\<lbrakk>p \<in> formula;
+ v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> iterates_MH_fm(p,v,x,y,z) \<in> formula"
by (simp add: iterates_MH_fm_def)
lemma sats_iterates_MH_fm:
assumes is_F_iff_sats:
- "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
- ==> is_F(a,b) \<longleftrightarrow>
+ "\<And>a b c d. \<lbrakk>a \<in> A; b \<in> A; c \<in> A; d \<in> A\<rbrakk>
+ \<Longrightarrow> is_F(a,b) \<longleftrightarrow>
sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))"
shows
- "[|v \<in> nat; x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|]
- ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>v \<in> nat; x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, iterates_MH_fm(p,v,x,y,z), env) \<longleftrightarrow>
iterates_MH(##A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))"
apply (frule lt_length_in_nat, assumption)
apply (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm
@@ -916,13 +916,13 @@
lemma iterates_MH_iff_sats:
assumes is_F_iff_sats:
- "!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|]
- ==> is_F(a,b) \<longleftrightarrow>
+ "\<And>a b c d. \<lbrakk>a \<in> A; b \<in> A; c \<in> A; d \<in> A\<rbrakk>
+ \<Longrightarrow> is_F(a,b) \<longleftrightarrow>
sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))"
shows
- "[| nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i' \<in> nat; i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|]
- ==> iterates_MH(##A, is_F, v, x, y, z) \<longleftrightarrow>
+ "\<lbrakk>nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i' \<in> nat; i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> iterates_MH(##A, is_F, v, x, y, z) \<longleftrightarrow>
sats(A, iterates_MH_fm(p,i',i,j,k), env)"
by (simp add: sats_iterates_MH_fm [OF is_F_iff_sats])
@@ -931,7 +931,7 @@
argument, we cannot prove reflection for \<^term>\<open>list_N\<close>.\<close>
theorem iterates_MH_reflection:
assumes p_reflection:
- "!!f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
+ "\<And>f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
\<lambda>i x. p(##Lset(i), h(x), f(x), g(x))]"
shows "REFLECTS[\<lambda>x. iterates_MH(L, p(L,x), e(x), f(x), g(x), h(x)),
\<lambda>i x. iterates_MH(##Lset(i), p(##Lset(i),x), e(x), f(x), g(x), h(x))]"
@@ -946,13 +946,13 @@
text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0;
\<^term>\<open>p\<close> is enclosed by 9 (??) quantifiers.\<close>
-(* "is_iterates(M,isF,v,n,Z) ==
+(* "is_iterates(M,isF,v,n,Z) \<equiv>
\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
1 0 is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*)
definition
is_iterates_fm :: "[i, i, i, i]=>i" where
- "is_iterates_fm(p,v,n,Z) ==
+ "is_iterates_fm(p,v,n,Z) \<equiv>
Exists(Exists(
And(succ_fm(n#+2,1),
And(Memrel_fm(1,0),
@@ -964,21 +964,21 @@
lemma is_iterates_type [TC]:
- "[| p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> is_iterates_fm(p,x,y,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> is_iterates_fm(p,x,y,z) \<in> formula"
by (simp add: is_iterates_fm_def)
lemma sats_is_iterates_fm:
assumes is_F_iff_sats:
- "!!a b c d e f g h i j k.
- [| a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A;
- g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A|]
- ==> is_F(a,b) \<longleftrightarrow>
+ "\<And>a b c d e f g h i j k.
+ \<lbrakk>a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A;
+ g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A\<rbrakk>
+ \<Longrightarrow> is_F(a,b) \<longleftrightarrow>
sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f,
Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))"
shows
- "[|x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
- ==> sats(A, is_iterates_fm(p,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_iterates_fm(p,x,y,z), env) \<longleftrightarrow>
is_iterates(##A, is_F, nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=z in lt_length_in_nat, assumption)
apply (frule lt_length_in_nat, assumption)
@@ -989,16 +989,16 @@
lemma is_iterates_iff_sats:
assumes is_F_iff_sats:
- "!!a b c d e f g h i j k.
- [| a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A;
- g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A|]
- ==> is_F(a,b) \<longleftrightarrow>
+ "\<And>a b c d e f g h i j k.
+ \<lbrakk>a \<in> A; b \<in> A; c \<in> A; d \<in> A; e \<in> A; f \<in> A;
+ g \<in> A; h \<in> A; i \<in> A; j \<in> A; k \<in> A\<rbrakk>
+ \<Longrightarrow> is_F(a,b) \<longleftrightarrow>
sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f,
Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))"
shows
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
- ==> is_iterates(##A, is_F, x, y, z) \<longleftrightarrow>
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j < length(env); k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_iterates(##A, is_F, x, y, z) \<longleftrightarrow>
sats(A, is_iterates_fm(p,i,j,k), env)"
by (simp add: sats_is_iterates_fm [OF is_F_iff_sats])
@@ -1007,7 +1007,7 @@
argument, we cannot prove reflection for \<^term>\<open>list_N\<close>.\<close>
theorem is_iterates_reflection:
assumes p_reflection:
- "!!f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
+ "\<And>f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
\<lambda>i x. p(##Lset(i), h(x), f(x), g(x))]"
shows "REFLECTS[\<lambda>x. is_iterates(L, p(L,x), f(x), g(x), h(x)),
\<lambda>i x. is_iterates(##Lset(i), p(##Lset(i),x), f(x), g(x), h(x))]"
@@ -1019,19 +1019,19 @@
subsubsection\<open>The Formula \<^term>\<open>is_eclose_n\<close>, Internalized\<close>
-(* is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z) *)
+(* is_eclose_n(M,A,n,Z) \<equiv> is_iterates(M, big_union(M), A, n, Z) *)
definition
eclose_n_fm :: "[i,i,i]=>i" where
- "eclose_n_fm(A,n,Z) == is_iterates_fm(big_union_fm(1,0), A, n, Z)"
+ "eclose_n_fm(A,n,Z) \<equiv> is_iterates_fm(big_union_fm(1,0), A, n, Z)"
lemma eclose_n_fm_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> eclose_n_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> eclose_n_fm(x,y,z) \<in> formula"
by (simp add: eclose_n_fm_def)
lemma sats_eclose_n_fm [simp]:
- "[| x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
- ==> sats(A, eclose_n_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, eclose_n_fm(x,y,z), env) \<longleftrightarrow>
is_eclose_n(##A, nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=z in lt_length_in_nat, assumption)
apply (frule_tac x=y in lt_length_in_nat, assumption)
@@ -1040,9 +1040,9 @@
done
lemma eclose_n_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
- ==> is_eclose_n(##A, x, y, z) \<longleftrightarrow> sats(A, eclose_n_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j < length(env); k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_eclose_n(##A, x, y, z) \<longleftrightarrow> sats(A, eclose_n_fm(i,j,k), env)"
by (simp)
theorem eclose_n_reflection:
@@ -1055,29 +1055,29 @@
subsubsection\<open>Membership in \<^term>\<open>eclose(A)\<close>\<close>
-(* mem_eclose(M,A,l) ==
+(* mem_eclose(M,A,l) \<equiv>
\<exists>n[M]. \<exists>eclosen[M].
finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen *)
definition
mem_eclose_fm :: "[i,i]=>i" where
- "mem_eclose_fm(x,y) ==
+ "mem_eclose_fm(x,y) \<equiv>
Exists(Exists(
And(finite_ordinal_fm(1),
And(eclose_n_fm(x#+2,1,0), Member(y#+2,0)))))"
lemma mem_eclose_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> mem_eclose_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> mem_eclose_fm(x,y) \<in> formula"
by (simp add: mem_eclose_fm_def)
lemma sats_mem_eclose_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, mem_eclose_fm(x,y), env) \<longleftrightarrow> mem_eclose(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, mem_eclose_fm(x,y), env) \<longleftrightarrow> mem_eclose(##A, nth(x,env), nth(y,env))"
by (simp add: mem_eclose_fm_def mem_eclose_def)
lemma mem_eclose_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> mem_eclose(##A, x, y) \<longleftrightarrow> sats(A, mem_eclose_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> mem_eclose(##A, x, y) \<longleftrightarrow> sats(A, mem_eclose_fm(i,j), env)"
by simp
theorem mem_eclose_reflection:
@@ -1090,25 +1090,25 @@
subsubsection\<open>The Predicate ``Is \<^term>\<open>eclose(A)\<close>''\<close>
-(* is_eclose(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_eclose(M,A,l) *)
+(* is_eclose(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_eclose(M,A,l) *)
definition
is_eclose_fm :: "[i,i]=>i" where
- "is_eclose_fm(A,Z) ==
+ "is_eclose_fm(A,Z) \<equiv>
Forall(Iff(Member(0,succ(Z)), mem_eclose_fm(succ(A),0)))"
lemma is_eclose_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> is_eclose_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> is_eclose_fm(x,y) \<in> formula"
by (simp add: is_eclose_fm_def)
lemma sats_is_eclose_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, is_eclose_fm(x,y), env) \<longleftrightarrow> is_eclose(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_eclose_fm(x,y), env) \<longleftrightarrow> is_eclose(##A, nth(x,env), nth(y,env))"
by (simp add: is_eclose_fm_def is_eclose_def)
lemma is_eclose_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_eclose(##A, x, y) \<longleftrightarrow> sats(A, is_eclose_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_eclose(##A, x, y) \<longleftrightarrow> sats(A, is_eclose_fm(i,j), env)"
by simp
theorem is_eclose_reflection:
@@ -1123,28 +1123,28 @@
definition
list_functor_fm :: "[i,i,i]=>i" where
-(* "is_list_functor(M,A,X,Z) ==
+(* "is_list_functor(M,A,X,Z) \<equiv>
\<exists>n1[M]. \<exists>AX[M].
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *)
- "list_functor_fm(A,X,Z) ==
+ "list_functor_fm(A,X,Z) \<equiv>
Exists(Exists(
And(number1_fm(1),
And(cartprod_fm(A#+2,X#+2,0), sum_fm(1,0,Z#+2)))))"
lemma list_functor_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> list_functor_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> list_functor_fm(x,y,z) \<in> formula"
by (simp add: list_functor_fm_def)
lemma sats_list_functor_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, list_functor_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, list_functor_fm(x,y,z), env) \<longleftrightarrow>
is_list_functor(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: list_functor_fm_def is_list_functor_def)
lemma list_functor_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_list_functor(##A, x, y, z) \<longleftrightarrow> sats(A, list_functor_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_list_functor(##A, x, y, z) \<longleftrightarrow> sats(A, list_functor_fm(i,j,k), env)"
by simp
theorem list_functor_reflection:
@@ -1158,24 +1158,24 @@
subsubsection\<open>The Formula \<^term>\<open>is_list_N\<close>, Internalized\<close>
-(* "is_list_N(M,A,n,Z) ==
+(* "is_list_N(M,A,n,Z) \<equiv>
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_list_functor(M,A), zero, n, Z)" *)
definition
list_N_fm :: "[i,i,i]=>i" where
- "list_N_fm(A,n,Z) ==
+ "list_N_fm(A,n,Z) \<equiv>
Exists(
And(empty_fm(0),
is_iterates_fm(list_functor_fm(A#+9#+3,1,0), 0, n#+1, Z#+1)))"
lemma list_N_fm_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> list_N_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> list_N_fm(x,y,z) \<in> formula"
by (simp add: list_N_fm_def)
lemma sats_list_N_fm [simp]:
- "[| x \<in> nat; y < length(env); z < length(env); env \<in> list(A)|]
- ==> sats(A, list_N_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, list_N_fm(x,y,z), env) \<longleftrightarrow>
is_list_N(##A, nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=z in lt_length_in_nat, assumption)
apply (frule_tac x=y in lt_length_in_nat, assumption)
@@ -1183,9 +1183,9 @@
done
lemma list_N_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j < length(env); k < length(env); env \<in> list(A)|]
- ==> is_list_N(##A, x, y, z) \<longleftrightarrow> sats(A, list_N_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j < length(env); k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_list_N(##A, x, y, z) \<longleftrightarrow> sats(A, list_N_fm(i,j,k), env)"
by (simp)
theorem list_N_reflection:
@@ -1200,29 +1200,29 @@
subsubsection\<open>The Predicate ``Is A List''\<close>
-(* mem_list(M,A,l) ==
+(* mem_list(M,A,l) \<equiv>
\<exists>n[M]. \<exists>listn[M].
finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn *)
definition
mem_list_fm :: "[i,i]=>i" where
- "mem_list_fm(x,y) ==
+ "mem_list_fm(x,y) \<equiv>
Exists(Exists(
And(finite_ordinal_fm(1),
And(list_N_fm(x#+2,1,0), Member(y#+2,0)))))"
lemma mem_list_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> mem_list_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> mem_list_fm(x,y) \<in> formula"
by (simp add: mem_list_fm_def)
lemma sats_mem_list_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, mem_list_fm(x,y), env) \<longleftrightarrow> mem_list(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, mem_list_fm(x,y), env) \<longleftrightarrow> mem_list(##A, nth(x,env), nth(y,env))"
by (simp add: mem_list_fm_def mem_list_def)
lemma mem_list_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> mem_list(##A, x, y) \<longleftrightarrow> sats(A, mem_list_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> mem_list(##A, x, y) \<longleftrightarrow> sats(A, mem_list_fm(i,j), env)"
by simp
theorem mem_list_reflection:
@@ -1235,25 +1235,25 @@
subsubsection\<open>The Predicate ``Is \<^term>\<open>list(A)\<close>''\<close>
-(* is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l) *)
+(* is_list(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l) *)
definition
is_list_fm :: "[i,i]=>i" where
- "is_list_fm(A,Z) ==
+ "is_list_fm(A,Z) \<equiv>
Forall(Iff(Member(0,succ(Z)), mem_list_fm(succ(A),0)))"
lemma is_list_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> is_list_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> is_list_fm(x,y) \<in> formula"
by (simp add: is_list_fm_def)
lemma sats_is_list_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, is_list_fm(x,y), env) \<longleftrightarrow> is_list(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_list_fm(x,y), env) \<longleftrightarrow> is_list(##A, nth(x,env), nth(y,env))"
by (simp add: is_list_fm_def is_list_def)
lemma is_list_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_list(##A, x, y) \<longleftrightarrow> sats(A, is_list_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_list(##A, x, y) \<longleftrightarrow> sats(A, is_list_fm(i,j), env)"
by simp
theorem is_list_reflection:
@@ -1267,14 +1267,14 @@
subsubsection\<open>The Formula Functor, Internalized\<close>
definition formula_functor_fm :: "[i,i]=>i" where
-(* "is_formula_functor(M,X,Z) ==
+(* "is_formula_functor(M,X,Z) \<equiv>
\<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
4 3 2 1 0
omega(M,nat') & cartprod(M,nat',nat',natnat) &
is_sum(M,natnat,natnat,natnatsum) &
cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) &
is_sum(M,natnatsum,X3,Z)" *)
- "formula_functor_fm(X,Z) ==
+ "formula_functor_fm(X,Z) \<equiv>
Exists(Exists(Exists(Exists(Exists(
And(omega_fm(4),
And(cartprod_fm(4,4,3),
@@ -1283,19 +1283,19 @@
And(sum_fm(1,X#+5,0), sum_fm(2,0,Z#+5)))))))))))"
lemma formula_functor_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> formula_functor_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> formula_functor_fm(x,y) \<in> formula"
by (simp add: formula_functor_fm_def)
lemma sats_formula_functor_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, formula_functor_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, formula_functor_fm(x,y), env) \<longleftrightarrow>
is_formula_functor(##A, nth(x,env), nth(y,env))"
by (simp add: formula_functor_fm_def is_formula_functor_def)
lemma formula_functor_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_formula_functor(##A, x, y) \<longleftrightarrow> sats(A, formula_functor_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_formula_functor(##A, x, y) \<longleftrightarrow> sats(A, formula_functor_fm(i,j), env)"
by simp
theorem formula_functor_reflection:
@@ -1309,23 +1309,23 @@
subsubsection\<open>The Formula \<^term>\<open>is_formula_N\<close>, Internalized\<close>
-(* "is_formula_N(M,n,Z) ==
+(* "is_formula_N(M,n,Z) \<equiv>
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_formula_functor(M), zero, n, Z)" *)
definition
formula_N_fm :: "[i,i]=>i" where
- "formula_N_fm(n,Z) ==
+ "formula_N_fm(n,Z) \<equiv>
Exists(
And(empty_fm(0),
is_iterates_fm(formula_functor_fm(1,0), 0, n#+1, Z#+1)))"
lemma formula_N_fm_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> formula_N_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> formula_N_fm(x,y) \<in> formula"
by (simp add: formula_N_fm_def)
lemma sats_formula_N_fm [simp]:
- "[| x < length(env); y < length(env); env \<in> list(A)|]
- ==> sats(A, formula_N_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x < length(env); y < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, formula_N_fm(x,y), env) \<longleftrightarrow>
is_formula_N(##A, nth(x,env), nth(y,env))"
apply (frule_tac x=y in lt_length_in_nat, assumption)
apply (frule lt_length_in_nat, assumption)
@@ -1333,9 +1333,9 @@
done
lemma formula_N_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i < length(env); j < length(env); env \<in> list(A)|]
- ==> is_formula_N(##A, x, y) \<longleftrightarrow> sats(A, formula_N_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i < length(env); j < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_formula_N(##A, x, y) \<longleftrightarrow> sats(A, formula_N_fm(i,j), env)"
by (simp)
theorem formula_N_reflection:
@@ -1350,28 +1350,28 @@
subsubsection\<open>The Predicate ``Is A Formula''\<close>
-(* mem_formula(M,p) ==
+(* mem_formula(M,p) \<equiv>
\<exists>n[M]. \<exists>formn[M].
finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn *)
definition
mem_formula_fm :: "i=>i" where
- "mem_formula_fm(x) ==
+ "mem_formula_fm(x) \<equiv>
Exists(Exists(
And(finite_ordinal_fm(1),
And(formula_N_fm(1,0), Member(x#+2,0)))))"
lemma mem_formula_type [TC]:
- "x \<in> nat ==> mem_formula_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> mem_formula_fm(x) \<in> formula"
by (simp add: mem_formula_fm_def)
lemma sats_mem_formula_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, mem_formula_fm(x), env) \<longleftrightarrow> mem_formula(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, mem_formula_fm(x), env) \<longleftrightarrow> mem_formula(##A, nth(x,env))"
by (simp add: mem_formula_fm_def mem_formula_def)
lemma mem_formula_iff_sats:
- "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
- ==> mem_formula(##A, x) \<longleftrightarrow> sats(A, mem_formula_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> mem_formula(##A, x) \<longleftrightarrow> sats(A, mem_formula_fm(i), env)"
by simp
theorem mem_formula_reflection:
@@ -1385,23 +1385,23 @@
subsubsection\<open>The Predicate ``Is \<^term>\<open>formula\<close>''\<close>
-(* is_formula(M,Z) == \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p) *)
+(* is_formula(M,Z) \<equiv> \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p) *)
definition
is_formula_fm :: "i=>i" where
- "is_formula_fm(Z) == Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))"
+ "is_formula_fm(Z) \<equiv> Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))"
lemma is_formula_type [TC]:
- "x \<in> nat ==> is_formula_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> is_formula_fm(x) \<in> formula"
by (simp add: is_formula_fm_def)
lemma sats_is_formula_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, is_formula_fm(x), env) \<longleftrightarrow> is_formula(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_formula_fm(x), env) \<longleftrightarrow> is_formula(##A, nth(x,env))"
by (simp add: is_formula_fm_def is_formula_def)
lemma is_formula_iff_sats:
- "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
- ==> is_formula(##A, x) \<longleftrightarrow> sats(A, is_formula_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_formula(##A, x) \<longleftrightarrow> sats(A, is_formula_fm(i), env)"
by simp
theorem is_formula_reflection:
@@ -1420,14 +1420,14 @@
the corresponding quantified variables with de Bruijn indices 2, 1, 0.\<close>
(* is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o"
- "is_transrec(M,MH,a,z) ==
+ "is_transrec(M,MH,a,z) \<equiv>
\<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
2 1 0
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
is_wfrec(M,MH,mesa,a,z)" *)
definition
is_transrec_fm :: "[i, i, i]=>i" where
- "is_transrec_fm(p,a,z) ==
+ "is_transrec_fm(p,a,z) \<equiv>
Exists(Exists(Exists(
And(upair_fm(a#+3,a#+3,2),
And(is_eclose_fm(2,1),
@@ -1435,20 +1435,20 @@
lemma is_transrec_type [TC]:
- "[| p \<in> formula; x \<in> nat; z \<in> nat |]
- ==> is_transrec_fm(p,x,z) \<in> formula"
+ "\<lbrakk>p \<in> formula; x \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> is_transrec_fm(p,x,z) \<in> formula"
by (simp add: is_transrec_fm_def)
lemma sats_is_transrec_fm:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4 a5 a6 a7.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow>
+ "\<And>a0 a1 a2 a3 a4 a5 a6 a7.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow>
sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))"
shows
- "[|x < length(env); z < length(env); env \<in> list(A)|]
- ==> sats(A, is_transrec_fm(p,x,z), env) \<longleftrightarrow>
+ "\<lbrakk>x < length(env); z < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, is_transrec_fm(p,x,z), env) \<longleftrightarrow>
is_transrec(##A, MH, nth(x,env), nth(z,env))"
apply (frule_tac x=z in lt_length_in_nat, assumption)
apply (frule_tac x=x in lt_length_in_nat, assumption)
@@ -1458,20 +1458,20 @@
lemma is_transrec_iff_sats:
assumes MH_iff_sats:
- "!!a0 a1 a2 a3 a4 a5 a6 a7.
- [|a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A|]
- ==> MH(a2, a1, a0) \<longleftrightarrow>
+ "\<And>a0 a1 a2 a3 a4 a5 a6 a7.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A; a3\<in>A; a4\<in>A; a5\<in>A; a6\<in>A; a7\<in>A\<rbrakk>
+ \<Longrightarrow> MH(a2, a1, a0) \<longleftrightarrow>
sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,
Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))"
shows
- "[|nth(i,env) = x; nth(k,env) = z;
- i < length(env); k < length(env); env \<in> list(A)|]
- ==> is_transrec(##A, MH, x, z) \<longleftrightarrow> sats(A, is_transrec_fm(p,i,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(k,env) = z;
+ i < length(env); k < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_transrec(##A, MH, x, z) \<longleftrightarrow> sats(A, is_transrec_fm(p,i,k), env)"
by (simp add: sats_is_transrec_fm [OF MH_iff_sats])
theorem is_transrec_reflection:
assumes MH_reflection:
- "!!f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
+ "\<And>f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
\<lambda>i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]"
shows "REFLECTS[\<lambda>x. is_transrec(L, MH(L,x), f(x), h(x)),
\<lambda>i x. is_transrec(##Lset(i), MH(##Lset(i),x), f(x), h(x))]"
--- a/src/ZF/Constructible/L_axioms.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 27 16:51:35 2022 +0100
@@ -8,7 +8,7 @@
text \<open>The class L satisfies the premises of locale \<open>M_trivial\<close>\<close>
-lemma transL: "[| y\<in>x; L(x) |] ==> L(y)"
+lemma transL: "\<lbrakk>y\<in>x; L(x)\<rbrakk> \<Longrightarrow> L(y)"
apply (insert Transset_Lset)
apply (simp add: Transset_def L_def, blast)
done
@@ -52,8 +52,8 @@
there too!*)
lemma LReplace_in_Lset:
- "[|X \<in> Lset(i); univalent(L,X,Q); Ord(i)|]
- ==> \<exists>j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Lset(j)"
+ "\<lbrakk>X \<in> Lset(i); univalent(L,X,Q); Ord(i)\<rbrakk>
+ \<Longrightarrow> \<exists>j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Lset(j)"
apply (rule_tac x="\<Union>y \<in> Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))"
in exI)
apply simp
@@ -64,8 +64,8 @@
done
lemma LReplace_in_L:
- "[|L(X); univalent(L,X,Q)|]
- ==> \<exists>Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Y"
+ "\<lbrakk>L(X); univalent(L,X,Q)\<rbrakk>
+ \<Longrightarrow> \<exists>Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Y"
apply (drule L_D, clarify)
apply (drule LReplace_in_Lset, assumption+)
apply (blast intro: L_I Lset_in_Lset_succ)
@@ -79,8 +79,8 @@
done
lemma strong_replacementI [rule_format]:
- "[| \<forall>B[L]. separation(L, %u. \<exists>x[L]. x\<in>B & P(x,u)) |]
- ==> strong_replacement(L,P)"
+ "\<lbrakk>\<forall>B[L]. separation(L, %u. \<exists>x[L]. x\<in>B & P(x,u))\<rbrakk>
+ \<Longrightarrow> strong_replacement(L,P)"
apply (simp add: strong_replacement_def, clarify)
apply (frule replacementD [OF replacement], assumption, clarify)
apply (drule_tac x=A in rspec, clarify)
@@ -127,22 +127,22 @@
definition
L_F0 :: "[i=>o,i] => i" where
- "L_F0(P,y) == \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(<y,z>))"
+ "L_F0(P,y) \<equiv> \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(<y,z>))"
definition
L_FF :: "[i=>o,i] => i" where
- "L_FF(P) == \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)"
+ "L_FF(P) \<equiv> \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)"
definition
L_ClEx :: "[i=>o,i] => o" where
- "L_ClEx(P) == \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
+ "L_ClEx(P) \<equiv> \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
text\<open>We must use the meta-existential quantifier; otherwise the reflection
terms become enormous!\<close>
definition
L_Reflects :: "[i=>o,[i,i]=>o] => prop" (\<open>(3REFLECTS/ [_,/ _])\<close>) where
- "REFLECTS[P,Q] == (\<Or>Cl. Closed_Unbounded(Cl) &
+ "REFLECTS[P,Q] \<equiv> (\<Or>Cl. Closed_Unbounded(Cl) &
(\<forall>a. Cl(a) \<longrightarrow> (\<forall>x \<in> Lset(a). P(x) \<longleftrightarrow> Q(a,x))))"
@@ -154,15 +154,15 @@
done
theorem Not_reflection:
- "REFLECTS[P,Q] ==> REFLECTS[\<lambda>x. ~P(x), \<lambda>a x. ~Q(a,x)]"
+ "REFLECTS[P,Q] \<Longrightarrow> REFLECTS[\<lambda>x. \<not>P(x), \<lambda>a x. \<not>Q(a,x)]"
apply (unfold L_Reflects_def)
apply (erule meta_exE)
apply (rule_tac x=Cl in meta_exI, simp)
done
theorem And_reflection:
- "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
- ==> REFLECTS[\<lambda>x. P(x) \<and> P'(x), \<lambda>a x. Q(a,x) \<and> Q'(a,x)]"
+ "\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk>
+ \<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<and> P'(x), \<lambda>a x. Q(a,x) \<and> Q'(a,x)]"
apply (unfold L_Reflects_def)
apply (elim meta_exE)
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -170,8 +170,8 @@
done
theorem Or_reflection:
- "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
- ==> REFLECTS[\<lambda>x. P(x) \<or> P'(x), \<lambda>a x. Q(a,x) \<or> Q'(a,x)]"
+ "\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk>
+ \<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<or> P'(x), \<lambda>a x. Q(a,x) \<or> Q'(a,x)]"
apply (unfold L_Reflects_def)
apply (elim meta_exE)
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -179,8 +179,8 @@
done
theorem Imp_reflection:
- "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
- ==> REFLECTS[\<lambda>x. P(x) \<longrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x)]"
+ "\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk>
+ \<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<longrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x)]"
apply (unfold L_Reflects_def)
apply (elim meta_exE)
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -188,8 +188,8 @@
done
theorem Iff_reflection:
- "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |]
- ==> REFLECTS[\<lambda>x. P(x) \<longleftrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x)]"
+ "\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk>
+ \<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<longleftrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x)]"
apply (unfold L_Reflects_def)
apply (elim meta_exE)
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI)
@@ -204,7 +204,7 @@
theorem Ex_reflection:
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<exists>z. L(z) \<and> P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z. L(z) \<and> P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def)
apply (elim meta_exE)
apply (rule meta_exI)
@@ -213,7 +213,7 @@
theorem All_reflection:
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<forall>z. L(z) \<longrightarrow> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z. L(z) \<longrightarrow> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def)
apply (elim meta_exE)
apply (rule meta_exI)
@@ -222,14 +222,14 @@
theorem Rex_reflection:
"REFLECTS[ \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
apply (unfold rex_def)
apply (intro And_reflection Ex_reflection, assumption)
done
theorem Rall_reflection:
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]"
apply (unfold rall_def)
apply (intro Imp_reflection All_reflection, assumption)
done
@@ -238,7 +238,7 @@
in the second argument of \<open>REFLECTS\<close>.\<close>
theorem Rex_reflection':
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z[##Lset(a)]. Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z[##Lset(a)]. Q(a,x,z)]"
apply (unfold setclass_def rex_def)
apply (erule Rex_reflection [unfolded rex_def Bex_def])
done
@@ -246,7 +246,7 @@
text\<open>As above.\<close>
theorem Rall_reflection':
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
- ==> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z[##Lset(a)]. Q(a,x,z)]"
+ \<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z[##Lset(a)]. Q(a,x,z)]"
apply (unfold setclass_def rall_def)
apply (erule Rall_reflection [unfolded rall_def Ball_def])
done
@@ -257,17 +257,17 @@
Rex_reflection Rall_reflection Rex_reflection' Rall_reflection'
lemma ReflectsD:
- "[|REFLECTS[P,Q]; Ord(i)|]
- ==> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))"
+ "\<lbrakk>REFLECTS[P,Q]; Ord(i)\<rbrakk>
+ \<Longrightarrow> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))"
apply (unfold L_Reflects_def Closed_Unbounded_def)
apply (elim meta_exE, clarify)
apply (blast dest!: UnboundedD)
done
lemma ReflectsE:
- "[| REFLECTS[P,Q]; Ord(i);
- !!j. [|i<j; \<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x)|] ==> R |]
- ==> R"
+ "\<lbrakk>REFLECTS[P,Q]; Ord(i);
+ \<And>j. \<lbrakk>i<j; \<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x)\<rbrakk> \<Longrightarrow> R\<rbrakk>
+ \<Longrightarrow> R"
by (drule ReflectsD, assumption, blast)
lemma Collect_mem_eq: "{x\<in>A. x\<in>B} = A \<inter> B"
@@ -279,46 +279,46 @@
subsubsection\<open>Some numbers to help write de Bruijn indices\<close>
abbreviation
- digit3 :: i (\<open>3\<close>) where "3 == succ(2)"
+ digit3 :: i (\<open>3\<close>) where "3 \<equiv> succ(2)"
abbreviation
- digit4 :: i (\<open>4\<close>) where "4 == succ(3)"
+ digit4 :: i (\<open>4\<close>) where "4 \<equiv> succ(3)"
abbreviation
- digit5 :: i (\<open>5\<close>) where "5 == succ(4)"
+ digit5 :: i (\<open>5\<close>) where "5 \<equiv> succ(4)"
abbreviation
- digit6 :: i (\<open>6\<close>) where "6 == succ(5)"
+ digit6 :: i (\<open>6\<close>) where "6 \<equiv> succ(5)"
abbreviation
- digit7 :: i (\<open>7\<close>) where "7 == succ(6)"
+ digit7 :: i (\<open>7\<close>) where "7 \<equiv> succ(6)"
abbreviation
- digit8 :: i (\<open>8\<close>) where "8 == succ(7)"
+ digit8 :: i (\<open>8\<close>) where "8 \<equiv> succ(7)"
abbreviation
- digit9 :: i (\<open>9\<close>) where "9 == succ(8)"
+ digit9 :: i (\<open>9\<close>) where "9 \<equiv> succ(8)"
subsubsection\<open>The Empty Set, Internalized\<close>
definition
empty_fm :: "i=>i" where
- "empty_fm(x) == Forall(Neg(Member(0,succ(x))))"
+ "empty_fm(x) \<equiv> Forall(Neg(Member(0,succ(x))))"
lemma empty_type [TC]:
- "x \<in> nat ==> empty_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> empty_fm(x) \<in> formula"
by (simp add: empty_fm_def)
lemma sats_empty_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, empty_fm(x), env) \<longleftrightarrow> empty(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, empty_fm(x), env) \<longleftrightarrow> empty(##A, nth(x,env))"
by (simp add: empty_fm_def empty_def)
lemma empty_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> empty(##A, x) \<longleftrightarrow> sats(A, empty_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> empty(##A, x) \<longleftrightarrow> sats(A, empty_fm(i), env)"
by simp
theorem empty_reflection:
@@ -330,8 +330,8 @@
text\<open>Not used. But maybe useful?\<close>
lemma Transset_sats_empty_fm_eq_0:
- "[| n \<in> nat; env \<in> list(A); Transset(A)|]
- ==> sats(A, empty_fm(n), env) \<longleftrightarrow> nth(n,env) = 0"
+ "\<lbrakk>n \<in> nat; env \<in> list(A); Transset(A)\<rbrakk>
+ \<Longrightarrow> sats(A, empty_fm(n), env) \<longleftrightarrow> nth(n,env) = 0"
apply (simp add: empty_fm_def empty_def Transset_def, auto)
apply (case_tac "n < length(env)")
apply (frule nth_type, assumption+, blast)
@@ -343,32 +343,32 @@
definition
upair_fm :: "[i,i,i]=>i" where
- "upair_fm(x,y,z) ==
+ "upair_fm(x,y,z) \<equiv>
And(Member(x,z),
And(Member(y,z),
Forall(Implies(Member(0,succ(z)),
Or(Equal(0,succ(x)), Equal(0,succ(y)))))))"
lemma upair_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> upair_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> upair_fm(x,y,z) \<in> formula"
by (simp add: upair_fm_def)
lemma sats_upair_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
upair(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: upair_fm_def upair_def)
lemma upair_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> upair(##A, x, y, z) \<longleftrightarrow> sats(A, upair_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> upair(##A, x, y, z) \<longleftrightarrow> sats(A, upair_fm(i,j,k), env)"
by (simp)
text\<open>Useful? At least it refers to "real" unordered pairs\<close>
lemma sats_upair_fm2 [simp]:
- "[| x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)|]
- ==> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)\<rbrakk>
+ \<Longrightarrow> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
nth(z,env) = {nth(x,env), nth(y,env)}"
apply (frule lt_length_in_nat, assumption)
apply (simp add: upair_fm_def Transset_def, auto)
@@ -386,25 +386,25 @@
definition
pair_fm :: "[i,i,i]=>i" where
- "pair_fm(x,y,z) ==
+ "pair_fm(x,y,z) \<equiv>
Exists(And(upair_fm(succ(x),succ(x),0),
Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0),
upair_fm(1,0,succ(succ(z)))))))"
lemma pair_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> pair_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> pair_fm(x,y,z) \<in> formula"
by (simp add: pair_fm_def)
lemma sats_pair_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, pair_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, pair_fm(x,y,z), env) \<longleftrightarrow>
pair(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: pair_fm_def pair_def)
lemma pair_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> pair(##A, x, y, z) \<longleftrightarrow> sats(A, pair_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> pair(##A, x, y, z) \<longleftrightarrow> sats(A, pair_fm(i,j,k), env)"
by (simp)
theorem pair_reflection:
@@ -419,24 +419,24 @@
definition
union_fm :: "[i,i,i]=>i" where
- "union_fm(x,y,z) ==
+ "union_fm(x,y,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Or(Member(0,succ(x)),Member(0,succ(y)))))"
lemma union_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> union_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> union_fm(x,y,z) \<in> formula"
by (simp add: union_fm_def)
lemma sats_union_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, union_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, union_fm(x,y,z), env) \<longleftrightarrow>
union(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: union_fm_def union_def)
lemma union_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> union(##A, x, y, z) \<longleftrightarrow> sats(A, union_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> union(##A, x, y, z) \<longleftrightarrow> sats(A, union_fm(i,j,k), env)"
by (simp)
theorem union_reflection:
@@ -451,25 +451,25 @@
definition
cons_fm :: "[i,i,i]=>i" where
- "cons_fm(x,y,z) ==
+ "cons_fm(x,y,z) \<equiv>
Exists(And(upair_fm(succ(x),succ(x),0),
union_fm(0,succ(y),succ(z))))"
lemma cons_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> cons_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> cons_fm(x,y,z) \<in> formula"
by (simp add: cons_fm_def)
lemma sats_cons_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, cons_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, cons_fm(x,y,z), env) \<longleftrightarrow>
is_cons(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: cons_fm_def is_cons_def)
lemma cons_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_cons(##A, x, y, z) \<longleftrightarrow> sats(A, cons_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_cons(##A, x, y, z) \<longleftrightarrow> sats(A, cons_fm(i,j,k), env)"
by simp
theorem cons_reflection:
@@ -484,22 +484,22 @@
definition
succ_fm :: "[i,i]=>i" where
- "succ_fm(x,y) == cons_fm(x,x,y)"
+ "succ_fm(x,y) \<equiv> cons_fm(x,x,y)"
lemma succ_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> succ_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> succ_fm(x,y) \<in> formula"
by (simp add: succ_fm_def)
lemma sats_succ_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, succ_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, succ_fm(x,y), env) \<longleftrightarrow>
successor(##A, nth(x,env), nth(y,env))"
by (simp add: succ_fm_def successor_def)
lemma successor_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> successor(##A, x, y) \<longleftrightarrow> sats(A, succ_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> successor(##A, x, y) \<longleftrightarrow> sats(A, succ_fm(i,j), env)"
by simp
theorem successor_reflection:
@@ -512,24 +512,24 @@
subsubsection\<open>The Number 1, Internalized\<close>
-(* "number1(M,a) == (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
+(* "number1(M,a) \<equiv> (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
definition
number1_fm :: "i=>i" where
- "number1_fm(a) == Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
+ "number1_fm(a) \<equiv> Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
lemma number1_type [TC]:
- "x \<in> nat ==> number1_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> number1_fm(x) \<in> formula"
by (simp add: number1_fm_def)
lemma sats_number1_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, number1_fm(x), env) \<longleftrightarrow> number1(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, number1_fm(x), env) \<longleftrightarrow> number1(##A, nth(x,env))"
by (simp add: number1_fm_def number1_def)
lemma number1_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> number1(##A, x) \<longleftrightarrow> sats(A, number1_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> number1(##A, x) \<longleftrightarrow> sats(A, number1_fm(i), env)"
by simp
theorem number1_reflection:
@@ -542,27 +542,27 @@
subsubsection\<open>Big Union, Internalized\<close>
-(* "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
+(* "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
definition
big_union_fm :: "[i,i]=>i" where
- "big_union_fm(A,z) ==
+ "big_union_fm(A,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(A))), Member(1,0)))))"
lemma big_union_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> big_union_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> big_union_fm(x,y) \<in> formula"
by (simp add: big_union_fm_def)
lemma sats_big_union_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, big_union_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, big_union_fm(x,y), env) \<longleftrightarrow>
big_union(##A, nth(x,env), nth(y,env))"
by (simp add: big_union_fm_def big_union_def)
lemma big_union_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> big_union(##A, x, y) \<longleftrightarrow> sats(A, big_union_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> big_union(##A, x, y) \<longleftrightarrow> sats(A, big_union_fm(i,j), env)"
by simp
theorem big_union_reflection:
@@ -582,8 +582,8 @@
\<open>M_trivial\<close>, we no longer require the earlier versions.\<close>
lemma sats_subset_fm':
- "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, subset_fm(x,y), env) \<longleftrightarrow> subset(##A, nth(x,env), nth(y,env))"
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, subset_fm(x,y), env) \<longleftrightarrow> subset(##A, nth(x,env), nth(y,env))"
by (simp add: subset_fm_def Relative.subset_def)
theorem subset_reflection:
@@ -594,8 +594,8 @@
done
lemma sats_transset_fm':
- "[|x \<in> nat; env \<in> list(A)|]
- ==> sats(A, transset_fm(x), env) \<longleftrightarrow> transitive_set(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, transset_fm(x), env) \<longleftrightarrow> transitive_set(##A, nth(x,env))"
by (simp add: sats_subset_fm' transset_fm_def transitive_set_def)
theorem transitive_set_reflection:
@@ -606,13 +606,13 @@
done
lemma sats_ordinal_fm':
- "[|x \<in> nat; env \<in> list(A)|]
- ==> sats(A, ordinal_fm(x), env) \<longleftrightarrow> ordinal(##A,nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, ordinal_fm(x), env) \<longleftrightarrow> ordinal(##A,nth(x,env))"
by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def)
lemma ordinal_iff_sats:
- "[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|]
- ==> ordinal(##A, x) \<longleftrightarrow> sats(A, ordinal_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> ordinal(##A, x) \<longleftrightarrow> sats(A, ordinal_fm(i), env)"
by (simp add: sats_ordinal_fm')
theorem ordinal_reflection:
@@ -626,7 +626,7 @@
definition
Memrel_fm :: "[i,i]=>i" where
- "Memrel_fm(A,r) ==
+ "Memrel_fm(A,r) \<equiv>
Forall(Iff(Member(0,succ(r)),
Exists(And(Member(0,succ(succ(A))),
Exists(And(Member(0,succ(succ(succ(A)))),
@@ -634,19 +634,19 @@
pair_fm(1,0,2))))))))"
lemma Memrel_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> Memrel_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> Memrel_fm(x,y) \<in> formula"
by (simp add: Memrel_fm_def)
lemma sats_Memrel_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, Memrel_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, Memrel_fm(x,y), env) \<longleftrightarrow>
membership(##A, nth(x,env), nth(y,env))"
by (simp add: Memrel_fm_def membership_def)
lemma Memrel_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> membership(##A, x, y) \<longleftrightarrow> sats(A, Memrel_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> membership(##A, x, y) \<longleftrightarrow> sats(A, Memrel_fm(i,j), env)"
by simp
theorem membership_reflection:
@@ -660,7 +660,7 @@
definition
pred_set_fm :: "[i,i,i,i]=>i" where
- "pred_set_fm(A,x,r,B) ==
+ "pred_set_fm(A,x,r,B) \<equiv>
Forall(Iff(Member(0,succ(B)),
Exists(And(Member(0,succ(succ(r))),
And(Member(1,succ(succ(A))),
@@ -668,20 +668,20 @@
lemma pred_set_type [TC]:
- "[| A \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat |]
- ==> pred_set_fm(A,x,r,B) \<in> formula"
+ "\<lbrakk>A \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat\<rbrakk>
+ \<Longrightarrow> pred_set_fm(A,x,r,B) \<in> formula"
by (simp add: pred_set_fm_def)
lemma sats_pred_set_fm [simp]:
- "[| U \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat; env \<in> list(A)|]
- ==> sats(A, pred_set_fm(U,x,r,B), env) \<longleftrightarrow>
+ "\<lbrakk>U \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, pred_set_fm(U,x,r,B), env) \<longleftrightarrow>
pred_set(##A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))"
by (simp add: pred_set_fm_def pred_set_def)
lemma pred_set_iff_sats:
- "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B;
- i \<in> nat; j \<in> nat; k \<in> nat; l \<in> nat; env \<in> list(A)|]
- ==> pred_set(##A,U,x,r,B) \<longleftrightarrow> sats(A, pred_set_fm(i,j,k,l), env)"
+ "\<lbrakk>nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B;
+ i \<in> nat; j \<in> nat; k \<in> nat; l \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> pred_set(##A,U,x,r,B) \<longleftrightarrow> sats(A, pred_set_fm(i,j,k,l), env)"
by (simp)
theorem pred_set_reflection:
@@ -695,29 +695,29 @@
subsubsection\<open>Domain of a Relation, Internalized\<close>
-(* "is_domain(M,r,z) ==
+(* "is_domain(M,r,z) \<equiv>
\<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
definition
domain_fm :: "[i,i]=>i" where
- "domain_fm(r,z) ==
+ "domain_fm(r,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
Exists(pair_fm(2,0,1))))))"
lemma domain_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> domain_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> domain_fm(x,y) \<in> formula"
by (simp add: domain_fm_def)
lemma sats_domain_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, domain_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, domain_fm(x,y), env) \<longleftrightarrow>
is_domain(##A, nth(x,env), nth(y,env))"
by (simp add: domain_fm_def is_domain_def)
lemma domain_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_domain(##A, x, y) \<longleftrightarrow> sats(A, domain_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_domain(##A, x, y) \<longleftrightarrow> sats(A, domain_fm(i,j), env)"
by simp
theorem domain_reflection:
@@ -730,29 +730,29 @@
subsubsection\<open>Range of a Relation, Internalized\<close>
-(* "is_range(M,r,z) ==
+(* "is_range(M,r,z) \<equiv>
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
definition
range_fm :: "[i,i]=>i" where
- "range_fm(r,z) ==
+ "range_fm(r,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
Exists(pair_fm(0,2,1))))))"
lemma range_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> range_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> range_fm(x,y) \<in> formula"
by (simp add: range_fm_def)
lemma sats_range_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, range_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, range_fm(x,y), env) \<longleftrightarrow>
is_range(##A, nth(x,env), nth(y,env))"
by (simp add: range_fm_def is_range_def)
lemma range_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_range(##A, x, y) \<longleftrightarrow> sats(A, range_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_range(##A, x, y) \<longleftrightarrow> sats(A, range_fm(i,j), env)"
by simp
theorem range_reflection:
@@ -765,30 +765,30 @@
subsubsection\<open>Field of a Relation, Internalized\<close>
-(* "is_field(M,r,z) ==
+(* "is_field(M,r,z) \<equiv>
\<exists>dr[M]. is_domain(M,r,dr) &
(\<exists>rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *)
definition
field_fm :: "[i,i]=>i" where
- "field_fm(r,z) ==
+ "field_fm(r,z) \<equiv>
Exists(And(domain_fm(succ(r),0),
Exists(And(range_fm(succ(succ(r)),0),
union_fm(1,0,succ(succ(z)))))))"
lemma field_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> field_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> field_fm(x,y) \<in> formula"
by (simp add: field_fm_def)
lemma sats_field_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, field_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, field_fm(x,y), env) \<longleftrightarrow>
is_field(##A, nth(x,env), nth(y,env))"
by (simp add: field_fm_def is_field_def)
lemma field_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> is_field(##A, x, y) \<longleftrightarrow> sats(A, field_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_field(##A, x, y) \<longleftrightarrow> sats(A, field_fm(i,j), env)"
by simp
theorem field_reflection:
@@ -802,30 +802,30 @@
subsubsection\<open>Image under a Relation, Internalized\<close>
-(* "image(M,r,A,z) ==
+(* "image(M,r,A,z) \<equiv>
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
definition
image_fm :: "[i,i,i]=>i" where
- "image_fm(r,A,z) ==
+ "image_fm(r,A,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
Exists(And(Member(0,succ(succ(succ(A)))),
pair_fm(0,2,1)))))))"
lemma image_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> image_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> image_fm(x,y,z) \<in> formula"
by (simp add: image_fm_def)
lemma sats_image_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, image_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, image_fm(x,y,z), env) \<longleftrightarrow>
image(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: image_fm_def Relative.image_def)
lemma image_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> image(##A, x, y, z) \<longleftrightarrow> sats(A, image_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> image(##A, x, y, z) \<longleftrightarrow> sats(A, image_fm(i,j,k), env)"
by (simp)
theorem image_reflection:
@@ -838,30 +838,30 @@
subsubsection\<open>Pre-Image under a Relation, Internalized\<close>
-(* "pre_image(M,r,A,z) ==
+(* "pre_image(M,r,A,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
definition
pre_image_fm :: "[i,i,i]=>i" where
- "pre_image_fm(r,A,z) ==
+ "pre_image_fm(r,A,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
Exists(And(Member(0,succ(succ(succ(A)))),
pair_fm(2,0,1)))))))"
lemma pre_image_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> pre_image_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> pre_image_fm(x,y,z) \<in> formula"
by (simp add: pre_image_fm_def)
lemma sats_pre_image_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, pre_image_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, pre_image_fm(x,y,z), env) \<longleftrightarrow>
pre_image(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: pre_image_fm_def Relative.pre_image_def)
lemma pre_image_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> pre_image(##A, x, y, z) \<longleftrightarrow> sats(A, pre_image_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> pre_image(##A, x, y, z) \<longleftrightarrow> sats(A, pre_image_fm(i,j,k), env)"
by (simp)
theorem pre_image_reflection:
@@ -874,30 +874,30 @@
subsubsection\<open>Function Application, Internalized\<close>
-(* "fun_apply(M,f,x,y) ==
+(* "fun_apply(M,f,x,y) \<equiv>
(\<exists>xs[M]. \<exists>fxs[M].
upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *)
definition
fun_apply_fm :: "[i,i,i]=>i" where
- "fun_apply_fm(f,x,y) ==
+ "fun_apply_fm(f,x,y) \<equiv>
Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1),
And(image_fm(succ(succ(f)), 1, 0),
big_union_fm(0,succ(succ(y)))))))"
lemma fun_apply_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> fun_apply_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> fun_apply_fm(x,y,z) \<in> formula"
by (simp add: fun_apply_fm_def)
lemma sats_fun_apply_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, fun_apply_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, fun_apply_fm(x,y,z), env) \<longleftrightarrow>
fun_apply(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: fun_apply_fm_def fun_apply_def)
lemma fun_apply_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> fun_apply(##A, x, y, z) \<longleftrightarrow> sats(A, fun_apply_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> fun_apply(##A, x, y, z) \<longleftrightarrow> sats(A, fun_apply_fm(i,j,k), env)"
by simp
theorem fun_apply_reflection:
@@ -911,26 +911,26 @@
subsubsection\<open>The Concept of Relation, Internalized\<close>
-(* "is_relation(M,r) ==
+(* "is_relation(M,r) \<equiv>
(\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
definition
relation_fm :: "i=>i" where
- "relation_fm(r) ==
+ "relation_fm(r) \<equiv>
Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))"
lemma relation_type [TC]:
- "[| x \<in> nat |] ==> relation_fm(x) \<in> formula"
+ "\<lbrakk>x \<in> nat\<rbrakk> \<Longrightarrow> relation_fm(x) \<in> formula"
by (simp add: relation_fm_def)
lemma sats_relation_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, relation_fm(x), env) \<longleftrightarrow> is_relation(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, relation_fm(x), env) \<longleftrightarrow> is_relation(##A, nth(x,env))"
by (simp add: relation_fm_def is_relation_def)
lemma relation_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> is_relation(##A, x) \<longleftrightarrow> sats(A, relation_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_relation(##A, x) \<longleftrightarrow> sats(A, relation_fm(i), env)"
by simp
theorem is_relation_reflection:
@@ -943,12 +943,12 @@
subsubsection\<open>The Concept of Function, Internalized\<close>
-(* "is_function(M,r) ==
+(* "is_function(M,r) \<equiv>
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'" *)
definition
function_fm :: "i=>i" where
- "function_fm(r) ==
+ "function_fm(r) \<equiv>
Forall(Forall(Forall(Forall(Forall(
Implies(pair_fm(4,3,1),
Implies(pair_fm(4,2,0),
@@ -956,18 +956,18 @@
Implies(Member(0,r#+5), Equal(3,2))))))))))"
lemma function_type [TC]:
- "[| x \<in> nat |] ==> function_fm(x) \<in> formula"
+ "\<lbrakk>x \<in> nat\<rbrakk> \<Longrightarrow> function_fm(x) \<in> formula"
by (simp add: function_fm_def)
lemma sats_function_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, function_fm(x), env) \<longleftrightarrow> is_function(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, function_fm(x), env) \<longleftrightarrow> is_function(##A, nth(x,env))"
by (simp add: function_fm_def is_function_def)
lemma is_function_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> is_function(##A, x) \<longleftrightarrow> sats(A, function_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_function(##A, x) \<longleftrightarrow> sats(A, function_fm(i), env)"
by simp
theorem is_function_reflection:
@@ -980,13 +980,13 @@
subsubsection\<open>Typed Functions, Internalized\<close>
-(* "typed_function(M,A,B,r) ==
+(* "typed_function(M,A,B,r) \<equiv>
is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
(\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))" *)
definition
typed_function_fm :: "[i,i,i]=>i" where
- "typed_function_fm(A,B,r) ==
+ "typed_function_fm(A,B,r) \<equiv>
And(function_fm(r),
And(relation_fm(r),
And(domain_fm(r,A),
@@ -994,19 +994,19 @@
Forall(Forall(Implies(pair_fm(1,0,2),Member(0,B#+3)))))))))"
lemma typed_function_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> typed_function_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> typed_function_fm(x,y,z) \<in> formula"
by (simp add: typed_function_fm_def)
lemma sats_typed_function_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, typed_function_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, typed_function_fm(x,y,z), env) \<longleftrightarrow>
typed_function(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: typed_function_fm_def typed_function_def)
lemma typed_function_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> typed_function(##A, x, y, z) \<longleftrightarrow> sats(A, typed_function_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> typed_function(##A, x, y, z) \<longleftrightarrow> sats(A, typed_function_fm(i,j,k), env)"
by simp
lemmas function_reflections =
@@ -1039,14 +1039,14 @@
subsubsection\<open>Composition of Relations, Internalized\<close>
-(* "composition(M,r,s,t) ==
+(* "composition(M,r,s,t) \<equiv>
\<forall>p[M]. p \<in> t \<longleftrightarrow>
(\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
xy \<in> s & yz \<in> r)" *)
definition
composition_fm :: "[i,i,i]=>i" where
- "composition_fm(r,s,t) ==
+ "composition_fm(r,s,t) \<equiv>
Forall(Iff(Member(0,succ(t)),
Exists(Exists(Exists(Exists(Exists(
And(pair_fm(4,2,5),
@@ -1055,19 +1055,19 @@
And(Member(1,s#+6), Member(0,r#+6))))))))))))"
lemma composition_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> composition_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> composition_fm(x,y,z) \<in> formula"
by (simp add: composition_fm_def)
lemma sats_composition_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, composition_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, composition_fm(x,y,z), env) \<longleftrightarrow>
composition(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: composition_fm_def composition_def)
lemma composition_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> composition(##A, x, y, z) \<longleftrightarrow> sats(A, composition_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> composition(##A, x, y, z) \<longleftrightarrow> sats(A, composition_fm(i,j,k), env)"
by simp
theorem composition_reflection:
@@ -1080,13 +1080,13 @@
subsubsection\<open>Injections, Internalized\<close>
-(* "injection(M,A,B,f) ==
+(* "injection(M,A,B,f) \<equiv>
typed_function(M,A,B,f) &
(\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')" *)
definition
injection_fm :: "[i,i,i]=>i" where
- "injection_fm(A,B,f) ==
+ "injection_fm(A,B,f) \<equiv>
And(typed_function_fm(A,B,f),
Forall(Forall(Forall(Forall(Forall(
Implies(pair_fm(4,2,1),
@@ -1096,19 +1096,19 @@
lemma injection_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> injection_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> injection_fm(x,y,z) \<in> formula"
by (simp add: injection_fm_def)
lemma sats_injection_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, injection_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, injection_fm(x,y,z), env) \<longleftrightarrow>
injection(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: injection_fm_def injection_def)
lemma injection_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> injection(##A, x, y, z) \<longleftrightarrow> sats(A, injection_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> injection(##A, x, y, z) \<longleftrightarrow> sats(A, injection_fm(i,j,k), env)"
by simp
theorem injection_reflection:
@@ -1122,31 +1122,31 @@
subsubsection\<open>Surjections, Internalized\<close>
(* surjection :: "[i=>o,i,i,i] => o"
- "surjection(M,A,B,f) ==
+ "surjection(M,A,B,f) \<equiv>
typed_function(M,A,B,f) &
(\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
definition
surjection_fm :: "[i,i,i]=>i" where
- "surjection_fm(A,B,f) ==
+ "surjection_fm(A,B,f) \<equiv>
And(typed_function_fm(A,B,f),
Forall(Implies(Member(0,succ(B)),
Exists(And(Member(0,succ(succ(A))),
fun_apply_fm(succ(succ(f)),0,1))))))"
lemma surjection_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> surjection_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> surjection_fm(x,y,z) \<in> formula"
by (simp add: surjection_fm_def)
lemma sats_surjection_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, surjection_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, surjection_fm(x,y,z), env) \<longleftrightarrow>
surjection(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: surjection_fm_def surjection_def)
lemma surjection_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> surjection(##A, x, y, z) \<longleftrightarrow> sats(A, surjection_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> surjection(##A, x, y, z) \<longleftrightarrow> sats(A, surjection_fm(i,j,k), env)"
by simp
theorem surjection_reflection:
@@ -1161,25 +1161,25 @@
subsubsection\<open>Bijections, Internalized\<close>
(* bijection :: "[i=>o,i,i,i] => o"
- "bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)" *)
+ "bijection(M,A,B,f) \<equiv> injection(M,A,B,f) & surjection(M,A,B,f)" *)
definition
bijection_fm :: "[i,i,i]=>i" where
- "bijection_fm(A,B,f) == And(injection_fm(A,B,f), surjection_fm(A,B,f))"
+ "bijection_fm(A,B,f) \<equiv> And(injection_fm(A,B,f), surjection_fm(A,B,f))"
lemma bijection_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> bijection_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> bijection_fm(x,y,z) \<in> formula"
by (simp add: bijection_fm_def)
lemma sats_bijection_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, bijection_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, bijection_fm(x,y,z), env) \<longleftrightarrow>
bijection(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: bijection_fm_def bijection_def)
lemma bijection_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> bijection(##A, x, y, z) \<longleftrightarrow> sats(A, bijection_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> bijection(##A, x, y, z) \<longleftrightarrow> sats(A, bijection_fm(i,j,k), env)"
by simp
theorem bijection_reflection:
@@ -1193,30 +1193,30 @@
subsubsection\<open>Restriction of a Relation, Internalized\<close>
-(* "restriction(M,r,A,z) ==
+(* "restriction(M,r,A,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
definition
restriction_fm :: "[i,i,i]=>i" where
- "restriction_fm(r,A,z) ==
+ "restriction_fm(r,A,z) \<equiv>
Forall(Iff(Member(0,succ(z)),
And(Member(0,succ(r)),
Exists(And(Member(0,succ(succ(A))),
Exists(pair_fm(1,0,2)))))))"
lemma restriction_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> restriction_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> restriction_fm(x,y,z) \<in> formula"
by (simp add: restriction_fm_def)
lemma sats_restriction_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, restriction_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, restriction_fm(x,y,z), env) \<longleftrightarrow>
restriction(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: restriction_fm_def restriction_def)
lemma restriction_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> restriction(##A, x, y, z) \<longleftrightarrow> sats(A, restriction_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> restriction(##A, x, y, z) \<longleftrightarrow> sats(A, restriction_fm(i,j,k), env)"
by simp
theorem restriction_reflection:
@@ -1229,7 +1229,7 @@
subsubsection\<open>Order-Isomorphisms, Internalized\<close>
(* order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
- "order_isomorphism(M,A,r,B,s,f) ==
+ "order_isomorphism(M,A,r,B,s,f) \<equiv>
bijection(M,A,B,f) &
(\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
(\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
@@ -1239,7 +1239,7 @@
definition
order_isomorphism_fm :: "[i,i,i,i,i]=>i" where
- "order_isomorphism_fm(A,r,B,s,f) ==
+ "order_isomorphism_fm(A,r,B,s,f) \<equiv>
And(bijection_fm(A,B,f),
Forall(Implies(Member(0,succ(A)),
Forall(Implies(Member(0,succ(succ(A))),
@@ -1251,22 +1251,22 @@
Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))"
lemma order_isomorphism_type [TC]:
- "[| A \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat |]
- ==> order_isomorphism_fm(A,r,B,s,f) \<in> formula"
+ "\<lbrakk>A \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat\<rbrakk>
+ \<Longrightarrow> order_isomorphism_fm(A,r,B,s,f) \<in> formula"
by (simp add: order_isomorphism_fm_def)
lemma sats_order_isomorphism_fm [simp]:
- "[| U \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat; env \<in> list(A)|]
- ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) \<longleftrightarrow>
+ "\<lbrakk>U \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, order_isomorphism_fm(U,r,B,s,f), env) \<longleftrightarrow>
order_isomorphism(##A, nth(U,env), nth(r,env), nth(B,env),
nth(s,env), nth(f,env))"
by (simp add: order_isomorphism_fm_def order_isomorphism_def)
lemma order_isomorphism_iff_sats:
- "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s;
+ "\<lbrakk>nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s;
nth(k',env) = f;
- i \<in> nat; j \<in> nat; k \<in> nat; j' \<in> nat; k' \<in> nat; env \<in> list(A)|]
- ==> order_isomorphism(##A,U,r,B,s,f) \<longleftrightarrow>
+ i \<in> nat; j \<in> nat; k \<in> nat; j' \<in> nat; k' \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> order_isomorphism(##A,U,r,B,s,f) \<longleftrightarrow>
sats(A, order_isomorphism_fm(i,j,k,j',k'), env)"
by simp
@@ -1281,13 +1281,13 @@
text\<open>A limit ordinal is a non-empty, successor-closed ordinal\<close>
-(* "limit_ordinal(M,a) ==
- ordinal(M,a) & ~ empty(M,a) &
+(* "limit_ordinal(M,a) \<equiv>
+ ordinal(M,a) & \<not> empty(M,a) &
(\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
definition
limit_ordinal_fm :: "i=>i" where
- "limit_ordinal_fm(x) ==
+ "limit_ordinal_fm(x) \<equiv>
And(ordinal_fm(x),
And(Neg(empty_fm(x)),
Forall(Implies(Member(0,succ(x)),
@@ -1295,18 +1295,18 @@
succ_fm(1,0)))))))"
lemma limit_ordinal_type [TC]:
- "x \<in> nat ==> limit_ordinal_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> limit_ordinal_fm(x) \<in> formula"
by (simp add: limit_ordinal_fm_def)
lemma sats_limit_ordinal_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, limit_ordinal_fm(x), env) \<longleftrightarrow> limit_ordinal(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, limit_ordinal_fm(x), env) \<longleftrightarrow> limit_ordinal(##A, nth(x,env))"
by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm')
lemma limit_ordinal_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> limit_ordinal(##A, x) \<longleftrightarrow> sats(A, limit_ordinal_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> limit_ordinal(##A, x) \<longleftrightarrow> sats(A, limit_ordinal_fm(i), env)"
by simp
theorem limit_ordinal_reflection:
@@ -1319,30 +1319,30 @@
subsubsection\<open>Finite Ordinals: The Predicate ``Is A Natural Number''\<close>
-(* "finite_ordinal(M,a) ==
- ordinal(M,a) & ~ limit_ordinal(M,a) &
- (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))" *)
+(* "finite_ordinal(M,a) \<equiv>
+ ordinal(M,a) & \<not> limit_ordinal(M,a) &
+ (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))" *)
definition
finite_ordinal_fm :: "i=>i" where
- "finite_ordinal_fm(x) ==
+ "finite_ordinal_fm(x) \<equiv>
And(ordinal_fm(x),
And(Neg(limit_ordinal_fm(x)),
Forall(Implies(Member(0,succ(x)),
Neg(limit_ordinal_fm(0))))))"
lemma finite_ordinal_type [TC]:
- "x \<in> nat ==> finite_ordinal_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> finite_ordinal_fm(x) \<in> formula"
by (simp add: finite_ordinal_fm_def)
lemma sats_finite_ordinal_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, finite_ordinal_fm(x), env) \<longleftrightarrow> finite_ordinal(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, finite_ordinal_fm(x), env) \<longleftrightarrow> finite_ordinal(##A, nth(x,env))"
by (simp add: finite_ordinal_fm_def sats_ordinal_fm' finite_ordinal_def)
lemma finite_ordinal_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> finite_ordinal(##A, x) \<longleftrightarrow> sats(A, finite_ordinal_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> finite_ordinal(##A, x) \<longleftrightarrow> sats(A, finite_ordinal_fm(i), env)"
by simp
theorem finite_ordinal_reflection:
@@ -1355,27 +1355,27 @@
subsubsection\<open>Omega: The Set of Natural Numbers\<close>
-(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x)) *)
+(* omega(M,a) \<equiv> limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x)) *)
definition
omega_fm :: "i=>i" where
- "omega_fm(x) ==
+ "omega_fm(x) \<equiv>
And(limit_ordinal_fm(x),
Forall(Implies(Member(0,succ(x)),
Neg(limit_ordinal_fm(0)))))"
lemma omega_type [TC]:
- "x \<in> nat ==> omega_fm(x) \<in> formula"
+ "x \<in> nat \<Longrightarrow> omega_fm(x) \<in> formula"
by (simp add: omega_fm_def)
lemma sats_omega_fm [simp]:
- "[| x \<in> nat; env \<in> list(A)|]
- ==> sats(A, omega_fm(x), env) \<longleftrightarrow> omega(##A, nth(x,env))"
+ "\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, omega_fm(x), env) \<longleftrightarrow> omega(##A, nth(x,env))"
by (simp add: omega_fm_def omega_def)
lemma omega_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; env \<in> list(A)|]
- ==> omega(##A, x) \<longleftrightarrow> sats(A, omega_fm(i), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> omega(##A, x) \<longleftrightarrow> sats(A, omega_fm(i), env)"
by simp
theorem omega_reflection:
--- a/src/ZF/Constructible/MetaExists.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/MetaExists.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,9 +11,9 @@
definition
ex :: "(('a::{}) \<Rightarrow> prop) \<Rightarrow> prop" (binder \<open>\<Or>\<close> 0) where
- "ex(P) == (\<And>Q. (\<And>x. PROP P(x) \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q)"
+ "ex(P) \<equiv> (\<And>Q. (\<And>x. PROP P(x) \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q)"
-lemma meta_exI: "PROP P(x) ==> (\<Or>x. PROP P(x))"
+lemma meta_exI: "PROP P(x) \<Longrightarrow> (\<Or>x. PROP P(x))"
proof (unfold ex_def)
assume P: "PROP P(x)"
fix Q
@@ -21,7 +21,7 @@
from P show "PROP Q" by (rule PQ)
qed
-lemma meta_exE: "[|\<Or>x. PROP P(x); \<And>x. PROP P(x) ==> PROP R |] ==> PROP R"
+lemma meta_exE: "\<lbrakk>\<Or>x. PROP P(x); \<And>x. PROP P(x) \<Longrightarrow> PROP R\<rbrakk> \<Longrightarrow> PROP R"
proof (unfold ex_def)
assume QPQ: "\<And>Q. (\<And>x. PROP P(x) \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q"
assume PR: "\<And>x. PROP P(x) \<Longrightarrow> PROP R"
--- a/src/ZF/Constructible/Normal.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Normal.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,34 +19,34 @@
definition
Closed :: "(i=>o) => o" where
- "Closed(P) == \<forall>I. I \<noteq> 0 \<longrightarrow> (\<forall>i\<in>I. Ord(i) \<and> P(i)) \<longrightarrow> P(\<Union>(I))"
+ "Closed(P) \<equiv> \<forall>I. I \<noteq> 0 \<longrightarrow> (\<forall>i\<in>I. Ord(i) \<and> P(i)) \<longrightarrow> P(\<Union>(I))"
definition
Unbounded :: "(i=>o) => o" where
- "Unbounded(P) == \<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i<j \<and> P(j))"
+ "Unbounded(P) \<equiv> \<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i<j \<and> P(j))"
definition
Closed_Unbounded :: "(i=>o) => o" where
- "Closed_Unbounded(P) == Closed(P) \<and> Unbounded(P)"
+ "Closed_Unbounded(P) \<equiv> Closed(P) \<and> Unbounded(P)"
subsubsection\<open>Simple facts about c.u. classes\<close>
lemma ClosedI:
- "[| !!I. [| I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i) |] ==> P(\<Union>(I)) |]
- ==> Closed(P)"
+ "\<lbrakk>\<And>I. \<lbrakk>I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i)\<rbrakk> \<Longrightarrow> P(\<Union>(I))\<rbrakk>
+ \<Longrightarrow> Closed(P)"
by (simp add: Closed_def)
lemma ClosedD:
- "[| Closed(P); I \<noteq> 0; !!i. i\<in>I ==> Ord(i); !!i. i\<in>I ==> P(i) |]
- ==> P(\<Union>(I))"
+ "\<lbrakk>Closed(P); I \<noteq> 0; \<And>i. i\<in>I \<Longrightarrow> Ord(i); \<And>i. i\<in>I \<Longrightarrow> P(i)\<rbrakk>
+ \<Longrightarrow> P(\<Union>(I))"
by (simp add: Closed_def)
lemma UnboundedD:
- "[| Unbounded(P); Ord(i) |] ==> \<exists>j. i<j \<and> P(j)"
+ "\<lbrakk>Unbounded(P); Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. i<j \<and> P(j)"
by (simp add: Unbounded_def)
-lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) ==> Unbounded(C)"
+lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) \<Longrightarrow> Unbounded(C)"
by (simp add: Closed_Unbounded_def)
@@ -82,11 +82,11 @@
fixes P and A
fixes next_greater \<comment> \<open>the next ordinal satisfying class \<^term>\<open>A\<close>\<close>
fixes sup_greater \<comment> \<open>sup of those ordinals over all \<^term>\<open>A\<close>\<close>
- assumes closed: "a\<in>A ==> Closed(P(a))"
- and unbounded: "a\<in>A ==> Unbounded(P(a))"
+ assumes closed: "a\<in>A \<Longrightarrow> Closed(P(a))"
+ and unbounded: "a\<in>A \<Longrightarrow> Unbounded(P(a))"
and A_non0: "A\<noteq>0"
- defines "next_greater(a,x) == \<mu> y. x<y \<and> P(a,y)"
- and "sup_greater(x) == \<Union>a\<in>A. next_greater(a,x)"
+ defines "next_greater(a,x) \<equiv> \<mu> y. x<y \<and> P(a,y)"
+ and "sup_greater(x) \<equiv> \<Union>a\<in>A. next_greater(a,x)"
text\<open>Trivial that the intersection is closed.\<close>
@@ -106,7 +106,7 @@
text\<open>\<^term>\<open>next_greater\<close> works as expected: it returns a larger value
and one that belongs to class \<^term>\<open>P(a)\<close>.\<close>
lemma (in cub_family) next_greater_lemma:
- "[| Ord(x); a\<in>A |] ==> P(a, next_greater(a,x)) \<and> x < next_greater(a,x)"
+ "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x)) \<and> x < next_greater(a,x)"
apply (simp add: next_greater_def)
apply (rule exE [OF UnboundedD [OF unbounded]])
apply assumption+
@@ -114,29 +114,29 @@
done
lemma (in cub_family) next_greater_in_P:
- "[| Ord(x); a\<in>A |] ==> P(a, next_greater(a,x))"
+ "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x))"
by (blast dest: next_greater_lemma)
lemma (in cub_family) next_greater_gt:
- "[| Ord(x); a\<in>A |] ==> x < next_greater(a,x)"
+ "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> x < next_greater(a,x)"
by (blast dest: next_greater_lemma)
lemma (in cub_family) sup_greater_gt:
- "Ord(x) ==> x < sup_greater(x)"
+ "Ord(x) \<Longrightarrow> x < sup_greater(x)"
apply (simp add: sup_greater_def)
apply (insert A_non0)
apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater)
done
lemma (in cub_family) next_greater_le_sup_greater:
- "a\<in>A ==> next_greater(a,x) \<le> sup_greater(x)"
+ "a\<in>A \<Longrightarrow> next_greater(a,x) \<le> sup_greater(x)"
apply (simp add: sup_greater_def)
apply (blast intro: UN_upper_le Ord_next_greater)
done
lemma (in cub_family) omega_sup_greater_eq_UN:
- "[| Ord(x); a\<in>A |]
- ==> sup_greater^\<omega> (x) =
+ "\<lbrakk>Ord(x); a\<in>A\<rbrakk>
+ \<Longrightarrow> sup_greater^\<omega> (x) =
(\<Union>n\<in>nat. next_greater(a, sup_greater^n (x)))"
apply (simp add: iterates_omega_def)
apply (rule le_anti_sym)
@@ -153,7 +153,7 @@
done
lemma (in cub_family) P_omega_sup_greater:
- "[| Ord(x); a\<in>A |] ==> P(a, sup_greater^\<omega> (x))"
+ "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, sup_greater^\<omega> (x))"
apply (simp add: omega_sup_greater_eq_UN)
apply (rule ClosedD [OF closed])
apply (blast intro: ltD, auto)
@@ -162,7 +162,7 @@
done
lemma (in cub_family) omega_sup_greater_gt:
- "Ord(x) ==> x < sup_greater^\<omega> (x)"
+ "Ord(x) \<Longrightarrow> x < sup_greater^\<omega> (x)"
apply (simp add: iterates_omega_def)
apply (rule UN_upper_lt [of 1], simp_all)
apply (blast intro: sup_greater_gt)
@@ -180,8 +180,8 @@
theorem Closed_Unbounded_INT:
- "(!!a. a\<in>A ==> Closed_Unbounded(P(a)))
- ==> Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a, x))"
+ "(\<And>a. a\<in>A \<Longrightarrow> Closed_Unbounded(P(a)))
+ \<Longrightarrow> Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a, x))"
apply (case_tac "A=0", simp)
apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro])
apply (simp_all add: Closed_Unbounded_def)
@@ -192,8 +192,8 @@
by auto
theorem Closed_Unbounded_Int:
- "[| Closed_Unbounded(P); Closed_Unbounded(Q) |]
- ==> Closed_Unbounded(\<lambda>x. P(x) \<and> Q(x))"
+ "\<lbrakk>Closed_Unbounded(P); Closed_Unbounded(Q)\<rbrakk>
+ \<Longrightarrow> Closed_Unbounded(\<lambda>x. P(x) \<and> Q(x))"
apply (simp only: Int_iff_INT2)
apply (rule Closed_Unbounded_INT, auto)
done
@@ -203,43 +203,43 @@
definition
mono_le_subset :: "(i=>i) => o" where
- "mono_le_subset(M) == \<forall>i j. i\<le>j \<longrightarrow> M(i) \<subseteq> M(j)"
+ "mono_le_subset(M) \<equiv> \<forall>i j. i\<le>j \<longrightarrow> M(i) \<subseteq> M(j)"
definition
mono_Ord :: "(i=>i) => o" where
- "mono_Ord(F) == \<forall>i j. i<j \<longrightarrow> F(i) < F(j)"
+ "mono_Ord(F) \<equiv> \<forall>i j. i<j \<longrightarrow> F(i) < F(j)"
definition
cont_Ord :: "(i=>i) => o" where
- "cont_Ord(F) == \<forall>l. Limit(l) \<longrightarrow> F(l) = (\<Union>i<l. F(i))"
+ "cont_Ord(F) \<equiv> \<forall>l. Limit(l) \<longrightarrow> F(l) = (\<Union>i<l. F(i))"
definition
Normal :: "(i=>i) => o" where
- "Normal(F) == mono_Ord(F) \<and> cont_Ord(F)"
+ "Normal(F) \<equiv> mono_Ord(F) \<and> cont_Ord(F)"
subsubsection\<open>Immediate properties of the definitions\<close>
lemma NormalI:
- "[|!!i j. i<j ==> F(i) < F(j); !!l. Limit(l) ==> F(l) = (\<Union>i<l. F(i))|]
- ==> Normal(F)"
+ "\<lbrakk>\<And>i j. i<j \<Longrightarrow> F(i) < F(j); \<And>l. Limit(l) \<Longrightarrow> F(l) = (\<Union>i<l. F(i))\<rbrakk>
+ \<Longrightarrow> Normal(F)"
by (simp add: Normal_def mono_Ord_def cont_Ord_def)
-lemma mono_Ord_imp_Ord: "[| Ord(i); mono_Ord(F) |] ==> Ord(F(i))"
+lemma mono_Ord_imp_Ord: "\<lbrakk>Ord(i); mono_Ord(F)\<rbrakk> \<Longrightarrow> Ord(F(i))"
apply (auto simp add: mono_Ord_def)
apply (blast intro: lt_Ord)
done
-lemma mono_Ord_imp_mono: "[| i<j; mono_Ord(F) |] ==> F(i) < F(j)"
+lemma mono_Ord_imp_mono: "\<lbrakk>i<j; mono_Ord(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)"
by (simp add: mono_Ord_def)
-lemma Normal_imp_Ord [simp]: "[| Normal(F); Ord(i) |] ==> Ord(F(i))"
+lemma Normal_imp_Ord [simp]: "\<lbrakk>Normal(F); Ord(i)\<rbrakk> \<Longrightarrow> Ord(F(i))"
by (simp add: Normal_def mono_Ord_imp_Ord)
-lemma Normal_imp_cont: "[| Normal(F); Limit(l) |] ==> F(l) = (\<Union>i<l. F(i))"
+lemma Normal_imp_cont: "\<lbrakk>Normal(F); Limit(l)\<rbrakk> \<Longrightarrow> F(l) = (\<Union>i<l. F(i))"
by (simp add: Normal_def cont_Ord_def)
-lemma Normal_imp_mono: "[| i<j; Normal(F) |] ==> F(i) < F(j)"
+lemma Normal_imp_mono: "\<lbrakk>i<j; Normal(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)"
by (simp add: Normal_def mono_Ord_def)
lemma Normal_increasing:
@@ -271,7 +271,7 @@
text\<open>The proof is from Drake, pages 113--114.\<close>
-lemma mono_Ord_imp_le_subset: "mono_Ord(F) ==> mono_le_subset(F)"
+lemma mono_Ord_imp_le_subset: "mono_Ord(F) \<Longrightarrow> mono_le_subset(F)"
apply (simp add: mono_le_subset_def, clarify)
apply (subgoal_tac "F(i)\<le>F(j)", blast dest: le_imp_subset)
apply (simp add: le_iff)
@@ -280,8 +280,8 @@
text\<open>The following equation is taken for granted in any set theory text.\<close>
lemma cont_Ord_Union:
- "[| cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x) |]
- ==> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
+ "\<lbrakk>cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x)\<rbrakk>
+ \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
apply (frule Ord_set_cases)
apply (erule disjE, force)
apply (thin_tac "X=0 \<longrightarrow> Q" for Q, auto)
@@ -306,12 +306,12 @@
done
lemma Normal_Union:
- "[| X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F) |] ==> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
+ "\<lbrakk>X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F)\<rbrakk> \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
apply (simp add: Normal_def)
apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union)
done
-lemma Normal_imp_fp_Closed: "Normal(F) ==> Closed(\<lambda>i. F(i) = i)"
+lemma Normal_imp_fp_Closed: "Normal(F) \<Longrightarrow> Closed(\<lambda>i. F(i) = i)"
apply (simp add: Closed_def ball_conj_distrib, clarify)
apply (frule Ord_set_cases)
apply (auto simp add: Normal_Union)
@@ -319,19 +319,19 @@
lemma iterates_Normal_increasing:
- "[| n\<in>nat; x < F(x); Normal(F) |]
- ==> F^n (x) < F^(succ(n)) (x)"
+ "\<lbrakk>n\<in>nat; x < F(x); Normal(F)\<rbrakk>
+ \<Longrightarrow> F^n (x) < F^(succ(n)) (x)"
apply (induct n rule: nat_induct)
apply (simp_all add: Normal_imp_mono)
done
lemma Ord_iterates_Normal:
- "[| n\<in>nat; Normal(F); Ord(x) |] ==> Ord(F^n (x))"
+ "\<lbrakk>n\<in>nat; Normal(F); Ord(x)\<rbrakk> \<Longrightarrow> Ord(F^n (x))"
by (simp)
text\<open>THIS RESULT IS UNUSED\<close>
lemma iterates_omega_Limit:
- "[| Normal(F); x < F(x) |] ==> Limit(F^\<omega> (x))"
+ "\<lbrakk>Normal(F); x < F(x)\<rbrakk> \<Longrightarrow> Limit(F^\<omega> (x))"
apply (frule lt_Ord)
apply (simp add: iterates_omega_def)
apply (rule increasing_LimitI)
@@ -347,7 +347,7 @@
done
lemma iterates_omega_fixedpoint:
- "[| Normal(F); Ord(a) |] ==> F(F^\<omega> (a)) = F^\<omega> (a)"
+ "\<lbrakk>Normal(F); Ord(a)\<rbrakk> \<Longrightarrow> F(F^\<omega> (a)) = F^\<omega> (a)"
apply (frule Normal_increasing, assumption)
apply (erule leE)
apply (simp_all add: iterates_omega_triv [OF sym]) (*for subgoal 2*)
@@ -363,12 +363,12 @@
done
lemma iterates_omega_increasing:
- "[| Normal(F); Ord(a) |] ==> a \<le> F^\<omega> (a)"
+ "\<lbrakk>Normal(F); Ord(a)\<rbrakk> \<Longrightarrow> a \<le> F^\<omega> (a)"
apply (unfold iterates_omega_def)
apply (rule UN_upper_le [of 0], simp_all)
done
-lemma Normal_imp_fp_Unbounded: "Normal(F) ==> Unbounded(\<lambda>i. F(i) = i)"
+lemma Normal_imp_fp_Unbounded: "Normal(F) \<Longrightarrow> Unbounded(\<lambda>i. F(i) = i)"
apply (unfold Unbounded_def, clarify)
apply (rule_tac x="F^\<omega> (succ(i))" in exI)
apply (simp add: iterates_omega_fixedpoint)
@@ -377,7 +377,7 @@
theorem Normal_imp_fp_Closed_Unbounded:
- "Normal(F) ==> Closed_Unbounded(\<lambda>i. F(i) = i)"
+ "Normal(F) \<Longrightarrow> Closed_Unbounded(\<lambda>i. F(i) = i)"
by (simp add: Closed_Unbounded_def Normal_imp_fp_Closed
Normal_imp_fp_Unbounded)
@@ -391,17 +391,17 @@
\<close>
definition
normalize :: "[i=>i, i] => i" where
- "normalize(F,a) == transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))"
+ "normalize(F,a) \<equiv> transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))"
lemma Ord_normalize [simp, intro]:
- "[| Ord(a); !!x. Ord(x) ==> Ord(F(x)) |] ==> Ord(normalize(F, a))"
+ "\<lbrakk>Ord(a); \<And>x. Ord(x) \<Longrightarrow> Ord(F(x))\<rbrakk> \<Longrightarrow> Ord(normalize(F, a))"
apply (induct a rule: trans_induct3)
apply (simp_all add: ltD def_transrec2 [OF normalize_def])
done
lemma normalize_increasing:
- assumes ab: "a < b" and F: "!!x. Ord(x) ==> Ord(F(x))"
+ assumes ab: "a < b" and F: "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))"
shows "normalize(F,a) < normalize(F,b)"
proof -
{ fix x
@@ -428,14 +428,14 @@
qed
theorem Normal_normalize:
- "(!!x. Ord(x) ==> Ord(F(x))) ==> Normal(normalize(F))"
+ "(\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))) \<Longrightarrow> Normal(normalize(F))"
apply (rule NormalI)
apply (blast intro!: normalize_increasing)
apply (simp add: def_transrec2 [OF normalize_def])
done
theorem le_normalize:
- assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "!!x. Ord(x) ==> Ord(F(x))"
+ assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))"
shows "F(a) \<le> normalize(F,a)"
using a
proof (induct a rule: trans_induct3)
@@ -457,10 +457,10 @@
definition
Aleph :: "i => i" (\<open>\<aleph>_\<close> [90] 90) where
- "Aleph(a) == transrec2(a, nat, \<lambda>x r. csucc(r))"
+ "Aleph(a) \<equiv> transrec2(a, nat, \<lambda>x r. csucc(r))"
lemma Card_Aleph [simp, intro]:
- "Ord(a) ==> Card(Aleph(a))"
+ "Ord(a) \<Longrightarrow> Card(Aleph(a))"
apply (erule trans_induct3)
apply (simp_all add: Card_csucc Card_nat Card_is_Ord
def_transrec2 [OF Aleph_def])
--- a/src/ZF/Constructible/Rank.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Rank.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,24 +11,24 @@
locale M_ordertype = M_basic +
assumes well_ord_iso_separation:
- "[| M(A); M(f); M(r) |]
- ==> separation (M, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[M]. (\<exists>p[M].
+ "\<lbrakk>M(A); M(f); M(r)\<rbrakk>
+ \<Longrightarrow> separation (M, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[M]. (\<exists>p[M].
fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
and obase_separation:
\<comment> \<open>part of the order type formalization\<close>
- "[| M(A); M(r) |]
- ==> separation(M, \<lambda>a. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
+ "\<lbrakk>M(A); M(r)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>a. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
order_isomorphism(M,par,r,x,mx,g))"
and obase_equals_separation:
- "[| M(A); M(r) |]
- ==> separation (M, \<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[M]. \<exists>g[M].
+ "\<lbrakk>M(A); M(r)\<rbrakk>
+ \<Longrightarrow> separation (M, \<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[M]. \<exists>g[M].
ordinal(M,y) & (\<exists>my[M]. \<exists>pxr[M].
membership(M,y,my) & pred_set(M,A,x,r,pxr) &
order_isomorphism(M,pxr,r,y,my,g))))"
and omap_replacement:
- "[| M(A); M(r) |]
- ==> strong_replacement(M,
+ "\<lbrakk>M(A); M(r)\<rbrakk>
+ \<Longrightarrow> strong_replacement(M,
\<lambda>a z. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) &
pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
@@ -37,8 +37,8 @@
text\<open>Inductive argument for Kunen's Lemma I 6.1, etc.
Simple proof from Halmos, page 72\<close>
lemma (in M_ordertype) wellordered_iso_subset_lemma:
- "[| wellordered(M,A,r); f \<in> ord_iso(A,r, A',r); A'<= A; y \<in> A;
- M(A); M(f); M(r) |] ==> ~ <f`y, y> \<in> r"
+ "\<lbrakk>wellordered(M,A,r); f \<in> ord_iso(A,r, A',r); A'<= A; y \<in> A;
+ M(A); M(f); M(r)\<rbrakk> \<Longrightarrow> \<not> <f`y, y> \<in> r"
apply (unfold wellordered_def ord_iso_def)
apply (elim conjE CollectE)
apply (erule wellfounded_on_induct, assumption+)
@@ -51,12 +51,12 @@
text\<open>Kunen's Lemma I 6.1, page 14:
there's no order-isomorphism to an initial segment of a well-ordering\<close>
lemma (in M_ordertype) wellordered_iso_predD:
- "[| wellordered(M,A,r); f \<in> ord_iso(A, r, Order.pred(A,x,r), r);
- M(A); M(f); M(r) |] ==> x \<notin> A"
+ "\<lbrakk>wellordered(M,A,r); f \<in> ord_iso(A, r, Order.pred(A,x,r), r);
+ M(A); M(f); M(r)\<rbrakk> \<Longrightarrow> x \<notin> A"
apply (rule notI)
apply (frule wellordered_iso_subset_lemma, assumption)
apply (auto elim: predE)
-(*Now we know ~ (f`x < x) *)
+(*Now we know \<not> (f`x < x) *)
apply (drule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption)
(*Now we also know f`x \<in> pred(A,x,r); contradiction! *)
apply (simp add: Order.pred_def)
@@ -64,8 +64,8 @@
lemma (in M_ordertype) wellordered_iso_pred_eq_lemma:
- "[| f \<in> \<langle>Order.pred(A,y,r), r\<rangle> \<cong> \<langle>Order.pred(A,x,r), r\<rangle>;
- wellordered(M,A,r); x\<in>A; y\<in>A; M(A); M(f); M(r) |] ==> <x,y> \<notin> r"
+ "\<lbrakk>f \<in> \<langle>Order.pred(A,y,r), r\<rangle> \<cong> \<langle>Order.pred(A,x,r), r\<rangle>;
+ wellordered(M,A,r); x\<in>A; y\<in>A; M(A); M(f); M(r)\<rbrakk> \<Longrightarrow> <x,y> \<notin> r"
apply (frule wellordered_is_trans_on, assumption)
apply (rule notI)
apply (drule_tac x2=y and x=x and r2=r in
@@ -77,9 +77,9 @@
text\<open>Simple consequence of Lemma 6.1\<close>
lemma (in M_ordertype) wellordered_iso_pred_eq:
- "[| wellordered(M,A,r);
+ "\<lbrakk>wellordered(M,A,r);
f \<in> ord_iso(Order.pred(A,a,r), r, Order.pred(A,c,r), r);
- M(A); M(f); M(r); a\<in>A; c\<in>A |] ==> a=c"
+ M(A); M(f); M(r); a\<in>A; c\<in>A\<rbrakk> \<Longrightarrow> a=c"
apply (frule wellordered_is_trans_on, assumption)
apply (frule wellordered_is_linear, assumption)
apply (erule_tac x=a and y=c in linearE, auto)
@@ -95,11 +95,11 @@
text\<open>Can't use \<open>well_ord_iso_preserving\<close> because it needs the
strong premise \<^term>\<open>well_ord(A,r)\<close>\<close>
lemma (in M_ordertype) ord_iso_pred_imp_lt:
- "[| f \<in> ord_iso(Order.pred(A,x,r), r, i, Memrel(i));
+ "\<lbrakk>f \<in> ord_iso(Order.pred(A,x,r), r, i, Memrel(i));
g \<in> ord_iso(Order.pred(A,y,r), r, j, Memrel(j));
wellordered(M,A,r); x \<in> A; y \<in> A; M(A); M(r); M(f); M(g); M(j);
- Ord(i); Ord(j); \<langle>x,y\<rangle> \<in> r |]
- ==> i < j"
+ Ord(i); Ord(j); \<langle>x,y\<rangle> \<in> r\<rbrakk>
+ \<Longrightarrow> i < j"
apply (frule wellordered_is_trans_on, assumption)
apply (frule_tac y=y in transM, assumption)
apply (rule_tac i=i and j=j in Ord_linear_lt, auto)
@@ -127,8 +127,8 @@
lemma ord_iso_converse1:
- "[| f: ord_iso(A,r,B,s); <b, f`a>: s; a:A; b:B |]
- ==> <converse(f) ` b, a> \<in> r"
+ "\<lbrakk>f: ord_iso(A,r,B,s); <b, f`a>: s; a:A; b:B\<rbrakk>
+ \<Longrightarrow> <converse(f) ` b, a> \<in> r"
apply (frule ord_iso_converse, assumption+)
apply (blast intro: ord_iso_is_bij [THEN bij_is_fun, THEN apply_funtype])
apply (simp add: left_inverse_bij [OF ord_iso_is_bij])
@@ -138,28 +138,28 @@
definition
obase :: "[i=>o,i,i] => i" where
\<comment> \<open>the domain of \<open>om\<close>, eventually shown to equal \<open>A\<close>\<close>
- "obase(M,A,r) == {a\<in>A. \<exists>x[M]. \<exists>g[M]. Ord(x) &
+ "obase(M,A,r) \<equiv> {a\<in>A. \<exists>x[M]. \<exists>g[M]. Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}"
definition
omap :: "[i=>o,i,i,i] => o" where
\<comment> \<open>the function that maps wosets to order types\<close>
- "omap(M,A,r,f) ==
+ "omap(M,A,r,f) \<equiv>
\<forall>z[M].
z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
definition
otype :: "[i=>o,i,i,i] => o" where \<comment> \<open>the order types themselves\<close>
- "otype(M,A,r,i) == \<exists>f[M]. omap(M,A,r,f) & is_range(M,f,i)"
+ "otype(M,A,r,i) \<equiv> \<exists>f[M]. omap(M,A,r,f) & is_range(M,f,i)"
text\<open>Can also be proved with the premise \<^term>\<open>M(z)\<close> instead of
\<^term>\<open>M(f)\<close>, but that version is less useful. This lemma
is also more useful than the definition, \<open>omap_def\<close>.\<close>
lemma (in M_ordertype) omap_iff:
- "[| omap(M,A,r,f); M(A); M(f) |]
- ==> z \<in> f \<longleftrightarrow>
+ "\<lbrakk>omap(M,A,r,f); M(A); M(f)\<rbrakk>
+ \<Longrightarrow> z \<in> f \<longleftrightarrow>
(\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
apply (simp add: omap_def)
@@ -170,18 +170,18 @@
done
lemma (in M_ordertype) omap_unique:
- "[| omap(M,A,r,f); omap(M,A,r,f'); M(A); M(r); M(f); M(f') |] ==> f' = f"
+ "\<lbrakk>omap(M,A,r,f); omap(M,A,r,f'); M(A); M(r); M(f); M(f')\<rbrakk> \<Longrightarrow> f' = f"
apply (rule equality_iffI)
apply (simp add: omap_iff)
done
lemma (in M_ordertype) omap_yields_Ord:
- "[| omap(M,A,r,f); \<langle>a,x\<rangle> \<in> f; M(a); M(x) |] ==> Ord(x)"
+ "\<lbrakk>omap(M,A,r,f); \<langle>a,x\<rangle> \<in> f; M(a); M(x)\<rbrakk> \<Longrightarrow> Ord(x)"
by (simp add: omap_def)
lemma (in M_ordertype) otype_iff:
- "[| otype(M,A,r,i); M(A); M(r); M(i) |]
- ==> x \<in> i \<longleftrightarrow>
+ "\<lbrakk>otype(M,A,r,i); M(A); M(r); M(i)\<rbrakk>
+ \<Longrightarrow> x \<in> i \<longleftrightarrow>
(M(x) & Ord(x) &
(\<exists>a\<in>A. \<exists>g[M]. g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))))"
apply (auto simp add: omap_iff otype_def)
@@ -192,15 +192,15 @@
done
lemma (in M_ordertype) otype_eq_range:
- "[| omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i) |]
- ==> i = range(f)"
+ "\<lbrakk>omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\<rbrakk>
+ \<Longrightarrow> i = range(f)"
apply (auto simp add: otype_def omap_iff)
apply (blast dest: omap_unique)
done
lemma (in M_ordertype) Ord_otype:
- "[| otype(M,A,r,i); trans[A](r); M(A); M(r); M(i) |] ==> Ord(i)"
+ "\<lbrakk>otype(M,A,r,i); trans[A](r); M(A); M(r); M(i)\<rbrakk> \<Longrightarrow> Ord(i)"
apply (rule OrdI)
prefer 2
apply (simp add: Ord_def otype_def omap_def)
@@ -220,32 +220,32 @@
done
lemma (in M_ordertype) domain_omap:
- "[| omap(M,A,r,f); M(A); M(r); M(B); M(f) |]
- ==> domain(f) = obase(M,A,r)"
+ "\<lbrakk>omap(M,A,r,f); M(A); M(r); M(B); M(f)\<rbrakk>
+ \<Longrightarrow> domain(f) = obase(M,A,r)"
apply (simp add: obase_def)
apply (rule equality_iffI)
apply (simp add: domain_iff omap_iff, blast)
done
lemma (in M_ordertype) omap_subset:
- "[| omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(B); M(i) |] ==> f \<subseteq> obase(M,A,r) * i"
+ "\<lbrakk>omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(B); M(i)\<rbrakk> \<Longrightarrow> f \<subseteq> obase(M,A,r) * i"
apply clarify
apply (simp add: omap_iff obase_def)
apply (force simp add: otype_iff)
done
lemma (in M_ordertype) omap_funtype:
- "[| omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i) |] ==> f \<in> obase(M,A,r) -> i"
+ "\<lbrakk>omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i)\<rbrakk> \<Longrightarrow> f \<in> obase(M,A,r) -> i"
apply (simp add: domain_omap omap_subset Pi_iff function_def omap_iff)
apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans)
done
lemma (in M_ordertype) wellordered_omap_bij:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i) |] ==> f \<in> bij(obase(M,A,r),i)"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i)\<rbrakk> \<Longrightarrow> f \<in> bij(obase(M,A,r),i)"
apply (insert omap_funtype [of A r f i])
apply (auto simp add: bij_def inj_def)
prefer 2 apply (blast intro: fun_is_surj dest: otype_eq_range)
@@ -258,8 +258,8 @@
text\<open>This is not the final result: we must show \<^term>\<open>oB(A,r) = A\<close>\<close>
lemma (in M_ordertype) omap_ord_iso:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i) |] ==> f \<in> ord_iso(obase(M,A,r),r,i,Memrel(i))"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i)\<rbrakk> \<Longrightarrow> f \<in> ord_iso(obase(M,A,r),r,i,Memrel(i))"
apply (rule ord_isoI)
apply (erule wellordered_omap_bij, assumption+)
apply (insert omap_funtype [of A r f i], simp)
@@ -279,8 +279,8 @@
done
lemma (in M_ordertype) Ord_omap_image_pred:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i); b \<in> A |] ==> Ord(f `` Order.pred(A,b,r))"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i); b \<in> A\<rbrakk> \<Longrightarrow> Ord(f `` Order.pred(A,b,r))"
apply (frule wellordered_is_trans_on, assumption)
apply (rule OrdI)
prefer 2 apply (simp add: image_iff omap_iff Ord_def, blast)
@@ -304,9 +304,9 @@
done
lemma (in M_ordertype) restrict_omap_ord_iso:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- D \<subseteq> obase(M,A,r); M(A); M(r); M(f); M(i) |]
- ==> restrict(f,D) \<in> (\<langle>D,r\<rangle> \<cong> \<langle>f``D, Memrel(f``D)\<rangle>)"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ D \<subseteq> obase(M,A,r); M(A); M(r); M(f); M(i)\<rbrakk>
+ \<Longrightarrow> restrict(f,D) \<in> (\<langle>D,r\<rangle> \<cong> \<langle>f``D, Memrel(f``D)\<rangle>)"
apply (frule ord_iso_restrict_image [OF omap_ord_iso [of A r f i]],
assumption+)
apply (drule ord_iso_sym [THEN subset_ord_iso_Memrel])
@@ -315,8 +315,8 @@
done
lemma (in M_ordertype) obase_equals:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i) |] ==> obase(M,A,r) = A"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i)\<rbrakk> \<Longrightarrow> obase(M,A,r) = A"
apply (rule equalityI, force simp add: obase_def, clarify)
apply (unfold obase_def, simp)
apply (frule wellordered_is_wellfounded_on, assumption)
@@ -334,21 +334,21 @@
text\<open>Main result: \<^term>\<open>om\<close> gives the order-isomorphism
\<^term>\<open>\<langle>A,r\<rangle> \<cong> \<langle>i, Memrel(i)\<rangle>\<close>\<close>
theorem (in M_ordertype) omap_ord_iso_otype:
- "[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
- M(A); M(r); M(f); M(i) |] ==> f \<in> ord_iso(A, r, i, Memrel(i))"
+ "\<lbrakk>wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
+ M(A); M(r); M(f); M(i)\<rbrakk> \<Longrightarrow> f \<in> ord_iso(A, r, i, Memrel(i))"
apply (frule omap_ord_iso, assumption+)
apply (simp add: obase_equals)
done
lemma (in M_ordertype) obase_exists:
- "[| M(A); M(r) |] ==> M(obase(M,A,r))"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> M(obase(M,A,r))"
apply (simp add: obase_def)
apply (insert obase_separation [of A r])
apply (simp add: separation_def)
done
lemma (in M_ordertype) omap_exists:
- "[| M(A); M(r) |] ==> \<exists>z[M]. omap(M,A,r,z)"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> \<exists>z[M]. omap(M,A,r,z)"
apply (simp add: omap_def)
apply (insert omap_replacement [of A r])
apply (simp add: strong_replacement_def)
@@ -363,7 +363,7 @@
done
lemma (in M_ordertype) otype_exists:
- "[| wellordered(M,A,r); M(A); M(r) |] ==> \<exists>i[M]. otype(M,A,r,i)"
+ "\<lbrakk>wellordered(M,A,r); M(A); M(r)\<rbrakk> \<Longrightarrow> \<exists>i[M]. otype(M,A,r,i)"
apply (insert omap_exists [of A r])
apply (simp add: otype_def, safe)
apply (rule_tac x="range(x)" in rexI)
@@ -371,8 +371,8 @@
done
lemma (in M_ordertype) ordertype_exists:
- "[| wellordered(M,A,r); M(A); M(r) |]
- ==> \<exists>f[M]. (\<exists>i[M]. Ord(i) & f \<in> ord_iso(A, r, i, Memrel(i)))"
+ "\<lbrakk>wellordered(M,A,r); M(A); M(r)\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. (\<exists>i[M]. Ord(i) & f \<in> ord_iso(A, r, i, Memrel(i)))"
apply (insert obase_exists [of A r] omap_exists [of A r] otype_exists [of A r], simp, clarify)
apply (rename_tac i)
apply (subgoal_tac "Ord(i)", blast intro: omap_ord_iso_otype)
@@ -383,7 +383,7 @@
lemma (in M_ordertype) relativized_imp_well_ord:
- "[| wellordered(M,A,r); M(A); M(r) |] ==> well_ord(A,r)"
+ "\<lbrakk>wellordered(M,A,r); M(A); M(r)\<rbrakk> \<Longrightarrow> well_ord(A,r)"
apply (insert ordertype_exists [of A r], simp)
apply (blast intro: well_ord_ord_iso well_ord_Memrel)
done
@@ -392,14 +392,14 @@
text\<open>(a) The notion of Wellordering is absolute\<close>
theorem (in M_ordertype) well_ord_abs [simp]:
- "[| M(A); M(r) |] ==> wellordered(M,A,r) \<longleftrightarrow> well_ord(A,r)"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> wellordered(M,A,r) \<longleftrightarrow> well_ord(A,r)"
by (blast intro: well_ord_imp_relativized relativized_imp_well_ord)
text\<open>(b) Order types are absolute\<close>
theorem (in M_ordertype) ordertypes_are_absolute:
- "[| wellordered(M,A,r); f \<in> ord_iso(A, r, i, Memrel(i));
- M(A); M(r); M(f); M(i); Ord(i) |] ==> i = ordertype(A,r)"
+ "\<lbrakk>wellordered(M,A,r); f \<in> ord_iso(A, r, i, Memrel(i));
+ M(A); M(r); M(f); M(i); Ord(i)\<rbrakk> \<Longrightarrow> i = ordertype(A,r)"
by (blast intro: Ord_ordertype relativized_imp_well_ord ordertype_ord_iso
Ord_iso_implies_eq ord_iso_sym ord_iso_trans)
@@ -410,14 +410,14 @@
subsubsection\<open>Ordinal Addition\<close>
-(*FIXME: update to use new techniques!!*)
+(*FIXME: update to use new techniques\<And>*)
(*This expresses ordinal addition in the language of ZF. It also
provides an abbreviation that can be used in the instance of strong
replacement below. Here j is used to define the relation, namely
Memrel(succ(j)), while x determines the domain of f.*)
definition
is_oadd_fun :: "[i=>o,i,i,i,i] => o" where
- "is_oadd_fun(M,i,j,x,f) ==
+ "is_oadd_fun(M,i,j,x,f) \<equiv>
(\<forall>sj msj. M(sj) \<longrightarrow> M(msj) \<longrightarrow>
successor(M,j,sj) \<longrightarrow> membership(M,sj,msj) \<longrightarrow>
M_is_recfun(M,
@@ -426,10 +426,10 @@
definition
is_oadd :: "[i=>o,i,i,i] => o" where
- "is_oadd(M,i,j,k) ==
- (~ ordinal(M,i) & ~ ordinal(M,j) & k=0) |
- (~ ordinal(M,i) & ordinal(M,j) & k=j) |
- (ordinal(M,i) & ~ ordinal(M,j) & k=i) |
+ "is_oadd(M,i,j,k) \<equiv>
+ (\<not> ordinal(M,i) & \<not> ordinal(M,j) & k=0) |
+ (\<not> ordinal(M,i) & ordinal(M,j) & k=j) |
+ (ordinal(M,i) & \<not> ordinal(M,j) & k=i) |
(ordinal(M,i) & ordinal(M,j) &
(\<exists>f fj sj. M(f) & M(fj) & M(sj) &
successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) &
@@ -438,7 +438,7 @@
definition
(*NEEDS RELATIVIZATION*)
omult_eqns :: "[i,i,i,i] => o" where
- "omult_eqns(i,x,g,z) ==
+ "omult_eqns(i,x,g,z) \<equiv>
Ord(x) &
(x=0 \<longrightarrow> z=0) &
(\<forall>j. x = succ(j) \<longrightarrow> z = g`j ++ i) &
@@ -446,14 +446,14 @@
definition
is_omult_fun :: "[i=>o,i,i,i] => o" where
- "is_omult_fun(M,i,j,f) ==
+ "is_omult_fun(M,i,j,f) \<equiv>
(\<exists>df. M(df) & is_function(M,f) &
is_domain(M,f,df) & subset(M, j, df)) &
(\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
definition
is_omult :: "[i=>o,i,i,i] => o" where
- "is_omult(M,i,j,k) ==
+ "is_omult(M,i,j,k) \<equiv>
\<exists>f fj sj. M(f) & M(fj) & M(sj) &
successor(M,j,sj) & is_omult_fun(M,i,sj,f) &
fun_apply(M,f,j,fj) & fj = k"
@@ -461,14 +461,14 @@
locale M_ord_arith = M_ordertype +
assumes oadd_strong_replacement:
- "[| M(i); M(j) |] ==>
+ "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
strong_replacement(M,
\<lambda>x z. \<exists>y[M]. pair(M,x,y,z) &
(\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) &
image(M,f,x,fx) & y = i \<union> fx))"
and omult_strong_replacement':
- "[| M(i); M(j) |] ==>
+ "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
strong_replacement(M,
\<lambda>x z. \<exists>y[M]. z = <x,y> &
(\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) &
@@ -478,8 +478,8 @@
text\<open>\<open>is_oadd_fun\<close>: Relating the pure "language of set theory" to Isabelle/ZF\<close>
lemma (in M_ord_arith) is_oadd_fun_iff:
- "[| a\<le>j; M(i); M(j); M(a); M(f) |]
- ==> is_oadd_fun(M,i,j,a,f) \<longleftrightarrow>
+ "\<lbrakk>a\<le>j; M(i); M(j); M(a); M(f)\<rbrakk>
+ \<Longrightarrow> is_oadd_fun(M,i,j,a,f) \<longleftrightarrow>
f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) \<longrightarrow> x < a \<longrightarrow> f`x = i \<union> f``x)"
apply (frule lt_Ord)
apply (simp add: is_oadd_fun_def
@@ -492,7 +492,7 @@
lemma (in M_ord_arith) oadd_strong_replacement':
- "[| M(i); M(j) |] ==>
+ "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
strong_replacement(M,
\<lambda>x z. \<exists>y[M]. z = <x,y> &
(\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \<union> g``x,g) &
@@ -504,14 +504,14 @@
lemma (in M_ord_arith) exists_oadd:
- "[| Ord(j); M(i); M(j) |]
- ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i \<union> g``x, f)"
+ "\<lbrakk>Ord(j); M(i); M(j)\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i \<union> g``x, f)"
apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
apply (simp_all add: Memrel_type oadd_strong_replacement')
done
lemma (in M_ord_arith) exists_oadd_fun:
- "[| Ord(j); M(i); M(j) |] ==> \<exists>f[M]. is_oadd_fun(M,i,succ(j),succ(j),f)"
+ "\<lbrakk>Ord(j); M(i); M(j)\<rbrakk> \<Longrightarrow> \<exists>f[M]. is_oadd_fun(M,i,succ(j),succ(j),f)"
apply (rule exists_oadd [THEN rexE])
apply (erule Ord_succ, assumption, simp)
apply (rename_tac f)
@@ -522,8 +522,8 @@
done
lemma (in M_ord_arith) is_oadd_fun_apply:
- "[| x < j; M(i); M(j); M(f); is_oadd_fun(M,i,j,j,f) |]
- ==> f`x = i \<union> (\<Union>k\<in>x. {f ` k})"
+ "\<lbrakk>x < j; M(i); M(j); M(f); is_oadd_fun(M,i,j,j,f)\<rbrakk>
+ \<Longrightarrow> f`x = i \<union> (\<Union>k\<in>x. {f ` k})"
apply (simp add: is_oadd_fun_iff lt_Ord2, clarify)
apply (frule lt_closed, simp)
apply (frule leI [THEN le_imp_subset])
@@ -531,8 +531,8 @@
done
lemma (in M_ord_arith) is_oadd_fun_iff_oadd [rule_format]:
- "[| is_oadd_fun(M,i,J,J,f); M(i); M(J); M(f); Ord(i); Ord(j) |]
- ==> j<J \<longrightarrow> f`j = i++j"
+ "\<lbrakk>is_oadd_fun(M,i,J,J,f); M(i); M(J); M(f); Ord(i); Ord(j)\<rbrakk>
+ \<Longrightarrow> j<J \<longrightarrow> f`j = i++j"
apply (erule_tac i=j in trans_induct, clarify)
apply (subgoal_tac "\<forall>k\<in>x. k<J")
apply (simp (no_asm_simp) add: is_oadd_def oadd_unfold is_oadd_fun_apply)
@@ -540,20 +540,20 @@
done
lemma (in M_ord_arith) Ord_oadd_abs:
- "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
+ "\<lbrakk>M(i); M(j); M(k); Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
apply (simp add: is_oadd_def is_oadd_fun_iff_oadd)
apply (frule exists_oadd_fun [of j i], blast+)
done
lemma (in M_ord_arith) oadd_abs:
- "[| M(i); M(j); M(k) |] ==> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
+ "\<lbrakk>M(i); M(j); M(k)\<rbrakk> \<Longrightarrow> is_oadd(M,i,j,k) \<longleftrightarrow> k = i++j"
apply (case_tac "Ord(i) & Ord(j)")
apply (simp add: Ord_oadd_abs)
apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd)
done
lemma (in M_ord_arith) oadd_closed [intro,simp]:
- "[| M(i); M(j) |] ==> M(i++j)"
+ "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow> M(i++j)"
apply (simp add: oadd_eq_if_raw_oadd, clarify)
apply (simp add: raw_oadd_eq_oadd)
apply (frule exists_oadd_fun [of j i], auto)
@@ -564,7 +564,7 @@
subsubsection\<open>Ordinal Multiplication\<close>
lemma omult_eqns_unique:
- "[| omult_eqns(i,x,g,z); omult_eqns(i,x,g,z') |] ==> z=z'"
+ "\<lbrakk>omult_eqns(i,x,g,z); omult_eqns(i,x,g,z')\<rbrakk> \<Longrightarrow> z=z'"
apply (simp add: omult_eqns_def, clarify)
apply (erule Ord_cases, simp_all)
done
@@ -579,26 +579,26 @@
by (simp add: omult_eqns_def)
lemma the_omult_eqns_succ:
- "Ord(j) ==> (THE z. omult_eqns(i,succ(j),g,z)) = g`j ++ i"
+ "Ord(j) \<Longrightarrow> (THE z. omult_eqns(i,succ(j),g,z)) = g`j ++ i"
by (simp add: omult_eqns_succ)
lemma omult_eqns_Limit:
- "Limit(x) ==> omult_eqns(i,x,g,z) \<longleftrightarrow> z = \<Union>(g``x)"
+ "Limit(x) \<Longrightarrow> omult_eqns(i,x,g,z) \<longleftrightarrow> z = \<Union>(g``x)"
apply (simp add: omult_eqns_def)
apply (blast intro: Limit_is_Ord)
done
lemma the_omult_eqns_Limit:
- "Limit(x) ==> (THE z. omult_eqns(i,x,g,z)) = \<Union>(g``x)"
+ "Limit(x) \<Longrightarrow> (THE z. omult_eqns(i,x,g,z)) = \<Union>(g``x)"
by (simp add: omult_eqns_Limit)
-lemma omult_eqns_Not: "~ Ord(x) ==> ~ omult_eqns(i,x,g,z)"
+lemma omult_eqns_Not: "\<not> Ord(x) \<Longrightarrow> \<not> omult_eqns(i,x,g,z)"
by (simp add: omult_eqns_def)
lemma (in M_ord_arith) the_omult_eqns_closed:
- "[| M(i); M(x); M(g); function(g) |]
- ==> M(THE z. omult_eqns(i, x, g, z))"
+ "\<lbrakk>M(i); M(x); M(g); function(g)\<rbrakk>
+ \<Longrightarrow> M(THE z. omult_eqns(i, x, g, z))"
apply (case_tac "Ord(x)")
prefer 2 apply (simp add: omult_eqns_Not) \<comment> \<open>trivial, non-Ord case\<close>
apply (erule Ord_cases)
@@ -608,15 +608,15 @@
done
lemma (in M_ord_arith) exists_omult:
- "[| Ord(j); M(i); M(j) |]
- ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)"
+ "\<lbrakk>Ord(j); M(i); M(j)\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)"
apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
apply (simp_all add: Memrel_type omult_strong_replacement')
apply (blast intro: the_omult_eqns_closed)
done
lemma (in M_ord_arith) exists_omult_fun:
- "[| Ord(j); M(i); M(j) |] ==> \<exists>f[M]. is_omult_fun(M,i,succ(j),f)"
+ "\<lbrakk>Ord(j); M(i); M(j)\<rbrakk> \<Longrightarrow> \<exists>f[M]. is_omult_fun(M,i,succ(j),f)"
apply (rule exists_omult [THEN rexE])
apply (erule Ord_succ, assumption, simp)
apply (rename_tac f)
@@ -630,24 +630,24 @@
done
lemma (in M_ord_arith) is_omult_fun_apply_0:
- "[| 0 < j; is_omult_fun(M,i,j,f) |] ==> f`0 = 0"
+ "\<lbrakk>0 < j; is_omult_fun(M,i,j,f)\<rbrakk> \<Longrightarrow> f`0 = 0"
by (simp add: is_omult_fun_def omult_eqns_def lt_def ball_conj_distrib)
lemma (in M_ord_arith) is_omult_fun_apply_succ:
- "[| succ(x) < j; is_omult_fun(M,i,j,f) |] ==> f`succ(x) = f`x ++ i"
+ "\<lbrakk>succ(x) < j; is_omult_fun(M,i,j,f)\<rbrakk> \<Longrightarrow> f`succ(x) = f`x ++ i"
by (simp add: is_omult_fun_def omult_eqns_def lt_def, blast)
lemma (in M_ord_arith) is_omult_fun_apply_Limit:
- "[| x < j; Limit(x); M(j); M(f); is_omult_fun(M,i,j,f) |]
- ==> f ` x = (\<Union>y\<in>x. f`y)"
+ "\<lbrakk>x < j; Limit(x); M(j); M(f); is_omult_fun(M,i,j,f)\<rbrakk>
+ \<Longrightarrow> f ` x = (\<Union>y\<in>x. f`y)"
apply (simp add: is_omult_fun_def omult_eqns_def lt_def, clarify)
apply (drule subset_trans [OF OrdmemD], assumption+)
apply (simp add: ball_conj_distrib omult_Limit image_function)
done
lemma (in M_ord_arith) is_omult_fun_eq_omult:
- "[| is_omult_fun(M,i,J,f); M(J); M(f); Ord(i); Ord(j) |]
- ==> j<J \<longrightarrow> f`j = i**j"
+ "\<lbrakk>is_omult_fun(M,i,J,f); M(J); M(f); Ord(i); Ord(j)\<rbrakk>
+ \<Longrightarrow> j<J \<longrightarrow> f`j = i**j"
apply (erule_tac i=j in trans_induct3)
apply (safe del: impCE)
apply (simp add: is_omult_fun_apply_0)
@@ -660,7 +660,7 @@
done
lemma (in M_ord_arith) omult_abs:
- "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) \<longleftrightarrow> k = i**j"
+ "\<lbrakk>M(i); M(j); M(k); Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> is_omult(M,i,j,k) \<longleftrightarrow> k = i**j"
apply (simp add: is_omult_def is_omult_fun_eq_omult)
apply (frule exists_omult_fun [of j i], blast+)
done
@@ -675,22 +675,22 @@
locale M_wfrank = M_trancl +
assumes wfrank_separation:
- "M(r) ==>
+ "M(r) \<Longrightarrow>
separation (M, \<lambda>x.
\<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
- ~ (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))"
+ \<not> (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))"
and wfrank_strong_replacement:
- "M(r) ==>
+ "M(r) \<Longrightarrow>
strong_replacement(M, \<lambda>x z.
\<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
(\<exists>y[M]. \<exists>f[M]. pair(M,x,y,z) &
M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) &
is_range(M,f,y)))"
and Ord_wfrank_separation:
- "M(r) ==>
+ "M(r) \<Longrightarrow>
separation (M, \<lambda>x.
\<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
- ~ (\<forall>f[M]. \<forall>rangef[M].
+ \<not> (\<forall>f[M]. \<forall>rangef[M].
is_range(M,f,rangef) \<longrightarrow>
M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) \<longrightarrow>
ordinal(M,rangef)))"
@@ -700,15 +700,15 @@
agree with the "real" ones.\<close>
lemma (in M_wfrank) wfrank_separation':
- "M(r) ==>
+ "M(r) \<Longrightarrow>
separation
- (M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
+ (M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
apply (insert wfrank_separation [of r])
apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
done
lemma (in M_wfrank) wfrank_strong_replacement':
- "M(r) ==>
+ "M(r) \<Longrightarrow>
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>f[M].
pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) &
y = range(f))"
@@ -717,9 +717,9 @@
done
lemma (in M_wfrank) Ord_wfrank_separation':
- "M(r) ==>
+ "M(r) \<Longrightarrow>
separation (M, \<lambda>x.
- ~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) \<longrightarrow> Ord(range(f))))"
+ \<not> (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) \<longrightarrow> Ord(range(f))))"
apply (insert Ord_wfrank_separation [of r])
apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
done
@@ -728,14 +728,14 @@
well-founded relations within the class M.\<close>
definition
wellfoundedrank :: "[i=>o,i,i] => i" where
- "wellfoundedrank(M,r,A) ==
+ "wellfoundedrank(M,r,A) \<equiv>
{p. x\<in>A, \<exists>y[M]. \<exists>f[M].
p = <x,y> & is_recfun(r^+, x, %x f. range(f), f) &
y = range(f)}"
lemma (in M_wfrank) exists_wfrank:
- "[| wellfounded(M,r); M(a); M(r) |]
- ==> \<exists>f[M]. is_recfun(r^+, a, %x f. range(f), f)"
+ "\<lbrakk>wellfounded(M,r); M(a); M(r)\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. is_recfun(r^+, a, %x f. range(f), f)"
apply (rule wellfounded_exists_is_recfun)
apply (blast intro: wellfounded_trancl)
apply (rule trans_trancl)
@@ -745,7 +745,7 @@
done
lemma (in M_wfrank) M_wellfoundedrank:
- "[| wellfounded(M,r); M(r); M(A) |] ==> M(wellfoundedrank(M,r,A))"
+ "\<lbrakk>wellfounded(M,r); M(r); M(A)\<rbrakk> \<Longrightarrow> M(wellfoundedrank(M,r,A))"
apply (insert wfrank_strong_replacement' [of r])
apply (simp add: wellfoundedrank_def)
apply (rule strong_replacement_closed)
@@ -758,8 +758,8 @@
done
lemma (in M_wfrank) Ord_wfrank_range [rule_format]:
- "[| wellfounded(M,r); a\<in>A; M(r); M(A) |]
- ==> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) \<longrightarrow> Ord(range(f))"
+ "\<lbrakk>wellfounded(M,r); a\<in>A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) \<longrightarrow> Ord(range(f))"
apply (drule wellfounded_trancl, assumption)
apply (rule wellfounded_induct, assumption, erule (1) transM)
apply simp
@@ -790,8 +790,8 @@
done
lemma (in M_wfrank) Ord_range_wellfoundedrank:
- "[| wellfounded(M,r); r \<subseteq> A*A; M(r); M(A) |]
- ==> Ord (range(wellfoundedrank(M,r,A)))"
+ "\<lbrakk>wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> Ord (range(wellfoundedrank(M,r,A)))"
apply (frule wellfounded_trancl, assumption)
apply (frule trancl_subset_times)
apply (simp add: wellfoundedrank_def)
@@ -819,8 +819,8 @@
done
lemma (in M_wfrank) function_wellfoundedrank:
- "[| wellfounded(M,r); M(r); M(A)|]
- ==> function(wellfoundedrank(M,r,A))"
+ "\<lbrakk>wellfounded(M,r); M(r); M(A)\<rbrakk>
+ \<Longrightarrow> function(wellfoundedrank(M,r,A))"
apply (simp add: wellfoundedrank_def function_def, clarify)
txt\<open>Uniqueness: repeated below!\<close>
apply (drule is_recfun_functional, assumption)
@@ -829,8 +829,8 @@
done
lemma (in M_wfrank) domain_wellfoundedrank:
- "[| wellfounded(M,r); M(r); M(A)|]
- ==> domain(wellfoundedrank(M,r,A)) = A"
+ "\<lbrakk>wellfounded(M,r); M(r); M(A)\<rbrakk>
+ \<Longrightarrow> domain(wellfoundedrank(M,r,A)) = A"
apply (simp add: wellfoundedrank_def function_def)
apply (rule equalityI, auto)
apply (frule transM, assumption)
@@ -847,8 +847,8 @@
done
lemma (in M_wfrank) wellfoundedrank_type:
- "[| wellfounded(M,r); M(r); M(A)|]
- ==> wellfoundedrank(M,r,A) \<in> A -> range(wellfoundedrank(M,r,A))"
+ "\<lbrakk>wellfounded(M,r); M(r); M(A)\<rbrakk>
+ \<Longrightarrow> wellfoundedrank(M,r,A) \<in> A -> range(wellfoundedrank(M,r,A))"
apply (frule function_wellfoundedrank [of r A], assumption+)
apply (frule function_imp_Pi)
apply (simp add: wellfoundedrank_def relation_def)
@@ -857,15 +857,15 @@
done
lemma (in M_wfrank) Ord_wellfoundedrank:
- "[| wellfounded(M,r); a \<in> A; r \<subseteq> A*A; M(r); M(A) |]
- ==> Ord(wellfoundedrank(M,r,A) ` a)"
+ "\<lbrakk>wellfounded(M,r); a \<in> A; r \<subseteq> A*A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> Ord(wellfoundedrank(M,r,A) ` a)"
by (blast intro: apply_funtype [OF wellfoundedrank_type]
Ord_in_Ord [OF Ord_range_wellfoundedrank])
lemma (in M_wfrank) wellfoundedrank_eq:
- "[| is_recfun(r^+, a, %x. range, f);
- wellfounded(M,r); a \<in> A; M(f); M(r); M(A)|]
- ==> wellfoundedrank(M,r,A) ` a = range(f)"
+ "\<lbrakk>is_recfun(r^+, a, %x. range, f);
+ wellfounded(M,r); a \<in> A; M(f); M(r); M(A)\<rbrakk>
+ \<Longrightarrow> wellfoundedrank(M,r,A) ` a = range(f)"
apply (rule apply_equality)
prefer 2 apply (blast intro: wellfoundedrank_type)
apply (simp add: wellfoundedrank_def)
@@ -882,9 +882,9 @@
lemma (in M_wfrank) wellfoundedrank_lt:
- "[| <a,b> \<in> r;
- wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|]
- ==> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b"
+ "\<lbrakk><a,b> \<in> r;
+ wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b"
apply (frule wellfounded_trancl, assumption)
apply (subgoal_tac "a\<in>A & b\<in>A")
prefer 2 apply blast
@@ -908,8 +908,8 @@
lemma (in M_wfrank) wellfounded_imp_subset_rvimage:
- "[|wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|]
- ==> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
+ "\<lbrakk>wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI)
apply (rule_tac x="wellfoundedrank(M,r,A)" in exI)
apply (simp add: Ord_range_wellfoundedrank, clarify)
@@ -919,12 +919,12 @@
done
lemma (in M_wfrank) wellfounded_imp_wf:
- "[|wellfounded(M,r); relation(r); M(r)|] ==> wf(r)"
+ "\<lbrakk>wellfounded(M,r); relation(r); M(r)\<rbrakk> \<Longrightarrow> wf(r)"
by (blast dest!: relation_field_times_field wellfounded_imp_subset_rvimage
intro: wf_rvimage_Ord [THEN wf_subset])
lemma (in M_wfrank) wellfounded_on_imp_wf_on:
- "[|wellfounded_on(M,A,r); relation(r); M(r); M(A)|] ==> wf[A](r)"
+ "\<lbrakk>wellfounded_on(M,A,r); relation(r); M(r); M(A)\<rbrakk> \<Longrightarrow> wf[A](r)"
apply (simp add: wellfounded_on_iff_wellfounded wf_on_def)
apply (rule wellfounded_imp_wf)
apply (simp_all add: relation_def)
@@ -932,11 +932,11 @@
theorem (in M_wfrank) wf_abs:
- "[|relation(r); M(r)|] ==> wellfounded(M,r) \<longleftrightarrow> wf(r)"
+ "\<lbrakk>relation(r); M(r)\<rbrakk> \<Longrightarrow> wellfounded(M,r) \<longleftrightarrow> wf(r)"
by (blast intro: wellfounded_imp_wf wf_imp_relativized)
theorem (in M_wfrank) wf_on_abs:
- "[|relation(r); M(r); M(A)|] ==> wellfounded_on(M,A,r) \<longleftrightarrow> wf[A](r)"
+ "\<lbrakk>relation(r); M(r); M(A)\<rbrakk> \<Longrightarrow> wellfounded_on(M,A,r) \<longleftrightarrow> wf[A](r)"
by (blast intro: wellfounded_on_imp_wf_on wf_on_imp_relativized)
end
--- a/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 16:51:35 2022 +0100
@@ -24,8 +24,8 @@
by (intro FOL_reflections function_reflections)
lemma well_ord_iso_separation:
- "[| L(A); L(f); L(r) |]
- ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[L]. (\<exists>p[L].
+ "\<lbrakk>L(A); L(f); L(r)\<rbrakk>
+ \<Longrightarrow> separation (L, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[L]. (\<exists>p[L].
fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"],
auto)
@@ -47,8 +47,8 @@
lemma obase_separation:
\<comment> \<open>part of the order type formalization\<close>
- "[| L(A); L(r) |]
- ==> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
+ "\<lbrakk>L(A); L(r)\<rbrakk>
+ \<Longrightarrow> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
order_isomorphism(L,par,r,x,mx,g))"
apply (rule gen_separation_multi [OF obase_reflects, of "{A,r}"], auto)
@@ -60,19 +60,19 @@
subsubsection\<open>Separation for a Theorem about \<^term>\<open>obase\<close>\<close>
lemma obase_equals_reflects:
- "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
+ "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[L]. \<exists>g[L].
ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
membership(L,y,my) & pred_set(L,A,x,r,pxr) &
order_isomorphism(L,pxr,r,y,my,g))),
- \<lambda>i x. x\<in>A \<longrightarrow> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
+ \<lambda>i x. x\<in>A \<longrightarrow> \<not>(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
ordinal(##Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) &
order_isomorphism(##Lset(i),pxr,r,y,my,g)))]"
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma obase_equals_separation:
- "[| L(A); L(r) |]
- ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
+ "\<lbrakk>L(A); L(r)\<rbrakk>
+ \<Longrightarrow> separation (L, \<lambda>x. x\<in>A \<longrightarrow> \<not>(\<exists>y[L]. \<exists>g[L].
ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
membership(L,y,my) & pred_set(L,A,x,r,pxr) &
order_isomorphism(L,pxr,r,y,my,g))))"
@@ -96,8 +96,8 @@
by (intro FOL_reflections function_reflections fun_plus_reflections)
lemma omap_replacement:
- "[| L(A); L(r) |]
- ==> strong_replacement(L,
+ "\<lbrakk>L(A); L(r)\<rbrakk>
+ \<Longrightarrow> strong_replacement(L,
\<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
@@ -133,17 +133,17 @@
lemma wfrank_Reflects:
"REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
- ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
+ \<not> (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
\<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
- ~ (\<exists>f \<in> Lset(i).
+ \<not> (\<exists>f \<in> Lset(i).
M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y),
rplus, x, f))]"
by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection)
lemma wfrank_separation:
- "L(r) ==>
+ "L(r) \<Longrightarrow>
separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
- ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
+ \<not> (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
apply (rule gen_separation [OF wfrank_Reflects], simp)
apply (rule_tac env="[r]" in DPow_LsetI)
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
@@ -167,7 +167,7 @@
is_recfun_reflection tran_closure_reflection)
lemma wfrank_strong_replacement:
- "L(r) ==>
+ "L(r) \<Longrightarrow>
strong_replacement(L, \<lambda>x z.
\<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
(\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) &
@@ -186,12 +186,12 @@
lemma Ord_wfrank_Reflects:
"REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
- ~ (\<forall>f[L]. \<forall>rangef[L].
+ \<not> (\<forall>f[L]. \<forall>rangef[L].
is_range(L,f,rangef) \<longrightarrow>
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
ordinal(L,rangef)),
\<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
- ~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
+ \<not> (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
is_range(##Lset(i),f,rangef) \<longrightarrow>
M_is_recfun(##Lset(i), \<lambda>x f y. is_range(##Lset(i),f,y),
rplus, x, f) \<longrightarrow>
@@ -200,10 +200,10 @@
tran_closure_reflection ordinal_reflection)
lemma Ord_wfrank_separation:
- "L(r) ==>
+ "L(r) \<Longrightarrow>
separation (L, \<lambda>x.
\<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
- ~ (\<forall>f[L]. \<forall>rangef[L].
+ \<not> (\<forall>f[L]. \<forall>rangef[L].
is_range(L,f,rangef) \<longrightarrow>
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
ordinal(L,rangef)))"
--- a/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 16:51:35 2022 +0100
@@ -8,7 +8,7 @@
text\<open>This theory proves all instances needed for locales \<open>M_trancl\<close> and \<open>M_datatypes\<close>\<close>
-lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> j<i"
+lemma eq_succ_imp_lt: "\<lbrakk>i = succ(j); Ord(i)\<rbrakk> \<Longrightarrow> j<i"
by simp
@@ -18,7 +18,7 @@
text\<open>First, The Defining Formula\<close>
-(* "rtran_closure_mem(M,A,r,p) ==
+(* "rtran_closure_mem(M,A,r,p) \<equiv>
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
(\<exists>f[M]. typed_function(M,n',A,f) &
@@ -30,7 +30,7 @@
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
definition
rtran_closure_mem_fm :: "[i,i,i]=>i" where
- "rtran_closure_mem_fm(A,r,p) ==
+ "rtran_closure_mem_fm(A,r,p) \<equiv>
Exists(Exists(Exists(
And(omega_fm(2),
And(Member(1,2),
@@ -49,19 +49,19 @@
lemma rtran_closure_mem_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> rtran_closure_mem_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> rtran_closure_mem_fm(x,y,z) \<in> formula"
by (simp add: rtran_closure_mem_fm_def)
lemma sats_rtran_closure_mem_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, rtran_closure_mem_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, rtran_closure_mem_fm(x,y,z), env) \<longleftrightarrow>
rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def)
lemma rtran_closure_mem_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> rtran_closure_mem(##A, x, y, z) \<longleftrightarrow> sats(A, rtran_closure_mem_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> rtran_closure_mem(##A, x, y, z) \<longleftrightarrow> sats(A, rtran_closure_mem_fm(i,j,k), env)"
by (simp)
lemma rtran_closure_mem_reflection:
@@ -73,7 +73,7 @@
text\<open>Separation for \<^term>\<open>rtrancl(r)\<close>.\<close>
lemma rtrancl_separation:
- "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))"
+ "\<lbrakk>L(r); L(A)\<rbrakk> \<Longrightarrow> separation (L, rtran_closure_mem(L,A,r))"
apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"],
auto)
apply (rule_tac env="[r,A]" in DPow_LsetI)
@@ -83,30 +83,30 @@
subsubsection\<open>Reflexive/Transitive Closure, Internalized\<close>
-(* "rtran_closure(M,r,s) ==
+(* "rtran_closure(M,r,s) \<equiv>
\<forall>A[M]. is_field(M,r,A) \<longrightarrow>
(\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))" *)
definition
rtran_closure_fm :: "[i,i]=>i" where
- "rtran_closure_fm(r,s) ==
+ "rtran_closure_fm(r,s) \<equiv>
Forall(Implies(field_fm(succ(r),0),
Forall(Iff(Member(0,succ(succ(s))),
rtran_closure_mem_fm(1,succ(succ(r)),0)))))"
lemma rtran_closure_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> rtran_closure_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> rtran_closure_fm(x,y) \<in> formula"
by (simp add: rtran_closure_fm_def)
lemma sats_rtran_closure_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, rtran_closure_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, rtran_closure_fm(x,y), env) \<longleftrightarrow>
rtran_closure(##A, nth(x,env), nth(y,env))"
by (simp add: rtran_closure_fm_def rtran_closure_def)
lemma rtran_closure_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> rtran_closure(##A, x, y) \<longleftrightarrow> sats(A, rtran_closure_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> rtran_closure(##A, x, y) \<longleftrightarrow> sats(A, rtran_closure_fm(i,j), env)"
by simp
theorem rtran_closure_reflection:
@@ -119,27 +119,27 @@
subsubsection\<open>Transitive Closure of a Relation, Internalized\<close>
-(* "tran_closure(M,r,t) ==
+(* "tran_closure(M,r,t) \<equiv>
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *)
definition
tran_closure_fm :: "[i,i]=>i" where
- "tran_closure_fm(r,s) ==
+ "tran_closure_fm(r,s) \<equiv>
Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"
lemma tran_closure_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> tran_closure_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> tran_closure_fm(x,y) \<in> formula"
by (simp add: tran_closure_fm_def)
lemma sats_tran_closure_fm [simp]:
- "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
- ==> sats(A, tran_closure_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, tran_closure_fm(x,y), env) \<longleftrightarrow>
tran_closure(##A, nth(x,env), nth(y,env))"
by (simp add: tran_closure_fm_def tran_closure_def)
lemma tran_closure_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j \<in> nat; env \<in> list(A)|]
- ==> tran_closure(##A, x, y) \<longleftrightarrow> sats(A, tran_closure_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> tran_closure(##A, x, y) \<longleftrightarrow> sats(A, tran_closure_fm(i,j), env)"
by simp
theorem tran_closure_reflection:
@@ -163,7 +163,7 @@
tran_closure_reflection)
lemma wellfounded_trancl_separation:
- "[| L(r); L(Z) |] ==>
+ "\<lbrakk>L(r); L(Z)\<rbrakk> \<Longrightarrow>
separation (L, \<lambda>x.
\<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)"
@@ -204,7 +204,7 @@
lemma list_replacement1:
- "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)"
+ "L(A) \<Longrightarrow> iterates_replacement(L, is_list_functor(L,A), 0)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
apply (rule strong_replacementI)
apply (rule_tac u="{B,A,n,0,Memrel(succ(n))}"
@@ -226,7 +226,7 @@
is_iterates_reflection list_functor_reflection)
lemma list_replacement2:
- "L(A) ==> strong_replacement(L,
+ "L(A) \<Longrightarrow> strong_replacement(L,
\<lambda>n y. n\<in>nat & is_iterates(L, is_list_functor(L,A), 0, n, y))"
apply (rule strong_replacementI)
apply (rule_tac u="{A,B,0,nat}"
@@ -292,30 +292,30 @@
subsubsection\<open>The Formula \<^term>\<open>is_nth\<close>, Internalized\<close>
-(* "is_nth(M,n,l,Z) ==
+(* "is_nth(M,n,l,Z) \<equiv>
\<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *)
definition
nth_fm :: "[i,i,i]=>i" where
- "nth_fm(n,l,Z) ==
+ "nth_fm(n,l,Z) \<equiv>
Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0),
hd_fm(0,succ(Z))))"
lemma nth_fm_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> nth_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> nth_fm(x,y,z) \<in> formula"
by (simp add: nth_fm_def)
lemma sats_nth_fm [simp]:
- "[| x < length(env); y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, nth_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x < length(env); y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, nth_fm(x,y,z), env) \<longleftrightarrow>
is_nth(##A, nth(x,env), nth(y,env), nth(z,env))"
apply (frule lt_length_in_nat, assumption)
apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm)
done
lemma nth_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i < length(env); j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_nth(##A, x, y, z) \<longleftrightarrow> sats(A, nth_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i < length(env); j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_nth(##A, x, y, z) \<longleftrightarrow> sats(A, nth_fm(i,j,k), env)"
by (simp)
theorem nth_reflection:
@@ -343,7 +343,7 @@
iterates_MH_reflection tl_reflection)
lemma nth_replacement:
- "L(w) ==> iterates_replacement(L, is_tl(L), w)"
+ "L(w) \<Longrightarrow> iterates_replacement(L, is_tl(L), w)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
apply (rule strong_replacementI)
apply (rule_tac u="{B,w,Memrel(succ(n))}"
@@ -390,7 +390,7 @@
iterates_MH_reflection)
lemma eclose_replacement1:
- "L(A) ==> iterates_replacement(L, big_union(L), A)"
+ "L(A) \<Longrightarrow> iterates_replacement(L, big_union(L), A)"
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
apply (rule strong_replacementI)
apply (rule_tac u="{B,A,n,Memrel(succ(n))}"
@@ -410,7 +410,7 @@
by (intro FOL_reflections function_reflections is_iterates_reflection)
lemma eclose_replacement2:
- "L(A) ==> strong_replacement(L,
+ "L(A) \<Longrightarrow> strong_replacement(L,
\<lambda>n y. n\<in>nat & is_iterates(L, big_union(L), A, n, y))"
apply (rule strong_replacementI)
apply (rule_tac u="{A,B,nat}"
--- a/src/ZF/Constructible/Reflection.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Reflection.thy Tue Sep 27 16:51:35 2022 +0100
@@ -6,11 +6,11 @@
theory Reflection imports Normal begin
-lemma all_iff_not_ex_not: "(\<forall>x. P(x)) \<longleftrightarrow> (~ (\<exists>x. ~ P(x)))"
-by blast
+lemma all_iff_not_ex_not: "(\<forall>x. P(x)) \<longleftrightarrow> (\<not> (\<exists>x. \<not> P(x)))"
+ by blast
-lemma ball_iff_not_bex_not: "(\<forall>x\<in>A. P(x)) \<longleftrightarrow> (~ (\<exists>x\<in>A. ~ P(x)))"
-by blast
+lemma ball_iff_not_bex_not: "(\<forall>x\<in>A. P(x)) \<longleftrightarrow> (\<not> (\<exists>x\<in>A. \<not> P(x)))"
+ by blast
text\<open>From the notes of A. S. Kechris, page 6, and from
Andrzej Mostowski, \emph{Constructible Sets with Applications},
@@ -33,105 +33,106 @@
fixes Mset and M and Reflects
assumes Mset_mono_le : "mono_le_subset(Mset)"
and Mset_cont : "cont_Ord(Mset)"
- and Pair_in_Mset : "[| x \<in> Mset(a); y \<in> Mset(a); Limit(a) |]
- ==> <x,y> \<in> Mset(a)"
- defines "M(x) == \<exists>a. Ord(a) & x \<in> Mset(a)"
- and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) &
+ and Pair_in_Mset : "\<lbrakk>x \<in> Mset(a); y \<in> Mset(a); Limit(a)\<rbrakk>
+ \<Longrightarrow> <x,y> \<in> Mset(a)"
+ defines "M(x) \<equiv> \<exists>a. Ord(a) & x \<in> Mset(a)"
+ and "Reflects(Cl,P,Q) \<equiv> Closed_Unbounded(Cl) &
(\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))"
fixes F0 \<comment> \<open>ordinal for a specific value \<^term>\<open>y\<close>\<close>
fixes FF \<comment> \<open>sup over the whole level, \<^term>\<open>y\<in>Mset(a)\<close>\<close>
fixes ClEx \<comment> \<open>Reflecting ordinals for the formula \<^term>\<open>\<exists>z. P\<close>\<close>
- defines "F0(P,y) == \<mu> b. (\<exists>z. M(z) & P(<y,z>)) \<longrightarrow>
+ defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) & P(<y,z>)) \<longrightarrow>
(\<exists>z\<in>Mset(b). P(<y,z>))"
- and "FF(P) == \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)"
- and "ClEx(P,a) == Limit(a) & normalize(FF(P),a) = a"
+ and "FF(P) \<equiv> \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)"
+ and "ClEx(P,a) \<equiv> Limit(a) & normalize(FF(P),a) = a"
-lemma (in reflection) Mset_mono: "i\<le>j ==> Mset(i) \<subseteq> Mset(j)"
-apply (insert Mset_mono_le)
-apply (simp add: mono_le_subset_def leI)
-done
+begin
+
+lemma Mset_mono: "i\<le>j \<Longrightarrow> Mset(i) \<subseteq> Mset(j)"
+ using Mset_mono_le by (simp add: mono_le_subset_def leI)
text\<open>Awkward: we need a version of \<open>ClEx_def\<close> as an equality
at the level of classes, which do not really exist\<close>
-lemma (in reflection) ClEx_eq:
- "ClEx(P) == \<lambda>a. Limit(a) & normalize(FF(P),a) = a"
+lemma ClEx_eq:
+ "ClEx(P) \<equiv> \<lambda>a. Limit(a) & normalize(FF(P),a) = a"
by (simp add: ClEx_def [symmetric])
subsection\<open>Easy Cases of the Reflection Theorem\<close>
-theorem (in reflection) Triv_reflection [intro]:
- "Reflects(Ord, P, \<lambda>a x. P(x))"
-by (simp add: Reflects_def)
+theorem Triv_reflection [intro]:
+ "Reflects(Ord, P, \<lambda>a x. P(x))"
+ by (simp add: Reflects_def)
-theorem (in reflection) Not_reflection [intro]:
- "Reflects(Cl,P,Q) ==> Reflects(Cl, \<lambda>x. ~P(x), \<lambda>a x. ~Q(a,x))"
-by (simp add: Reflects_def)
+theorem Not_reflection [intro]:
+ "Reflects(Cl,P,Q) \<Longrightarrow> Reflects(Cl, \<lambda>x. \<not>P(x), \<lambda>a x. \<not>Q(a,x))"
+ by (simp add: Reflects_def)
-theorem (in reflection) And_reflection [intro]:
- "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
- ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x),
+theorem And_reflection [intro]:
+ "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) & P'(x),
\<lambda>a x. Q(a,x) & Q'(a,x))"
-by (simp add: Reflects_def Closed_Unbounded_Int, blast)
+ by (simp add: Reflects_def Closed_Unbounded_Int, blast)
-theorem (in reflection) Or_reflection [intro]:
- "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
- ==> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x),
+theorem Or_reflection [intro]:
+ "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a), \<lambda>x. P(x) | P'(x),
\<lambda>a x. Q(a,x) | Q'(a,x))"
by (simp add: Reflects_def Closed_Unbounded_Int, blast)
-theorem (in reflection) Imp_reflection [intro]:
- "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
- ==> Reflects(\<lambda>a. Cl(a) & C'(a),
+theorem Imp_reflection [intro]:
+ "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a),
\<lambda>x. P(x) \<longrightarrow> P'(x),
\<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x))"
by (simp add: Reflects_def Closed_Unbounded_Int, blast)
-theorem (in reflection) Iff_reflection [intro]:
- "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |]
- ==> Reflects(\<lambda>a. Cl(a) & C'(a),
+theorem Iff_reflection [intro]:
+ "\<lbrakk>Reflects(Cl,P,Q); Reflects(C',P',Q')\<rbrakk>
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & C'(a),
\<lambda>x. P(x) \<longleftrightarrow> P'(x),
\<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x))"
by (simp add: Reflects_def Closed_Unbounded_Int, blast)
subsection\<open>Reflection for Existential Quantifiers\<close>
-lemma (in reflection) F0_works:
- "[| y\<in>Mset(a); Ord(a); M(z); P(<y,z>) |] ==> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)"
+lemma F0_works:
+ "\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(<y,z>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)"
apply (unfold F0_def M_def, clarify)
apply (rule LeastI2)
apply (blast intro: Mset_mono [THEN subsetD])
apply (blast intro: lt_Ord2, blast)
done
-lemma (in reflection) Ord_F0 [intro,simp]: "Ord(F0(P,y))"
+lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))"
by (simp add: F0_def)
-lemma (in reflection) Ord_FF [intro,simp]: "Ord(FF(P,y))"
+lemma Ord_FF [intro,simp]: "Ord(FF(P,y))"
by (simp add: FF_def)
-lemma (in reflection) cont_Ord_FF: "cont_Ord(FF(P))"
+lemma cont_Ord_FF: "cont_Ord(FF(P))"
apply (insert Mset_cont)
apply (simp add: cont_Ord_def FF_def, blast)
done
text\<open>Recall that \<^term>\<open>F0\<close> depends upon \<^term>\<open>y\<in>Mset(a)\<close>,
while \<^term>\<open>FF\<close> depends only upon \<^term>\<open>a\<close>.\<close>
-lemma (in reflection) FF_works:
- "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |] ==> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)"
+lemma FF_works:
+ "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)"
apply (simp add: FF_def)
apply (simp_all add: cont_Ord_Union [of concl: Mset]
Mset_cont Mset_mono_le not_emptyI)
apply (blast intro: F0_works)
done
-lemma (in reflection) FFN_works:
- "[| M(z); y\<in>Mset(a); P(<y,z>); Ord(a) |]
- ==> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)"
+lemma FFN_works:
+ "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk>
+ \<Longrightarrow> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)"
apply (drule FF_works [of concl: P], assumption+)
apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD])
done
+end
text\<open>Locale for the induction hypothesis\<close>
@@ -139,11 +140,13 @@
fixes P \<comment> \<open>the original formula\<close>
fixes Q \<comment> \<open>the reflected formula\<close>
fixes Cl \<comment> \<open>the class of reflecting ordinals\<close>
- assumes Cl_reflects: "[| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)"
+ assumes Cl_reflects: "\<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)"
+
+begin
-lemma (in ex_reflection) ClEx_downward:
- "[| M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a) |]
- ==> \<exists>z\<in>Mset(a). Q(a,<y,z>)"
+lemma ClEx_downward:
+ "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a)\<rbrakk>
+ \<Longrightarrow> \<exists>z\<in>Mset(a). Q(a,<y,z>)"
apply (simp add: ClEx_def, clarify)
apply (frule Limit_is_Ord)
apply (frule FFN_works [of concl: P], assumption+)
@@ -151,35 +154,40 @@
apply (auto simp add: Limit_is_Ord Pair_in_Mset)
done
-lemma (in ex_reflection) ClEx_upward:
- "[| z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a) |]
- ==> \<exists>z. M(z) & P(<y,z>)"
+lemma ClEx_upward:
+ "\<lbrakk>z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a)\<rbrakk>
+ \<Longrightarrow> \<exists>z. M(z) & P(<y,z>)"
apply (simp add: ClEx_def M_def)
apply (blast dest: Cl_reflects
intro: Limit_is_Ord Pair_in_Mset)
done
text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close>
-lemma (in ex_reflection) ZF_ClEx_iff:
- "[| y\<in>Mset(a); Cl(a); ClEx(P,a) |]
- ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+lemma ZF_ClEx_iff:
+ "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a)\<rbrakk>
+ \<Longrightarrow> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
by (blast intro: dest: ClEx_downward ClEx_upward)
text\<open>...and it is closed and unbounded\<close>
-lemma (in ex_reflection) ZF_Closed_Unbounded_ClEx:
+lemma ZF_Closed_Unbounded_ClEx:
"Closed_Unbounded(ClEx(P))"
apply (simp add: ClEx_eq)
apply (fast intro: Closed_Unbounded_Int Normal_imp_fp_Closed_Unbounded
Closed_Unbounded_Limit Normal_normalize)
done
+end
+
text\<open>The same two theorems, exported to locale \<open>reflection\<close>.\<close>
+context reflection
+begin
+
text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close>
-lemma (in reflection) ClEx_iff:
- "[| y\<in>Mset(a); Cl(a); ClEx(P,a);
- !!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x) |]
- ==> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+lemma ClEx_iff:
+ "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a);
+ \<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)\<rbrakk>
+ \<Longrightarrow> (\<exists>z. M(z) & P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
apply (unfold ClEx_def FF_def F0_def M_def)
apply (rule ex_reflection.ZF_ClEx_iff
[OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro,
@@ -194,9 +202,9 @@
apply (simp add: ex_reflection_axioms.intro, assumption+)
*)
-lemma (in reflection) Closed_Unbounded_ClEx:
- "(!!a. [| Cl(a); Ord(a) |] ==> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x))
- ==> Closed_Unbounded(ClEx(P))"
+lemma Closed_Unbounded_ClEx:
+ "(\<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x))
+ \<Longrightarrow> Closed_Unbounded(ClEx(P))"
apply (unfold ClEx_eq FF_def F0_def M_def)
apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl])
apply (rule ex_reflection.intro, rule reflection_axioms)
@@ -205,9 +213,9 @@
subsection\<open>Packaging the Quantifier Reflection Rules\<close>
-lemma (in reflection) Ex_reflection_0:
+lemma Ex_reflection_0:
"Reflects(Cl,P0,Q0)
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(P0,a),
\<lambda>x. \<exists>z. M(z) & P0(<x,z>),
\<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))"
apply (simp add: Reflects_def)
@@ -217,9 +225,9 @@
apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast)
done
-lemma (in reflection) All_reflection_0:
+lemma All_reflection_0:
"Reflects(Cl,P0,Q0)
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.~P0(x), a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x.\<not>P0(x), a),
\<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(<x,z>),
\<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))"
apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not)
@@ -227,17 +235,17 @@
apply (erule Ex_reflection_0)
done
-theorem (in reflection) Ex_reflection [intro]:
+theorem Ex_reflection [intro]:
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
\<lambda>x. \<exists>z. M(z) & P(x,z),
\<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
by (rule Ex_reflection_0 [of _ " \<lambda>x. P(fst(x),snd(x))"
"\<lambda>a x. Q(a,fst(x),snd(x))", simplified])
-theorem (in reflection) All_reflection [intro]:
+theorem All_reflection [intro]:
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
\<lambda>x. \<forall>z. M(z) \<longrightarrow> P(x,z),
\<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
by (rule All_reflection_0 [of _ "\<lambda>x. P(fst(x),snd(x))"
@@ -245,16 +253,16 @@
text\<open>And again, this time using class-bounded quantifiers\<close>
-theorem (in reflection) Rex_reflection [intro]:
+theorem Rex_reflection [intro]:
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. P(fst(x),snd(x)), a),
\<lambda>x. \<exists>z[M]. P(x,z),
\<lambda>a x. \<exists>z\<in>Mset(a). Q(a,x,z))"
by (unfold rex_def, blast)
-theorem (in reflection) Rall_reflection [intro]:
+theorem Rall_reflection [intro]:
"Reflects(Cl, \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x)))
- ==> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. ~P(fst(x),snd(x)), a),
+ \<Longrightarrow> Reflects(\<lambda>a. Cl(a) & ClEx(\<lambda>x. \<not>P(fst(x),snd(x)), a),
\<lambda>x. \<forall>z[M]. P(x,z),
\<lambda>a x. \<forall>z\<in>Mset(a). Q(a,x,z))"
by (unfold rall_def, blast)
@@ -268,7 +276,7 @@
text\<open>Example 1: reflecting a simple formula. The reflecting class is first
given as the variable \<open>?Cl\<close> and later retrieved from the final
proof state.\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & x \<in> y,
\<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
@@ -277,7 +285,7 @@
text\<open>Problem here: there needs to be a conjunction (class intersection)
in the class of reflecting ordinals. The \<^term>\<open>Ord(a)\<close> is redundant,
though harmless.\<close>
-lemma (in reflection)
+lemma
"Reflects(\<lambda>a. Ord(a) & ClEx(\<lambda>x. fst(x) \<in> snd(x), a),
\<lambda>x. \<exists>y. M(y) & x \<in> y,
\<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
@@ -285,31 +293,31 @@
text\<open>Example 2\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
by fast
text\<open>Example 2'. We give the reflecting class explicitly.\<close>
-lemma (in reflection)
+lemma
"Reflects
(\<lambda>a. (Ord(a) &
- ClEx(\<lambda>x. ~ (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) &
+ ClEx(\<lambda>x. \<not> (snd(x) \<subseteq> fst(fst(x)) \<longrightarrow> snd(x) \<in> snd(fst(x))), a)) &
ClEx(\<lambda>x. \<forall>z. M(z) \<longrightarrow> z \<subseteq> fst(x) \<longrightarrow> z \<in> snd(x), a),
\<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
by fast
text\<open>Example 2''. We expand the subset relation.\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> (\<forall>w. M(w) \<longrightarrow> w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y),
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z \<longrightarrow> w\<in>x) \<longrightarrow> z\<in>y)"
by fast
text\<open>Example 2'''. Single-step version, to reveal the reflecting class.\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<subseteq> x \<longrightarrow> z \<in> y),
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x \<longrightarrow> z \<in> y)"
@@ -329,21 +337,21 @@
text\<open>Example 3. Warning: the following examples make sense only
if \<^term>\<open>P\<close> is quantifier-free, since it is not being relativized.\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) \<longrightarrow> z \<in> y \<longleftrightarrow> z \<in> x & P(z)),
\<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y \<longleftrightarrow> z \<in> x & P(z))"
by fast
text\<open>Example 3'\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & y = Collect(x,P),
\<lambda>a x. \<exists>y\<in>Mset(a). y = Collect(x,P))"
by fast
text\<open>Example 3''\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>x. \<exists>y. M(y) & y = Replace(x,P),
\<lambda>a x. \<exists>y\<in>Mset(a). y = Replace(x,P))"
@@ -351,7 +359,7 @@
text\<open>Example 4: Axiom of Choice. Possibly wrong, since \<open>\<Pi>\<close> needs
to be relativized.\<close>
-schematic_goal (in reflection)
+schematic_goal
"Reflects(?Cl,
\<lambda>A. 0\<notin>A \<longrightarrow> (\<exists>f. M(f) & f \<in> (\<Prod>X \<in> A. X)),
\<lambda>a A. 0\<notin>A \<longrightarrow> (\<exists>f\<in>Mset(a). f \<in> (\<Prod>X \<in> A. X)))"
@@ -359,3 +367,5 @@
end
+end
+
--- a/src/ZF/Constructible/Relative.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Relative.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,115 +12,115 @@
definition
empty :: "[i=>o,i] => o" where
- "empty(M,z) == \<forall>x[M]. x \<notin> z"
+ "empty(M,z) \<equiv> \<forall>x[M]. x \<notin> z"
definition
subset :: "[i=>o,i,i] => o" where
- "subset(M,A,B) == \<forall>x[M]. x\<in>A \<longrightarrow> x \<in> B"
+ "subset(M,A,B) \<equiv> \<forall>x[M]. x\<in>A \<longrightarrow> x \<in> B"
definition
upair :: "[i=>o,i,i,i] => o" where
- "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z \<longrightarrow> x = a | x = b)"
+ "upair(M,a,b,z) \<equiv> a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z \<longrightarrow> x = a | x = b)"
definition
pair :: "[i=>o,i,i,i] => o" where
- "pair(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) &
+ "pair(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) &
(\<exists>y[M]. upair(M,a,b,y) & upair(M,x,y,z))"
definition
union :: "[i=>o,i,i,i] => o" where
- "union(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a | x \<in> b"
+ "union(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a | x \<in> b"
definition
is_cons :: "[i=>o,i,i,i] => o" where
- "is_cons(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) & union(M,x,b,z)"
+ "is_cons(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) & union(M,x,b,z)"
definition
successor :: "[i=>o,i,i] => o" where
- "successor(M,a,z) == is_cons(M,a,a,z)"
+ "successor(M,a,z) \<equiv> is_cons(M,a,a,z)"
definition
number1 :: "[i=>o,i] => o" where
- "number1(M,a) == \<exists>x[M]. empty(M,x) & successor(M,x,a)"
+ "number1(M,a) \<equiv> \<exists>x[M]. empty(M,x) & successor(M,x,a)"
definition
number2 :: "[i=>o,i] => o" where
- "number2(M,a) == \<exists>x[M]. number1(M,x) & successor(M,x,a)"
+ "number2(M,a) \<equiv> \<exists>x[M]. number1(M,x) & successor(M,x,a)"
definition
number3 :: "[i=>o,i] => o" where
- "number3(M,a) == \<exists>x[M]. number2(M,x) & successor(M,x,a)"
+ "number3(M,a) \<equiv> \<exists>x[M]. number2(M,x) & successor(M,x,a)"
definition
powerset :: "[i=>o,i,i] => o" where
- "powerset(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> subset(M,x,A)"
+ "powerset(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> subset(M,x,A)"
definition
is_Collect :: "[i=>o,i,i=>o,i] => o" where
- "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)"
+ "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)"
definition
is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where
- "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))"
+ "is_Replace(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))"
definition
inter :: "[i=>o,i,i,i] => o" where
- "inter(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<in> b"
+ "inter(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<in> b"
definition
setdiff :: "[i=>o,i,i,i] => o" where
- "setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<notin> b"
+ "setdiff(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<notin> b"
definition
big_union :: "[i=>o,i,i] => o" where
- "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)"
+ "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)"
definition
big_inter :: "[i=>o,i,i] => o" where
- "big_inter(M,A,z) ==
+ "big_inter(M,A,z) \<equiv>
(A=0 \<longrightarrow> z=0) &
(A\<noteq>0 \<longrightarrow> (\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> x \<in> y)))"
definition
cartprod :: "[i=>o,i,i,i] => o" where
- "cartprod(M,A,B,z) ==
+ "cartprod(M,A,B,z) \<equiv>
\<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
definition
is_sum :: "[i=>o,i,i,i] => o" where
- "is_sum(M,A,B,Z) ==
+ "is_sum(M,A,B,Z) \<equiv>
\<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)"
definition
is_Inl :: "[i=>o,i,i] => o" where
- "is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z)"
+ "is_Inl(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z)"
definition
is_Inr :: "[i=>o,i,i] => o" where
- "is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z)"
+ "is_Inr(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z)"
definition
is_converse :: "[i=>o,i,i] => o" where
- "is_converse(M,r,z) ==
+ "is_converse(M,r,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow>
(\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
definition
pre_image :: "[i=>o,i,i,i] => o" where
- "pre_image(M,r,A,z) ==
+ "pre_image(M,r,A,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
definition
is_domain :: "[i=>o,i,i] => o" where
- "is_domain(M,r,z) ==
+ "is_domain(M,r,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
definition
image :: "[i=>o,i,i,i] => o" where
- "image(M,r,A,z) ==
+ "image(M,r,A,z) \<equiv>
\<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
definition
@@ -129,46 +129,46 @@
\<^term>\<open>\<exists>r'[M]. is_converse(M,r,r') & is_domain(M,r',z)\<close>
unfortunately needs an instance of separation in order to prove
\<^term>\<open>M(converse(r))\<close>.\<close>
- "is_range(M,r,z) ==
+ "is_range(M,r,z) \<equiv>
\<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
definition
is_field :: "[i=>o,i,i] => o" where
- "is_field(M,r,z) ==
+ "is_field(M,r,z) \<equiv>
\<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
union(M,dr,rr,z)"
definition
is_relation :: "[i=>o,i] => o" where
- "is_relation(M,r) ==
+ "is_relation(M,r) \<equiv>
(\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
definition
is_function :: "[i=>o,i] => o" where
- "is_function(M,r) ==
+ "is_function(M,r) \<equiv>
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'"
definition
fun_apply :: "[i=>o,i,i,i] => o" where
- "fun_apply(M,f,x,y) ==
+ "fun_apply(M,f,x,y) \<equiv>
(\<exists>xs[M]. \<exists>fxs[M].
upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))"
definition
typed_function :: "[i=>o,i,i,i] => o" where
- "typed_function(M,A,B,r) ==
+ "typed_function(M,A,B,r) \<equiv>
is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
(\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))"
definition
is_funspace :: "[i=>o,i,i,i] => o" where
- "is_funspace(M,A,B,F) ==
+ "is_funspace(M,A,B,F) \<equiv>
\<forall>f[M]. f \<in> F \<longleftrightarrow> typed_function(M,A,B,f)"
definition
composition :: "[i=>o,i,i,i] => o" where
- "composition(M,r,s,t) ==
+ "composition(M,r,s,t) \<equiv>
\<forall>p[M]. p \<in> t \<longleftrightarrow>
(\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
@@ -176,104 +176,104 @@
definition
injection :: "[i=>o,i,i,i] => o" where
- "injection(M,A,B,f) ==
+ "injection(M,A,B,f) \<equiv>
typed_function(M,A,B,f) &
(\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')"
definition
surjection :: "[i=>o,i,i,i] => o" where
- "surjection(M,A,B,f) ==
+ "surjection(M,A,B,f) \<equiv>
typed_function(M,A,B,f) &
(\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
definition
bijection :: "[i=>o,i,i,i] => o" where
- "bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)"
+ "bijection(M,A,B,f) \<equiv> injection(M,A,B,f) & surjection(M,A,B,f)"
definition
restriction :: "[i=>o,i,i,i] => o" where
- "restriction(M,r,A,z) ==
+ "restriction(M,r,A,z) \<equiv>
\<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
definition
transitive_set :: "[i=>o,i] => o" where
- "transitive_set(M,a) == \<forall>x[M]. x\<in>a \<longrightarrow> subset(M,x,a)"
+ "transitive_set(M,a) \<equiv> \<forall>x[M]. x\<in>a \<longrightarrow> subset(M,x,a)"
definition
ordinal :: "[i=>o,i] => o" where
\<comment> \<open>an ordinal is a transitive set of transitive sets\<close>
- "ordinal(M,a) == transitive_set(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
+ "ordinal(M,a) \<equiv> transitive_set(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
definition
limit_ordinal :: "[i=>o,i] => o" where
\<comment> \<open>a limit ordinal is a non-empty, successor-closed ordinal\<close>
- "limit_ordinal(M,a) ==
- ordinal(M,a) & ~ empty(M,a) &
+ "limit_ordinal(M,a) \<equiv>
+ ordinal(M,a) & \<not> empty(M,a) &
(\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
definition
successor_ordinal :: "[i=>o,i] => o" where
\<comment> \<open>a successor ordinal is any ordinal that is neither empty nor limit\<close>
- "successor_ordinal(M,a) ==
- ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
+ "successor_ordinal(M,a) \<equiv>
+ ordinal(M,a) & \<not> empty(M,a) & \<not> limit_ordinal(M,a)"
definition
finite_ordinal :: "[i=>o,i] => o" where
\<comment> \<open>an ordinal is finite if neither it nor any of its elements are limit\<close>
- "finite_ordinal(M,a) ==
- ordinal(M,a) & ~ limit_ordinal(M,a) &
- (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
+ "finite_ordinal(M,a) \<equiv>
+ ordinal(M,a) & \<not> limit_ordinal(M,a) &
+ (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))"
definition
omega :: "[i=>o,i] => o" where
\<comment> \<open>omega is a limit ordinal none of whose elements are limit\<close>
- "omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
+ "omega(M,a) \<equiv> limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))"
definition
is_quasinat :: "[i=>o,i] => o" where
- "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))"
+ "is_quasinat(M,z) \<equiv> empty(M,z) | (\<exists>m[M]. successor(M,m,z))"
definition
is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where
- "is_nat_case(M, a, is_b, k, z) ==
+ "is_nat_case(M, a, is_b, k, z) \<equiv>
(empty(M,k) \<longrightarrow> z=a) &
(\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))"
definition
relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where
- "relation1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) \<longleftrightarrow> y = f(x)"
+ "relation1(M,is_f,f) \<equiv> \<forall>x[M]. \<forall>y[M]. is_f(x,y) \<longleftrightarrow> y = f(x)"
definition
Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where
\<comment> \<open>as above, but typed\<close>
- "Relation1(M,A,is_f,f) ==
+ "Relation1(M,A,is_f,f) \<equiv>
\<forall>x[M]. \<forall>y[M]. x\<in>A \<longrightarrow> is_f(x,y) \<longleftrightarrow> y = f(x)"
definition
relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where
- "relation2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
+ "relation2(M,is_f,f) \<equiv> \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
definition
Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where
- "Relation2(M,A,B,is_f,f) ==
+ "Relation2(M,A,B,is_f,f) \<equiv>
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
definition
relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
- "relation3(M,is_f,f) ==
+ "relation3(M,is_f,f) \<equiv>
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
definition
Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
- "Relation3(M,A,B,C,is_f,f) ==
+ "Relation3(M,A,B,C,is_f,f) \<equiv>
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M].
x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> z\<in>C \<longrightarrow> is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
definition
relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where
- "relation4(M,is_f,f) ==
+ "relation4(M,is_f,f) \<equiv>
\<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) \<longleftrightarrow> a = f(u,x,y,z)"
@@ -291,7 +291,7 @@
definition
extensionality :: "(i=>o) => o" where
- "extensionality(M) ==
+ "extensionality(M) \<equiv>
\<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x \<longleftrightarrow> z \<in> y) \<longrightarrow> x=y"
definition
@@ -300,42 +300,42 @@
belonging to \<open>M\<close> and all its quantifiers must be relativized
to \<open>M\<close>. We do not have separation as a scheme; every instance
that we need must be assumed (and later proved) separately.\<close>
- "separation(M,P) ==
+ "separation(M,P) \<equiv>
\<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
definition
upair_ax :: "(i=>o) => o" where
- "upair_ax(M) == \<forall>x[M]. \<forall>y[M]. \<exists>z[M]. upair(M,x,y,z)"
+ "upair_ax(M) \<equiv> \<forall>x[M]. \<forall>y[M]. \<exists>z[M]. upair(M,x,y,z)"
definition
Union_ax :: "(i=>o) => o" where
- "Union_ax(M) == \<forall>x[M]. \<exists>z[M]. big_union(M,x,z)"
+ "Union_ax(M) \<equiv> \<forall>x[M]. \<exists>z[M]. big_union(M,x,z)"
definition
power_ax :: "(i=>o) => o" where
- "power_ax(M) == \<forall>x[M]. \<exists>z[M]. powerset(M,x,z)"
+ "power_ax(M) \<equiv> \<forall>x[M]. \<exists>z[M]. powerset(M,x,z)"
definition
univalent :: "[i=>o, i, [i,i]=>o] => o" where
- "univalent(M,A,P) ==
+ "univalent(M,A,P) \<equiv>
\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) \<longrightarrow> y=z)"
definition
replacement :: "[i=>o, [i,i]=>o] => o" where
- "replacement(M,P) ==
+ "replacement(M,P) \<equiv>
\<forall>A[M]. univalent(M,A,P) \<longrightarrow>
(\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y)"
definition
strong_replacement :: "[i=>o, [i,i]=>o] => o" where
- "strong_replacement(M,P) ==
+ "strong_replacement(M,P) \<equiv>
\<forall>A[M]. univalent(M,A,P) \<longrightarrow>
(\<exists>Y[M]. \<forall>b[M]. b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b)))"
definition
foundation_ax :: "(i=>o) => o" where
- "foundation_ax(M) ==
- \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
+ "foundation_ax(M) \<equiv>
+ \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x & \<not>(\<exists>z[M]. z\<in>x & z \<in> y))"
subsection\<open>A trivial consistency proof for $V_\omega$\<close>
@@ -344,28 +344,28 @@
(or \<open>univ\<close> in Isabelle) satisfies some ZF axioms.
Kunen, Theorem IV 3.13, page 123.\<close>
-lemma univ0_downwards_mem: "[| y \<in> x; x \<in> univ(0) |] ==> y \<in> univ(0)"
+lemma univ0_downwards_mem: "\<lbrakk>y \<in> x; x \<in> univ(0)\<rbrakk> \<Longrightarrow> y \<in> univ(0)"
apply (insert Transset_univ [OF Transset_0])
apply (simp add: Transset_def, blast)
done
lemma univ0_Ball_abs [simp]:
- "A \<in> univ(0) ==> (\<forall>x\<in>A. x \<in> univ(0) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
+ "A \<in> univ(0) \<Longrightarrow> (\<forall>x\<in>A. x \<in> univ(0) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
by (blast intro: univ0_downwards_mem)
lemma univ0_Bex_abs [simp]:
- "A \<in> univ(0) ==> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
+ "A \<in> univ(0) \<Longrightarrow> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
by (blast intro: univ0_downwards_mem)
text\<open>Congruence rule for separation: can assume the variable is in \<open>M\<close>\<close>
lemma separation_cong [cong]:
- "(!!x. M(x) ==> P(x) \<longleftrightarrow> P'(x))
- ==> separation(M, %x. P(x)) \<longleftrightarrow> separation(M, %x. P'(x))"
+ "(\<And>x. M(x) \<Longrightarrow> P(x) \<longleftrightarrow> P'(x))
+ \<Longrightarrow> separation(M, %x. P(x)) \<longleftrightarrow> separation(M, %x. P'(x))"
by (simp add: separation_def)
lemma univalent_cong [cong]:
- "[| A=A'; !!x y. [| x\<in>A; M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
- ==> univalent(M, A, %x y. P(x,y)) \<longleftrightarrow> univalent(M, A', %x y. P'(x,y))"
+ "\<lbrakk>A=A'; \<And>x y. \<lbrakk>x\<in>A; M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y)\<rbrakk>
+ \<Longrightarrow> univalent(M, A, %x y. P(x,y)) \<longleftrightarrow> univalent(M, A', %x y. P'(x,y))"
by (simp add: univalent_def)
lemma univalent_triv [intro,simp]:
@@ -373,13 +373,13 @@
by (simp add: univalent_def)
lemma univalent_conjI2 [intro,simp]:
- "univalent(M,A,Q) ==> univalent(M, A, \<lambda>x y. P(x,y) & Q(x,y))"
+ "univalent(M,A,Q) \<Longrightarrow> univalent(M, A, \<lambda>x y. P(x,y) & Q(x,y))"
by (simp add: univalent_def, blast)
text\<open>Congruence rule for replacement\<close>
lemma strong_replacement_cong [cong]:
- "[| !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
- ==> strong_replacement(M, %x y. P(x,y)) \<longleftrightarrow>
+ "\<lbrakk>\<And>x y. \<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y)\<rbrakk>
+ \<Longrightarrow> strong_replacement(M, %x y. P(x,y)) \<longleftrightarrow>
strong_replacement(M, %x y. P'(x,y))"
by (simp add: strong_replacement_def)
@@ -391,21 +391,21 @@
text\<open>The separation axiom requires some lemmas\<close>
lemma Collect_in_Vfrom:
- "[| X \<in> Vfrom(A,j); Transset(A) |] ==> Collect(X,P) \<in> Vfrom(A, succ(j))"
+ "\<lbrakk>X \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> Collect(X,P) \<in> Vfrom(A, succ(j))"
apply (drule Transset_Vfrom)
apply (rule subset_mem_Vfrom)
apply (unfold Transset_def, blast)
done
lemma Collect_in_VLimit:
- "[| X \<in> Vfrom(A,i); Limit(i); Transset(A) |]
- ==> Collect(X,P) \<in> Vfrom(A,i)"
+ "\<lbrakk>X \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk>
+ \<Longrightarrow> Collect(X,P) \<in> Vfrom(A,i)"
apply (rule Limit_VfromE, assumption+)
apply (blast intro: Limit_has_succ VfromI Collect_in_Vfrom)
done
lemma Collect_in_univ:
- "[| X \<in> univ(A); Transset(A) |] ==> Collect(X,P) \<in> univ(A)"
+ "\<lbrakk>X \<in> univ(A); Transset(A)\<rbrakk> \<Longrightarrow> Collect(X,P) \<in> univ(A)"
by (simp add: univ_def Collect_in_VLimit)
lemma "separation(\<lambda>x. x \<in> univ(0), P)"
@@ -431,7 +431,7 @@
text\<open>Powerset axiom\<close>
lemma Pow_in_univ:
- "[| X \<in> univ(A); Transset(A) |] ==> Pow(X) \<in> univ(A)"
+ "\<lbrakk>X \<in> univ(A); Transset(A)\<rbrakk> \<Longrightarrow> Pow(X) \<in> univ(A)"
apply (simp add: univ_def Pow_in_VLimit)
done
@@ -486,17 +486,17 @@
done
lemma replacementD:
- "[| replacement(M,P); M(A); univalent(M,A,P) |]
- ==> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y))"
+ "\<lbrakk>replacement(M,P); M(A); univalent(M,A,P)\<rbrakk>
+ \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y))"
by (simp add: replacement_def)
lemma strong_replacementD:
- "[| strong_replacement(M,P); M(A); univalent(M,A,P) |]
- ==> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b))))"
+ "\<lbrakk>strong_replacement(M,P); M(A); univalent(M,A,P)\<rbrakk>
+ \<Longrightarrow> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b))))"
by (simp add: strong_replacement_def)
lemma separationD:
- "[| separation(M,P); M(z) |] ==> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
+ "\<lbrakk>separation(M,P); M(z)\<rbrakk> \<Longrightarrow> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
by (simp add: separation_def)
@@ -504,7 +504,7 @@
definition
order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
- "order_isomorphism(M,A,r,B,s,f) ==
+ "order_isomorphism(M,A,r,B,s,f) \<equiv>
bijection(M,A,B,f) &
(\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
(\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
@@ -513,12 +513,12 @@
definition
pred_set :: "[i=>o,i,i,i,i] => o" where
- "pred_set(M,A,x,r,B) ==
+ "pred_set(M,A,x,r,B) \<equiv>
\<forall>y[M]. y \<in> B \<longleftrightarrow> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
definition
membership :: "[i=>o,i,i] => o" where \<comment> \<open>membership relation\<close>
- "membership(M,A,r) ==
+ "membership(M,A,r) \<equiv>
\<forall>p[M]. p \<in> r \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
@@ -527,7 +527,7 @@
text\<open>The class M is assumed to be transitive and inhabited\<close>
locale M_trans =
fixes M
- assumes transM: "[| y\<in>x; M(x) |] ==> M(y)"
+ assumes transM: "\<lbrakk>y\<in>x; M(x)\<rbrakk> \<Longrightarrow> M(y)"
and M_inhabited: "\<exists>x . M(x)"
lemma (in M_trans) nonempty [simp]: "M(0)"
@@ -562,15 +562,15 @@
and Union_ax: "Union_ax(M)"
lemma (in M_trans) rall_abs [simp]:
- "M(A) ==> (\<forall>x[M]. x\<in>A \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
+ "M(A) \<Longrightarrow> (\<forall>x[M]. x\<in>A \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
by (blast intro: transM)
lemma (in M_trans) rex_abs [simp]:
- "M(A) ==> (\<exists>x[M]. x\<in>A & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
+ "M(A) \<Longrightarrow> (\<exists>x[M]. x\<in>A & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
by (blast intro: transM)
lemma (in M_trans) ball_iff_equiv:
- "M(A) ==> (\<forall>x[M]. (x\<in>A \<longleftrightarrow> P(x))) \<longleftrightarrow>
+ "M(A) \<Longrightarrow> (\<forall>x[M]. (x\<in>A \<longleftrightarrow> P(x))) \<longleftrightarrow>
(\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) \<longrightarrow> M(x) \<longrightarrow> x\<in>A)"
by (blast intro: transM)
@@ -579,26 +579,26 @@
But it's not the only way to prove such equalities: its
premises \<^term>\<open>M(A)\<close> and \<^term>\<open>M(B)\<close> can be too strong.\<close>
lemma (in M_trans) M_equalityI:
- "[| !!x. M(x) ==> x\<in>A \<longleftrightarrow> x\<in>B; M(A); M(B) |] ==> A=B"
+ "\<lbrakk>\<And>x. M(x) \<Longrightarrow> x\<in>A \<longleftrightarrow> x\<in>B; M(A); M(B)\<rbrakk> \<Longrightarrow> A=B"
by (blast dest: transM)
subsubsection\<open>Trivial Absoluteness Proofs: Empty Set, Pairs, etc.\<close>
lemma (in M_trans) empty_abs [simp]:
- "M(z) ==> empty(M,z) \<longleftrightarrow> z=0"
+ "M(z) \<Longrightarrow> empty(M,z) \<longleftrightarrow> z=0"
apply (simp add: empty_def)
apply (blast intro: transM)
done
lemma (in M_trans) subset_abs [simp]:
- "M(A) ==> subset(M,A,B) \<longleftrightarrow> A \<subseteq> B"
+ "M(A) \<Longrightarrow> subset(M,A,B) \<longleftrightarrow> A \<subseteq> B"
apply (simp add: subset_def)
apply (blast intro: transM)
done
lemma (in M_trans) upair_abs [simp]:
- "M(z) ==> upair(M,a,b,z) \<longleftrightarrow> z={a,b}"
+ "M(z) \<Longrightarrow> upair(M,a,b,z) \<longleftrightarrow> z={a,b}"
apply (simp add: upair_def)
apply (blast intro: transM)
done
@@ -628,7 +628,7 @@
by blast
lemma (in M_trans) pair_abs [simp]:
- "M(z) ==> pair(M,a,b,z) \<longleftrightarrow> z=<a,b>"
+ "M(z) \<Longrightarrow> pair(M,a,b,z) \<longleftrightarrow> z=<a,b>"
apply (simp add: pair_def Pair_def)
apply (blast intro: transM)
done
@@ -646,11 +646,11 @@
by blast
lemma (in M_trans) pair_components_in_M:
- "[| <x,y> \<in> A; M(A) |] ==> M(x) & M(y)"
+ "\<lbrakk><x,y> \<in> A; M(A)\<rbrakk> \<Longrightarrow> M(x) & M(y)"
by (blast dest: transM)
lemma (in M_trivial) cartprod_abs [simp]:
- "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) \<longleftrightarrow> z = A*B"
+ "\<lbrakk>M(A); M(B); M(z)\<rbrakk> \<Longrightarrow> cartprod(M,A,B,z) \<longleftrightarrow> z = A*B"
apply (simp add: cartprod_def)
apply (rule iffI)
apply (blast intro!: equalityI intro: transM dest!: rspec)
@@ -660,43 +660,43 @@
subsubsection\<open>Absoluteness for Unions and Intersections\<close>
lemma (in M_trans) union_abs [simp]:
- "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) \<longleftrightarrow> z = a \<union> b"
+ "\<lbrakk>M(a); M(b); M(z)\<rbrakk> \<Longrightarrow> union(M,a,b,z) \<longleftrightarrow> z = a \<union> b"
unfolding union_def
by (blast intro: transM )
lemma (in M_trans) inter_abs [simp]:
- "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) \<longleftrightarrow> z = a \<inter> b"
+ "\<lbrakk>M(a); M(b); M(z)\<rbrakk> \<Longrightarrow> inter(M,a,b,z) \<longleftrightarrow> z = a \<inter> b"
unfolding inter_def
by (blast intro: transM)
lemma (in M_trans) setdiff_abs [simp]:
- "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) \<longleftrightarrow> z = a-b"
+ "\<lbrakk>M(a); M(b); M(z)\<rbrakk> \<Longrightarrow> setdiff(M,a,b,z) \<longleftrightarrow> z = a-b"
unfolding setdiff_def
by (blast intro: transM)
lemma (in M_trans) Union_abs [simp]:
- "[| M(A); M(z) |] ==> big_union(M,A,z) \<longleftrightarrow> z = \<Union>(A)"
+ "\<lbrakk>M(A); M(z)\<rbrakk> \<Longrightarrow> big_union(M,A,z) \<longleftrightarrow> z = \<Union>(A)"
unfolding big_union_def
by (blast dest: transM)
lemma (in M_trivial) Union_closed [intro,simp]:
- "M(A) ==> M(\<Union>(A))"
+ "M(A) \<Longrightarrow> M(\<Union>(A))"
by (insert Union_ax, simp add: Union_ax_def)
lemma (in M_trivial) Un_closed [intro,simp]:
- "[| M(A); M(B) |] ==> M(A \<union> B)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> M(A \<union> B)"
by (simp only: Un_eq_Union, blast)
lemma (in M_trivial) cons_closed [intro,simp]:
- "[| M(a); M(A) |] ==> M(cons(a,A))"
+ "\<lbrakk>M(a); M(A)\<rbrakk> \<Longrightarrow> M(cons(a,A))"
by (subst cons_eq [symmetric], blast)
lemma (in M_trivial) cons_abs [simp]:
- "[| M(b); M(z) |] ==> is_cons(M,a,b,z) \<longleftrightarrow> z = cons(a,b)"
+ "\<lbrakk>M(b); M(z)\<rbrakk> \<Longrightarrow> is_cons(M,a,b,z) \<longleftrightarrow> z = cons(a,b)"
by (simp add: is_cons_def, blast intro: transM)
lemma (in M_trivial) successor_abs [simp]:
- "[| M(a); M(z) |] ==> successor(M,a,z) \<longleftrightarrow> z = succ(a)"
+ "\<lbrakk>M(a); M(z)\<rbrakk> \<Longrightarrow> successor(M,a,z) \<longleftrightarrow> z = succ(a)"
by (simp add: successor_def, blast)
lemma (in M_trans) succ_in_MD [dest!]:
@@ -716,7 +716,7 @@
subsubsection\<open>Absoluteness for Separation and Replacement\<close>
lemma (in M_trans) separation_closed [intro,simp]:
- "[| separation(M,P); M(A) |] ==> M(Collect(A,P))"
+ "\<lbrakk>separation(M,P); M(A)\<rbrakk> \<Longrightarrow> M(Collect(A,P))"
apply (insert separation, simp add: separation_def)
apply (drule rspec, assumption, clarify)
apply (subgoal_tac "y = Collect(A,P)", blast)
@@ -728,7 +728,7 @@
by (simp add: separation_def is_Collect_def)
lemma (in M_trans) Collect_abs [simp]:
- "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) \<longleftrightarrow> z = Collect(A,P)"
+ "\<lbrakk>M(A); M(z)\<rbrakk> \<Longrightarrow> is_Collect(M,A,P,z) \<longleftrightarrow> z = Collect(A,P)"
unfolding is_Collect_def
by (blast dest: transM)
@@ -736,24 +736,24 @@
lemma is_Replace_cong [cong]:
- "[| A=A';
- !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y);
- z=z' |]
- ==> is_Replace(M, A, %x y. P(x,y), z) \<longleftrightarrow>
+ "\<lbrakk>A=A';
+ \<And>x y. \<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y);
+ z=z'\<rbrakk>
+ \<Longrightarrow> is_Replace(M, A, %x y. P(x,y), z) \<longleftrightarrow>
is_Replace(M, A', %x y. P'(x,y), z')"
by (simp add: is_Replace_def)
lemma (in M_trans) univalent_Replace_iff:
- "[| M(A); univalent(M,A,P);
- !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |]
- ==> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A & P(x,u))"
+ "\<lbrakk>M(A); univalent(M,A,P);
+ \<And>x y. \<lbrakk>x\<in>A; P(x,y)\<rbrakk> \<Longrightarrow> M(y)\<rbrakk>
+ \<Longrightarrow> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A & P(x,u))"
unfolding Replace_iff univalent_def
by (blast dest: transM)
(*The last premise expresses that P takes M to M*)
lemma (in M_trans) strong_replacement_closed [intro,simp]:
- "[| strong_replacement(M,P); M(A); univalent(M,A,P);
- !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |] ==> M(Replace(A,P))"
+ "\<lbrakk>strong_replacement(M,P); M(A); univalent(M,A,P);
+ \<And>x y. \<lbrakk>x\<in>A; P(x,y)\<rbrakk> \<Longrightarrow> M(y)\<rbrakk> \<Longrightarrow> M(Replace(A,P))"
apply (simp add: strong_replacement_def)
apply (drule_tac x=A in rspec, safe)
apply (subgoal_tac "Replace(A,P) = Y")
@@ -764,9 +764,9 @@
done
lemma (in M_trans) Replace_abs:
- "[| M(A); M(z); univalent(M,A,P);
- !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |]
- ==> is_Replace(M,A,P,z) \<longleftrightarrow> z = Replace(A,P)"
+ "\<lbrakk>M(A); M(z); univalent(M,A,P);
+ \<And>x y. \<lbrakk>x\<in>A; P(x,y)\<rbrakk> \<Longrightarrow> M(y)\<rbrakk>
+ \<Longrightarrow> is_Replace(M,A,P,z) \<longleftrightarrow> z = Replace(A,P)"
apply (simp add: is_Replace_def)
apply (rule iffI)
apply (rule equality_iffI)
@@ -783,8 +783,8 @@
even for f \<in> M -> M.
*)
lemma (in M_trans) RepFun_closed:
- "[| strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
- ==> M(RepFun(A,f))"
+ "\<lbrakk>strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x\<in>A. M(f(x))\<rbrakk>
+ \<Longrightarrow> M(RepFun(A,f))"
apply (simp add: RepFun_def)
done
@@ -794,8 +794,8 @@
text\<open>Better than \<open>RepFun_closed\<close> when having the formula \<^term>\<open>x\<in>A\<close>
makes relativization easier.\<close>
lemma (in M_trans) RepFun_closed2:
- "[| strong_replacement(M, \<lambda>x y. x\<in>A & y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
- ==> M(RepFun(A, %x. f(x)))"
+ "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A & y = f(x)); M(A); \<forall>x\<in>A. M(f(x))\<rbrakk>
+ \<Longrightarrow> M(RepFun(A, %x. f(x)))"
apply (simp add: RepFun_def)
apply (frule strong_replacement_closed, assumption)
apply (auto dest: transM simp add: Replace_conj_eq univalent_def)
@@ -805,26 +805,26 @@
definition
is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where
- "is_lambda(M, A, is_b, z) ==
+ "is_lambda(M, A, is_b, z) \<equiv>
\<forall>p[M]. p \<in> z \<longleftrightarrow>
(\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))"
lemma (in M_trivial) lam_closed:
- "[| strong_replacement(M, \<lambda>x y. y = <x,b(x)>); M(A); \<forall>x\<in>A. M(b(x)) |]
- ==> M(\<lambda>x\<in>A. b(x))"
+ "\<lbrakk>strong_replacement(M, \<lambda>x y. y = <x,b(x)>); M(A); \<forall>x\<in>A. M(b(x))\<rbrakk>
+ \<Longrightarrow> M(\<lambda>x\<in>A. b(x))"
by (simp add: lam_def, blast intro: RepFun_closed dest: transM)
text\<open>Better than \<open>lam_closed\<close>: has the formula \<^term>\<open>x\<in>A\<close>\<close>
lemma (in M_trivial) lam_closed2:
- "[|strong_replacement(M, \<lambda>x y. x\<in>A & y = \<langle>x, b(x)\<rangle>);
- M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m))|] ==> M(Lambda(A,b))"
+ "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A & y = \<langle>x, b(x)\<rangle>);
+ M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m))\<rbrakk> \<Longrightarrow> M(Lambda(A,b))"
apply (simp add: lam_def)
apply (blast intro: RepFun_closed2 dest: transM)
done
lemma (in M_trivial) lambda_abs2:
- "[| Relation1(M,A,is_b,b); M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m)); M(z) |]
- ==> is_lambda(M,A,is_b,z) \<longleftrightarrow> z = Lambda(A,b)"
+ "\<lbrakk>Relation1(M,A,is_b,b); M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m)); M(z)\<rbrakk>
+ \<Longrightarrow> is_lambda(M,A,is_b,z) \<longleftrightarrow> z = Lambda(A,b)"
apply (simp add: Relation1_def is_lambda_def)
apply (rule iffI)
prefer 2 apply (simp add: lam_def)
@@ -836,14 +836,14 @@
done
lemma is_lambda_cong [cong]:
- "[| A=A'; z=z';
- !!x y. [| x\<in>A; M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
- ==> is_lambda(M, A, %x y. is_b(x,y), z) \<longleftrightarrow>
+ "\<lbrakk>A=A'; z=z';
+ \<And>x y. \<lbrakk>x\<in>A; M(x); M(y)\<rbrakk> \<Longrightarrow> is_b(x,y) \<longleftrightarrow> is_b'(x,y)\<rbrakk>
+ \<Longrightarrow> is_lambda(M, A, %x y. is_b(x,y), z) \<longleftrightarrow>
is_lambda(M, A', %x y. is_b'(x,y), z')"
by (simp add: is_lambda_def)
lemma (in M_trans) image_abs [simp]:
- "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) \<longleftrightarrow> z = r``A"
+ "\<lbrakk>M(r); M(A); M(z)\<rbrakk> \<Longrightarrow> image(M,r,A,z) \<longleftrightarrow> z = r``A"
apply (simp add: image_def)
apply (rule iffI)
apply (blast intro!: equalityI dest: transM, blast)
@@ -862,7 +862,7 @@
real powerset.\<close>
lemma (in M_trans) powerset_imp_subset_Pow:
- "[| powerset(M,x,y); M(y) |] ==> y \<subseteq> Pow(x)"
+ "\<lbrakk>powerset(M,x,y); M(y)\<rbrakk> \<Longrightarrow> y \<subseteq> Pow(x)"
apply (simp add: powerset_def)
apply (blast dest: transM)
done
@@ -892,23 +892,23 @@
subsubsection\<open>Absoluteness for the Natural Numbers\<close>
lemma (in M_trivial) nat_into_M [intro]:
- "n \<in> nat ==> M(n)"
+ "n \<in> nat \<Longrightarrow> M(n)"
by (induct n rule: nat_induct, simp_all)
lemma (in M_trans) nat_case_closed [intro,simp]:
- "[|M(k); M(a); \<forall>m[M]. M(b(m))|] ==> M(nat_case(a,b,k))"
+ "\<lbrakk>M(k); M(a); \<forall>m[M]. M(b(m))\<rbrakk> \<Longrightarrow> M(nat_case(a,b,k))"
apply (case_tac "k=0", simp)
apply (case_tac "\<exists>m. k = succ(m)", force)
apply (simp add: nat_case_def)
done
lemma (in M_trivial) quasinat_abs [simp]:
- "M(z) ==> is_quasinat(M,z) \<longleftrightarrow> quasinat(z)"
+ "M(z) \<Longrightarrow> is_quasinat(M,z) \<longleftrightarrow> quasinat(z)"
by (auto simp add: is_quasinat_def quasinat_def)
lemma (in M_trivial) nat_case_abs [simp]:
- "[| relation1(M,is_b,b); M(k); M(z) |]
- ==> is_nat_case(M,a,is_b,k,z) \<longleftrightarrow> z = nat_case(a,b,k)"
+ "\<lbrakk>relation1(M,is_b,b); M(k); M(z)\<rbrakk>
+ \<Longrightarrow> is_nat_case(M,a,is_b,k,z) \<longleftrightarrow> z = nat_case(a,b,k)"
apply (case_tac "quasinat(k)")
prefer 2
apply (simp add: is_nat_case_def non_nat_case)
@@ -922,9 +922,9 @@
causes the error "Failed congruence proof!" It may be better to replace
is_nat_case by nat_case before attempting congruence reasoning.*)
lemma is_nat_case_cong:
- "[| a = a'; k = k'; z = z'; M(z');
- !!x y. [| M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
- ==> is_nat_case(M, a, is_b, k, z) \<longleftrightarrow> is_nat_case(M, a', is_b', k', z')"
+ "\<lbrakk>a = a'; k = k'; z = z'; M(z');
+ \<And>x y. \<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> is_b(x,y) \<longleftrightarrow> is_b'(x,y)\<rbrakk>
+ \<Longrightarrow> is_nat_case(M, a, is_b, k, z) \<longleftrightarrow> is_nat_case(M, a', is_b', k', z')"
by (simp add: is_nat_case_def)
@@ -932,43 +932,43 @@
text\<open>These results constitute Theorem IV 5.1 of Kunen (page 126).\<close>
lemma (in M_trans) lt_closed:
- "[| j<i; M(i) |] ==> M(j)"
+ "\<lbrakk>j<i; M(i)\<rbrakk> \<Longrightarrow> M(j)"
by (blast dest: ltD intro: transM)
lemma (in M_trans) transitive_set_abs [simp]:
- "M(a) ==> transitive_set(M,a) \<longleftrightarrow> Transset(a)"
+ "M(a) \<Longrightarrow> transitive_set(M,a) \<longleftrightarrow> Transset(a)"
by (simp add: transitive_set_def Transset_def)
lemma (in M_trans) ordinal_abs [simp]:
- "M(a) ==> ordinal(M,a) \<longleftrightarrow> Ord(a)"
+ "M(a) \<Longrightarrow> ordinal(M,a) \<longleftrightarrow> Ord(a)"
by (simp add: ordinal_def Ord_def)
lemma (in M_trivial) limit_ordinal_abs [simp]:
- "M(a) ==> limit_ordinal(M,a) \<longleftrightarrow> Limit(a)"
+ "M(a) \<Longrightarrow> limit_ordinal(M,a) \<longleftrightarrow> Limit(a)"
apply (unfold Limit_def limit_ordinal_def)
apply (simp add: Ord_0_lt_iff)
apply (simp add: lt_def, blast)
done
lemma (in M_trivial) successor_ordinal_abs [simp]:
- "M(a) ==> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) & (\<exists>b[M]. a = succ(b))"
+ "M(a) \<Longrightarrow> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) & (\<exists>b[M]. a = succ(b))"
apply (simp add: successor_ordinal_def, safe)
apply (drule Ord_cases_disj, auto)
done
lemma finite_Ord_is_nat:
- "[| Ord(a); ~ Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a \<in> nat"
+ "\<lbrakk>Ord(a); \<not> Limit(a); \<forall>x\<in>a. \<not> Limit(x)\<rbrakk> \<Longrightarrow> a \<in> nat"
by (induct a rule: trans_induct3, simp_all)
lemma (in M_trivial) finite_ordinal_abs [simp]:
- "M(a) ==> finite_ordinal(M,a) \<longleftrightarrow> a \<in> nat"
+ "M(a) \<Longrightarrow> finite_ordinal(M,a) \<longleftrightarrow> a \<in> nat"
apply (simp add: finite_ordinal_def)
apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord
dest: Ord_trans naturals_not_limit)
done
lemma Limit_non_Limit_implies_nat:
- "[| Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a = nat"
+ "\<lbrakk>Limit(a); \<forall>x\<in>a. \<not> Limit(x)\<rbrakk> \<Longrightarrow> a = nat"
apply (rule le_anti_sym)
apply (rule all_lt_imp_le, blast, blast intro: Limit_is_Ord)
apply (simp add: lt_def)
@@ -977,21 +977,21 @@
done
lemma (in M_trivial) omega_abs [simp]:
- "M(a) ==> omega(M,a) \<longleftrightarrow> a = nat"
+ "M(a) \<Longrightarrow> omega(M,a) \<longleftrightarrow> a = nat"
apply (simp add: omega_def)
apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit)
done
lemma (in M_trivial) number1_abs [simp]:
- "M(a) ==> number1(M,a) \<longleftrightarrow> a = 1"
+ "M(a) \<Longrightarrow> number1(M,a) \<longleftrightarrow> a = 1"
by (simp add: number1_def)
lemma (in M_trivial) number2_abs [simp]:
- "M(a) ==> number2(M,a) \<longleftrightarrow> a = succ(1)"
+ "M(a) \<Longrightarrow> number2(M,a) \<longleftrightarrow> a = succ(1)"
by (simp add: number2_def)
lemma (in M_trivial) number3_abs [simp]:
- "M(a) ==> number3(M,a) \<longleftrightarrow> a = succ(succ(1))"
+ "M(a) \<Longrightarrow> number3(M,a) \<longleftrightarrow> a = succ(succ(1))"
by (simp add: number3_def)
text\<open>Kunen continued to 20...\<close>
@@ -1011,48 +1011,48 @@
definition
natnumber :: "[i=>o,i,i] => o"
- "natnumber(M,n,x) == natnumber_aux(M,n)`x = 1"
+ "natnumber(M,n,x) \<equiv> natnumber_aux(M,n)`x = 1"
lemma (in M_trivial) [simp]:
- "natnumber(M,0,x) == x=0"
+ "natnumber(M,0,x) \<equiv> x=0"
*)
subsection\<open>Some instances of separation and strong replacement\<close>
locale M_basic = M_trivial +
assumes Inter_separation:
- "M(A) ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A \<longrightarrow> x\<in>y)"
+ "M(A) \<Longrightarrow> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A \<longrightarrow> x\<in>y)"
and Diff_separation:
- "M(B) ==> separation(M, \<lambda>x. x \<notin> B)"
+ "M(B) \<Longrightarrow> separation(M, \<lambda>x. x \<notin> B)"
and cartprod_separation:
- "[| M(A); M(B) |]
- ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,z)))"
+ "\<lbrakk>M(A); M(B)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,z)))"
and image_separation:
- "[| M(A); M(r) |]
- ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,p)))"
+ "\<lbrakk>M(A); M(r)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,p)))"
and converse_separation:
- "M(r) ==> separation(M,
+ "M(r) \<Longrightarrow> separation(M,
\<lambda>z. \<exists>p[M]. p\<in>r & (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) & pair(M,y,x,z)))"
and restrict_separation:
- "M(A) ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. pair(M,x,y,z)))"
+ "M(A) \<Longrightarrow> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. pair(M,x,y,z)))"
and comp_separation:
- "[| M(r); M(s) |]
- ==> separation(M, \<lambda>xz. \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
+ "\<lbrakk>M(r); M(s)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>xz. \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) &
xy\<in>s & yz\<in>r)"
and pred_separation:
- "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
+ "\<lbrakk>M(r); M(x)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
and Memrel_separation:
"separation(M, \<lambda>z. \<exists>x[M]. \<exists>y[M]. pair(M,x,y,z) & x \<in> y)"
and funspace_succ_replacement:
- "M(n) ==>
+ "M(n) \<Longrightarrow>
strong_replacement(M, \<lambda>p z. \<exists>f[M]. \<exists>b[M]. \<exists>nb[M]. \<exists>cnbf[M].
pair(M,f,b,p) & pair(M,n,b,nb) & is_cons(M,nb,f,cnbf) &
upair(M,cnbf,cnbf,z))"
and is_recfun_separation:
\<comment> \<open>for well-founded recursion: used to prove \<open>is_recfun_equal\<close>\<close>
- "[| M(r); M(f); M(g); M(a); M(b) |]
- ==> separation(M,
+ "\<lbrakk>M(r); M(f); M(g); M(a); M(b)\<rbrakk>
+ \<Longrightarrow> separation(M,
\<lambda>x. \<exists>xa[M]. \<exists>xb[M].
pair(M,x,a,xa) & xa \<in> r & pair(M,x,b,xb) & xb \<in> r &
(\<exists>fx[M]. \<exists>gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) &
@@ -1060,9 +1060,9 @@
and power_ax: "power_ax(M)"
lemma (in M_trivial) cartprod_iff_lemma:
- "[| M(C); \<forall>u[M]. u \<in> C \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}});
- powerset(M, A \<union> B, p1); powerset(M, p1, p2); M(p2) |]
- ==> C = {u \<in> p2 . \<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}}"
+ "\<lbrakk>M(C); \<forall>u[M]. u \<in> C \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}});
+ powerset(M, A \<union> B, p1); powerset(M, p1, p2); M(p2)\<rbrakk>
+ \<Longrightarrow> C = {u \<in> p2 . \<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}}"
apply (simp add: powerset_def)
apply (rule equalityI, clarify, simp)
apply (frule transM, assumption)
@@ -1073,8 +1073,8 @@
done
lemma (in M_basic) cartprod_iff:
- "[| M(A); M(B); M(C) |]
- ==> cartprod(M,A,B,C) \<longleftrightarrow>
+ "\<lbrakk>M(A); M(B); M(C)\<rbrakk>
+ \<Longrightarrow> cartprod(M,A,B,C) \<longleftrightarrow>
(\<exists>p1[M]. \<exists>p2[M]. powerset(M,A \<union> B,p1) & powerset(M,p1,p2) &
C = {z \<in> p2. \<exists>x\<in>A. \<exists>y\<in>B. z = <x,y>})"
apply (simp add: Pair_def cartprod_def, safe)
@@ -1091,7 +1091,7 @@
done
lemma (in M_basic) cartprod_closed_lemma:
- "[| M(A); M(B) |] ==> \<exists>C[M]. cartprod(M,A,B,C)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> \<exists>C[M]. cartprod(M,A,B,C)"
apply (simp del: cartprod_abs add: cartprod_iff)
apply (insert power_ax, simp add: power_ax_def)
apply (frule_tac x="A \<union> B" and P="\<lambda>x. rex(M,Q(x))" for Q in rspec)
@@ -1104,15 +1104,15 @@
text\<open>All the lemmas above are necessary because Powerset is not absolute.
I should have used Replacement instead!\<close>
lemma (in M_basic) cartprod_closed [intro,simp]:
- "[| M(A); M(B) |] ==> M(A*B)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> M(A*B)"
by (frule cartprod_closed_lemma, assumption, force)
lemma (in M_basic) sum_closed [intro,simp]:
- "[| M(A); M(B) |] ==> M(A+B)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> M(A+B)"
by (simp add: sum_def)
lemma (in M_basic) sum_abs [simp]:
- "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) \<longleftrightarrow> (Z = A+B)"
+ "\<lbrakk>M(A); M(B); M(Z)\<rbrakk> \<Longrightarrow> is_sum(M,A,B,Z) \<longleftrightarrow> (Z = A+B)"
by (simp add: is_sum_def sum_def singleton_0 nat_into_M)
lemma (in M_trivial) Inl_in_M_iff [iff]:
@@ -1120,7 +1120,7 @@
by (simp add: Inl_def)
lemma (in M_trivial) Inl_abs [simp]:
- "M(Z) ==> is_Inl(M,a,Z) \<longleftrightarrow> (Z = Inl(a))"
+ "M(Z) \<Longrightarrow> is_Inl(M,a,Z) \<longleftrightarrow> (Z = Inl(a))"
by (simp add: is_Inl_def Inl_def)
lemma (in M_trivial) Inr_in_M_iff [iff]:
@@ -1128,14 +1128,14 @@
by (simp add: Inr_def)
lemma (in M_trivial) Inr_abs [simp]:
- "M(Z) ==> is_Inr(M,a,Z) \<longleftrightarrow> (Z = Inr(a))"
+ "M(Z) \<Longrightarrow> is_Inr(M,a,Z) \<longleftrightarrow> (Z = Inr(a))"
by (simp add: is_Inr_def Inr_def)
subsubsection \<open>converse of a relation\<close>
lemma (in M_basic) M_converse_iff:
- "M(r) ==>
+ "M(r) \<Longrightarrow>
converse(r) =
{z \<in> \<Union>(\<Union>(r)) * \<Union>(\<Union>(r)).
\<exists>p\<in>r. \<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> & z = \<langle>y,x\<rangle>}"
@@ -1146,13 +1146,13 @@
done
lemma (in M_basic) converse_closed [intro,simp]:
- "M(r) ==> M(converse(r))"
+ "M(r) \<Longrightarrow> M(converse(r))"
apply (simp add: M_converse_iff)
apply (insert converse_separation [of r], simp)
done
lemma (in M_basic) converse_abs [simp]:
- "[| M(r); M(z) |] ==> is_converse(M,r,z) \<longleftrightarrow> z = converse(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> is_converse(M,r,z) \<longleftrightarrow> z = converse(r)"
apply (simp add: is_converse_def)
apply (rule iffI)
prefer 2 apply blast
@@ -1165,106 +1165,106 @@
subsubsection \<open>image, preimage, domain, range\<close>
lemma (in M_basic) image_closed [intro,simp]:
- "[| M(A); M(r) |] ==> M(r``A)"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> M(r``A)"
apply (simp add: image_iff_Collect)
apply (insert image_separation [of A r], simp)
done
lemma (in M_basic) vimage_abs [simp]:
- "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) \<longleftrightarrow> z = r-``A"
+ "\<lbrakk>M(r); M(A); M(z)\<rbrakk> \<Longrightarrow> pre_image(M,r,A,z) \<longleftrightarrow> z = r-``A"
apply (simp add: pre_image_def)
apply (rule iffI)
apply (blast intro!: equalityI dest: transM, blast)
done
lemma (in M_basic) vimage_closed [intro,simp]:
- "[| M(A); M(r) |] ==> M(r-``A)"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> M(r-``A)"
by (simp add: vimage_def)
subsubsection\<open>Domain, range and field\<close>
lemma (in M_trans) domain_abs [simp]:
- "[| M(r); M(z) |] ==> is_domain(M,r,z) \<longleftrightarrow> z = domain(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> is_domain(M,r,z) \<longleftrightarrow> z = domain(r)"
apply (simp add: is_domain_def)
apply (blast intro!: equalityI dest: transM)
done
lemma (in M_basic) domain_closed [intro,simp]:
- "M(r) ==> M(domain(r))"
+ "M(r) \<Longrightarrow> M(domain(r))"
apply (simp add: domain_eq_vimage)
done
lemma (in M_trans) range_abs [simp]:
- "[| M(r); M(z) |] ==> is_range(M,r,z) \<longleftrightarrow> z = range(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> is_range(M,r,z) \<longleftrightarrow> z = range(r)"
apply (simp add: is_range_def)
apply (blast intro!: equalityI dest: transM)
done
lemma (in M_basic) range_closed [intro,simp]:
- "M(r) ==> M(range(r))"
+ "M(r) \<Longrightarrow> M(range(r))"
apply (simp add: range_eq_image)
done
lemma (in M_basic) field_abs [simp]:
- "[| M(r); M(z) |] ==> is_field(M,r,z) \<longleftrightarrow> z = field(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> is_field(M,r,z) \<longleftrightarrow> z = field(r)"
by (simp add: is_field_def field_def)
lemma (in M_basic) field_closed [intro,simp]:
- "M(r) ==> M(field(r))"
+ "M(r) \<Longrightarrow> M(field(r))"
by (simp add: field_def)
subsubsection\<open>Relations, functions and application\<close>
lemma (in M_trans) relation_abs [simp]:
- "M(r) ==> is_relation(M,r) \<longleftrightarrow> relation(r)"
+ "M(r) \<Longrightarrow> is_relation(M,r) \<longleftrightarrow> relation(r)"
apply (simp add: is_relation_def relation_def)
apply (blast dest!: bspec dest: pair_components_in_M)+
done
lemma (in M_trivial) function_abs [simp]:
- "M(r) ==> is_function(M,r) \<longleftrightarrow> function(r)"
+ "M(r) \<Longrightarrow> is_function(M,r) \<longleftrightarrow> function(r)"
apply (simp add: is_function_def function_def, safe)
apply (frule transM, assumption)
apply (blast dest: pair_components_in_M)+
done
lemma (in M_basic) apply_closed [intro,simp]:
- "[|M(f); M(a)|] ==> M(f`a)"
+ "\<lbrakk>M(f); M(a)\<rbrakk> \<Longrightarrow> M(f`a)"
by (simp add: apply_def)
lemma (in M_basic) apply_abs [simp]:
- "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) \<longleftrightarrow> f`x = y"
+ "\<lbrakk>M(f); M(x); M(y)\<rbrakk> \<Longrightarrow> fun_apply(M,f,x,y) \<longleftrightarrow> f`x = y"
apply (simp add: fun_apply_def apply_def, blast)
done
lemma (in M_trivial) typed_function_abs [simp]:
- "[| M(A); M(f) |] ==> typed_function(M,A,B,f) \<longleftrightarrow> f \<in> A -> B"
+ "\<lbrakk>M(A); M(f)\<rbrakk> \<Longrightarrow> typed_function(M,A,B,f) \<longleftrightarrow> f \<in> A -> B"
apply (auto simp add: typed_function_def relation_def Pi_iff)
apply (blast dest: pair_components_in_M)+
done
lemma (in M_basic) injection_abs [simp]:
- "[| M(A); M(f) |] ==> injection(M,A,B,f) \<longleftrightarrow> f \<in> inj(A,B)"
+ "\<lbrakk>M(A); M(f)\<rbrakk> \<Longrightarrow> injection(M,A,B,f) \<longleftrightarrow> f \<in> inj(A,B)"
apply (simp add: injection_def apply_iff inj_def)
apply (blast dest: transM [of _ A])
done
lemma (in M_basic) surjection_abs [simp]:
- "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) \<longleftrightarrow> f \<in> surj(A,B)"
+ "\<lbrakk>M(A); M(B); M(f)\<rbrakk> \<Longrightarrow> surjection(M,A,B,f) \<longleftrightarrow> f \<in> surj(A,B)"
by (simp add: surjection_def surj_def)
lemma (in M_basic) bijection_abs [simp]:
- "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) \<longleftrightarrow> f \<in> bij(A,B)"
+ "\<lbrakk>M(A); M(B); M(f)\<rbrakk> \<Longrightarrow> bijection(M,A,B,f) \<longleftrightarrow> f \<in> bij(A,B)"
by (simp add: bijection_def bij_def)
subsubsection\<open>Composition of relations\<close>
lemma (in M_basic) M_comp_iff:
- "[| M(r); M(s) |]
- ==> r O s =
+ "\<lbrakk>M(r); M(s)\<rbrakk>
+ \<Longrightarrow> r O s =
{xz \<in> domain(s) * range(r).
\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. xz = \<langle>x,z\<rangle> & \<langle>x,y\<rangle> \<in> s & \<langle>y,z\<rangle> \<in> r}"
apply (simp add: comp_def)
@@ -1275,13 +1275,13 @@
done
lemma (in M_basic) comp_closed [intro,simp]:
- "[| M(r); M(s) |] ==> M(r O s)"
+ "\<lbrakk>M(r); M(s)\<rbrakk> \<Longrightarrow> M(r O s)"
apply (simp add: M_comp_iff)
apply (insert comp_separation [of r s], simp)
done
lemma (in M_basic) composition_abs [simp]:
- "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) \<longleftrightarrow> t = r O s"
+ "\<lbrakk>M(r); M(s); M(t)\<rbrakk> \<Longrightarrow> composition(M,r,s,t) \<longleftrightarrow> t = r O s"
apply safe
txt\<open>Proving \<^term>\<open>composition(M, r, s, r O s)\<close>\<close>
prefer 2
@@ -1295,54 +1295,54 @@
text\<open>no longer needed\<close>
lemma (in M_basic) restriction_is_function:
- "[| restriction(M,f,A,z); function(f); M(f); M(A); M(z) |]
- ==> function(z)"
+ "\<lbrakk>restriction(M,f,A,z); function(f); M(f); M(A); M(z)\<rbrakk>
+ \<Longrightarrow> function(z)"
apply (simp add: restriction_def ball_iff_equiv)
apply (unfold function_def, blast)
done
lemma (in M_trans) restriction_abs [simp]:
- "[| M(f); M(A); M(z) |]
- ==> restriction(M,f,A,z) \<longleftrightarrow> z = restrict(f,A)"
+ "\<lbrakk>M(f); M(A); M(z)\<rbrakk>
+ \<Longrightarrow> restriction(M,f,A,z) \<longleftrightarrow> z = restrict(f,A)"
apply (simp add: ball_iff_equiv restriction_def restrict_def)
apply (blast intro!: equalityI dest: transM)
done
lemma (in M_trans) M_restrict_iff:
- "M(r) ==> restrict(r,A) = {z \<in> r . \<exists>x\<in>A. \<exists>y[M]. z = \<langle>x, y\<rangle>}"
+ "M(r) \<Longrightarrow> restrict(r,A) = {z \<in> r . \<exists>x\<in>A. \<exists>y[M]. z = \<langle>x, y\<rangle>}"
by (simp add: restrict_def, blast dest: transM)
lemma (in M_basic) restrict_closed [intro,simp]:
- "[| M(A); M(r) |] ==> M(restrict(r,A))"
+ "\<lbrakk>M(A); M(r)\<rbrakk> \<Longrightarrow> M(restrict(r,A))"
apply (simp add: M_restrict_iff)
apply (insert restrict_separation [of A], simp)
done
lemma (in M_trans) Inter_abs [simp]:
- "[| M(A); M(z) |] ==> big_inter(M,A,z) \<longleftrightarrow> z = \<Inter>(A)"
+ "\<lbrakk>M(A); M(z)\<rbrakk> \<Longrightarrow> big_inter(M,A,z) \<longleftrightarrow> z = \<Inter>(A)"
apply (simp add: big_inter_def Inter_def)
apply (blast intro!: equalityI dest: transM)
done
lemma (in M_basic) Inter_closed [intro,simp]:
- "M(A) ==> M(\<Inter>(A))"
+ "M(A) \<Longrightarrow> M(\<Inter>(A))"
by (insert Inter_separation, simp add: Inter_def)
lemma (in M_basic) Int_closed [intro,simp]:
- "[| M(A); M(B) |] ==> M(A \<inter> B)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> M(A \<inter> B)"
apply (subgoal_tac "M({A,B})")
apply (frule Inter_closed, force+)
done
lemma (in M_basic) Diff_closed [intro,simp]:
- "[|M(A); M(B)|] ==> M(A-B)"
+ "\<lbrakk>M(A); M(B)\<rbrakk> \<Longrightarrow> M(A-B)"
by (insert Diff_separation, simp add: Diff_def)
subsubsection\<open>Some Facts About Separation Axioms\<close>
lemma (in M_basic) separation_conj:
- "[|separation(M,P); separation(M,Q)|] ==> separation(M, \<lambda>z. P(z) & Q(z))"
+ "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) & Q(z))"
by (simp del: separation_closed
add: separation_iff Collect_Int_Collect_eq [symmetric])
@@ -1352,37 +1352,37 @@
by blast
lemma Diff_Collect_eq:
- "A - Collect(A,P) = Collect(A, %x. ~ P(x))"
+ "A - Collect(A,P) = Collect(A, %x. \<not> P(x))"
by blast
lemma (in M_trans) Collect_rall_eq:
- "M(Y) ==> Collect(A, %x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
+ "M(Y) \<Longrightarrow> Collect(A, %x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
(if Y=0 then A else (\<Inter>y \<in> Y. {x \<in> A. P(x,y)}))"
by (simp,blast dest: transM)
lemma (in M_basic) separation_disj:
- "[|separation(M,P); separation(M,Q)|] ==> separation(M, \<lambda>z. P(z) | Q(z))"
+ "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk> \<Longrightarrow> separation(M, \<lambda>z. P(z) | Q(z))"
by (simp del: separation_closed
add: separation_iff Collect_Un_Collect_eq [symmetric])
lemma (in M_basic) separation_neg:
- "separation(M,P) ==> separation(M, \<lambda>z. ~P(z))"
+ "separation(M,P) \<Longrightarrow> separation(M, \<lambda>z. \<not>P(z))"
by (simp del: separation_closed
add: separation_iff Diff_Collect_eq [symmetric])
lemma (in M_basic) separation_imp:
- "[|separation(M,P); separation(M,Q)|]
- ==> separation(M, \<lambda>z. P(z) \<longrightarrow> Q(z))"
+ "\<lbrakk>separation(M,P); separation(M,Q)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>z. P(z) \<longrightarrow> Q(z))"
by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric])
text\<open>This result is a hint of how little can be done without the Reflection
Theorem. The quantifier has to be bounded by a set. We also need another
instance of Separation!\<close>
lemma (in M_basic) separation_rall:
- "[|M(Y); \<forall>y[M]. separation(M, \<lambda>x. P(x,y));
- \<forall>z[M]. strong_replacement(M, \<lambda>x y. y = {u \<in> z . P(u,x)})|]
- ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y))"
+ "\<lbrakk>M(Y); \<forall>y[M]. separation(M, \<lambda>x. P(x,y));
+ \<forall>z[M]. strong_replacement(M, \<lambda>x y. y = {u \<in> z . P(u,x)})\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y))"
apply (simp del: separation_closed rall_abs
add: separation_iff Collect_rall_eq)
apply (blast intro!: RepFun_closed dest: transM)
@@ -1394,7 +1394,7 @@
text\<open>The assumption \<^term>\<open>M(A->B)\<close> is unusual, but essential: in
all but trivial cases, A->B cannot be expected to belong to \<^term>\<open>M\<close>.\<close>
lemma (in M_trivial) is_funspace_abs [simp]:
- "[|M(A); M(B); M(F); M(A->B)|] ==> is_funspace(M,A,B,F) \<longleftrightarrow> F = A->B"
+ "\<lbrakk>M(A); M(B); M(F); M(A->B)\<rbrakk> \<Longrightarrow> is_funspace(M,A,B,F) \<longleftrightarrow> F = A->B"
apply (simp add: is_funspace_def)
apply (rule iffI)
prefer 2 apply blast
@@ -1403,7 +1403,7 @@
done
lemma (in M_basic) succ_fun_eq2:
- "[|M(B); M(n->B)|] ==>
+ "\<lbrakk>M(B); M(n->B)\<rbrakk> \<Longrightarrow>
succ(n) -> B =
\<Union>{z. p \<in> (n->B)*B, \<exists>f[M]. \<exists>b[M]. p = <f,b> & z = {cons(<n,b>, f)}}"
apply (simp add: succ_fun_eq)
@@ -1411,7 +1411,7 @@
done
lemma (in M_basic) funspace_succ:
- "[|M(n); M(B); M(n->B) |] ==> M(succ(n) -> B)"
+ "\<lbrakk>M(n); M(B); M(n->B)\<rbrakk> \<Longrightarrow> M(succ(n) -> B)"
apply (insert funspace_succ_replacement [of n], simp)
apply (force simp add: succ_fun_eq2 univalent_def)
done
@@ -1420,7 +1420,7 @@
absoluteness of transitive closure. See the definition of
\<open>rtrancl_alt\<close> in in \<open>WF_absolute.thy\<close>.\<close>
lemma (in M_basic) finite_funspace_closed [intro,simp]:
- "[|n\<in>nat; M(B)|] ==> M(n->B)"
+ "\<lbrakk>n\<in>nat; M(B)\<rbrakk> \<Longrightarrow> M(n->B)"
apply (induct_tac n, simp)
apply (simp add: funspace_succ nat_into_M)
done
@@ -1430,38 +1430,38 @@
definition
is_bool_of_o :: "[i=>o, o, i] => o" where
- "is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))"
+ "is_bool_of_o(M,P,z) \<equiv> (P & number1(M,z)) | (\<not>P & empty(M,z))"
definition
is_not :: "[i=>o, i, i] => o" where
- "is_not(M,a,z) == (number1(M,a) & empty(M,z)) |
- (~number1(M,a) & number1(M,z))"
+ "is_not(M,a,z) \<equiv> (number1(M,a) & empty(M,z)) |
+ (\<not>number1(M,a) & number1(M,z))"
definition
is_and :: "[i=>o, i, i, i] => o" where
- "is_and(M,a,b,z) == (number1(M,a) & z=b) |
- (~number1(M,a) & empty(M,z))"
+ "is_and(M,a,b,z) \<equiv> (number1(M,a) & z=b) |
+ (\<not>number1(M,a) & empty(M,z))"
definition
is_or :: "[i=>o, i, i, i] => o" where
- "is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) |
- (~number1(M,a) & z=b)"
+ "is_or(M,a,b,z) \<equiv> (number1(M,a) & number1(M,z)) |
+ (\<not>number1(M,a) & z=b)"
lemma (in M_trivial) bool_of_o_abs [simp]:
- "M(z) ==> is_bool_of_o(M,P,z) \<longleftrightarrow> z = bool_of_o(P)"
+ "M(z) \<Longrightarrow> is_bool_of_o(M,P,z) \<longleftrightarrow> z = bool_of_o(P)"
by (simp add: is_bool_of_o_def bool_of_o_def)
lemma (in M_trivial) not_abs [simp]:
- "[| M(a); M(z)|] ==> is_not(M,a,z) \<longleftrightarrow> z = not(a)"
+ "\<lbrakk>M(a); M(z)\<rbrakk> \<Longrightarrow> is_not(M,a,z) \<longleftrightarrow> z = not(a)"
by (simp add: Bool.not_def cond_def is_not_def)
lemma (in M_trivial) and_abs [simp]:
- "[| M(a); M(b); M(z)|] ==> is_and(M,a,b,z) \<longleftrightarrow> z = a and b"
+ "\<lbrakk>M(a); M(b); M(z)\<rbrakk> \<Longrightarrow> is_and(M,a,b,z) \<longleftrightarrow> z = a and b"
by (simp add: Bool.and_def cond_def is_and_def)
lemma (in M_trivial) or_abs [simp]:
- "[| M(a); M(b); M(z)|] ==> is_or(M,a,b,z) \<longleftrightarrow> z = a or b"
+ "\<lbrakk>M(a); M(b); M(z)\<rbrakk> \<Longrightarrow> is_or(M,a,b,z) \<longleftrightarrow> z = a or b"
by (simp add: Bool.or_def cond_def is_or_def)
@@ -1470,15 +1470,15 @@
by (simp add: bool_of_o_def)
lemma (in M_trivial) and_closed [intro,simp]:
- "[| M(p); M(q) |] ==> M(p and q)"
+ "\<lbrakk>M(p); M(q)\<rbrakk> \<Longrightarrow> M(p and q)"
by (simp add: and_def cond_def)
lemma (in M_trivial) or_closed [intro,simp]:
- "[| M(p); M(q) |] ==> M(p or q)"
+ "\<lbrakk>M(p); M(q)\<rbrakk> \<Longrightarrow> M(p or q)"
by (simp add: or_def cond_def)
lemma (in M_trivial) not_closed [intro,simp]:
- "M(p) ==> M(not(p))"
+ "M(p) \<Longrightarrow> M(not(p))"
by (simp add: Bool.not_def cond_def)
@@ -1487,46 +1487,46 @@
definition
is_Nil :: "[i=>o, i] => o" where
\<comment> \<open>because \<^prop>\<open>[] \<equiv> Inl(0)\<close>\<close>
- "is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs)"
+ "is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs)"
definition
is_Cons :: "[i=>o,i,i,i] => o" where
\<comment> \<open>because \<^prop>\<open>Cons(a, l) \<equiv> Inr(\<langle>a,l\<rangle>)\<close>\<close>
- "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)"
+ "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)"
lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)"
by (simp add: Nil_def)
-lemma (in M_trivial) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) \<longleftrightarrow> (Z = Nil)"
+lemma (in M_trivial) Nil_abs [simp]: "M(Z) \<Longrightarrow> is_Nil(M,Z) \<longleftrightarrow> (Z = Nil)"
by (simp add: is_Nil_def Nil_def)
lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \<longleftrightarrow> M(a) & M(l)"
by (simp add: Cons_def)
lemma (in M_trivial) Cons_abs [simp]:
- "[|M(a); M(l); M(Z)|] ==> is_Cons(M,a,l,Z) \<longleftrightarrow> (Z = Cons(a,l))"
+ "\<lbrakk>M(a); M(l); M(Z)\<rbrakk> \<Longrightarrow> is_Cons(M,a,l,Z) \<longleftrightarrow> (Z = Cons(a,l))"
by (simp add: is_Cons_def Cons_def)
definition
quasilist :: "i => o" where
- "quasilist(xs) == xs=Nil | (\<exists>x l. xs = Cons(x,l))"
+ "quasilist(xs) \<equiv> xs=Nil | (\<exists>x l. xs = Cons(x,l))"
definition
is_quasilist :: "[i=>o,i] => o" where
- "is_quasilist(M,z) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
+ "is_quasilist(M,z) \<equiv> is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
definition
list_case' :: "[i, [i,i]=>i, i] => i" where
\<comment> \<open>A version of \<^term>\<open>list_case\<close> that's always defined.\<close>
- "list_case'(a,b,xs) ==
+ "list_case'(a,b,xs) \<equiv>
if quasilist(xs) then list_case(a,b,xs) else 0"
definition
is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where
\<comment> \<open>Returns 0 for non-lists\<close>
- "is_list_case(M, a, is_b, xs, z) ==
+ "is_list_case(M, a, is_b, xs, z) \<equiv>
(is_Nil(M,xs) \<longrightarrow> z=a) &
(\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) \<longrightarrow> is_b(x,l,z)) &
(is_quasilist(M,xs) | empty(M,z))"
@@ -1534,28 +1534,28 @@
definition
hd' :: "i => i" where
\<comment> \<open>A version of \<^term>\<open>hd\<close> that's always defined.\<close>
- "hd'(xs) == if quasilist(xs) then hd(xs) else 0"
+ "hd'(xs) \<equiv> if quasilist(xs) then hd(xs) else 0"
definition
tl' :: "i => i" where
\<comment> \<open>A version of \<^term>\<open>tl\<close> that's always defined.\<close>
- "tl'(xs) == if quasilist(xs) then tl(xs) else 0"
+ "tl'(xs) \<equiv> if quasilist(xs) then tl(xs) else 0"
definition
is_hd :: "[i=>o,i,i] => o" where
\<comment> \<open>\<^term>\<open>hd([]) = 0\<close> no constraints if not a list.
Avoiding implication prevents the simplifier's looping.\<close>
- "is_hd(M,xs,H) ==
+ "is_hd(M,xs,H) \<equiv>
(is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
- (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
+ (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | H=x) &
(is_quasilist(M,xs) | empty(M,H))"
definition
is_tl :: "[i=>o,i,i] => o" where
\<comment> \<open>\<^term>\<open>tl([]) = []\<close>; see comments about \<^term>\<open>is_hd\<close>\<close>
- "is_tl(M,xs,T) ==
+ "is_tl(M,xs,T) \<equiv>
(is_Nil(M,xs) \<longrightarrow> T=xs) &
- (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
+ (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | T=l) &
(is_quasilist(M,xs) | empty(M,T))"
subsubsection\<open>\<^term>\<open>quasilist\<close>: For Case-Splitting with \<^term>\<open>list_case'\<close>\<close>
@@ -1566,7 +1566,7 @@
lemma [iff]: "quasilist(Cons(x,l))"
by (simp add: quasilist_def)
-lemma list_imp_quasilist: "l \<in> list(A) ==> quasilist(l)"
+lemma list_imp_quasilist: "l \<in> list(A) \<Longrightarrow> quasilist(l)"
by (erule list.cases, simp_all)
subsubsection\<open>\<^term>\<open>list_case'\<close>, the Modified Version of \<^term>\<open>list_case\<close>\<close>
@@ -1577,27 +1577,27 @@
lemma list_case'_Cons [simp]: "list_case'(a,b,Cons(x,l)) = b(x,l)"
by (simp add: list_case'_def quasilist_def)
-lemma non_list_case: "~ quasilist(x) ==> list_case'(a,b,x) = 0"
+lemma non_list_case: "\<not> quasilist(x) \<Longrightarrow> list_case'(a,b,x) = 0"
by (simp add: quasilist_def list_case'_def)
lemma list_case'_eq_list_case [simp]:
- "xs \<in> list(A) ==>list_case'(a,b,xs) = list_case(a,b,xs)"
+ "xs \<in> list(A) \<Longrightarrow>list_case'(a,b,xs) = list_case(a,b,xs)"
by (erule list.cases, simp_all)
lemma (in M_basic) list_case'_closed [intro,simp]:
- "[|M(k); M(a); \<forall>x[M]. \<forall>y[M]. M(b(x,y))|] ==> M(list_case'(a,b,k))"
+ "\<lbrakk>M(k); M(a); \<forall>x[M]. \<forall>y[M]. M(b(x,y))\<rbrakk> \<Longrightarrow> M(list_case'(a,b,k))"
apply (case_tac "quasilist(k)")
apply (simp add: quasilist_def, force)
apply (simp add: non_list_case)
done
lemma (in M_trivial) quasilist_abs [simp]:
- "M(z) ==> is_quasilist(M,z) \<longleftrightarrow> quasilist(z)"
+ "M(z) \<Longrightarrow> is_quasilist(M,z) \<longleftrightarrow> quasilist(z)"
by (auto simp add: is_quasilist_def quasilist_def)
lemma (in M_trivial) list_case_abs [simp]:
- "[| relation2(M,is_b,b); M(k); M(z) |]
- ==> is_list_case(M,a,is_b,k,z) \<longleftrightarrow> z = list_case'(a,b,k)"
+ "\<lbrakk>relation2(M,is_b,b); M(k); M(z)\<rbrakk>
+ \<Longrightarrow> is_list_case(M,a,is_b,k,z) \<longleftrightarrow> z = list_case'(a,b,k)"
apply (case_tac "quasilist(k)")
prefer 2
apply (simp add: is_list_case_def non_list_case)
@@ -1614,11 +1614,11 @@
by (simp add: is_hd_def)
lemma (in M_trivial) is_hd_Cons:
- "[|M(a); M(l)|] ==> is_hd(M,Cons(a,l),Z) \<longleftrightarrow> Z = a"
+ "\<lbrakk>M(a); M(l)\<rbrakk> \<Longrightarrow> is_hd(M,Cons(a,l),Z) \<longleftrightarrow> Z = a"
by (force simp add: is_hd_def)
lemma (in M_trivial) hd_abs [simp]:
- "[|M(x); M(y)|] ==> is_hd(M,x,y) \<longleftrightarrow> y = hd'(x)"
+ "\<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> is_hd(M,x,y) \<longleftrightarrow> y = hd'(x)"
apply (simp add: hd'_def)
apply (intro impI conjI)
prefer 2 apply (force simp add: is_hd_def)
@@ -1630,11 +1630,11 @@
by (simp add: is_tl_def)
lemma (in M_trivial) is_tl_Cons:
- "[|M(a); M(l)|] ==> is_tl(M,Cons(a,l),Z) \<longleftrightarrow> Z = l"
+ "\<lbrakk>M(a); M(l)\<rbrakk> \<Longrightarrow> is_tl(M,Cons(a,l),Z) \<longleftrightarrow> Z = l"
by (force simp add: is_tl_def)
lemma (in M_trivial) tl_abs [simp]:
- "[|M(x); M(y)|] ==> is_tl(M,x,y) \<longleftrightarrow> y = tl'(x)"
+ "\<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> is_tl(M,x,y) \<longleftrightarrow> y = tl'(x)"
apply (simp add: tl'_def)
apply (intro impI conjI)
prefer 2 apply (force simp add: is_tl_def)
@@ -1657,12 +1657,12 @@
lemma tl'_Cons: "tl'(Cons(a,l)) = l"
by (simp add: tl'_def)
-lemma iterates_tl_Nil: "n \<in> nat ==> tl'^n ([]) = []"
+lemma iterates_tl_Nil: "n \<in> nat \<Longrightarrow> tl'^n ([]) = []"
apply (induct_tac n)
apply (simp_all add: tl'_Nil)
done
-lemma (in M_basic) tl'_closed: "M(x) ==> M(tl'(x))"
+lemma (in M_basic) tl'_closed: "M(x) \<Longrightarrow> M(tl'(x))"
apply (simp add: tl'_def)
apply (force simp add: quasilist_def)
done
--- a/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,14 +11,14 @@
subsubsection\<open>The Formula \<^term>\<open>is_depth\<close>, Internalized\<close>
-(* "is_depth(M,p,n) ==
+(* "is_depth(M,p,n) \<equiv>
\<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M].
2 1 0
is_formula_N(M,n,formula_n) & p \<notin> formula_n &
successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn" *)
definition
depth_fm :: "[i,i]=>i" where
- "depth_fm(p,n) ==
+ "depth_fm(p,n) \<equiv>
Exists(Exists(Exists(
And(formula_N_fm(n#+3,1),
And(Neg(Member(p#+3,1)),
@@ -26,21 +26,21 @@
And(formula_N_fm(2,0), Member(p#+3,0))))))))"
lemma depth_fm_type [TC]:
- "[| x \<in> nat; y \<in> nat |] ==> depth_fm(x,y) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> depth_fm(x,y) \<in> formula"
by (simp add: depth_fm_def)
lemma sats_depth_fm [simp]:
- "[| x \<in> nat; y < length(env); env \<in> list(A)|]
- ==> sats(A, depth_fm(x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, depth_fm(x,y), env) \<longleftrightarrow>
is_depth(##A, nth(x,env), nth(y,env))"
apply (frule_tac x=y in lt_length_in_nat, assumption)
apply (simp add: depth_fm_def is_depth_def)
done
lemma depth_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j < length(env); env \<in> list(A)|]
- ==> is_depth(##A, x, y) \<longleftrightarrow> sats(A, depth_fm(i,j), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_depth(##A, x, y) \<longleftrightarrow> sats(A, depth_fm(i,j), env)"
by (simp)
theorem depth_reflection:
@@ -59,7 +59,7 @@
(* is_formula_case ::
"[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o"
- "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) ==
+ "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) \<equiv>
(\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Member(M,x,y,v) \<longrightarrow> is_a(x,y,z)) &
(\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Equal(M,x,y,v) \<longrightarrow> is_b(x,y,z)) &
(\<forall>x[M]. \<forall>y[M]. x\<in>formula \<longrightarrow> y\<in>formula \<longrightarrow>
@@ -68,7 +68,7 @@
definition
formula_case_fm :: "[i, i, i, i, i, i]=>i" where
- "formula_case_fm(is_a, is_b, is_c, is_d, v, z) ==
+ "formula_case_fm(is_a, is_b, is_c, is_d, v, z) \<equiv>
And(Forall(Forall(Implies(finite_ordinal_fm(1),
Implies(finite_ordinal_fm(0),
Implies(Member_fm(1,0,v#+2),
@@ -87,31 +87,31 @@
lemma is_formula_case_type [TC]:
- "[| is_a \<in> formula; is_b \<in> formula; is_c \<in> formula; is_d \<in> formula;
- x \<in> nat; y \<in> nat |]
- ==> formula_case_fm(is_a, is_b, is_c, is_d, x, y) \<in> formula"
+ "\<lbrakk>is_a \<in> formula; is_b \<in> formula; is_c \<in> formula; is_d \<in> formula;
+ x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> formula_case_fm(is_a, is_b, is_c, is_d, x, y) \<in> formula"
by (simp add: formula_case_fm_def)
lemma sats_formula_case_fm:
assumes is_a_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_b_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_c_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_d_iff_sats:
- "!!a0 a1.
- [|a0\<in>A; a1\<in>A|]
- ==> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
+ "\<And>a0 a1.
+ \<lbrakk>a0\<in>A; a1\<in>A\<rbrakk>
+ \<Longrightarrow> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
shows
- "[|x \<in> nat; y < length(env); env \<in> list(A)|]
- ==> sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) \<longleftrightarrow>
is_formula_case(##A, ISA, ISB, ISC, ISD, nth(x,env), nth(y,env))"
apply (frule_tac x=y in lt_length_in_nat, assumption)
apply (simp add: formula_case_fm_def is_formula_case_def
@@ -121,25 +121,25 @@
lemma formula_case_iff_sats:
assumes is_a_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISA(a2, a1, a0) \<longleftrightarrow> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_b_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISB(a2, a1, a0) \<longleftrightarrow> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_c_iff_sats:
- "!!a0 a1 a2.
- [|a0\<in>A; a1\<in>A; a2\<in>A|]
- ==> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
+ "\<And>a0 a1 a2.
+ \<lbrakk>a0\<in>A; a1\<in>A; a2\<in>A\<rbrakk>
+ \<Longrightarrow> ISC(a2, a1, a0) \<longleftrightarrow> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))"
and is_d_iff_sats:
- "!!a0 a1.
- [|a0\<in>A; a1\<in>A|]
- ==> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
+ "\<And>a0 a1.
+ \<lbrakk>a0\<in>A; a1\<in>A\<rbrakk>
+ \<Longrightarrow> ISD(a1, a0) \<longleftrightarrow> sats(A, is_d, Cons(a0,Cons(a1,env)))"
shows
- "[|nth(i,env) = x; nth(j,env) = y;
- i \<in> nat; j < length(env); env \<in> list(A)|]
- ==> is_formula_case(##A, ISA, ISB, ISC, ISD, x, y) \<longleftrightarrow>
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y;
+ i \<in> nat; j < length(env); env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_formula_case(##A, ISA, ISB, ISC, ISD, x, y) \<longleftrightarrow>
sats(A, formula_case_fm(is_a,is_b,is_c,is_d,i,j), env)"
by (simp add: sats_formula_case_fm [OF is_a_iff_sats is_b_iff_sats
is_c_iff_sats is_d_iff_sats])
@@ -150,16 +150,16 @@
based on that of \<open>is_nat_case_reflection\<close>.\<close>
theorem is_formula_case_reflection:
assumes is_a_reflection:
- "!!h f g g'. REFLECTS[\<lambda>x. is_a(L, h(x), f(x), g(x), g'(x)),
+ "\<And>h f g g'. REFLECTS[\<lambda>x. is_a(L, h(x), f(x), g(x), g'(x)),
\<lambda>i x. is_a(##Lset(i), h(x), f(x), g(x), g'(x))]"
and is_b_reflection:
- "!!h f g g'. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x), g'(x)),
+ "\<And>h f g g'. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x), g'(x)),
\<lambda>i x. is_b(##Lset(i), h(x), f(x), g(x), g'(x))]"
and is_c_reflection:
- "!!h f g g'. REFLECTS[\<lambda>x. is_c(L, h(x), f(x), g(x), g'(x)),
+ "\<And>h f g g'. REFLECTS[\<lambda>x. is_c(L, h(x), f(x), g(x), g'(x)),
\<lambda>i x. is_c(##Lset(i), h(x), f(x), g(x), g'(x))]"
and is_d_reflection:
- "!!h f g g'. REFLECTS[\<lambda>x. is_d(L, h(x), f(x), g(x)),
+ "\<And>h f g g'. REFLECTS[\<lambda>x. is_d(L, h(x), f(x), g(x)),
\<lambda>i x. is_d(##Lset(i), h(x), f(x), g(x))]"
shows "REFLECTS[\<lambda>x. is_formula_case(L, is_a(L,x), is_b(L,x), is_c(L,x), is_d(L,x), g(x), h(x)),
\<lambda>i x. is_formula_case(##Lset(i), is_a(##Lset(i), x), is_b(##Lset(i), x), is_c(##Lset(i), x), is_d(##Lset(i), x), g(x), h(x))]"
@@ -177,14 +177,14 @@
definition
is_depth_apply :: "[i=>o,i,i,i] => o" where
\<comment> \<open>Merely a useful abbreviation for the sequel.\<close>
- "is_depth_apply(M,h,p,z) ==
+ "is_depth_apply(M,h,p,z) \<equiv>
\<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M].
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)"
lemma (in M_datatypes) is_depth_apply_abs [simp]:
- "[|M(h); p \<in> formula; M(z)|]
- ==> is_depth_apply(M,h,p,z) \<longleftrightarrow> z = h ` succ(depth(p)) ` p"
+ "\<lbrakk>M(h); p \<in> formula; M(z)\<rbrakk>
+ \<Longrightarrow> is_depth_apply(M,h,p,z) \<longleftrightarrow> z = h ` succ(depth(p)) ` p"
by (simp add: is_depth_apply_def formula_into_M depth_type eq_commute)
@@ -196,12 +196,12 @@
\<^term>\<open>c\<close>, \<^term>\<open>d\<close>, etc., of the locale \<open>Formula_Rec\<close>.\<close>
definition
satisfies_a :: "[i,i,i]=>i" where
- "satisfies_a(A) ==
+ "satisfies_a(A) \<equiv>
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env))"
definition
satisfies_is_a :: "[i=>o,i,i,i,i]=>o" where
- "satisfies_is_a(M,A) ==
+ "satisfies_is_a(M,A) \<equiv>
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. is_bool_of_o(M,
@@ -211,14 +211,14 @@
definition
satisfies_b :: "[i,i,i]=>i" where
- "satisfies_b(A) ==
+ "satisfies_b(A) \<equiv>
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env))"
definition
satisfies_is_b :: "[i=>o,i,i,i,i]=>o" where
\<comment> \<open>We simplify the formula to have just \<^term>\<open>nx\<close> rather than
introducing \<^term>\<open>ny\<close> with \<^term>\<open>nx=ny\<close>\<close>
- "satisfies_is_b(M,A) ==
+ "satisfies_is_b(M,A) \<equiv>
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. is_bool_of_o(M,
@@ -227,11 +227,11 @@
definition
satisfies_c :: "[i,i,i,i,i]=>i" where
- "satisfies_c(A) == \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)"
+ "satisfies_c(A) \<equiv> \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)"
definition
satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" where
- "satisfies_is_c(M,A,h) ==
+ "satisfies_is_c(M,A,h) \<equiv>
\<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M].
(\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) &
@@ -242,11 +242,11 @@
definition
satisfies_d :: "[i,i,i]=>i" where
"satisfies_d(A)
- == \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)"
+ \<equiv> \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)"
definition
satisfies_is_d :: "[i=>o,i,i,i,i]=>o" where
- "satisfies_is_d(M,A,h) ==
+ "satisfies_is_d(M,A,h) \<equiv>
\<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) &
@@ -261,7 +261,7 @@
satisfies_MH :: "[i=>o,i,i,i,i]=>o" where
\<comment> \<open>The variable \<^term>\<open>u\<close> is unused, but gives \<^term>\<open>satisfies_MH\<close>
the correct arity.\<close>
- "satisfies_MH ==
+ "satisfies_MH \<equiv>
\<lambda>M A u f z.
\<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
is_lambda (M, fml,
@@ -272,7 +272,7 @@
definition
is_satisfies :: "[i=>o,i,i,i]=>o" where
- "is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))"
+ "is_satisfies(M,A) \<equiv> is_formula_rec (M, satisfies_MH(M,A))"
text\<open>This lemma relates the fragments defined above to the original primitive
@@ -291,32 +291,32 @@
locale M_satisfies = M_eclose +
assumes
Member_replacement:
- "[|M(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M].
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) &
is_bool_of_o(M, nx \<in> ny, bo) &
pair(M, env, bo, z))"
and
Equal_replacement:
- "[|M(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M].
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) &
is_bool_of_o(M, nx = ny, bo) &
pair(M, env, bo, z))"
and
Nand_replacement:
- "[|M(A); M(rp); M(rq)|]
- ==> strong_replacement
+ "\<lbrakk>M(A); M(rp); M(rq)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M].
fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) &
is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) &
env \<in> list(A) & pair(M, env, notpq, z))"
and
Forall_replacement:
- "[|M(A); M(rp)|]
- ==> strong_replacement
+ "\<lbrakk>M(A); M(rp)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. \<exists>bo[M].
env \<in> list(A) &
is_bool_of_o (M,
@@ -328,11 +328,11 @@
and
formula_rec_replacement:
\<comment> \<open>For the \<^term>\<open>transrec\<close>\<close>
- "[|n \<in> nat; M(A)|] ==> transrec_replacement(M, satisfies_MH(M,A), n)"
+ "\<lbrakk>n \<in> nat; M(A)\<rbrakk> \<Longrightarrow> transrec_replacement(M, satisfies_MH(M,A), n)"
and
formula_rec_lambda_replacement:
\<comment> \<open>For the \<open>\<lambda>-abstraction\<close> in the \<^term>\<open>transrec\<close> body\<close>
- "[|M(g); M(A)|] ==>
+ "\<lbrakk>M(g); M(A)\<rbrakk> \<Longrightarrow>
strong_replacement (M,
\<lambda>x y. mem_formula(M,x) &
(\<exists>c[M]. is_formula_case(M, satisfies_is_a(M,A),
@@ -343,60 +343,60 @@
lemma (in M_satisfies) Member_replacement':
- "[|M(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. env \<in> list(A) &
z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)"
by (insert Member_replacement, simp)
lemma (in M_satisfies) Equal_replacement':
- "[|M(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>M(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. env \<in> list(A) &
z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)"
by (insert Equal_replacement, simp)
lemma (in M_satisfies) Nand_replacement':
- "[|M(A); M(rp); M(rq)|]
- ==> strong_replacement
+ "\<lbrakk>M(A); M(rp); M(rq)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)"
by (insert Nand_replacement, simp)
lemma (in M_satisfies) Forall_replacement':
- "[|M(A); M(rp)|]
- ==> strong_replacement
+ "\<lbrakk>M(A); M(rp)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(M, \<lambda>env z.
env \<in> list(A) &
z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)"
by (insert Forall_replacement, simp)
lemma (in M_satisfies) a_closed:
- "[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_a(A,x,y))"
+ "\<lbrakk>M(A); x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> M(satisfies_a(A,x,y))"
apply (simp add: satisfies_a_def)
apply (blast intro: lam_closed2 Member_replacement')
done
lemma (in M_satisfies) a_rel:
- "M(A) ==> Relation2(M, nat, nat, satisfies_is_a(M,A), satisfies_a(A))"
+ "M(A) \<Longrightarrow> Relation2(M, nat, nat, satisfies_is_a(M,A), satisfies_a(A))"
apply (simp add: Relation2_def satisfies_is_a_def satisfies_a_def)
apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def)
done
lemma (in M_satisfies) b_closed:
- "[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_b(A,x,y))"
+ "\<lbrakk>M(A); x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> M(satisfies_b(A,x,y))"
apply (simp add: satisfies_b_def)
apply (blast intro: lam_closed2 Equal_replacement')
done
lemma (in M_satisfies) b_rel:
- "M(A) ==> Relation2(M, nat, nat, satisfies_is_b(M,A), satisfies_b(A))"
+ "M(A) \<Longrightarrow> Relation2(M, nat, nat, satisfies_is_b(M,A), satisfies_b(A))"
apply (simp add: Relation2_def satisfies_is_b_def satisfies_b_def)
apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def)
done
lemma (in M_satisfies) c_closed:
- "[|M(A); x \<in> formula; y \<in> formula; M(rx); M(ry)|]
- ==> M(satisfies_c(A,x,y,rx,ry))"
+ "\<lbrakk>M(A); x \<in> formula; y \<in> formula; M(rx); M(ry)\<rbrakk>
+ \<Longrightarrow> M(satisfies_c(A,x,y,rx,ry))"
apply (simp add: satisfies_c_def)
apply (rule lam_closed2)
apply (rule Nand_replacement')
@@ -404,7 +404,7 @@
done
lemma (in M_satisfies) c_rel:
- "[|M(A); M(f)|] ==>
+ "\<lbrakk>M(A); M(f)\<rbrakk> \<Longrightarrow>
Relation2 (M, formula, formula,
satisfies_is_c(M,A,f),
\<lambda>u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u,
@@ -415,7 +415,7 @@
done
lemma (in M_satisfies) d_closed:
- "[|M(A); x \<in> formula; M(rx)|] ==> M(satisfies_d(A,x,rx))"
+ "\<lbrakk>M(A); x \<in> formula; M(rx)\<rbrakk> \<Longrightarrow> M(satisfies_d(A,x,rx))"
apply (simp add: satisfies_d_def)
apply (rule lam_closed2)
apply (rule Forall_replacement')
@@ -423,7 +423,7 @@
done
lemma (in M_satisfies) d_rel:
- "[|M(A); M(f)|] ==>
+ "\<lbrakk>M(A); M(f)\<rbrakk> \<Longrightarrow>
Relation1(M, formula, satisfies_is_d(M,A,f),
\<lambda>u. satisfies_d(A, u, f ` succ(depth(u)) ` u))"
apply (simp del: rall_abs
@@ -433,11 +433,11 @@
lemma (in M_satisfies) fr_replace:
- "[|n \<in> nat; M(A)|] ==> transrec_replacement(M,satisfies_MH(M,A),n)"
+ "\<lbrakk>n \<in> nat; M(A)\<rbrakk> \<Longrightarrow> transrec_replacement(M,satisfies_MH(M,A),n)"
by (blast intro: formula_rec_replacement)
lemma (in M_satisfies) formula_case_satisfies_closed:
- "[|M(g); M(A); x \<in> formula|] ==>
+ "\<lbrakk>M(g); M(A); x \<in> formula\<rbrakk> \<Longrightarrow>
M(formula_case (satisfies_a(A), satisfies_b(A),
\<lambda>u v. satisfies_c(A, u, v,
g ` succ(depth(u)) ` u, g ` succ(depth(v)) ` v),
@@ -446,7 +446,7 @@
by (blast intro: a_closed b_closed c_closed d_closed)
lemma (in M_satisfies) fr_lam_replace:
- "[|M(g); M(A)|] ==>
+ "\<lbrakk>M(g); M(A)\<rbrakk> \<Longrightarrow>
strong_replacement (M, \<lambda>x y. x \<in> formula &
y = \<langle>x,
formula_rec_case(satisfies_a(A),
@@ -464,7 +464,7 @@
Function \<^term>\<open>satisfies\<close>\<close>
lemma (in M_satisfies) Formula_Rec_axioms_M:
- "M(A) ==>
+ "M(A) \<Longrightarrow>
Formula_Rec_axioms(M, satisfies_a(A), satisfies_is_a(M,A),
satisfies_b(A), satisfies_is_b(M,A),
satisfies_c(A), satisfies_is_c(M,A),
@@ -478,7 +478,7 @@
theorem (in M_satisfies) Formula_Rec_M:
- "M(A) ==>
+ "M(A) \<Longrightarrow>
Formula_Rec(M, satisfies_a(A), satisfies_is_a(M,A),
satisfies_b(A), satisfies_is_b(M,A),
satisfies_c(A), satisfies_is_c(M,A),
@@ -494,13 +494,13 @@
lemma (in M_satisfies) satisfies_closed:
- "[|M(A); p \<in> formula|] ==> M(satisfies(A,p))"
+ "\<lbrakk>M(A); p \<in> formula\<rbrakk> \<Longrightarrow> M(satisfies(A,p))"
by (simp add: Formula_Rec.formula_rec_closed [OF Formula_Rec_M]
satisfies_eq)
lemma (in M_satisfies) satisfies_abs:
- "[|M(A); M(z); p \<in> formula|]
- ==> is_satisfies(M,A,p,z) \<longleftrightarrow> z = satisfies(A,p)"
+ "\<lbrakk>M(A); M(z); p \<in> formula\<rbrakk>
+ \<Longrightarrow> is_satisfies(M,A,p,z) \<longleftrightarrow> z = satisfies(A,p)"
by (simp only: Formula_Rec.formula_rec_abs [OF Formula_Rec_M]
satisfies_eq is_satisfies_def satisfies_MH_def)
@@ -509,14 +509,14 @@
subsubsection\<open>The Operator \<^term>\<open>is_depth_apply\<close>, Internalized\<close>
-(* is_depth_apply(M,h,p,z) ==
+(* is_depth_apply(M,h,p,z) \<equiv>
\<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M].
2 1 0
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *)
definition
depth_apply_fm :: "[i,i,i]=>i" where
- "depth_apply_fm(h,p,z) ==
+ "depth_apply_fm(h,p,z) \<equiv>
Exists(Exists(Exists(
And(finite_ordinal_fm(2),
And(depth_fm(p#+3,2),
@@ -524,19 +524,19 @@
And(fun_apply_fm(h#+3,1,0), fun_apply_fm(0,p#+3,z#+3))))))))"
lemma depth_apply_type [TC]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> depth_apply_fm(x,y,z) \<in> formula"
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> depth_apply_fm(x,y,z) \<in> formula"
by (simp add: depth_apply_fm_def)
lemma sats_depth_apply_fm [simp]:
- "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, depth_apply_fm(x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, depth_apply_fm(x,y,z), env) \<longleftrightarrow>
is_depth_apply(##A, nth(x,env), nth(y,env), nth(z,env))"
by (simp add: depth_apply_fm_def is_depth_apply_def)
lemma depth_apply_iff_sats:
- "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
- i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
- ==> is_depth_apply(##A, x, y, z) \<longleftrightarrow> sats(A, depth_apply_fm(i,j,k), env)"
+ "\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
+ i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> is_depth_apply(##A, x, y, z) \<longleftrightarrow> sats(A, depth_apply_fm(i,j,k), env)"
by simp
lemma depth_apply_reflection:
@@ -550,7 +550,7 @@
subsubsection\<open>The Operator \<^term>\<open>satisfies_is_a\<close>, Internalized\<close>
-(* satisfies_is_a(M,A) ==
+(* satisfies_is_a(M,A) \<equiv>
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. is_bool_of_o(M,
@@ -560,7 +560,7 @@
definition
satisfies_is_a_fm :: "[i,i,i,i]=>i" where
- "satisfies_is_a_fm(A,x,y,z) ==
+ "satisfies_is_a_fm(A,x,y,z) \<equiv>
Forall(
Implies(is_list_fm(succ(A),0),
lambda_fm(
@@ -571,13 +571,13 @@
0, succ(z))))"
lemma satisfies_is_a_type [TC]:
- "[| A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> satisfies_is_a_fm(A,x,y,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> satisfies_is_a_fm(A,x,y,z) \<in> formula"
by (simp add: satisfies_is_a_fm_def)
lemma sats_satisfies_is_a_fm [simp]:
- "[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_is_a_fm(u,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_is_a_fm(u,x,y,z), env) \<longleftrightarrow>
satisfies_is_a(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=x in lt_length_in_nat, assumption)
apply (frule_tac x=y in lt_length_in_nat, assumption)
@@ -586,9 +586,9 @@
done
lemma satisfies_is_a_iff_sats:
- "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
- u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
- ==> satisfies_is_a(##A,nu,nx,ny,nz) \<longleftrightarrow>
+ "\<lbrakk>nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
+ u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> satisfies_is_a(##A,nu,nx,ny,nz) \<longleftrightarrow>
sats(A, satisfies_is_a_fm(u,x,y,z), env)"
by simp
@@ -603,7 +603,7 @@
subsubsection\<open>The Operator \<^term>\<open>satisfies_is_b\<close>, Internalized\<close>
-(* satisfies_is_b(M,A) ==
+(* satisfies_is_b(M,A) \<equiv>
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. is_bool_of_o(M,
@@ -612,7 +612,7 @@
definition
satisfies_is_b_fm :: "[i,i,i,i]=>i" where
- "satisfies_is_b_fm(A,x,y,z) ==
+ "satisfies_is_b_fm(A,x,y,z) \<equiv>
Forall(
Implies(is_list_fm(succ(A),0),
lambda_fm(
@@ -620,13 +620,13 @@
0, succ(z))))"
lemma satisfies_is_b_type [TC]:
- "[| A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> satisfies_is_b_fm(A,x,y,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> satisfies_is_b_fm(A,x,y,z) \<in> formula"
by (simp add: satisfies_is_b_fm_def)
lemma sats_satisfies_is_b_fm [simp]:
- "[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_is_b_fm(u,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_is_b_fm(u,x,y,z), env) \<longleftrightarrow>
satisfies_is_b(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
apply (frule_tac x=x in lt_length_in_nat, assumption)
apply (frule_tac x=y in lt_length_in_nat, assumption)
@@ -635,9 +635,9 @@
done
lemma satisfies_is_b_iff_sats:
- "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
- u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|]
- ==> satisfies_is_b(##A,nu,nx,ny,nz) \<longleftrightarrow>
+ "\<lbrakk>nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
+ u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> satisfies_is_b(##A,nu,nx,ny,nz) \<longleftrightarrow>
sats(A, satisfies_is_b_fm(u,x,y,z), env)"
by simp
@@ -652,7 +652,7 @@
subsubsection\<open>The Operator \<^term>\<open>satisfies_is_c\<close>, Internalized\<close>
-(* satisfies_is_c(M,A,h) ==
+(* satisfies_is_c(M,A,h) \<equiv>
\<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M].
(\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) &
@@ -662,7 +662,7 @@
definition
satisfies_is_c_fm :: "[i,i,i,i,i]=>i" where
- "satisfies_is_c_fm(A,h,p,q,zz) ==
+ "satisfies_is_c_fm(A,h,p,q,zz) \<equiv>
Forall(
Implies(is_list_fm(succ(A),0),
lambda_fm(
@@ -673,22 +673,22 @@
0, succ(zz))))"
lemma satisfies_is_c_type [TC]:
- "[| A \<in> nat; h \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |]
- ==> satisfies_is_c_fm(A,h,x,y,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; h \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> satisfies_is_c_fm(A,h,x,y,z) \<in> formula"
by (simp add: satisfies_is_c_fm_def)
lemma sats_satisfies_is_c_fm [simp]:
- "[| u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_is_c_fm(u,v,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_is_c_fm(u,v,x,y,z), env) \<longleftrightarrow>
satisfies_is_c(##A, nth(u,env), nth(v,env), nth(x,env),
nth(y,env), nth(z,env))"
by (simp add: satisfies_is_c_fm_def satisfies_is_c_def sats_lambda_fm)
lemma satisfies_is_c_iff_sats:
- "[| nth(u,env) = nu; nth(v,env) = nv; nth(x,env) = nx; nth(y,env) = ny;
+ "\<lbrakk>nth(u,env) = nu; nth(v,env) = nv; nth(x,env) = nx; nth(y,env) = ny;
nth(z,env) = nz;
- u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> satisfies_is_c(##A,nu,nv,nx,ny,nz) \<longleftrightarrow>
+ u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> satisfies_is_c(##A,nu,nv,nx,ny,nz) \<longleftrightarrow>
sats(A, satisfies_is_c_fm(u,v,x,y,z), env)"
by simp
@@ -703,7 +703,7 @@
subsubsection\<open>The Operator \<^term>\<open>satisfies_is_d\<close>, Internalized\<close>
-(* satisfies_is_d(M,A,h) ==
+(* satisfies_is_d(M,A,h) \<equiv>
\<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
is_lambda(M, lA,
\<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) &
@@ -716,7 +716,7 @@
definition
satisfies_is_d_fm :: "[i,i,i,i]=>i" where
- "satisfies_is_d_fm(A,h,p,zz) ==
+ "satisfies_is_d_fm(A,h,p,zz) \<equiv>
Forall(
Implies(is_list_fm(succ(A),0),
lambda_fm(
@@ -730,21 +730,21 @@
0, succ(zz))))"
lemma satisfies_is_d_type [TC]:
- "[| A \<in> nat; h \<in> nat; x \<in> nat; z \<in> nat |]
- ==> satisfies_is_d_fm(A,h,x,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; h \<in> nat; x \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> satisfies_is_d_fm(A,h,x,z) \<in> formula"
by (simp add: satisfies_is_d_fm_def)
lemma sats_satisfies_is_d_fm [simp]:
- "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_is_d_fm(u,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_is_d_fm(u,x,y,z), env) \<longleftrightarrow>
satisfies_is_d(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
by (simp add: satisfies_is_d_fm_def satisfies_is_d_def sats_lambda_fm
sats_bool_of_o_fm)
lemma satisfies_is_d_iff_sats:
- "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
- u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> satisfies_is_d(##A,nu,nx,ny,nz) \<longleftrightarrow>
+ "\<lbrakk>nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
+ u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> satisfies_is_d(##A,nu,nx,ny,nz) \<longleftrightarrow>
sats(A, satisfies_is_d_fm(u,x,y,z), env)"
by simp
@@ -760,7 +760,7 @@
subsubsection\<open>The Operator \<^term>\<open>satisfies_MH\<close>, Internalized\<close>
-(* satisfies_MH ==
+(* satisfies_MH \<equiv>
\<lambda>M A u f zz.
\<forall>fml[M]. is_formula(M,fml) \<longrightarrow>
is_lambda (M, fml,
@@ -771,7 +771,7 @@
definition
satisfies_MH_fm :: "[i,i,i,i]=>i" where
- "satisfies_MH_fm(A,u,f,zz) ==
+ "satisfies_MH_fm(A,u,f,zz) \<equiv>
Forall(
Implies(is_formula_fm(0),
lambda_fm(
@@ -783,21 +783,21 @@
0, succ(zz))))"
lemma satisfies_MH_type [TC]:
- "[| A \<in> nat; u \<in> nat; x \<in> nat; z \<in> nat |]
- ==> satisfies_MH_fm(A,u,x,z) \<in> formula"
+ "\<lbrakk>A \<in> nat; u \<in> nat; x \<in> nat; z \<in> nat\<rbrakk>
+ \<Longrightarrow> satisfies_MH_fm(A,u,x,z) \<in> formula"
by (simp add: satisfies_MH_fm_def)
lemma sats_satisfies_MH_fm [simp]:
- "[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> sats(A, satisfies_MH_fm(u,x,y,z), env) \<longleftrightarrow>
+ "\<lbrakk>u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> sats(A, satisfies_MH_fm(u,x,y,z), env) \<longleftrightarrow>
satisfies_MH(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))"
by (simp add: satisfies_MH_fm_def satisfies_MH_def sats_lambda_fm
sats_formula_case_fm)
lemma satisfies_MH_iff_sats:
- "[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
- u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
- ==> satisfies_MH(##A,nu,nx,ny,nz) \<longleftrightarrow>
+ "\<lbrakk>nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz;
+ u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk>
+ \<Longrightarrow> satisfies_MH(##A,nu,nx,ny,nz) \<longleftrightarrow>
sats(A, satisfies_MH_fm(u,x,y,z), env)"
by simp
@@ -833,8 +833,8 @@
lemma Member_replacement:
- "[|L(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>L(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L].
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) &
is_bool_of_o(L, nx \<in> ny, bo) &
@@ -863,8 +863,8 @@
lemma Equal_replacement:
- "[|L(A); x \<in> nat; y \<in> nat|]
- ==> strong_replacement
+ "\<lbrakk>L(A); x \<in> nat; y \<in> nat\<rbrakk>
+ \<Longrightarrow> strong_replacement
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L].
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) &
is_bool_of_o(L, nx = ny, bo) &
@@ -895,8 +895,8 @@
done
lemma Nand_replacement:
- "[|L(A); L(rp); L(rq)|]
- ==> strong_replacement
+ "\<lbrakk>L(A); L(rp); L(rq)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L].
fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) &
is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) &
@@ -930,8 +930,8 @@
done
lemma Forall_replacement:
- "[|L(A); L(rp)|]
- ==> strong_replacement
+ "\<lbrakk>L(A); L(rp)\<rbrakk>
+ \<Longrightarrow> strong_replacement
(L, \<lambda>env z. \<exists>bo[L].
env \<in> list(A) &
is_bool_of_o (L,
@@ -960,7 +960,7 @@
lemma formula_rec_replacement:
\<comment> \<open>For the \<^term>\<open>transrec\<close>\<close>
- "[|n \<in> nat; L(A)|] ==> transrec_replacement(L, satisfies_MH(L,A), n)"
+ "\<lbrakk>n \<in> nat; L(A)\<rbrakk> \<Longrightarrow> transrec_replacement(L, satisfies_MH(L,A), n)"
apply (rule L.transrec_replacementI, simp add: L.nat_into_M)
apply (rule strong_replacementI)
apply (rule_tac u="{B,A,n,Memrel(eclose({n}))}"
@@ -996,7 +996,7 @@
lemma formula_rec_lambda_replacement:
\<comment> \<open>For the \<^term>\<open>transrec\<close>\<close>
- "[|L(g); L(A)|] ==>
+ "\<lbrakk>L(g); L(A)\<rbrakk> \<Longrightarrow>
strong_replacement (L,
\<lambda>x y. mem_formula(L,x) &
(\<exists>c[L]. is_formula_case(L, satisfies_is_a(L,A),
--- a/src/ZF/Constructible/Separation.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Separation.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,7 +9,7 @@
text\<open>This theory proves all instances needed for locale \<open>M_basic\<close>\<close>
text\<open>Helps us solve for de Bruijn indices!\<close>
-lemma nth_ConsI: "[|nth(n,l) = x; n \<in> nat|] ==> nth(succ(n), Cons(a,l)) = x"
+lemma nth_ConsI: "\<lbrakk>nth(n,l) = x; n \<in> nat\<rbrakk> \<Longrightarrow> nth(succ(n), Cons(a,l)) = x"
by simp
lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI
@@ -17,20 +17,20 @@
fun_plus_iff_sats
lemma Collect_conj_in_DPow:
- "[| {x\<in>A. P(x)} \<in> DPow(A); {x\<in>A. Q(x)} \<in> DPow(A) |]
- ==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
+ "\<lbrakk>{x\<in>A. P(x)} \<in> DPow(A); {x\<in>A. Q(x)} \<in> DPow(A)\<rbrakk>
+ \<Longrightarrow> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric])
lemma Collect_conj_in_DPow_Lset:
- "[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
- ==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
+ "\<lbrakk>z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))\<rbrakk>
+ \<Longrightarrow> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
apply (frule mem_Lset_imp_subset_Lset)
apply (simp add: Collect_conj_in_DPow Collect_mem_eq
subset_Int_iff2 elem_subset_in_DPow)
done
lemma separation_CollectI:
- "(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
+ "(\<And>z. L(z) \<Longrightarrow> L({x \<in> z . P(x)})) \<Longrightarrow> separation(L, \<lambda>x. P(x))"
apply (unfold separation_def, clarify)
apply (rule_tac x="{x\<in>z. P(x)}" in rexI)
apply simp_all
@@ -38,9 +38,9 @@
text\<open>Reduces the original comprehension to the reflected one\<close>
lemma reflection_imp_L_separation:
- "[| \<forall>x\<in>Lset(j). P(x) \<longleftrightarrow> Q(x);
+ "\<lbrakk>\<forall>x\<in>Lset(j). P(x) \<longleftrightarrow> Q(x);
{x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j));
- Ord(j); z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
+ Ord(j); z \<in> Lset(j)\<rbrakk> \<Longrightarrow> L({x \<in> z . P(x)})"
apply (rule_tac i = "succ(j)" in L_I)
prefer 2 apply simp
apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
@@ -54,7 +54,7 @@
lemma gen_separation:
assumes reflection: "REFLECTS [P,Q]"
and Lu: "L(u)"
- and collI: "!!j. u \<in> Lset(j)
+ and collI: "\<And>j. u \<in> Lset(j)
\<Longrightarrow> Collect(Lset(j), Q(j)) \<in> DPow(Lset(j))"
shows "separation(L,P)"
apply (rule separation_CollectI)
@@ -73,7 +73,7 @@
lemma gen_separation_multi:
assumes reflection: "REFLECTS [P,Q]"
and Lu: "L(u)"
- and collI: "!!j. u \<subseteq> Lset(j)
+ and collI: "\<And>j. u \<subseteq> Lset(j)
\<Longrightarrow> Collect(Lset(j), Q(j)) \<in> DPow(Lset(j))"
shows "separation(L,P)"
apply (rule gen_separation [OF reflection Lu])
@@ -90,7 +90,7 @@
by (intro FOL_reflections)
lemma Inter_separation:
- "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x\<in>y)"
+ "L(A) \<Longrightarrow> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A \<longrightarrow> x\<in>y)"
apply (rule gen_separation [OF Inter_Reflects], simp)
apply (rule DPow_LsetI)
txt\<open>I leave this one example of a manual proof. The tedium of manually
@@ -109,7 +109,7 @@
by (intro FOL_reflections)
lemma Diff_separation:
- "L(B) ==> separation(L, \<lambda>x. x \<notin> B)"
+ "L(B) \<Longrightarrow> separation(L, \<lambda>x. x \<notin> B)"
apply (rule gen_separation [OF Diff_Reflects], simp)
apply (rule_tac env="[B]" in DPow_LsetI)
apply (rule sep_rules | simp)+
@@ -124,8 +124,8 @@
by (intro FOL_reflections function_reflections)
lemma cartprod_separation:
- "[| L(A); L(B) |]
- ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
+ "\<lbrakk>L(A); L(B)\<rbrakk>
+ \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
apply (rule gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto)
apply (rule_tac env="[A,B]" in DPow_LsetI)
apply (rule sep_rules | simp)+
@@ -139,8 +139,8 @@
by (intro FOL_reflections function_reflections)
lemma image_separation:
- "[| L(A); L(r) |]
- ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
+ "\<lbrakk>L(A); L(r)\<rbrakk>
+ \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
apply (rule gen_separation_multi [OF image_Reflects, of "{A,r}"], auto)
apply (rule_tac env="[A,r]" in DPow_LsetI)
apply (rule sep_rules | simp)+
@@ -156,7 +156,7 @@
by (intro FOL_reflections function_reflections)
lemma converse_separation:
- "L(r) ==> separation(L,
+ "L(r) \<Longrightarrow> separation(L,
\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
apply (rule gen_separation [OF converse_Reflects], simp)
apply (rule_tac env="[r]" in DPow_LsetI)
@@ -172,7 +172,7 @@
by (intro FOL_reflections function_reflections)
lemma restrict_separation:
- "L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
+ "L(A) \<Longrightarrow> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
apply (rule gen_separation [OF restrict_Reflects], simp)
apply (rule_tac env="[A]" in DPow_LsetI)
apply (rule sep_rules | simp)+
@@ -191,8 +191,8 @@
by (intro FOL_reflections function_reflections)
lemma comp_separation:
- "[| L(r); L(s) |]
- ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
+ "\<lbrakk>L(r); L(s)\<rbrakk>
+ \<Longrightarrow> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L].
pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) &
xy\<in>s & yz\<in>r)"
apply (rule gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto)
@@ -213,7 +213,7 @@
by (intro FOL_reflections function_reflections)
lemma pred_separation:
- "[| L(r); L(x) |] ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
+ "\<lbrakk>L(r); L(x)\<rbrakk> \<Longrightarrow> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto)
apply (rule_tac env="[r,x]" in DPow_LsetI)
apply (rule sep_rules | simp)+
@@ -248,7 +248,7 @@
by (intro FOL_reflections function_reflections)
lemma funspace_succ_replacement:
- "L(n) ==>
+ "L(n) \<Longrightarrow>
strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L].
pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
upair(L,cnbf,cnbf,z))"
@@ -275,8 +275,8 @@
lemma is_recfun_separation:
\<comment> \<open>for well-founded recursion\<close>
- "[| L(r); L(f); L(g); L(a); L(b) |]
- ==> separation(L,
+ "\<lbrakk>L(r); L(f); L(g); L(a); L(b)\<rbrakk>
+ \<Longrightarrow> separation(L,
\<lambda>x. \<exists>xa[L]. \<exists>xb[L].
pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r &
(\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) &
--- a/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,14 +10,14 @@
definition
rtrancl_alt :: "[i,i]=>i" where
- "rtrancl_alt(A,r) ==
+ "rtrancl_alt(A,r) \<equiv>
{p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
(\<exists>x y. p = <x,y> & f`0 = x & f`n = y) &
(\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
lemma alt_rtrancl_lemma1 [rule_format]:
"n \<in> nat
- ==> \<forall>f \<in> succ(n) -> field(r).
+ \<Longrightarrow> \<forall>f \<in> succ(n) -> field(r).
(\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<langle>f`0, f`n\<rangle> \<in> r^*"
apply (induct_tac n)
apply (simp_all add: apply_funtype rtrancl_refl, clarify)
@@ -61,7 +61,7 @@
definition
rtran_closure_mem :: "[i=>o,i,i,i] => o" where
\<comment> \<open>The property of belonging to \<open>rtran_closure(r)\<close>\<close>
- "rtran_closure_mem(M,A,r,p) ==
+ "rtran_closure_mem(M,A,r,p) \<equiv>
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
(\<exists>f[M]. typed_function(M,n',A,f) &
@@ -74,28 +74,28 @@
definition
rtran_closure :: "[i=>o,i,i] => o" where
- "rtran_closure(M,r,s) ==
+ "rtran_closure(M,r,s) \<equiv>
\<forall>A[M]. is_field(M,r,A) \<longrightarrow>
(\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
definition
tran_closure :: "[i=>o,i,i] => o" where
- "tran_closure(M,r,t) ==
+ "tran_closure(M,r,t) \<equiv>
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
locale M_trancl = M_basic +
assumes rtrancl_separation:
- "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
+ "\<lbrakk>M(r); M(A)\<rbrakk> \<Longrightarrow> separation (M, rtran_closure_mem(M,A,r))"
and wellfounded_trancl_separation:
- "[| M(r); M(Z) |] ==>
+ "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow>
separation (M, \<lambda>x.
\<exists>w[M]. \<exists>wx[M]. \<exists>rp[M].
w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
and M_nat [iff] : "M(nat)"
lemma (in M_trancl) rtran_closure_mem_iff:
- "[|M(A); M(r); M(p)|]
- ==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
+ "\<lbrakk>M(A); M(r); M(p)\<rbrakk>
+ \<Longrightarrow> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
(\<exists>n[M]. n\<in>nat &
(\<exists>f[M]. f \<in> succ(n) -> A &
(\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) &
@@ -104,21 +104,21 @@
done
lemma (in M_trancl) rtran_closure_rtrancl:
- "M(r) ==> rtran_closure(M,r,rtrancl(r))"
+ "M(r) \<Longrightarrow> rtran_closure(M,r,rtrancl(r))"
apply (simp add: rtran_closure_def rtran_closure_mem_iff
rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def)
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype)
done
lemma (in M_trancl) rtrancl_closed [intro,simp]:
- "M(r) ==> M(rtrancl(r))"
+ "M(r) \<Longrightarrow> M(rtrancl(r))"
apply (insert rtrancl_separation [of r "field(r)"])
apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
rtrancl_alt_def rtran_closure_mem_iff)
done
lemma (in M_trancl) rtrancl_abs [simp]:
- "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
apply (rule iffI)
txt\<open>Proving the right-to-left implication\<close>
prefer 2 apply (blast intro: rtran_closure_rtrancl)
@@ -129,20 +129,20 @@
done
lemma (in M_trancl) trancl_closed [intro,simp]:
- "M(r) ==> M(trancl(r))"
+ "M(r) \<Longrightarrow> M(trancl(r))"
by (simp add: trancl_def)
lemma (in M_trancl) trancl_abs [simp]:
- "[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
+ "\<lbrakk>M(r); M(z)\<rbrakk> \<Longrightarrow> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
by (simp add: tran_closure_def trancl_def)
lemma (in M_trancl) wellfounded_trancl_separation':
- "[| M(r); M(Z) |] ==> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)"
+ "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)"
by (insert wellfounded_trancl_separation [of r Z], simp)
text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
relativized version. Original version is on theory WF.\<close>
-lemma "[| wf[A](r); r-``A \<subseteq> A |] ==> wf[A](r^+)"
+lemma "\<lbrakk>wf[A](r); r-``A \<subseteq> A\<rbrakk> \<Longrightarrow> wf[A](r^+)"
apply (simp add: wf_on_def wf_def)
apply (safe)
apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
@@ -150,8 +150,8 @@
done
lemma (in M_trancl) wellfounded_on_trancl:
- "[| wellfounded_on(M,A,r); r-``A \<subseteq> A; M(r); M(A) |]
- ==> wellfounded_on(M,A,r^+)"
+ "\<lbrakk>wellfounded_on(M,A,r); r-``A \<subseteq> A; M(r); M(A)\<rbrakk>
+ \<Longrightarrow> wellfounded_on(M,A,r^+)"
apply (simp add: wellfounded_on_def)
apply (safe intro!: equalityI)
apply (rename_tac Z x)
@@ -168,7 +168,7 @@
done
lemma (in M_trancl) wellfounded_trancl:
- "[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)"
+ "\<lbrakk>wellfounded(M,r); M(r)\<rbrakk> \<Longrightarrow> wellfounded(M,r^+)"
apply (simp add: wellfounded_iff_wellfounded_on_field)
apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl)
apply blast
@@ -181,13 +181,13 @@
(*first use is_recfun, then M_is_recfun*)
lemma (in M_trancl) wfrec_relativize:
- "[|wf(r); M(a); M(r);
+ "\<lbrakk>wf(r); M(a); M(r);
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
pair(M,x,y,z) &
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) &
y = H(x, restrict(g, r -`` {x})));
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> wfrec(r,a,H) = z \<longleftrightarrow>
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow>
(\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) &
z = H(a,restrict(f,r-``{a})))"
apply (frule wf_trancl)
@@ -204,10 +204,10 @@
The premise \<^term>\<open>relation(r)\<close> is necessary
before we can replace \<^term>\<open>r^+\<close> by \<^term>\<open>r\<close>.\<close>
theorem (in M_trancl) trans_wfrec_relativize:
- "[|wf(r); trans(r); relation(r); M(r); M(a);
+ "\<lbrakk>wf(r); trans(r); relation(r); M(r); M(a);
wfrec_replacement(M,MH,r); relation2(M,MH,H);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))"
apply (frule wfrec_replacement', assumption+)
apply (simp cong: is_recfun_cong
add: wfrec_relativize trancl_eq_r
@@ -215,18 +215,18 @@
done
theorem (in M_trancl) trans_wfrec_abs:
- "[|wf(r); trans(r); relation(r); M(r); M(a); M(z);
+ "\<lbrakk>wf(r); trans(r); relation(r); M(r); M(a); M(z);
wfrec_replacement(M,MH,r); relation2(M,MH,H);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)"
by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast)
lemma (in M_trancl) trans_eq_pair_wfrec_iff:
- "[|wf(r); trans(r); relation(r); M(r); M(y);
+ "\<lbrakk>wf(r); trans(r); relation(r); M(r); M(y);
wfrec_replacement(M,MH,r); relation2(M,MH,H);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow>
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> y = <x, wfrec(r, x, H)> \<longleftrightarrow>
(\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
apply safe
apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x])
@@ -240,10 +240,10 @@
text\<open>Lemma with the awkward premise mentioning \<open>wfrec\<close>.\<close>
lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
- "[|wf(r); M(r);
+ "\<lbrakk>wf(r); M(r);
strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> M(a) \<longrightarrow> M(wfrec(r,a,H))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> M(a) \<longrightarrow> M(wfrec(r,a,H))"
apply (rule_tac a=a in wf_induct, assumption+)
apply (subst wfrec, assumption, clarify)
apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)"
@@ -264,10 +264,10 @@
text\<open>Useful version for transitive relations\<close>
theorem (in M_trancl) trans_wfrec_closed:
- "[|wf(r); trans(r); relation(r); M(r); M(a);
+ "\<lbrakk>wf(r); trans(r); relation(r); M(r); M(a);
wfrec_replacement(M,MH,r); relation2(M,MH,H);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> M(wfrec(r,a,H))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> M(wfrec(r,a,H))"
apply (frule wfrec_replacement', assumption+)
apply (frule wfrec_replacement_iff [THEN iffD1])
apply (rule wfrec_closed_lemma, assumption+)
@@ -276,13 +276,13 @@
subsection\<open>Absoluteness without assuming transitivity\<close>
lemma (in M_trancl) eq_pair_wfrec_iff:
- "[|wf(r); M(r); M(y);
+ "\<lbrakk>wf(r); M(r); M(y);
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
pair(M,x,y,z) &
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) &
y = H(x, restrict(g, r -`` {x})));
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow>
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> y = <x, wfrec(r, x, H)> \<longleftrightarrow>
(\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) &
y = <x, H(x,restrict(f,r-``{x}))>)"
apply safe
@@ -294,11 +294,11 @@
text\<open>Full version not assuming transitivity, but maybe not very useful.\<close>
theorem (in M_trancl) wfrec_closed:
- "[|wf(r); M(r); M(a);
+ "\<lbrakk>wf(r); M(r); M(a);
wfrec_replacement(M,MH,r^+);
relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> M(wfrec(r,a,H))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> M(wfrec(r,a,H))"
apply (frule wfrec_replacement'
[of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
prefer 4
--- a/src/ZF/Constructible/WFrec.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/WFrec.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,7 +12,7 @@
(*Many of these might be useful in WF.thy*)
lemma apply_recfun2:
- "[| is_recfun(r,a,H,f); <x,i>:f |] ==> i = H(x, restrict(f,r-``{x}))"
+ "\<lbrakk>is_recfun(r,a,H,f); <x,i>:f\<rbrakk> \<Longrightarrow> i = H(x, restrict(f,r-``{x}))"
apply (frule apply_recfun)
apply (blast dest: is_recfun_type fun_is_rel)
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
@@ -32,24 +32,24 @@
apply (fast intro: lam_type, simp)
done
-lemma is_recfun_imp_in_r: "[|is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f|] ==> \<langle>x, a\<rangle> \<in> r"
+lemma is_recfun_imp_in_r: "\<lbrakk>is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f\<rbrakk> \<Longrightarrow> \<langle>x, a\<rangle> \<in> r"
by (blast dest: is_recfun_type fun_is_rel)
lemma trans_Int_eq:
- "[| trans(r); <y,x> \<in> r |] ==> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
+ "\<lbrakk>trans(r); <y,x> \<in> r\<rbrakk> \<Longrightarrow> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
by (blast intro: transD)
lemma is_recfun_restrict_idem:
- "is_recfun(r,a,H,f) ==> restrict(f, r -`` {a}) = f"
+ "is_recfun(r,a,H,f) \<Longrightarrow> restrict(f, r -`` {a}) = f"
apply (drule is_recfun_type)
apply (auto simp add: Pi_iff subset_Sigma_imp_relation restrict_idem)
done
lemma is_recfun_cong_lemma:
- "[| is_recfun(r,a,H,f); r = r'; a = a'; f = f';
- !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x} |]
- ==> H(x,g) = H'(x,g) |]
- ==> is_recfun(r',a',H',f')"
+ "\<lbrakk>is_recfun(r,a,H,f); r = r'; a = a'; f = f';
+ \<And>x g. \<lbrakk><x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x}\<rbrakk>
+ \<Longrightarrow> H(x,g) = H'(x,g)\<rbrakk>
+ \<Longrightarrow> is_recfun(r',a',H',f')"
apply (simp add: is_recfun_def)
apply (erule trans)
apply (rule lam_cong)
@@ -59,10 +59,10 @@
text\<open>For \<open>is_recfun\<close> we need only pay attention to functions
whose domains are initial segments of \<^term>\<open>r\<close>.\<close>
lemma is_recfun_cong:
- "[| r = r'; a = a'; f = f';
- !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x} |]
- ==> H(x,g) = H'(x,g) |]
- ==> is_recfun(r,a,H,f) \<longleftrightarrow> is_recfun(r',a',H',f')"
+ "\<lbrakk>r = r'; a = a'; f = f';
+ \<And>x g. \<lbrakk><x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x}\<rbrakk>
+ \<Longrightarrow> H(x,g) = H'(x,g)\<rbrakk>
+ \<Longrightarrow> is_recfun(r,a,H,f) \<longleftrightarrow> is_recfun(r',a',H',f')"
apply (rule iffI)
txt\<open>Messy: fast and blast don't work for some reason\<close>
apply (erule is_recfun_cong_lemma, auto)
@@ -73,9 +73,9 @@
subsection\<open>Reworking of the Recursion Theory Within \<^term>\<open>M\<close>\<close>
lemma (in M_basic) is_recfun_separation':
- "[| f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g);
- M(r); M(f); M(g); M(a); M(b) |]
- ==> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))"
+ "\<lbrakk>f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g);
+ M(r); M(f); M(g); M(a); M(b)\<rbrakk>
+ \<Longrightarrow> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))"
apply (insert is_recfun_separation [of r f g a b])
apply (simp add: vimage_singleton_iff)
done
@@ -88,10 +88,10 @@
but without them we'd have to undertake
more work to set up the induction formula.\<close>
lemma (in M_basic) is_recfun_equal [rule_format]:
- "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g);
+ "\<lbrakk>is_recfun(r,a,H,f); is_recfun(r,b,H,g);
wellfounded(M,r); trans(r);
- M(f); M(g); M(r); M(x); M(a); M(b) |]
- ==> <x,a> \<in> r \<longrightarrow> <x,b> \<in> r \<longrightarrow> f`x=g`x"
+ M(f); M(g); M(r); M(x); M(a); M(b)\<rbrakk>
+ \<Longrightarrow> <x,a> \<in> r \<longrightarrow> <x,b> \<in> 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)
@@ -111,10 +111,10 @@
done
lemma (in M_basic) is_recfun_cut:
- "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g);
+ "\<lbrakk>is_recfun(r,a,H,f); is_recfun(r,b,H,g);
wellfounded(M,r); trans(r);
- M(f); M(g); M(r); <b,a> \<in> r |]
- ==> restrict(f, r-``{b}) = g"
+ M(f); M(g); M(r); <b,a> \<in> r\<rbrakk>
+ \<Longrightarrow> restrict(f, r-``{b}) = g"
apply (frule_tac f=f in is_recfun_type)
apply (rule fun_extension)
apply (blast intro: transD restrict_type2)
@@ -123,8 +123,8 @@
done
lemma (in M_basic) is_recfun_functional:
- "[|is_recfun(r,a,H,f); is_recfun(r,a,H,g);
- wellfounded(M,r); trans(r); M(f); M(g); M(r) |] ==> f=g"
+ "\<lbrakk>is_recfun(r,a,H,f); is_recfun(r,a,H,g);
+ wellfounded(M,r); trans(r); M(f); M(g); M(r)\<rbrakk> \<Longrightarrow> f=g"
apply (rule fun_extension)
apply (erule is_recfun_type)+
apply (blast intro!: is_recfun_equal dest: transM)
@@ -132,8 +132,8 @@
text\<open>Tells us that \<open>is_recfun\<close> can (in principle) be relativized.\<close>
lemma (in M_basic) is_recfun_relativize:
- "[| M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> is_recfun(r,a,H,f) \<longleftrightarrow>
+ "\<lbrakk>M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> is_recfun(r,a,H,f) \<longleftrightarrow>
(\<forall>z[M]. z \<in> f \<longleftrightarrow>
(\<exists>x[M]. <x,a> \<in> r & z = <x, H(x, restrict(f, r-``{x}))>))"
apply (simp add: is_recfun_def lam_def)
@@ -153,10 +153,10 @@
done
lemma (in M_basic) is_recfun_restrict:
- "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r;
+ "\<lbrakk>wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r;
M(r); M(f);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> is_recfun(r, y, H, restrict(f, r -`` {y}))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> is_recfun(r, y, H, restrict(f, r -`` {y}))"
apply (frule pair_components_in_M, assumption, clarify)
apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff
trans_Int_eq)
@@ -170,14 +170,14 @@
done
lemma (in M_basic) restrict_Y_lemma:
- "[| wellfounded(M,r); trans(r); M(r);
+ "\<lbrakk>wellfounded(M,r); trans(r); M(r);
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); M(Y);
\<forall>b[M].
b \<in> Y \<longleftrightarrow>
(\<exists>x[M]. <x,a1> \<in> r &
(\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
- \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f) |]
- ==> restrict(Y, r -`` {x}) = f"
+ \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f)\<rbrakk>
+ \<Longrightarrow> restrict(Y, r -`` {x}) = f"
apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y \<longleftrightarrow> <y,z>:f")
apply (simp (no_asm_simp) add: restrict_def)
apply (thin_tac "rall(M,P)" for P)+ \<comment> \<open>essential for efficiency\<close>
@@ -199,8 +199,8 @@
text\<open>For typical applications of Replacement for recursive definitions\<close>
lemma (in M_basic) univalent_is_recfun:
- "[|wellfounded(M,r); trans(r); M(r)|]
- ==> univalent (M, A, \<lambda>x p.
+ "\<lbrakk>wellfounded(M,r); trans(r); M(r)\<rbrakk>
+ \<Longrightarrow> univalent (M, A, \<lambda>x p.
\<exists>y[M]. p = \<langle>x,y\<rangle> & (\<exists>f[M]. is_recfun(r,x,H,f) & y = H(x,f)))"
apply (simp add: univalent_def)
apply (blast dest: is_recfun_functional)
@@ -210,12 +210,12 @@
text\<open>Proof of the inductive step for \<open>exists_is_recfun\<close>, since
we must prove two versions.\<close>
lemma (in M_basic) exists_is_recfun_indstep:
- "[|\<forall>y. \<langle>y, a1\<rangle> \<in> r \<longrightarrow> (\<exists>f[M]. is_recfun(r, y, H, f));
+ "\<lbrakk>\<forall>y. \<langle>y, a1\<rangle> \<in> r \<longrightarrow> (\<exists>f[M]. is_recfun(r, y, H, f));
wellfounded(M,r); trans(r); M(r); M(a1);
strong_replacement(M, \<lambda>x z.
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g));
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|]
- ==> \<exists>f[M]. is_recfun(r,a1,H,f)"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. is_recfun(r,a1,H,f)"
apply (drule_tac A="r-``{a1}" in strong_replacementD)
apply blast
txt\<open>Discharge the "univalent" obligation of Replacement\<close>
@@ -243,23 +243,23 @@
text\<open>Relativized version, when we have the (currently weaker) premise
\<^term>\<open>wellfounded(M,r)\<close>\<close>
lemma (in M_basic) wellfounded_exists_is_recfun:
- "[|wellfounded(M,r); trans(r);
- separation(M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r, x, H, f)));
+ "\<lbrakk>wellfounded(M,r); trans(r);
+ separation(M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r, x, H, f)));
strong_replacement(M, \<lambda>x z.
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g));
M(r); M(a);
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> \<exists>f[M]. is_recfun(r,a,H,f)"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> \<exists>f[M]. is_recfun(r,a,H,f)"
apply (rule wellfounded_induct, assumption+, clarify)
apply (rule exists_is_recfun_indstep, assumption+)
done
lemma (in M_basic) wf_exists_is_recfun [rule_format]:
- "[|wf(r); trans(r); M(r);
+ "\<lbrakk>wf(r); trans(r); M(r);
strong_replacement(M, \<lambda>x z.
\<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g));
- \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |]
- ==> M(a) \<longrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f))"
+ \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
+ \<Longrightarrow> M(a) \<longrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f))"
apply (rule wf_induct, assumption+)
apply (frule wf_imp_relativized)
apply (intro impI)
@@ -272,7 +272,7 @@
definition
M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
- "M_is_recfun(M,MH,r,a,f) ==
+ "M_is_recfun(M,MH,r,a,f) \<equiv>
\<forall>z[M]. z \<in> f \<longleftrightarrow>
(\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M].
pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
@@ -281,50 +281,50 @@
definition
is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
- "is_wfrec(M,MH,r,a,z) ==
+ "is_wfrec(M,MH,r,a,z) \<equiv>
\<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)"
definition
wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
- "wfrec_replacement(M,MH,r) ==
+ "wfrec_replacement(M,MH,r) \<equiv>
strong_replacement(M,
\<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
lemma (in M_basic) is_recfun_abs:
- "[| \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); M(r); M(a); M(f);
- relation2(M,MH,H) |]
- ==> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> is_recfun(r,a,H,f)"
+ "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); M(r); M(a); M(f);
+ relation2(M,MH,H)\<rbrakk>
+ \<Longrightarrow> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> is_recfun(r,a,H,f)"
apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize)
apply (rule rall_cong)
apply (blast dest: transM)
done
lemma M_is_recfun_cong [cong]:
- "[| r = r'; a = a'; f = f';
- !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) \<longleftrightarrow> MH'(x,g,y) |]
- ==> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> M_is_recfun(M,MH',r',a',f')"
+ "\<lbrakk>r = r'; a = a'; f = f';
+ \<And>x g y. \<lbrakk>M(x); M(g); M(y)\<rbrakk> \<Longrightarrow> MH(x,g,y) \<longleftrightarrow> MH'(x,g,y)\<rbrakk>
+ \<Longrightarrow> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> M_is_recfun(M,MH',r',a',f')"
by (simp add: M_is_recfun_def)
lemma (in M_basic) is_wfrec_abs:
- "[| \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));
- relation2(M,MH,H); M(r); M(a); M(z) |]
- ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow>
+ "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));
+ relation2(M,MH,H); M(r); M(a); M(z)\<rbrakk>
+ \<Longrightarrow> is_wfrec(M,MH,r,a,z) \<longleftrightarrow>
(\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))"
by (simp add: is_wfrec_def relation2_def is_recfun_abs)
text\<open>Relating \<^term>\<open>wfrec_replacement\<close> to native constructs\<close>
lemma (in M_basic) wfrec_replacement':
- "[|wfrec_replacement(M,MH,r);
+ "\<lbrakk>wfrec_replacement(M,MH,r);
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));
- relation2(M,MH,H); M(r)|]
- ==> strong_replacement(M, \<lambda>x z. \<exists>y[M].
+ relation2(M,MH,H); M(r)\<rbrakk>
+ \<Longrightarrow> strong_replacement(M, \<lambda>x z. \<exists>y[M].
pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))"
by (simp add: wfrec_replacement_def is_wfrec_abs)
lemma wfrec_replacement_cong [cong]:
- "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) \<longleftrightarrow> MH'(x,y,z);
- r=r' |]
- ==> wfrec_replacement(M, %x y. MH(x,y), r) \<longleftrightarrow>
+ "\<lbrakk>\<And>x y z. \<lbrakk>M(x); M(y); M(z)\<rbrakk> \<Longrightarrow> MH(x,y,z) \<longleftrightarrow> MH'(x,y,z);
+ r=r'\<rbrakk>
+ \<Longrightarrow> wfrec_replacement(M, %x y. MH(x,y), r) \<longleftrightarrow>
wfrec_replacement(M, %x y. MH'(x,y), r')"
by (simp add: is_wfrec_def wfrec_replacement_def)
--- a/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,69 +17,69 @@
definition
irreflexive :: "[i=>o,i,i]=>o" where
- "irreflexive(M,A,r) == \<forall>x[M]. x\<in>A \<longrightarrow> <x,x> \<notin> r"
+ "irreflexive(M,A,r) \<equiv> \<forall>x[M]. x\<in>A \<longrightarrow> <x,x> \<notin> r"
definition
transitive_rel :: "[i=>o,i,i]=>o" where
- "transitive_rel(M,A,r) ==
+ "transitive_rel(M,A,r) \<equiv>
\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> (\<forall>z[M]. z\<in>A \<longrightarrow>
<x,y>\<in>r \<longrightarrow> <y,z>\<in>r \<longrightarrow> <x,z>\<in>r))"
definition
linear_rel :: "[i=>o,i,i]=>o" where
- "linear_rel(M,A,r) ==
+ "linear_rel(M,A,r) \<equiv>
\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> <x,y>\<in>r | x=y | <y,x>\<in>r)"
definition
wellfounded :: "[i=>o,i]=>o" where
\<comment> \<open>EVERY non-empty set has an \<open>r\<close>-minimal element\<close>
- "wellfounded(M,r) ==
- \<forall>x[M]. x\<noteq>0 \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+ "wellfounded(M,r) \<equiv>
+ \<forall>x[M]. x\<noteq>0 \<longrightarrow> (\<exists>y[M]. y\<in>x & \<not>(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
definition
wellfounded_on :: "[i=>o,i,i]=>o" where
\<comment> \<open>every non-empty SUBSET OF \<open>A\<close> has an \<open>r\<close>-minimal element\<close>
- "wellfounded_on(M,A,r) ==
- \<forall>x[M]. x\<noteq>0 \<longrightarrow> x\<subseteq>A \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
+ "wellfounded_on(M,A,r) \<equiv>
+ \<forall>x[M]. x\<noteq>0 \<longrightarrow> x\<subseteq>A \<longrightarrow> (\<exists>y[M]. y\<in>x & \<not>(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
definition
wellordered :: "[i=>o,i,i]=>o" where
\<comment> \<open>linear and wellfounded on \<open>A\<close>\<close>
- "wellordered(M,A,r) ==
+ "wellordered(M,A,r) \<equiv>
transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
subsubsection \<open>Trivial absoluteness proofs\<close>
lemma (in M_basic) irreflexive_abs [simp]:
- "M(A) ==> irreflexive(M,A,r) \<longleftrightarrow> irrefl(A,r)"
+ "M(A) \<Longrightarrow> irreflexive(M,A,r) \<longleftrightarrow> irrefl(A,r)"
by (simp add: irreflexive_def irrefl_def)
lemma (in M_basic) transitive_rel_abs [simp]:
- "M(A) ==> transitive_rel(M,A,r) \<longleftrightarrow> trans[A](r)"
+ "M(A) \<Longrightarrow> transitive_rel(M,A,r) \<longleftrightarrow> trans[A](r)"
by (simp add: transitive_rel_def trans_on_def)
lemma (in M_basic) linear_rel_abs [simp]:
- "M(A) ==> linear_rel(M,A,r) \<longleftrightarrow> linear(A,r)"
+ "M(A) \<Longrightarrow> linear_rel(M,A,r) \<longleftrightarrow> linear(A,r)"
by (simp add: linear_rel_def linear_def)
lemma (in M_basic) wellordered_is_trans_on:
- "[| wellordered(M,A,r); M(A) |] ==> trans[A](r)"
+ "\<lbrakk>wellordered(M,A,r); M(A)\<rbrakk> \<Longrightarrow> trans[A](r)"
by (auto simp add: wellordered_def)
lemma (in M_basic) wellordered_is_linear:
- "[| wellordered(M,A,r); M(A) |] ==> linear(A,r)"
+ "\<lbrakk>wellordered(M,A,r); M(A)\<rbrakk> \<Longrightarrow> linear(A,r)"
by (auto simp add: wellordered_def)
lemma (in M_basic) wellordered_is_wellfounded_on:
- "[| wellordered(M,A,r); M(A) |] ==> wellfounded_on(M,A,r)"
+ "\<lbrakk>wellordered(M,A,r); M(A)\<rbrakk> \<Longrightarrow> wellfounded_on(M,A,r)"
by (auto simp add: wellordered_def)
lemma (in M_basic) wellfounded_imp_wellfounded_on:
- "[| wellfounded(M,r); M(A) |] ==> wellfounded_on(M,A,r)"
+ "\<lbrakk>wellfounded(M,r); M(A)\<rbrakk> \<Longrightarrow> wellfounded_on(M,A,r)"
by (auto simp add: wellfounded_def wellfounded_on_def)
lemma (in M_basic) wellfounded_on_subset_A:
- "[| wellfounded_on(M,A,r); B<=A |] ==> wellfounded_on(M,B,r)"
+ "\<lbrakk>wellfounded_on(M,A,r); B<=A\<rbrakk> \<Longrightarrow> wellfounded_on(M,B,r)"
by (simp add: wellfounded_on_def, blast)
@@ -93,35 +93,35 @@
done
lemma (in M_basic) wellfounded_on_imp_wellfounded:
- "[|wellfounded_on(M,A,r); r \<subseteq> A*A|] ==> wellfounded(M,r)"
+ "\<lbrakk>wellfounded_on(M,A,r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> wellfounded(M,r)"
by (simp add: wellfounded_on_iff_wellfounded subset_Int_iff)
lemma (in M_basic) wellfounded_on_field_imp_wellfounded:
- "wellfounded_on(M, field(r), r) ==> wellfounded(M,r)"
+ "wellfounded_on(M, field(r), r) \<Longrightarrow> wellfounded(M,r)"
by (simp add: wellfounded_def wellfounded_on_iff_wellfounded, fast)
lemma (in M_basic) wellfounded_iff_wellfounded_on_field:
- "M(r) ==> wellfounded(M,r) \<longleftrightarrow> wellfounded_on(M, field(r), r)"
+ "M(r) \<Longrightarrow> wellfounded(M,r) \<longleftrightarrow> wellfounded_on(M, field(r), r)"
by (blast intro: wellfounded_imp_wellfounded_on
wellfounded_on_field_imp_wellfounded)
(*Consider the least z in domain(r) such that P(z) does not hold...*)
lemma (in M_basic) wellfounded_induct:
- "[| wellfounded(M,r); M(a); M(r); separation(M, \<lambda>x. ~P(x));
- \<forall>x. M(x) & (\<forall>y. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x) |]
- ==> P(a)"
+ "\<lbrakk>wellfounded(M,r); M(a); M(r); separation(M, \<lambda>x. \<not>P(x));
+ \<forall>x. M(x) & (\<forall>y. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> P(a)"
apply (simp (no_asm_use) add: wellfounded_def)
-apply (drule_tac x="{z \<in> domain(r). ~P(z)}" in rspec)
+apply (drule_tac x="{z \<in> domain(r). \<not>P(z)}" in rspec)
apply (blast dest: transM)+
done
lemma (in M_basic) wellfounded_on_induct:
- "[| a\<in>A; wellfounded_on(M,A,r); M(A);
- separation(M, \<lambda>x. x\<in>A \<longrightarrow> ~P(x));
- \<forall>x\<in>A. M(x) & (\<forall>y\<in>A. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x) |]
- ==> P(a)"
+ "\<lbrakk>a\<in>A; wellfounded_on(M,A,r); M(A);
+ separation(M, \<lambda>x. x\<in>A \<longrightarrow> \<not>P(x));
+ \<forall>x\<in>A. M(x) & (\<forall>y\<in>A. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> P(a)"
apply (simp (no_asm_use) add: wellfounded_on_def)
-apply (drule_tac x="{z\<in>A. z\<in>A \<longrightarrow> ~P(z)}" in rspec)
+apply (drule_tac x="{z\<in>A. z\<in>A \<longrightarrow> \<not>P(z)}" in rspec)
apply (blast intro: transM)+
done
@@ -129,11 +129,11 @@
subsubsection \<open>Kunen's lemma IV 3.14, page 123\<close>
lemma (in M_basic) linear_imp_relativized:
- "linear(A,r) ==> linear_rel(M,A,r)"
+ "linear(A,r) \<Longrightarrow> linear_rel(M,A,r)"
by (simp add: linear_def linear_rel_def)
lemma (in M_basic) trans_on_imp_relativized:
- "trans[A](r) ==> transitive_rel(M,A,r)"
+ "trans[A](r) \<Longrightarrow> transitive_rel(M,A,r)"
by (unfold transitive_rel_def trans_on_def, blast)
lemma (in M_basic) wf_on_imp_relativized:
@@ -143,13 +143,13 @@
done
lemma (in M_basic) wf_imp_relativized:
- "wf(r) ==> wellfounded(M,r)"
+ "wf(r) \<Longrightarrow> wellfounded(M,r)"
apply (simp add: wellfounded_def wf_def, clarify)
apply (drule_tac x=x in spec, blast)
done
lemma (in M_basic) well_ord_imp_relativized:
- "well_ord(A,r) ==> wellordered(M,A,r)"
+ "well_ord(A,r) \<Longrightarrow> wellordered(M,A,r)"
by (simp add: wellordered_def well_ord_def tot_ord_def part_ord_def
linear_imp_relativized trans_on_imp_relativized wf_on_imp_relativized)
@@ -160,12 +160,12 @@
subsection\<open>Relativized versions of order-isomorphisms and order types\<close>
lemma (in M_basic) order_isomorphism_abs [simp]:
- "[| M(A); M(B); M(f) |]
- ==> order_isomorphism(M,A,r,B,s,f) \<longleftrightarrow> f \<in> ord_iso(A,r,B,s)"
+ "\<lbrakk>M(A); M(B); M(f)\<rbrakk>
+ \<Longrightarrow> order_isomorphism(M,A,r,B,s,f) \<longleftrightarrow> f \<in> ord_iso(A,r,B,s)"
by (simp add: order_isomorphism_def ord_iso_def)
lemma (in M_trans) pred_set_abs [simp]:
- "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) \<longleftrightarrow> B = Order.pred(A,x,r)"
+ "\<lbrakk>M(r); M(B)\<rbrakk> \<Longrightarrow> pred_set(M,A,x,r,B) \<longleftrightarrow> B = Order.pred(A,x,r)"
apply (simp add: pred_set_def Order.pred_def)
apply (blast dest: transM)
done
@@ -175,7 +175,7 @@
using pred_separation [of r x] by (simp add: Order.pred_def)
lemma (in M_basic) membership_abs [simp]:
- "[| M(r); M(A) |] ==> membership(M,A,r) \<longleftrightarrow> r = Memrel(A)"
+ "\<lbrakk>M(r); M(A)\<rbrakk> \<Longrightarrow> membership(M,A,r) \<longleftrightarrow> r = Memrel(A)"
apply (simp add: membership_def Memrel_def, safe)
apply (rule equalityI)
apply clarify
@@ -220,14 +220,14 @@
done
lemma (in M_basic) wellfounded_on_asym:
- "[| wellfounded_on(M,A,r); <a,x>\<in>r; a\<in>A; x\<in>A; M(A) |] ==> <x,a>\<notin>r"
+ "\<lbrakk>wellfounded_on(M,A,r); <a,x>\<in>r; a\<in>A; x\<in>A; M(A)\<rbrakk> \<Longrightarrow> <x,a>\<notin>r"
apply (simp add: wellfounded_on_def)
apply (drule_tac x="{x,a}" in rspec)
apply (blast dest: transM)+
done
lemma (in M_basic) wellordered_asym:
- "[| wellordered(M,A,r); <a,x>\<in>r; a\<in>A; x\<in>A; M(A) |] ==> <x,a>\<notin>r"
+ "\<lbrakk>wellordered(M,A,r); <a,x>\<in>r; a\<in>A; x\<in>A; M(A)\<rbrakk> \<Longrightarrow> <x,a>\<notin>r"
by (simp add: wellordered_def, blast dest: wellfounded_on_asym)
end
--- a/src/ZF/Epsilon.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Epsilon.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,19 +9,19 @@
definition
eclose :: "i=>i" where
- "eclose(A) == \<Union>n\<in>nat. nat_rec(n, A, %m r. \<Union>(r))"
+ "eclose(A) \<equiv> \<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)"
+ "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"
definition
rank :: "i=>i" where
- "rank(a) == transrec(a, %x f. \<Union>y\<in>x. succ(f`y))"
+ "rank(a) \<equiv> transrec(a, %x f. \<Union>y\<in>x. succ(f`y))"
definition
transrec2 :: "[i, i, [i,i]=>i] =>i" where
- "transrec2(k, a, b) ==
+ "transrec2(k, a, b) \<equiv>
transrec(k,
%i r. if(i=0, a,
if(\<exists>j. i=succ(j),
@@ -30,11 +30,11 @@
definition
recursor :: "[i, [i,i]=>i, i]=>i" where
- "recursor(a,b,k) == transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))"
+ "recursor(a,b,k) \<equiv> transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))"
definition
rec :: "[i, i, [i,i]=>i]=>i" where
- "rec(k,a,b) == recursor(a,b,k)"
+ "rec(k,a,b) \<equiv> recursor(a,b,k)"
subsection\<open>Basic Closure Properties\<close>
@@ -56,19 +56,19 @@
apply (erule UnionI, assumption)
done
-(* @{term"x \<in> eclose(A) ==> x \<subseteq> eclose(A)"} *)
+(* @{term"x \<in> eclose(A) \<Longrightarrow> x \<subseteq> eclose(A)"} *)
lemmas eclose_subset =
Transset_eclose [unfolded Transset_def, THEN bspec]
-(* @{term"[| A \<in> eclose(B); c \<in> A |] ==> c \<in> eclose(B)"} *)
+(* @{term"\<lbrakk>A \<in> eclose(B); c \<in> A\<rbrakk> \<Longrightarrow> 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 \<in> eclose(A); !!x. [| x \<in> eclose(A); \<forall>y\<in>x. P(y) |] ==> P(x)
- |] ==> P(a)
+ \<lbrakk>a \<in> eclose(A); \<And>x. \<lbrakk>x \<in> eclose(A); \<forall>y\<in>x. P(y)\<rbrakk> \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> P(a)
*)
lemmas eclose_induct =
Transset_induct [OF _ Transset_eclose, induct set: eclose]
@@ -76,7 +76,7 @@
(*Epsilon induction*)
lemma eps_induct:
- "[| !!x. \<forall>y\<in>x. P(y) ==> P(x) |] ==> P(a)"
+ "\<lbrakk>\<And>x. \<forall>y\<in>x. P(y) \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> P(a)"
by (rule arg_in_eclose_sing [THEN eclose_induct], blast)
@@ -85,7 +85,7 @@
(** eclose(A) is the least transitive set including A as a subset. **)
lemma eclose_least_lemma:
- "[| Transset(X); A<=X; n \<in> nat |] ==> nat_rec(n, A, %m r. \<Union>(r)) \<subseteq> X"
+ "\<lbrakk>Transset(X); A<=X; n \<in> nat\<rbrakk> \<Longrightarrow> nat_rec(n, A, %m r. \<Union>(r)) \<subseteq> X"
apply (unfold Transset_def)
apply (erule nat_induct)
apply (simp add: nat_rec_0)
@@ -93,17 +93,17 @@
done
lemma eclose_least:
- "[| Transset(X); A<=X |] ==> eclose(A) \<subseteq> X"
+ "\<lbrakk>Transset(X); A<=X\<rbrakk> \<Longrightarrow> 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!!*)
+(*COMPLETELY DIFFERENT induction principle from eclose_induct\<And>*)
lemma eclose_induct_down [consumes 1]:
- "[| a \<in> eclose(b);
- !!y. [| y \<in> b |] ==> P(y);
- !!y z. [| y \<in> eclose(b); P(y); z \<in> y |] ==> P(z)
- |] ==> P(a)"
+ "\<lbrakk>a \<in> eclose(b);
+ \<And>y. \<lbrakk>y \<in> b\<rbrakk> \<Longrightarrow> P(y);
+ \<And>y z. \<lbrakk>y \<in> eclose(b); P(y); z \<in> y\<rbrakk> \<Longrightarrow> P(z)
+\<rbrakk> \<Longrightarrow> P(a)"
apply (rule eclose_least [THEN subsetD, THEN CollectD2, of "eclose(b)"])
prefer 3 apply assumption
apply (unfold Transset_def)
@@ -111,49 +111,49 @@
apply (blast intro: arg_subset_eclose [THEN subsetD])
done
-lemma Transset_eclose_eq_arg: "Transset(X) ==> eclose(X) = X"
+lemma Transset_eclose_eq_arg: "Transset(X) \<Longrightarrow> eclose(X) = X"
apply (erule equalityI [OF eclose_least arg_subset_eclose])
apply (rule subset_refl)
done
text\<open>A transitive set either is empty or contains the empty set.\<close>
-lemma Transset_0_lemma [rule_format]: "Transset(A) ==> x\<in>A \<longrightarrow> 0\<in>A"
+lemma Transset_0_lemma [rule_format]: "Transset(A) \<Longrightarrow> 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
-lemma Transset_0_disj: "Transset(A) ==> A=0 | 0\<in>A"
+lemma Transset_0_disj: "Transset(A) \<Longrightarrow> A=0 | 0\<in>A"
by (blast dest: Transset_0_lemma)
subsection\<open>Epsilon Recursion\<close>
(*Unused...*)
-lemma mem_eclose_trans: "[| A \<in> eclose(B); B \<in> eclose(C) |] ==> A \<in> eclose(C)"
+lemma mem_eclose_trans: "\<lbrakk>A \<in> eclose(B); B \<in> eclose(C)\<rbrakk> \<Longrightarrow> A \<in> eclose(C)"
by (rule eclose_least [OF Transset_eclose eclose_subset, THEN subsetD],
assumption+)
(*Variant of the previous lemma in a useable form for the sequel*)
lemma mem_eclose_sing_trans:
- "[| A \<in> eclose({B}); B \<in> eclose({C}) |] ==> A \<in> eclose({C})"
+ "\<lbrakk>A \<in> eclose({B}); B \<in> eclose({C})\<rbrakk> \<Longrightarrow> A \<in> eclose({C})"
by (rule eclose_least [OF Transset_eclose singleton_subsetI, THEN subsetD],
assumption+)
-lemma under_Memrel: "[| Transset(i); j \<in> i |] ==> Memrel(i)-``{j} = j"
+lemma under_Memrel: "\<lbrakk>Transset(i); j \<in> i\<rbrakk> \<Longrightarrow> Memrel(i)-``{j} = j"
by (unfold Transset_def, blast)
-lemma lt_Memrel: "j < i ==> Memrel(i) -`` {j} = j"
+lemma lt_Memrel: "j < i \<Longrightarrow> Memrel(i) -`` {j} = j"
by (simp add: lt_def Ord_def under_Memrel)
-(* @{term"j \<in> eclose(A) ==> Memrel(eclose(A)) -`` j = j"} *)
+(* @{term"j \<in> eclose(A) \<Longrightarrow> 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 \<in> eclose({j}); j \<in> eclose({i}) |] ==>
+ "\<lbrakk>k \<in> eclose({j}); j \<in> eclose({i})\<rbrakk> \<Longrightarrow>
wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)"
apply (erule eclose_induct)
apply (rule wfrec_ssubst)
@@ -162,7 +162,7 @@
done
lemma wfrec_eclose_eq2:
- "k \<in> i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)"
+ "k \<in> i \<Longrightarrow> 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
@@ -175,20 +175,20 @@
(*Avoids explosions in proofs; resolve it with a meta-level definition.*)
lemma def_transrec:
- "[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, \<lambda>x\<in>a. f(x))"
+ "\<lbrakk>\<And>x. f(x)\<equiv>transrec(x,H)\<rbrakk> \<Longrightarrow> f(a) = H(a, \<lambda>x\<in>a. f(x))"
apply simp
apply (rule transrec)
done
lemma transrec_type:
- "[| !!x u. [| x \<in> eclose({a}); u \<in> Pi(x,B) |] ==> H(x,u) \<in> B(x) |]
- ==> transrec(a,H) \<in> B(a)"
+ "\<lbrakk>\<And>x u. \<lbrakk>x \<in> eclose({a}); u \<in> Pi(x,B)\<rbrakk> \<Longrightarrow> H(x,u) \<in> B(x)\<rbrakk>
+ \<Longrightarrow> 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)
done
-lemma eclose_sing_Ord: "Ord(i) ==> eclose({i}) \<subseteq> succ(i)"
+lemma eclose_sing_Ord: "Ord(i) \<Longrightarrow> eclose({i}) \<subseteq> succ(i)"
apply (erule Ord_is_Transset [THEN Transset_succ, THEN eclose_least])
apply (rule succI1 [THEN singleton_subsetI])
done
@@ -198,7 +198,7 @@
apply (frule eclose_subset, blast)
done
-lemma eclose_sing_Ord_eq: "Ord(i) ==> eclose({i}) = succ(i)"
+lemma eclose_sing_Ord_eq: "Ord(i) \<Longrightarrow> eclose({i}) = succ(i)"
apply (rule equalityI)
apply (erule eclose_sing_Ord)
apply (rule succ_subset_eclose_sing)
@@ -207,7 +207,7 @@
lemma Ord_transrec_type:
assumes jini: "j \<in> i"
and ordi: "Ord(i)"
- and minor: " !!x u. [| x \<in> i; u \<in> Pi(x,B) |] ==> H(x,u) \<in> B(x)"
+ and minor: " \<And>x u. \<lbrakk>x \<in> i; u \<in> Pi(x,B)\<rbrakk> \<Longrightarrow> H(x,u) \<in> B(x)"
shows "transrec(j,H) \<in> B(j)"
apply (rule transrec_type)
apply (insert jini ordi)
@@ -229,25 +229,25 @@
apply (erule bspec, assumption)
done
-lemma rank_of_Ord: "Ord(i) ==> rank(i) = i"
+lemma rank_of_Ord: "Ord(i) \<Longrightarrow> rank(i) = i"
apply (erule trans_induct)
apply (subst rank)
apply (simp add: Ord_equality)
done
-lemma rank_lt: "a \<in> b ==> rank(a) < rank(b)"
+lemma rank_lt: "a \<in> b \<Longrightarrow> rank(a) < rank(b)"
apply (rule_tac a1 = b in rank [THEN ssubst])
apply (erule UN_I [THEN ltI])
apply (rule_tac [2] Ord_UN, auto)
done
-lemma eclose_rank_lt: "a \<in> eclose(b) ==> rank(a) < rank(b)"
+lemma eclose_rank_lt: "a \<in> eclose(b) \<Longrightarrow> rank(a) < rank(b)"
apply (erule eclose_induct_down)
apply (erule rank_lt)
apply (erule rank_lt [THEN lt_trans], assumption)
done
-lemma rank_mono: "a<=b ==> rank(a) \<le> rank(b)"
+lemma rank_mono: "a<=b \<Longrightarrow> rank(a) \<le> rank(b)"
apply (rule subset_imp_le)
apply (auto simp add: rank [of a] rank [of b])
done
@@ -305,14 +305,14 @@
(*Not clear how to remove the P(a) condition, since the "then" part
must refer to "a"*)
lemma the_equality_if:
- "P(a) ==> (THE x. P(x)) = (if (\<exists>!x. P(x)) then a else 0)"
+ "P(a) \<Longrightarrow> (THE x. P(x)) = (if (\<exists>!x. P(x)) then a else 0)"
by (simp add: the_0 the_equality2)
(*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 \<in> domain(f); function(f)|] ==> rank(f`i) < rank(f)"
+lemma rank_apply: "\<lbrakk>i \<in> domain(f); function(f)\<rbrakk> \<Longrightarrow> rank(f`i) < rank(f)"
apply clarify
apply (simp add: function_apply_equality)
apply (blast intro: lt_trans rank_lt rank_pair2)
@@ -321,12 +321,12 @@
subsection\<open>Corollaries of Leastness\<close>
-lemma mem_eclose_subset: "A \<in> B ==> eclose(A)<=eclose(B)"
+lemma mem_eclose_subset: "A \<in> B \<Longrightarrow> eclose(A)<=eclose(B)"
apply (rule Transset_eclose [THEN eclose_least])
apply (erule arg_into_eclose [THEN eclose_subset])
done
-lemma eclose_mono: "A<=B ==> eclose(A) \<subseteq> eclose(B)"
+lemma eclose_mono: "A<=B \<Longrightarrow> eclose(A) \<subseteq> eclose(B)"
apply (rule Transset_eclose [THEN eclose_least])
apply (erule subset_trans)
apply (rule arg_subset_eclose)
@@ -352,14 +352,14 @@
done
lemma transrec2_Limit:
- "Limit(i) ==> transrec2(i,a,b) = (\<Union>j<i. transrec2(j,a,b))"
+ "Limit(i) \<Longrightarrow> 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)
done
lemma def_transrec2:
- "(!!x. f(x)==transrec2(x,a,b))
- ==> f(0) = a &
+ "(\<And>x. f(x)\<equiv>transrec2(x,a,b))
+ \<Longrightarrow> 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 \<in> nat;
+ "\<lbrakk>n \<in> nat;
a \<in> C(0);
- !!m z. [| m \<in> nat; z \<in> C(m) |] ==> b(m,z): C(succ(m)) |]
- ==> rec(n,a,b) \<in> C(n)"
+ \<And>m z. \<lbrakk>m \<in> nat; z \<in> C(m)\<rbrakk> \<Longrightarrow> b(m,z): C(succ(m))\<rbrakk>
+ \<Longrightarrow> rec(n,a,b) \<in> C(n)"
by (erule nat_induct, auto)
end
--- a/src/ZF/EquivClass.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/EquivClass.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,49 +9,49 @@
definition
quotient :: "[i,i]=>i" (infixl \<open>'/'/\<close> 90) (*set of equiv classes*) where
- "A//r == {r``{x} . x \<in> A}"
+ "A//r \<equiv> {r``{x} . x \<in> A}"
definition
congruent :: "[i,i=>i]=>o" where
- "congruent(r,b) == \<forall>y z. <y,z>:r \<longrightarrow> b(y)=b(z)"
+ "congruent(r,b) \<equiv> \<forall>y z. <y,z>:r \<longrightarrow> b(y)=b(z)"
definition
congruent2 :: "[i,i,[i,i]=>i]=>o" where
- "congruent2(r1,r2,b) == \<forall>y1 z1 y2 z2.
+ "congruent2(r1,r2,b) \<equiv> \<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 \<open>respects\<close> 80) where
- "f respects r == congruent(r,f)"
+ "f respects r \<equiv> congruent(r,f)"
abbreviation
RESPECTS2 ::"[i=>i=>i, i] => o" (infixr \<open>respects2 \<close> 80) where
- "f respects2 r == congruent2(r,r,f)"
+ "f respects2 r \<equiv> congruent2(r,r,f)"
\<comment> \<open>Abbreviation for the common case where the relations are identical\<close>
subsection\<open>Suppes, Theorem 70:
\<^term>\<open>r\<close> is an equiv relation iff \<^term>\<open>converse(r) O r = r\<close>\<close>
-(** first half: equiv(A,r) ==> converse(r) O r = r **)
+(** first half: equiv(A,r) \<Longrightarrow> converse(r) O r = r **)
lemma sym_trans_comp_subset:
- "[| sym(r); trans(r) |] ==> converse(r) O r \<subseteq> r"
+ "\<lbrakk>sym(r); trans(r)\<rbrakk> \<Longrightarrow> converse(r) O r \<subseteq> r"
by (unfold trans_def sym_def, blast)
lemma refl_comp_subset:
- "[| refl(A,r); r \<subseteq> A*A |] ==> r \<subseteq> converse(r) O r"
+ "\<lbrakk>refl(A,r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> r \<subseteq> converse(r) O r"
by (unfold refl_def, blast)
lemma equiv_comp_eq:
- "equiv(A,r) ==> converse(r) O r = r"
+ "equiv(A,r) \<Longrightarrow> converse(r) O r = r"
apply (unfold equiv_def)
apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset)
done
(*second half*)
lemma comp_equivI:
- "[| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"
+ "\<lbrakk>converse(r) O r = r; domain(r) = A\<rbrakk> \<Longrightarrow> equiv(A,r)"
apply (unfold equiv_def refl_def sym_def trans_def)
apply (erule equalityE)
apply (subgoal_tac "\<forall>x y. <x,y> \<in> r \<longrightarrow> <y,x> \<in> r", blast+)
@@ -61,63 +61,63 @@
(*Lemma for the next result*)
lemma equiv_class_subset:
- "[| sym(r); trans(r); <a,b>: r |] ==> r``{a} \<subseteq> r``{b}"
+ "\<lbrakk>sym(r); trans(r); <a,b>: r\<rbrakk> \<Longrightarrow> r``{a} \<subseteq> r``{b}"
by (unfold trans_def sym_def, blast)
lemma equiv_class_eq:
- "[| equiv(A,r); <a,b>: r |] ==> r``{a} = r``{b}"
+ "\<lbrakk>equiv(A,r); <a,b>: r\<rbrakk> \<Longrightarrow> r``{a} = r``{b}"
apply (unfold equiv_def)
apply (safe del: subsetI intro!: equalityI equiv_class_subset)
apply (unfold sym_def, blast)
done
lemma equiv_class_self:
- "[| equiv(A,r); a \<in> A |] ==> a \<in> r``{a}"
+ "\<lbrakk>equiv(A,r); a \<in> A\<rbrakk> \<Longrightarrow> a \<in> r``{a}"
by (unfold equiv_def refl_def, blast)
(*Lemma for the next result*)
lemma subset_equiv_class:
- "[| equiv(A,r); r``{b} \<subseteq> r``{a}; b \<in> A |] ==> <a,b>: r"
+ "\<lbrakk>equiv(A,r); r``{b} \<subseteq> r``{a}; b \<in> A\<rbrakk> \<Longrightarrow> <a,b>: r"
by (unfold equiv_def refl_def, blast)
-lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b \<in> A |] ==> <a,b>: r"
+lemma eq_equiv_class: "\<lbrakk>r``{a} = r``{b}; equiv(A,r); b \<in> A\<rbrakk> \<Longrightarrow> <a,b>: r"
by (assumption | rule equalityD2 subset_equiv_class)+
(*thus r``{a} = r``{b} as well*)
lemma equiv_class_nondisjoint:
- "[| equiv(A,r); x: (r``{a} \<inter> r``{b}) |] ==> <a,b>: r"
+ "\<lbrakk>equiv(A,r); x: (r``{a} \<inter> r``{b})\<rbrakk> \<Longrightarrow> <a,b>: r"
by (unfold equiv_def trans_def sym_def, blast)
-lemma equiv_type: "equiv(A,r) ==> r \<subseteq> A*A"
+lemma equiv_type: "equiv(A,r) \<Longrightarrow> r \<subseteq> A*A"
by (unfold equiv_def, blast)
lemma equiv_class_eq_iff:
- "equiv(A,r) ==> <x,y>: r \<longleftrightarrow> r``{x} = r``{y} & x \<in> A & y \<in> A"
+ "equiv(A,r) \<Longrightarrow> <x,y>: r \<longleftrightarrow> r``{x} = r``{y} & x \<in> A & y \<in> A"
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
lemma eq_equiv_class_iff:
- "[| equiv(A,r); x \<in> A; y \<in> A |] ==> r``{x} = r``{y} \<longleftrightarrow> <x,y>: r"
+ "\<lbrakk>equiv(A,r); x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> r``{x} = r``{y} \<longleftrightarrow> <x,y>: r"
by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
(*** Quotients ***)
(** Introduction/elimination rules -- needed? **)
-lemma quotientI [TC]: "x \<in> A ==> r``{x}: A//r"
+lemma quotientI [TC]: "x \<in> A \<Longrightarrow> r``{x}: A//r"
apply (unfold quotient_def)
apply (erule RepFunI)
done
lemma quotientE:
- "[| X \<in> A//r; !!x. [| X = r``{x}; x \<in> A |] ==> P |] ==> P"
+ "\<lbrakk>X \<in> A//r; \<And>x. \<lbrakk>X = r``{x}; x \<in> A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold quotient_def, blast)
lemma Union_quotient:
- "equiv(A,r) ==> \<Union>(A//r) = A"
+ "equiv(A,r) \<Longrightarrow> \<Union>(A//r) = A"
by (unfold equiv_def refl_def quotient_def, blast)
lemma quotient_disj:
- "[| equiv(A,r); X \<in> A//r; Y \<in> A//r |] ==> X=Y | (X \<inter> Y \<subseteq> 0)"
+ "\<lbrakk>equiv(A,r); X \<in> A//r; Y \<in> A//r\<rbrakk> \<Longrightarrow> 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,7 +130,7 @@
(*Conversion rule*)
lemma UN_equiv_class:
- "[| equiv(A,r); b respects r; a \<in> A |] ==> (\<Union>x\<in>r``{a}. b(x)) = b(a)"
+ "\<lbrakk>equiv(A,r); b respects r; a \<in> A\<rbrakk> \<Longrightarrow> (\<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)
@@ -139,20 +139,20 @@
(*type checking of @{term"\<Union>x\<in>r``{a}. b(x)"} *)
lemma UN_equiv_class_type:
- "[| equiv(A,r); b respects r; X \<in> A//r; !!x. x \<in> A ==> b(x) \<in> B |]
- ==> (\<Union>x\<in>X. b(x)) \<in> B"
+ "\<lbrakk>equiv(A,r); b respects r; X \<in> A//r; \<And>x. x \<in> A \<Longrightarrow> b(x) \<in> B\<rbrakk>
+ \<Longrightarrow> (\<Union>x\<in>X. b(x)) \<in> B"
apply (unfold quotient_def, safe)
apply (simp (no_asm_simp) add: UN_equiv_class)
done
(*Sufficient conditions for injectiveness. Could weaken premises!
- major premise could be an inclusion; bcong could be !!y. y \<in> A ==> b(y):B
+ major premise could be an inclusion; bcong could be \<And>y. y \<in> A \<Longrightarrow> b(y):B
*)
lemma UN_equiv_class_inject:
- "[| equiv(A,r); b respects r;
+ "\<lbrakk>equiv(A,r); b respects r;
(\<Union>x\<in>X. b(x))=(\<Union>y\<in>Y. b(y)); X \<in> A//r; Y \<in> A//r;
- !!x y. [| x \<in> A; y \<in> A; b(x)=b(y) |] ==> <x,y>:r |]
- ==> X=Y"
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> A; b(x)=b(y)\<rbrakk> \<Longrightarrow> <x,y>:r\<rbrakk>
+ \<Longrightarrow> X=Y"
apply (unfold quotient_def, safe)
apply (rule equiv_class_eq, assumption)
apply (simp add: UN_equiv_class [of A r b])
@@ -162,11 +162,11 @@
subsection\<open>Defining Binary Operations upon Equivalence Classes\<close>
lemma congruent2_implies_congruent:
- "[| equiv(A,r1); congruent2(r1,r2,b); a \<in> A |] ==> congruent(r2,b(a))"
+ "\<lbrakk>equiv(A,r1); congruent2(r1,r2,b); a \<in> A\<rbrakk> \<Longrightarrow> congruent(r2,b(a))"
by (unfold congruent_def congruent2_def equiv_def refl_def, blast)
lemma congruent2_implies_congruent_UN:
- "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \<in> A2 |] ==>
+ "\<lbrakk>equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \<in> A2\<rbrakk> \<Longrightarrow>
congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
apply (unfold congruent_def, safe)
apply (frule equiv_type [THEN subsetD], assumption)
@@ -176,17 +176,17 @@
done
lemma UN_equiv_class2:
- "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2 |]
- ==> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)"
+ "\<lbrakk>equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2\<rbrakk>
+ \<Longrightarrow> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)"
by (simp add: UN_equiv_class congruent2_implies_congruent
congruent2_implies_congruent_UN)
(*type checking*)
lemma UN_equiv_class_type2:
- "[| equiv(A,r); b respects2 r;
+ "\<lbrakk>equiv(A,r); b respects2 r;
X1: A//r; X2: A//r;
- !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) \<in> B
- |] ==> (\<Union>x1\<in>X1. \<Union>x2\<in>X2. b(x1,x2)) \<in> B"
+ \<And>x1 x2. \<lbrakk>x1: A; x2: A\<rbrakk> \<Longrightarrow> b(x1,x2) \<in> B
+\<rbrakk> \<Longrightarrow> (\<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
congruent2_implies_congruent quotientI)
@@ -196,18 +196,18 @@
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
than the direct proof*)
lemma congruent2I:
- "[| 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)"
+ "\<lbrakk>equiv(A1,r1); equiv(A2,r2);
+ \<And>y z w. \<lbrakk>w \<in> A2; <y,z> \<in> r1\<rbrakk> \<Longrightarrow> b(y,w) = b(z,w);
+ \<And>y z w. \<lbrakk>w \<in> A1; <y,z> \<in> r2\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)
+\<rbrakk> \<Longrightarrow> congruent2(r1,r2,b)"
apply (unfold congruent2_def equiv_def refl_def, safe)
apply (blast intro: trans)
done
lemma congruent2_commuteI:
assumes equivA: "equiv(A,r)"
- and commute: "!! y z. [| y \<in> A; z \<in> A |] ==> b(y,z) = b(z,y)"
- and congt: "!! y z w. [| w \<in> A; <y,z>: r |] ==> b(w,y) = b(w,z)"
+ and commute: "\<And>y z. \<lbrakk>y \<in> A; z \<in> A\<rbrakk> \<Longrightarrow> b(y,z) = b(z,y)"
+ and congt: "\<And>y z w. \<lbrakk>w \<in> A; <y,z>: r\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)"
shows "b respects2 r"
apply (insert equivA [THEN equiv_type, THEN subsetD])
apply (rule congruent2I [OF equivA equivA])
@@ -219,10 +219,10 @@
(*Obsolete?*)
lemma congruent_commuteI:
- "[| equiv(A,r); Z \<in> A//r;
- !!w. [| w \<in> A |] ==> congruent(r, %z. b(w,z));
- !!x y. [| x \<in> A; y \<in> A |] ==> b(y,x) = b(x,y)
- |] ==> congruent(r, %w. \<Union>z\<in>Z. b(w,z))"
+ "\<lbrakk>equiv(A,r); Z \<in> A//r;
+ \<And>w. \<lbrakk>w \<in> A\<rbrakk> \<Longrightarrow> congruent(r, %z. b(w,z));
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> b(y,x) = b(x,y)
+\<rbrakk> \<Longrightarrow> 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)
--- a/src/ZF/Finite.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Finite.thy Tue Sep 27 16:51:35 2022 +0100
@@ -2,7 +2,7 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1994 University of Cambridge
-prove: b \<in> Fin(A) ==> inj(b,b) \<subseteq> surj(b,b)
+prove: b \<in> Fin(A) \<Longrightarrow> inj(b,b) \<subseteq> surj(b,b)
*)
section\<open>Finite Powerset Operator and Finite Function Space\<close>
@@ -25,7 +25,7 @@
domains "Fin(A)" \<subseteq> "Pow(A)"
intros
emptyI: "0 \<in> Fin(A)"
- consI: "[| a \<in> A; b \<in> Fin(A) |] ==> cons(a,b) \<in> Fin(A)"
+ consI: "\<lbrakk>a \<in> A; b \<in> Fin(A)\<rbrakk> \<Longrightarrow> cons(a,b) \<in> Fin(A)"
type_intros empty_subsetI cons_subsetI PowI
type_elims PowD [elim_format]
@@ -33,31 +33,31 @@
domains "FiniteFun(A,B)" \<subseteq> "Fin(A*B)"
intros
emptyI: "0 \<in> A -||> B"
- consI: "[| a \<in> A; b \<in> B; h \<in> A -||> B; a \<notin> domain(h) |]
- ==> cons(<a,b>,h) \<in> A -||> B"
+ consI: "\<lbrakk>a \<in> A; b \<in> B; h \<in> A -||> B; a \<notin> domain(h)\<rbrakk>
+ \<Longrightarrow> cons(<a,b>,h) \<in> A -||> B"
type_intros Fin.intros
subsection \<open>Finite Powerset Operator\<close>
-lemma Fin_mono: "A<=B ==> Fin(A) \<subseteq> Fin(B)"
+lemma Fin_mono: "A<=B \<Longrightarrow> Fin(A) \<subseteq> Fin(B)"
apply (unfold Fin.defs)
apply (rule lfp_mono)
apply (rule Fin.bnd_mono)+
apply blast
done
-(* @{term"A \<in> Fin(B) ==> A \<subseteq> B"} *)
+(* @{term"A \<in> Fin(B) \<Longrightarrow> A \<subseteq> B"} *)
lemmas FinD = Fin.dom_subset [THEN subsetD, THEN PowD]
(** Induction on finite sets **)
(*Discharging @{term"x\<notin>y"} entails extra work*)
lemma Fin_induct [case_names 0 cons, induct set: Fin]:
- "[| b \<in> Fin(A);
+ "\<lbrakk>b \<in> Fin(A);
P(0);
- !!x y. [| x \<in> A; y \<in> Fin(A); x\<notin>y; P(y) |] ==> P(cons(x,y))
- |] ==> P(b)"
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> Fin(A); x\<notin>y; P(y)\<rbrakk> \<Longrightarrow> P(cons(x,y))
+\<rbrakk> \<Longrightarrow> P(b)"
apply (erule Fin.induct, simp)
apply (case_tac "a \<in> b")
apply (erule cons_absorb [THEN ssubst], assumption) (*backtracking!*)
@@ -72,53 +72,53 @@
by (blast intro: Fin.emptyI dest: FinD)
(*The union of two finite sets is finite.*)
-lemma Fin_UnI [simp]: "[| b \<in> Fin(A); c \<in> Fin(A) |] ==> b \<union> c \<in> Fin(A)"
+lemma Fin_UnI [simp]: "\<lbrakk>b \<in> Fin(A); c \<in> Fin(A)\<rbrakk> \<Longrightarrow> 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 \<in> Fin(Fin(A)) ==> \<Union>(C) \<in> Fin(A)"
+lemma Fin_UnionI: "C \<in> Fin(Fin(A)) \<Longrightarrow> \<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 \<in> Fin(A) ==> \<forall>z. z<=b \<longrightarrow> z \<in> Fin(A)"
+lemma Fin_subset_lemma [rule_format]: "b \<in> Fin(A) \<Longrightarrow> \<forall>z. z<=b \<longrightarrow> z \<in> Fin(A)"
apply (erule Fin_induct)
apply (simp add: subset_empty_iff)
apply (simp add: subset_cons_iff distrib_simps, safe)
apply (erule_tac b = z in cons_Diff [THEN subst], simp)
done
-lemma Fin_subset: "[| c<=b; b \<in> Fin(A) |] ==> c \<in> Fin(A)"
+lemma Fin_subset: "\<lbrakk>c<=b; b \<in> Fin(A)\<rbrakk> \<Longrightarrow> c \<in> Fin(A)"
by (blast intro: Fin_subset_lemma)
-lemma Fin_IntI1 [intro,simp]: "b \<in> Fin(A) ==> b \<inter> c \<in> Fin(A)"
+lemma Fin_IntI1 [intro,simp]: "b \<in> Fin(A) \<Longrightarrow> b \<inter> c \<in> Fin(A)"
by (blast intro: Fin_subset)
-lemma Fin_IntI2 [intro,simp]: "c \<in> Fin(A) ==> b \<inter> c \<in> Fin(A)"
+lemma Fin_IntI2 [intro,simp]: "c \<in> Fin(A) \<Longrightarrow> b \<inter> c \<in> Fin(A)"
by (blast intro: Fin_subset)
lemma Fin_0_induct_lemma [rule_format]:
- "[| c \<in> Fin(A); b \<in> Fin(A); P(b);
- !!x y. [| x \<in> A; y \<in> Fin(A); x \<in> y; P(y) |] ==> P(y-{x})
- |] ==> c<=b \<longrightarrow> P(b-c)"
+ "\<lbrakk>c \<in> Fin(A); b \<in> Fin(A); P(b);
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> Fin(A); x \<in> y; P(y)\<rbrakk> \<Longrightarrow> P(y-{x})
+\<rbrakk> \<Longrightarrow> 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])
done
lemma Fin_0_induct:
- "[| b \<in> Fin(A);
+ "\<lbrakk>b \<in> Fin(A);
P(b);
- !!x y. [| x \<in> A; y \<in> Fin(A); x \<in> y; P(y) |] ==> P(y-{x})
- |] ==> P(0)"
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> Fin(A); x \<in> y; P(y)\<rbrakk> \<Longrightarrow> P(y-{x})
+\<rbrakk> \<Longrightarrow> P(0)"
apply (rule Diff_cancel [THEN subst])
apply (blast intro: Fin_0_induct_lemma)
done
(*Functions from a finite ordinal*)
-lemma nat_fun_subset_Fin: "n \<in> nat ==> n->A \<subseteq> Fin(nat*A)"
+lemma nat_fun_subset_Fin: "n \<in> nat \<Longrightarrow> 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,29 +129,29 @@
subsection\<open>Finite Function Space\<close>
lemma FiniteFun_mono:
- "[| A<=C; B<=D |] ==> A -||> B \<subseteq> C -||> D"
+ "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> 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 \<subseteq> B -||> B"
+lemma FiniteFun_mono1: "A<=B \<Longrightarrow> A -||> A \<subseteq> B -||> B"
by (blast dest: FiniteFun_mono)
-lemma FiniteFun_is_fun: "h \<in> A -||>B ==> h \<in> domain(h) -> B"
+lemma FiniteFun_is_fun: "h \<in> A -||>B \<Longrightarrow> h \<in> domain(h) -> B"
apply (erule FiniteFun.induct, simp)
apply (simp add: fun_extend3)
done
-lemma FiniteFun_domain_Fin: "h \<in> A -||>B ==> domain(h) \<in> Fin(A)"
+lemma FiniteFun_domain_Fin: "h \<in> A -||>B \<Longrightarrow> 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 \<in> A-||>B ==> \<forall>z. z<=b \<longrightarrow> z \<in> A-||>B"
+ "b \<in> A-||>B \<Longrightarrow> \<forall>z. z<=b \<longrightarrow> z \<in> A-||>B"
apply (erule FiniteFun.induct)
apply (simp add: subset_empty_iff FiniteFun.intros)
apply (simp add: subset_cons_iff distrib_simps, safe)
@@ -160,12 +160,12 @@
apply (fast intro!: FiniteFun.intros)
done
-lemma FiniteFun_subset: "[| c<=b; b \<in> A-||>B |] ==> c \<in> A-||>B"
+lemma FiniteFun_subset: "\<lbrakk>c<=b; b \<in> A-||>B\<rbrakk> \<Longrightarrow> c \<in> A-||>B"
by (blast intro: FiniteFun_subset_lemma)
(** Some further results by Sidi O. Ehmety **)
-lemma fun_FiniteFunI [rule_format]: "A \<in> Fin(X) ==> \<forall>f. f \<in> A->B \<longrightarrow> f \<in> A-||>B"
+lemma fun_FiniteFunI [rule_format]: "A \<in> Fin(X) \<Longrightarrow> \<forall>f. f \<in> A->B \<longrightarrow> f \<in> A-||>B"
apply (erule Fin.induct)
apply (simp add: FiniteFun.intros, clarify)
apply (case_tac "a \<in> b")
@@ -178,7 +178,7 @@
FiniteFun_mono [THEN [2] rev_subsetD])
done
-lemma lam_FiniteFun: "A \<in> Fin(X) ==> (\<lambda>x\<in>A. b(x)) \<in> A -||> {b(x). x \<in> A}"
+lemma lam_FiniteFun: "A \<in> Fin(X) \<Longrightarrow> (\<lambda>x\<in>A. b(x)) \<in> A -||> {b(x). x \<in> A}"
by (blast intro: fun_FiniteFunI lam_funtype)
lemma FiniteFun_Collect_iff:
@@ -201,7 +201,7 @@
definition
contents :: "i=>i" where
- "contents(X) == THE x. X = {x}"
+ "contents(X) \<equiv> THE x. X = {x}"
lemma contents_eq [simp]: "contents ({x}) = x"
by (simp add: contents_def)
--- a/src/ZF/Fixedpt.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Fixedpt.thy Tue Sep 27 16:51:35 2022 +0100
@@ -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 & (\<forall>W X. W<=X \<longrightarrow> X<=D \<longrightarrow> h(W) \<subseteq> h(X))"
+ "bnd_mono(D,h) \<equiv> 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) \<subseteq> X})"
+ "lfp(D,h) \<equiv> \<Inter>({X: Pow(D). h(X) \<subseteq> X})"
definition
gfp :: "[i,i=>i]=>i" where
- "gfp(D,h) == \<Union>({X: Pow(D). X \<subseteq> h(X)})"
+ "gfp(D,h) \<equiv> \<Union>({X: Pow(D). X \<subseteq> h(X)})"
text\<open>The theorem is proved in the lattice of subsets of \<^term>\<open>D\<close>,
namely \<^term>\<open>Pow(D)\<close>, with Inter as the greatest lower bound.\<close>
@@ -26,33 +26,33 @@
subsection\<open>Monotone Operators\<close>
lemma bnd_monoI:
- "[| h(D)<=D;
- !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) \<subseteq> h(X)
- |] ==> bnd_mono(D,h)"
+ "\<lbrakk>h(D)<=D;
+ \<And>W X. \<lbrakk>W<=D; X<=D; W<=X\<rbrakk> \<Longrightarrow> h(W) \<subseteq> h(X)
+\<rbrakk> \<Longrightarrow> bnd_mono(D,h)"
by (unfold bnd_mono_def, clarify, blast)
-lemma bnd_monoD1: "bnd_mono(D,h) ==> h(D) \<subseteq> D"
+lemma bnd_monoD1: "bnd_mono(D,h) \<Longrightarrow> 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) \<subseteq> h(X)"
+lemma bnd_monoD2: "\<lbrakk>bnd_mono(D,h); W<=X; X<=D\<rbrakk> \<Longrightarrow> h(W) \<subseteq> h(X)"
by (unfold bnd_mono_def, blast)
lemma bnd_mono_subset:
- "[| bnd_mono(D,h); X<=D |] ==> h(X) \<subseteq> D"
+ "\<lbrakk>bnd_mono(D,h); X<=D\<rbrakk> \<Longrightarrow> h(X) \<subseteq> D"
by (unfold bnd_mono_def, clarify, blast)
lemma bnd_mono_Un:
- "[| bnd_mono(D,h); A \<subseteq> D; B \<subseteq> D |] ==> h(A) \<union> h(B) \<subseteq> h(A \<union> B)"
+ "\<lbrakk>bnd_mono(D,h); A \<subseteq> D; B \<subseteq> D\<rbrakk> \<Longrightarrow> 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) \<subseteq> D |]
- ==> (\<Union>i\<in>I. h(A(i))) \<subseteq> h((\<Union>i\<in>I. A(i)))"
+ "\<lbrakk>bnd_mono(D,h); \<forall>i\<in>I. A(i) \<subseteq> D\<rbrakk>
+ \<Longrightarrow> (\<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 \<subseteq> D; B \<subseteq> D |] ==> h(A \<inter> B) \<subseteq> h(A) \<inter> h(B)"
+ "\<lbrakk>bnd_mono(D,h); A \<subseteq> D; B \<subseteq> D\<rbrakk> \<Longrightarrow> 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,7 +73,7 @@
(*lfp is contained in each pre-fixedpoint*)
lemma lfp_lowerbound:
- "[| h(A) \<subseteq> A; A<=D |] ==> lfp(D,h) \<subseteq> A"
+ "\<lbrakk>h(A) \<subseteq> A; A<=D\<rbrakk> \<Longrightarrow> lfp(D,h) \<subseteq> A"
by (unfold lfp_def, blast)
(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*)
@@ -81,29 +81,29 @@
by (unfold lfp_def Inter_def, blast)
(*Used in datatype package*)
-lemma def_lfp_subset: "A == lfp(D,h) ==> A \<subseteq> D"
+lemma def_lfp_subset: "A \<equiv> lfp(D,h) \<Longrightarrow> A \<subseteq> D"
apply simp
apply (rule lfp_subset)
done
lemma lfp_greatest:
- "[| h(D) \<subseteq> D; !!X. [| h(X) \<subseteq> X; X<=D |] ==> A<=X |] ==> A \<subseteq> lfp(D,h)"
+ "\<lbrakk>h(D) \<subseteq> D; \<And>X. \<lbrakk>h(X) \<subseteq> X; X<=D\<rbrakk> \<Longrightarrow> A<=X\<rbrakk> \<Longrightarrow> 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)) \<subseteq> A"
+ "\<lbrakk>bnd_mono(D,h); h(A)<=A; A<=D\<rbrakk> \<Longrightarrow> 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)) \<subseteq> lfp(D,h)"
+lemma lfp_lemma2: "bnd_mono(D,h) \<Longrightarrow> 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) \<subseteq> h(lfp(D,h))"
+ "bnd_mono(D,h) \<Longrightarrow> lfp(D,h) \<subseteq> h(lfp(D,h))"
apply (rule lfp_lowerbound)
apply (rule bnd_monoD2, assumption)
apply (rule lfp_lemma2, assumption)
@@ -111,7 +111,7 @@
apply (rule lfp_subset)+
done
-lemma lfp_unfold: "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))"
+lemma lfp_unfold: "bnd_mono(D,h) \<Longrightarrow> lfp(D,h) = h(lfp(D,h))"
apply (rule equalityI)
apply (erule lfp_lemma3)
apply (erule lfp_lemma2)
@@ -119,7 +119,7 @@
(*Definition form, to control unfolding*)
lemma def_lfp_unfold:
- "[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"
+ "\<lbrakk>A\<equiv>lfp(D,h); bnd_mono(D,h)\<rbrakk> \<Longrightarrow> A = h(A)"
apply simp
apply (erule lfp_unfold)
done
@@ -127,17 +127,17 @@
subsection\<open>General Induction Rule for Least Fixedpoints\<close>
lemma Collect_is_pre_fixedpt:
- "[| 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)"
+ "\<lbrakk>bnd_mono(D,h); \<And>x. x \<in> h(Collect(lfp(D,h),P)) \<Longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> 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 \<in> lfp(D,h);
- !!x. x \<in> h(Collect(lfp(D,h),P)) ==> P(x)
- |] ==> P(a)"
+ "\<lbrakk>bnd_mono(D,h); a \<in> lfp(D,h);
+ \<And>x. x \<in> h(Collect(lfp(D,h),P)) \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> P(a)"
apply (rule Collect_is_pre_fixedpt
[THEN lfp_lowerbound, THEN subsetD, THEN CollectD2])
apply (rule_tac [3] lfp_subset [THEN Collect_subset [THEN subset_trans]],
@@ -146,15 +146,15 @@
(*Definition form, to control unfolding*)
lemma def_induct:
- "[| A == lfp(D,h); bnd_mono(D,h); a:A;
- !!x. x \<in> h(Collect(A,P)) ==> P(x)
- |] ==> P(a)"
+ "\<lbrakk>A \<equiv> lfp(D,h); bnd_mono(D,h); a:A;
+ \<And>x. x \<in> h(Collect(A,P)) \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> 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 \<inter> A) \<subseteq> D or !!X. X<=D ==> h(X)<=D *)
+ second premise could simply be h(D \<inter> A) \<subseteq> D or \<And>X. X<=D \<Longrightarrow> h(X)<=D *)
lemma lfp_Int_lowerbound:
- "[| h(D \<inter> A) \<subseteq> A; bnd_mono(D,h) |] ==> lfp(D,h) \<subseteq> A"
+ "\<lbrakk>h(D \<inter> A) \<subseteq> A; bnd_mono(D,h)\<rbrakk> \<Longrightarrow> lfp(D,h) \<subseteq> A"
apply (rule lfp_lowerbound [THEN subset_trans])
apply (erule bnd_mono_subset [THEN Int_greatest], blast+)
done
@@ -164,7 +164,7 @@
lemma lfp_mono:
assumes hmono: "bnd_mono(D,h)"
and imono: "bnd_mono(E,i)"
- and subhi: "!!X. X<=D ==> h(X) \<subseteq> i(X)"
+ and subhi: "\<And>X. X<=D \<Longrightarrow> h(X) \<subseteq> i(X)"
shows "lfp(D,h) \<subseteq> lfp(E,i)"
apply (rule bnd_monoD1 [THEN lfp_greatest])
apply (rule imono)
@@ -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) \<subseteq> D; !!X. X<=D ==> h(X) \<subseteq> i(X) |] ==> lfp(D,h) \<subseteq> lfp(D,i)"
+ "\<lbrakk>i(D) \<subseteq> D; \<And>X. X<=D \<Longrightarrow> h(X) \<subseteq> i(X)\<rbrakk> \<Longrightarrow> 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 \<subseteq> D' ==> h(X) = h'(X)|] ==> lfp(D,h) = lfp(D',h')"
+ "\<lbrakk>D=D'; \<And>X. X \<subseteq> D' \<Longrightarrow> h(X) = h'(X)\<rbrakk> \<Longrightarrow> 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,7 +192,7 @@
subsection\<open>Proof of Knaster-Tarski Theorem using \<^term>\<open>gfp\<close>\<close>
(*gfp contains each post-fixedpoint that is contained in D*)
-lemma gfp_upperbound: "[| A \<subseteq> h(A); A<=D |] ==> A \<subseteq> gfp(D,h)"
+lemma gfp_upperbound: "\<lbrakk>A \<subseteq> h(A); A<=D\<rbrakk> \<Longrightarrow> A \<subseteq> gfp(D,h)"
apply (unfold gfp_def)
apply (rule PowI [THEN CollectI, THEN Union_upper])
apply (assumption+)
@@ -202,41 +202,41 @@
by (unfold gfp_def, blast)
(*Used in datatype package*)
-lemma def_gfp_subset: "A==gfp(D,h) ==> A \<subseteq> D"
+lemma def_gfp_subset: "A\<equiv>gfp(D,h) \<Longrightarrow> A \<subseteq> D"
apply simp
apply (rule gfp_subset)
done
lemma gfp_least:
- "[| bnd_mono(D,h); !!X. [| X \<subseteq> h(X); X<=D |] ==> X<=A |] ==>
+ "\<lbrakk>bnd_mono(D,h); \<And>X. \<lbrakk>X \<subseteq> h(X); X<=D\<rbrakk> \<Longrightarrow> X<=A\<rbrakk> \<Longrightarrow>
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 \<subseteq> h(gfp(D,h))"
+ "\<lbrakk>bnd_mono(D,h); A<=h(A); A<=D\<rbrakk> \<Longrightarrow> 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) \<subseteq> h(gfp(D,h))"
+lemma gfp_lemma2: "bnd_mono(D,h) \<Longrightarrow> 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)) \<subseteq> gfp(D,h)"
+ "bnd_mono(D,h) \<Longrightarrow> h(gfp(D,h)) \<subseteq> gfp(D,h)"
apply (rule gfp_upperbound)
apply (rule bnd_monoD2, assumption)
apply (rule gfp_lemma2, assumption)
apply (erule bnd_mono_subset, rule gfp_subset)+
done
-lemma gfp_unfold: "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))"
+lemma gfp_unfold: "bnd_mono(D,h) \<Longrightarrow> gfp(D,h) = h(gfp(D,h))"
apply (rule equalityI)
apply (erule gfp_lemma2)
apply (erule gfp_lemma3)
@@ -244,7 +244,7 @@
(*Definition form, to control unfolding*)
lemma def_gfp_unfold:
- "[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"
+ "\<lbrakk>A\<equiv>gfp(D,h); bnd_mono(D,h)\<rbrakk> \<Longrightarrow> A = h(A)"
apply simp
apply (erule gfp_unfold)
done
@@ -253,11 +253,11 @@
subsection\<open>Coinduction Rules for Greatest Fixed Points\<close>
(*weak version*)
-lemma weak_coinduct: "[| a: X; X \<subseteq> h(X); X \<subseteq> D |] ==> a \<in> gfp(D,h)"
+lemma weak_coinduct: "\<lbrakk>a: X; X \<subseteq> h(X); X \<subseteq> D\<rbrakk> \<Longrightarrow> a \<in> gfp(D,h)"
by (blast intro: gfp_upperbound [THEN subsetD])
lemma coinduct_lemma:
- "[| X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D; bnd_mono(D,h) |] ==>
+ "\<lbrakk>X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D; bnd_mono(D,h)\<rbrakk> \<Longrightarrow>
X \<union> gfp(D,h) \<subseteq> h(X \<union> gfp(D,h))"
apply (erule Un_least)
apply (rule gfp_lemma2 [THEN subset_trans], assumption)
@@ -268,8 +268,8 @@
(*strong version*)
lemma coinduct:
- "[| bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D |]
- ==> a \<in> gfp(D,h)"
+ "\<lbrakk>bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> gfp(D,h)); X \<subseteq> D\<rbrakk>
+ \<Longrightarrow> 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,7 +277,7 @@
(*Definition form, to control unfolding*)
lemma def_coinduct:
- "[| A == gfp(D,h); bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> A); X \<subseteq> D |] ==>
+ "\<lbrakk>A \<equiv> gfp(D,h); bnd_mono(D,h); a: X; X \<subseteq> h(X \<union> A); X \<subseteq> D\<rbrakk> \<Longrightarrow>
a \<in> A"
apply simp
apply (rule coinduct, assumption+)
@@ -285,16 +285,16 @@
(*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 \<subseteq> D; !!z. z: X ==> P(X \<union> A, z) |] ==>
+ "\<lbrakk>A \<equiv> gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w)));
+ a: X; X \<subseteq> D; \<And>z. z: X \<Longrightarrow> P(X \<union> A, z)\<rbrakk> \<Longrightarrow>
a \<in> A"
apply (rule def_coinduct, assumption+, blast+)
done
(*Monotonicity of gfp!*)
lemma gfp_mono:
- "[| bnd_mono(D,h); D \<subseteq> E;
- !!X. X<=D ==> h(X) \<subseteq> i(X) |] ==> gfp(D,h) \<subseteq> gfp(E,i)"
+ "\<lbrakk>bnd_mono(D,h); D \<subseteq> E;
+ \<And>X. X<=D \<Longrightarrow> h(X) \<subseteq> i(X)\<rbrakk> \<Longrightarrow> 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/IMP/Com.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/IMP/Com.thy Tue Sep 27 16:51:35 2022 +0100
@@ -24,16 +24,16 @@
abbreviation
evala_syntax :: "[i, i] => o" (infixl \<open>-a->\<close> 50)
- where "p -a-> n == <p,n> \<in> evala"
+ where "p -a-> n \<equiv> <p,n> \<in> evala"
inductive
domains "evala" \<subseteq> "(aexp \<times> (loc -> nat)) \<times> nat"
intros
- N: "[| n \<in> nat; sigma \<in> loc->nat |] ==> <N(n),sigma> -a-> n"
- X: "[| x \<in> loc; sigma \<in> loc->nat |] ==> <X(x),sigma> -a-> sigma`x"
- Op1: "[| <e,sigma> -a-> n; f \<in> nat -> nat |] ==> <Op1(f,e),sigma> -a-> f`n"
- Op2: "[| <e0,sigma> -a-> n0; <e1,sigma> -a-> n1; f \<in> (nat\<times>nat) -> nat |]
- ==> <Op2(f,e0,e1),sigma> -a-> f`<n0,n1>"
+ N: "\<lbrakk>n \<in> nat; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <N(n),sigma> -a-> n"
+ X: "\<lbrakk>x \<in> loc; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <X(x),sigma> -a-> sigma`x"
+ Op1: "\<lbrakk><e,sigma> -a-> n; f \<in> nat -> nat\<rbrakk> \<Longrightarrow> <Op1(f,e),sigma> -a-> f`n"
+ Op2: "\<lbrakk><e0,sigma> -a-> n0; <e1,sigma> -a-> n1; f \<in> (nat\<times>nat) -> nat\<rbrakk>
+ \<Longrightarrow> <Op2(f,e0,e1),sigma> -a-> f`<n0,n1>"
type_intros aexp.intros apply_funtype
@@ -54,20 +54,20 @@
abbreviation
evalb_syntax :: "[i,i] => o" (infixl \<open>-b->\<close> 50)
- where "p -b-> b == <p,b> \<in> evalb"
+ where "p -b-> b \<equiv> <p,b> \<in> evalb"
inductive
domains "evalb" \<subseteq> "(bexp \<times> (loc -> nat)) \<times> bool"
intros
- true: "[| sigma \<in> loc -> nat |] ==> <true,sigma> -b-> 1"
- false: "[| sigma \<in> loc -> nat |] ==> <false,sigma> -b-> 0"
- ROp: "[| <a0,sigma> -a-> n0; <a1,sigma> -a-> n1; f \<in> (nat*nat)->bool |]
- ==> <ROp(f,a0,a1),sigma> -b-> f`<n0,n1> "
- noti: "[| <b,sigma> -b-> w |] ==> <noti(b),sigma> -b-> not(w)"
- andi: "[| <b0,sigma> -b-> w0; <b1,sigma> -b-> w1 |]
- ==> <b0 andi b1,sigma> -b-> (w0 and w1)"
- ori: "[| <b0,sigma> -b-> w0; <b1,sigma> -b-> w1 |]
- ==> <b0 ori b1,sigma> -b-> (w0 or w1)"
+ true: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <true,sigma> -b-> 1"
+ false: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <false,sigma> -b-> 0"
+ ROp: "\<lbrakk><a0,sigma> -a-> n0; <a1,sigma> -a-> n1; f \<in> (nat*nat)->bool\<rbrakk>
+ \<Longrightarrow> <ROp(f,a0,a1),sigma> -b-> f`<n0,n1> "
+ noti: "\<lbrakk><b,sigma> -b-> w\<rbrakk> \<Longrightarrow> <noti(b),sigma> -b-> not(w)"
+ andi: "\<lbrakk><b0,sigma> -b-> w0; <b1,sigma> -b-> w1\<rbrakk>
+ \<Longrightarrow> <b0 andi b1,sigma> -b-> (w0 and w1)"
+ ori: "\<lbrakk><b0,sigma> -b-> w0; <b1,sigma> -b-> w1\<rbrakk>
+ \<Longrightarrow> <b0 ori b1,sigma> -b-> (w0 or w1)"
type_intros bexp.intros
apply_funtype and_type or_type bool_1I bool_0I not_type
type_elims evala.dom_subset [THEN subsetD, elim_format]
@@ -88,33 +88,33 @@
abbreviation
evalc_syntax :: "[i, i] => o" (infixl \<open>-c->\<close> 50)
- where "p -c-> s == <p,s> \<in> evalc"
+ where "p -c-> s \<equiv> <p,s> \<in> evalc"
inductive
domains "evalc" \<subseteq> "(com \<times> (loc -> nat)) \<times> (loc -> nat)"
intros
- skip: "[| sigma \<in> loc -> nat |] ==> <\<SKIP>,sigma> -c-> sigma"
+ skip: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <\<SKIP>,sigma> -c-> sigma"
- assign: "[| m \<in> nat; x \<in> loc; <a,sigma> -a-> m |]
- ==> <x \<ASSN> a,sigma> -c-> sigma(x:=m)"
+ assign: "\<lbrakk>m \<in> nat; x \<in> loc; <a,sigma> -a-> m\<rbrakk>
+ \<Longrightarrow> <x \<ASSN> a,sigma> -c-> sigma(x:=m)"
- semi: "[| <c0,sigma> -c-> sigma2; <c1,sigma2> -c-> sigma1 |]
- ==> <c0\<SEQ> c1, sigma> -c-> sigma1"
+ semi: "\<lbrakk><c0,sigma> -c-> sigma2; <c1,sigma2> -c-> sigma1\<rbrakk>
+ \<Longrightarrow> <c0\<SEQ> c1, sigma> -c-> sigma1"
- if1: "[| b \<in> bexp; c1 \<in> com; sigma \<in> loc->nat;
- <b,sigma> -b-> 1; <c0,sigma> -c-> sigma1 |]
- ==> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
+ if1: "\<lbrakk>b \<in> bexp; c1 \<in> com; sigma \<in> loc->nat;
+ <b,sigma> -b-> 1; <c0,sigma> -c-> sigma1\<rbrakk>
+ \<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
- if0: "[| b \<in> bexp; c0 \<in> com; sigma \<in> loc->nat;
- <b,sigma> -b-> 0; <c1,sigma> -c-> sigma1 |]
- ==> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
+ if0: "\<lbrakk>b \<in> bexp; c0 \<in> com; sigma \<in> loc->nat;
+ <b,sigma> -b-> 0; <c1,sigma> -c-> sigma1\<rbrakk>
+ \<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
- while0: "[| c \<in> com; <b, sigma> -b-> 0 |]
- ==> <\<WHILE> b \<DO> c,sigma> -c-> sigma"
+ while0: "\<lbrakk>c \<in> com; <b, sigma> -b-> 0\<rbrakk>
+ \<Longrightarrow> <\<WHILE> b \<DO> c,sigma> -c-> sigma"
- while1: "[| c \<in> com; <b,sigma> -b-> 1; <c,sigma> -c-> sigma2;
- <\<WHILE> b \<DO> c, sigma2> -c-> sigma1 |]
- ==> <\<WHILE> b \<DO> c, sigma> -c-> sigma1"
+ while1: "\<lbrakk>c \<in> com; <b,sigma> -b-> 1; <c,sigma> -c-> sigma2;
+ <\<WHILE> b \<DO> c, sigma2> -c-> sigma1\<rbrakk>
+ \<Longrightarrow> <\<WHILE> b \<DO> c, sigma> -c-> sigma1"
type_intros com.intros update_type
type_elims evala.dom_subset [THEN subsetD, elim_format]
--- a/src/ZF/IMP/Denotation.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/IMP/Denotation.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,7 +15,7 @@
definition
Gamma :: "[i,i,i] => i" (\<open>\<Gamma>\<close>) where
- "\<Gamma>(b,cden) ==
+ "\<Gamma>(b,cden) \<equiv>
(\<lambda>phi. {io \<in> (phi O cden). B(b,fst(io))=1} \<union>
{io \<in> id(loc->nat). B(b,fst(io))=0})"
@@ -45,28 +45,28 @@
subsection \<open>Misc lemmas\<close>
-lemma A_type [TC]: "[|a \<in> aexp; sigma \<in> loc->nat|] ==> A(a,sigma) \<in> nat"
+lemma A_type [TC]: "\<lbrakk>a \<in> aexp; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> A(a,sigma) \<in> nat"
by (erule aexp.induct) simp_all
-lemma B_type [TC]: "[|b \<in> bexp; sigma \<in> loc->nat|] ==> B(b,sigma) \<in> bool"
+lemma B_type [TC]: "\<lbrakk>b \<in> bexp; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> B(b,sigma) \<in> bool"
by (erule bexp.induct, simp_all)
-lemma C_subset: "c \<in> com ==> C(c) \<subseteq> (loc->nat) \<times> (loc->nat)"
+lemma C_subset: "c \<in> com \<Longrightarrow> C(c) \<subseteq> (loc->nat) \<times> (loc->nat)"
apply (erule com.induct)
apply simp_all
apply (blast dest: lfp_subset [THEN subsetD])+
done
lemma C_type_D [dest]:
- "[| <x,y> \<in> C(c); c \<in> com |] ==> x \<in> loc->nat & y \<in> loc->nat"
+ "\<lbrakk><x,y> \<in> C(c); c \<in> com\<rbrakk> \<Longrightarrow> x \<in> loc->nat & y \<in> loc->nat"
by (blast dest: C_subset [THEN subsetD])
-lemma C_type_fst [dest]: "[| x \<in> C(c); c \<in> com |] ==> fst(x) \<in> loc->nat"
+lemma C_type_fst [dest]: "\<lbrakk>x \<in> C(c); c \<in> com\<rbrakk> \<Longrightarrow> fst(x) \<in> loc->nat"
by (auto dest!: C_subset [THEN subsetD])
lemma Gamma_bnd_mono:
"cden \<subseteq> (loc->nat) \<times> (loc->nat)
- ==> bnd_mono ((loc->nat) \<times> (loc->nat), \<Gamma>(b,cden))"
+ \<Longrightarrow> bnd_mono ((loc->nat) \<times> (loc->nat), \<Gamma>(b,cden))"
by (unfold bnd_mono_def Gamma_def) blast
end
--- a/src/ZF/IMP/Equiv.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/IMP/Equiv.thy Tue Sep 27 16:51:35 2022 +0100
@@ -7,8 +7,8 @@
theory Equiv imports Denotation Com begin
lemma aexp_iff [rule_format]:
- "[| a \<in> aexp; sigma: loc -> nat |]
- ==> \<forall>n. <a,sigma> -a-> n \<longleftrightarrow> A(a,sigma) = n"
+ "\<lbrakk>a \<in> aexp; sigma: loc -> nat\<rbrakk>
+ \<Longrightarrow> \<forall>n. <a,sigma> -a-> n \<longleftrightarrow> A(a,sigma) = n"
apply (erule aexp.induct)
apply (force intro!: evala.intros)+
done
@@ -26,8 +26,8 @@
lemma bexp_iff [rule_format]:
- "[| b \<in> bexp; sigma: loc -> nat |]
- ==> \<forall>w. <b,sigma> -b-> w \<longleftrightarrow> B(b,sigma) = w"
+ "\<lbrakk>b \<in> bexp; sigma: loc -> nat\<rbrakk>
+ \<Longrightarrow> \<forall>w. <b,sigma> -b-> w \<longleftrightarrow> B(b,sigma) = w"
apply (erule bexp.induct)
apply (auto intro!: evalb.intros)
done
@@ -35,7 +35,7 @@
declare bexp_iff [THEN iffD1, simp]
bexp_iff [THEN iffD2, intro!]
-lemma com1: "<c,sigma> -c-> sigma' ==> <sigma,sigma'> \<in> C(c)"
+lemma com1: "<c,sigma> -c-> sigma' \<Longrightarrow> <sigma,sigma'> \<in> C(c)"
apply (erule evalc.induct)
apply (simp_all (no_asm_simp))
txt \<open>\<open>assign\<close>\<close>
@@ -54,7 +54,7 @@
declare evalc.intros [intro]
-lemma com2 [rule_format]: "c \<in> com ==> \<forall>x \<in> C(c). <c,fst(x)> -c-> snd(x)"
+lemma com2 [rule_format]: "c \<in> com \<Longrightarrow> \<forall>x \<in> C(c). <c,fst(x)> -c-> snd(x)"
apply (erule com.induct)
txt \<open>\<open>skip\<close>\<close>
apply force
@@ -76,7 +76,7 @@
subsection \<open>Main theorem\<close>
theorem com_equivalence:
- "c \<in> com ==> C(c) = {io \<in> (loc->nat) \<times> (loc->nat). <c,fst(io)> -c-> snd(io)}"
+ "c \<in> com \<Longrightarrow> C(c) = {io \<in> (loc->nat) \<times> (loc->nat). <c,fst(io)> -c-> snd(io)}"
by (force intro: C_subset [THEN subsetD] elim: com2 dest: com1)
end
--- a/src/ZF/Induct/Acc.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Acc.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,7 +17,7 @@
inductive
domains "acc(r)" \<subseteq> "field(r)"
intros
- vimage: "[| r-``{a}: Pow(acc(r)); a \<in> field(r) |] ==> a \<in> acc(r)"
+ vimage: "\<lbrakk>r-``{a}: Pow(acc(r)); a \<in> field(r)\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
monos Pow_mono
text \<open>
@@ -28,16 +28,16 @@
The intended introduction rule:
\<close>
-lemma accI: "[| !!b. <b,a>:r ==> b \<in> acc(r); a \<in> field(r) |] ==> a \<in> acc(r)"
+lemma accI: "\<lbrakk>\<And>b. <b,a>:r \<Longrightarrow> b \<in> acc(r); a \<in> field(r)\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
by (blast intro: acc.intros)
-lemma acc_downward: "[| b \<in> acc(r); <a,b>: r |] ==> a \<in> acc(r)"
+lemma acc_downward: "\<lbrakk>b \<in> acc(r); <a,b>: r\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
by (erule acc.cases) blast
lemma acc_induct [consumes 1, case_names vimage, induct set: acc]:
- "[| a \<in> acc(r);
- !!x. [| x \<in> acc(r); \<forall>y. <y,x>:r \<longrightarrow> P(y) |] ==> P(x)
- |] ==> P(a)"
+ "\<lbrakk>a \<in> acc(r);
+ \<And>x. \<lbrakk>x \<in> acc(r); \<forall>y. <y,x>:r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> P(a)"
by (erule acc.induct) (blast intro: acc.intros)
lemma wf_on_acc: "wf[acc(r)](r)"
@@ -49,7 +49,7 @@
lemma acc_wfI: "field(r) \<subseteq> acc(r) \<Longrightarrow> wf(r)"
by (erule wf_on_acc [THEN wf_on_subset_A, THEN wf_on_field_imp_wf])
-lemma acc_wfD: "wf(r) ==> field(r) \<subseteq> acc(r)"
+lemma acc_wfD: "wf(r) \<Longrightarrow> field(r) \<subseteq> acc(r)"
apply (rule subsetI)
apply (erule wf_induct2, assumption)
apply (blast intro: accI)+
--- a/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,7 +17,7 @@
declare bt.intros [simp]
-lemma Br_neq_left: "l \<in> bt(A) ==> Br(x, l, r) \<noteq> l"
+lemma Br_neq_left: "l \<in> bt(A) \<Longrightarrow> Br(x, l, r) \<noteq> l"
by (induct arbitrary: x r set: bt) auto
lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \<longleftrightarrow> a = a' & l = l' & r = r'"
@@ -32,7 +32,7 @@
definitions.
\<close>
-lemma bt_mono: "A \<subseteq> B ==> bt(A) \<subseteq> bt(B)"
+lemma bt_mono: "A \<subseteq> B \<Longrightarrow> bt(A) \<subseteq> bt(B)"
apply (unfold bt.defs)
apply (rule lfp_mono)
apply (rule bt.bnd_mono)+
@@ -46,18 +46,18 @@
apply (fast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ)
done
-lemma bt_subset_univ: "A \<subseteq> univ(B) ==> bt(A) \<subseteq> univ(B)"
+lemma bt_subset_univ: "A \<subseteq> univ(B) \<Longrightarrow> bt(A) \<subseteq> univ(B)"
apply (rule subset_trans)
apply (erule bt_mono)
apply (rule bt_univ)
done
lemma bt_rec_type:
- "[| t \<in> bt(A);
+ "\<lbrakk>t \<in> bt(A);
c \<in> C(Lf);
- !!x y z r s. [| x \<in> A; y \<in> bt(A); z \<in> bt(A); r \<in> C(y); s \<in> C(z) |] ==>
+ \<And>x y z r s. \<lbrakk>x \<in> A; y \<in> bt(A); z \<in> bt(A); r \<in> C(y); s \<in> C(z)\<rbrakk> \<Longrightarrow>
h(x, y, z, r, s) \<in> C(Br(x, y, z))
- |] ==> bt_rec(c, h, t) \<in> C(t)"
+\<rbrakk> \<Longrightarrow> bt_rec(c, h, t) \<in> C(t)"
\<comment> \<open>Type checking for recursor -- example only; not really needed.\<close>
apply (induct_tac t)
apply simp_all
@@ -71,7 +71,7 @@
"n_nodes(Lf) = 0"
"n_nodes(Br(a, l, r)) = succ(n_nodes(l) #+ n_nodes(r))"
-lemma n_nodes_type [simp]: "t \<in> bt(A) ==> n_nodes(t) \<in> nat"
+lemma n_nodes_type [simp]: "t \<in> bt(A) \<Longrightarrow> n_nodes(t) \<in> nat"
by (induct set: bt) auto
consts n_nodes_aux :: "i => i"
@@ -81,7 +81,7 @@
(\<lambda>k \<in> nat. n_nodes_aux(r) ` (n_nodes_aux(l) ` succ(k)))"
lemma n_nodes_aux_eq:
- "t \<in> bt(A) ==> k \<in> nat ==> n_nodes_aux(t)`k = n_nodes(t) #+ k"
+ "t \<in> bt(A) \<Longrightarrow> k \<in> nat \<Longrightarrow> n_nodes_aux(t)`k = n_nodes(t) #+ k"
apply (induct arbitrary: k set: bt)
apply simp
apply (atomize, simp)
@@ -89,9 +89,9 @@
definition
n_nodes_tail :: "i => i" where
- "n_nodes_tail(t) == n_nodes_aux(t) ` 0"
+ "n_nodes_tail(t) \<equiv> n_nodes_aux(t) ` 0"
-lemma "t \<in> bt(A) ==> n_nodes_tail(t) = n_nodes(t)"
+lemma "t \<in> bt(A) \<Longrightarrow> n_nodes_tail(t) = n_nodes(t)"
by (simp add: n_nodes_tail_def n_nodes_aux_eq)
@@ -103,7 +103,7 @@
"n_leaves(Lf) = 1"
"n_leaves(Br(a, l, r)) = n_leaves(l) #+ n_leaves(r)"
-lemma n_leaves_type [simp]: "t \<in> bt(A) ==> n_leaves(t) \<in> nat"
+lemma n_leaves_type [simp]: "t \<in> bt(A) \<Longrightarrow> n_leaves(t) \<in> nat"
by (induct set: bt) auto
@@ -115,24 +115,24 @@
"bt_reflect(Lf) = Lf"
"bt_reflect(Br(a, l, r)) = Br(a, bt_reflect(r), bt_reflect(l))"
-lemma bt_reflect_type [simp]: "t \<in> bt(A) ==> bt_reflect(t) \<in> bt(A)"
+lemma bt_reflect_type [simp]: "t \<in> bt(A) \<Longrightarrow> bt_reflect(t) \<in> bt(A)"
by (induct set: bt) auto
text \<open>
\medskip Theorems about \<^term>\<open>n_leaves\<close>.
\<close>
-lemma n_leaves_reflect: "t \<in> bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)"
+lemma n_leaves_reflect: "t \<in> bt(A) \<Longrightarrow> n_leaves(bt_reflect(t)) = n_leaves(t)"
by (induct set: bt) (simp_all add: add_commute)
-lemma n_leaves_nodes: "t \<in> bt(A) ==> n_leaves(t) = succ(n_nodes(t))"
+lemma n_leaves_nodes: "t \<in> bt(A) \<Longrightarrow> n_leaves(t) = succ(n_nodes(t))"
by (induct set: bt) simp_all
text \<open>
Theorems about \<^term>\<open>bt_reflect\<close>.
\<close>
-lemma bt_reflect_bt_reflect_ident: "t \<in> bt(A) ==> bt_reflect(bt_reflect(t)) = t"
+lemma bt_reflect_bt_reflect_ident: "t \<in> bt(A) \<Longrightarrow> bt_reflect(bt_reflect(t)) = t"
by (induct set: bt) simp_all
end
--- a/src/ZF/Induct/Brouwer.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Brouwer.thy Tue Sep 27 16:51:35 2022 +0100
@@ -25,8 +25,8 @@
assumes b: "b \<in> brouwer"
and cases:
"P(Zero)"
- "!!b. [| b \<in> brouwer; P(b) |] ==> P(Suc(b))"
- "!!h. [| h \<in> nat -> brouwer; \<forall>i \<in> nat. P(h`i) |] ==> P(Lim(h))"
+ "\<And>b. \<lbrakk>b \<in> brouwer; P(b)\<rbrakk> \<Longrightarrow> P(Suc(b))"
+ "\<And>h. \<lbrakk>h \<in> nat -> brouwer; \<forall>i \<in> nat. P(h`i)\<rbrakk> \<Longrightarrow> P(Lim(h))"
shows "P(b)"
\<comment> \<open>A nicer induction rule than the standard one.\<close>
using b
@@ -57,7 +57,7 @@
lemma Well_induct2 [consumes 1, case_names step]:
assumes w: "w \<in> Well(A, B)"
- and step: "!!a f. [| a \<in> A; f \<in> B(a) -> Well(A,B); \<forall>y \<in> B(a). P(f`y) |] ==> P(Sup(a,f))"
+ and step: "\<And>a f. \<lbrakk>a \<in> A; f \<in> B(a) -> Well(A,B); \<forall>y \<in> B(a). P(f`y)\<rbrakk> \<Longrightarrow> P(Sup(a,f))"
shows "P(w)"
\<comment> \<open>A nicer induction rule than the standard one.\<close>
using w
--- a/src/ZF/Induct/Comb.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Comb.thy Tue Sep 27 16:51:35 2022 +0100
@@ -41,10 +41,10 @@
inductive
domains "contract" \<subseteq> "comb \<times> comb"
intros
- K: "[| p \<in> comb; q \<in> comb |] ==> K\<bullet>p\<bullet>q \<rightarrow>\<^sup>1 p"
- S: "[| p \<in> comb; q \<in> comb; r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r \<rightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)"
- Ap1: "[| p\<rightarrow>\<^sup>1q; r \<in> comb |] ==> p\<bullet>r \<rightarrow>\<^sup>1 q\<bullet>r"
- Ap2: "[| p\<rightarrow>\<^sup>1q; r \<in> comb |] ==> r\<bullet>p \<rightarrow>\<^sup>1 r\<bullet>q"
+ K: "\<lbrakk>p \<in> comb; q \<in> comb\<rbrakk> \<Longrightarrow> K\<bullet>p\<bullet>q \<rightarrow>\<^sup>1 p"
+ S: "\<lbrakk>p \<in> comb; q \<in> comb; r \<in> comb\<rbrakk> \<Longrightarrow> S\<bullet>p\<bullet>q\<bullet>r \<rightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)"
+ Ap1: "\<lbrakk>p\<rightarrow>\<^sup>1q; r \<in> comb\<rbrakk> \<Longrightarrow> p\<bullet>r \<rightarrow>\<^sup>1 q\<bullet>r"
+ Ap2: "\<lbrakk>p\<rightarrow>\<^sup>1q; r \<in> comb\<rbrakk> \<Longrightarrow> r\<bullet>p \<rightarrow>\<^sup>1 r\<bullet>q"
type_intros comb.intros
text \<open>
@@ -55,18 +55,18 @@
consts parcontract :: i
abbreviation parcontract_syntax :: "[i,i] => o" (infixl \<open>\<Rrightarrow>\<^sup>1\<close> 50)
- where "p \<Rrightarrow>\<^sup>1 q == <p,q> \<in> parcontract"
+ where "p \<Rrightarrow>\<^sup>1 q \<equiv> <p,q> \<in> parcontract"
abbreviation parcontract_multi :: "[i,i] => o" (infixl \<open>\<Rrightarrow>\<close> 50)
- where "p \<Rrightarrow> q == <p,q> \<in> parcontract^+"
+ where "p \<Rrightarrow> q \<equiv> <p,q> \<in> parcontract^+"
inductive
domains "parcontract" \<subseteq> "comb \<times> comb"
intros
- refl: "[| p \<in> comb |] ==> p \<Rrightarrow>\<^sup>1 p"
- K: "[| p \<in> comb; q \<in> comb |] ==> K\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 p"
- S: "[| p \<in> comb; q \<in> comb; r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r \<Rrightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)"
- Ap: "[| p\<Rrightarrow>\<^sup>1q; r\<Rrightarrow>\<^sup>1s |] ==> p\<bullet>r \<Rrightarrow>\<^sup>1 q\<bullet>s"
+ refl: "\<lbrakk>p \<in> comb\<rbrakk> \<Longrightarrow> p \<Rrightarrow>\<^sup>1 p"
+ K: "\<lbrakk>p \<in> comb; q \<in> comb\<rbrakk> \<Longrightarrow> K\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 p"
+ S: "\<lbrakk>p \<in> comb; q \<in> comb; r \<in> comb\<rbrakk> \<Longrightarrow> S\<bullet>p\<bullet>q\<bullet>r \<Rrightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)"
+ Ap: "\<lbrakk>p\<Rrightarrow>\<^sup>1q; r\<Rrightarrow>\<^sup>1s\<rbrakk> \<Longrightarrow> p\<bullet>r \<Rrightarrow>\<^sup>1 q\<bullet>s"
type_intros comb.intros
text \<open>
@@ -84,7 +84,7 @@
subsection \<open>Transitive closure preserves the Church-Rosser property\<close>
lemma diamond_strip_lemmaD [rule_format]:
- "[| diamond(r); <x,y>:r^+ |] ==>
+ "\<lbrakk>diamond(r); <x,y>:r^+\<rbrakk> \<Longrightarrow>
\<forall>y'. <x,y'>:r \<longrightarrow> (\<exists>z. <y',z>: r^+ & <y,z>: r)"
apply (unfold diamond_def)
apply (erule trancl_induct)
@@ -94,7 +94,7 @@
apply (blast intro: r_into_trancl trans_trancl [THEN transD])
done
-lemma diamond_trancl: "diamond(r) ==> diamond(r^+)"
+lemma diamond_trancl: "diamond(r) \<Longrightarrow> diamond(r^+)"
apply (simp (no_asm_simp) add: diamond_def)
apply (rule impI [THEN allI, THEN allI])
apply (erule trancl_induct)
@@ -134,7 +134,7 @@
contract.Ap1 [THEN rtrancl_into_rtrancl2]
contract.Ap2 [THEN rtrancl_into_rtrancl2]
-lemma "p \<in> comb ==> I\<bullet>p \<rightarrow> p"
+lemma "p \<in> comb \<Longrightarrow> I\<bullet>p \<rightarrow> p"
\<comment> \<open>Example only: not used\<close>
unfolding I_def by (blast intro: reduction_rls)
@@ -150,13 +150,13 @@
and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r"
and Ap_contractE [elim!]: "p\<bullet>q \<rightarrow>\<^sup>1 r"
-lemma I_contract_E: "I \<rightarrow>\<^sup>1 r ==> P"
+lemma I_contract_E: "I \<rightarrow>\<^sup>1 r \<Longrightarrow> P"
by (auto simp add: I_def)
-lemma K1_contractD: "K\<bullet>p \<rightarrow>\<^sup>1 r ==> (\<exists>q. r = K\<bullet>q & p \<rightarrow>\<^sup>1 q)"
+lemma K1_contractD: "K\<bullet>p \<rightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>q. r = K\<bullet>q & p \<rightarrow>\<^sup>1 q)"
by auto
-lemma Ap_reduce1: "[| p \<rightarrow> q; r \<in> comb |] ==> p\<bullet>r \<rightarrow> q\<bullet>r"
+lemma Ap_reduce1: "\<lbrakk>p \<rightarrow> q; r \<in> comb\<rbrakk> \<Longrightarrow> p\<bullet>r \<rightarrow> q\<bullet>r"
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
apply (erule rtrancl_induct)
@@ -165,7 +165,7 @@
apply (blast intro: contract_combD2 reduction_rls)
done
-lemma Ap_reduce2: "[| p \<rightarrow> q; r \<in> comb |] ==> r\<bullet>p \<rightarrow> r\<bullet>q"
+lemma Ap_reduce2: "\<lbrakk>p \<rightarrow> q; r \<in> comb\<rbrakk> \<Longrightarrow> r\<bullet>p \<rightarrow> r\<bullet>q"
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
apply (erule rtrancl_induct)
@@ -215,15 +215,15 @@
subsection \<open>Basic properties of parallel contraction\<close>
lemma K1_parcontractD [dest!]:
- "K\<bullet>p \<Rrightarrow>\<^sup>1 r ==> (\<exists>p'. r = K\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
+ "K\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = K\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
by auto
lemma S1_parcontractD [dest!]:
- "S\<bullet>p \<Rrightarrow>\<^sup>1 r ==> (\<exists>p'. r = S\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
+ "S\<bullet>p \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p'. r = S\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')"
by auto
lemma S2_parcontractD [dest!]:
- "S\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 r ==> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' & p \<Rrightarrow>\<^sup>1 p' & q \<Rrightarrow>\<^sup>1 q')"
+ "S\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 r \<Longrightarrow> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' & p \<Rrightarrow>\<^sup>1 p' & q \<Rrightarrow>\<^sup>1 q')"
by auto
lemma diamond_parcontract: "diamond(parcontract)"
@@ -238,10 +238,10 @@
\medskip Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>.
\<close>
-lemma contract_imp_parcontract: "p\<rightarrow>\<^sup>1q ==> p\<Rrightarrow>\<^sup>1q"
+lemma contract_imp_parcontract: "p\<rightarrow>\<^sup>1q \<Longrightarrow> p\<Rrightarrow>\<^sup>1q"
by (induct set: contract) auto
-lemma reduce_imp_parreduce: "p\<rightarrow>q ==> p\<Rrightarrow>q"
+lemma reduce_imp_parreduce: "p\<rightarrow>q \<Longrightarrow> p\<Rrightarrow>q"
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
apply (erule rtrancl_induct)
@@ -250,7 +250,7 @@
trans_trancl [THEN transD])
done
-lemma parcontract_imp_reduce: "p\<Rrightarrow>\<^sup>1q ==> p\<rightarrow>q"
+lemma parcontract_imp_reduce: "p\<Rrightarrow>\<^sup>1q \<Longrightarrow> p\<rightarrow>q"
apply (induct set: parcontract)
apply (blast intro: reduction_rls)
apply (blast intro: reduction_rls)
@@ -259,7 +259,7 @@
Ap_reduce1 Ap_reduce2 parcontract_combD1 parcontract_combD2)
done
-lemma parreduce_imp_reduce: "p\<Rrightarrow>q ==> p\<rightarrow>q"
+lemma parreduce_imp_reduce: "p\<Rrightarrow>q \<Longrightarrow> p\<rightarrow>q"
apply (frule trancl_type [THEN subsetD, THEN SigmaD1])
apply (drule field_parcontract_eq [THEN equalityD1, THEN subsetD])
apply (erule trancl_induct, erule parcontract_imp_reduce)
--- a/src/ZF/Induct/Datatypes.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Datatypes.thy Tue Sep 27 16:51:35 2022 +0100
@@ -31,7 +31,7 @@
type definitions.
\<close>
-lemma data_mono: "[| A \<subseteq> C; B \<subseteq> D |] ==> data(A, B) \<subseteq> data(C, D)"
+lemma data_mono: "\<lbrakk>A \<subseteq> C; B \<subseteq> D\<rbrakk> \<Longrightarrow> data(A, B) \<subseteq> data(C, D)"
apply (unfold data.defs)
apply (rule lfp_mono)
apply (rule data.bnd_mono)+
@@ -46,7 +46,7 @@
done
lemma data_subset_univ:
- "[| A \<subseteq> univ(C); B \<subseteq> univ(C) |] ==> data(A, B) \<subseteq> univ(C)"
+ "\<lbrakk>A \<subseteq> univ(C); B \<subseteq> univ(C)\<rbrakk> \<Longrightarrow> data(A, B) \<subseteq> univ(C)"
by (rule subset_trans [OF data_mono data_univ])
--- a/src/ZF/Induct/FoldSet.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/FoldSet.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,18 +13,18 @@
inductive
domains "fold_set(A, B, f,e)" \<subseteq> "Fin(A)*B"
intros
- emptyI: "e\<in>B ==> <0, e>\<in>fold_set(A, B, f,e)"
- consI: "[| x\<in>A; x \<notin>C; <C,y> \<in> fold_set(A, B,f,e); f(x,y):B |]
- ==> <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
+ emptyI: "e\<in>B \<Longrightarrow> <0, e>\<in>fold_set(A, B, f,e)"
+ consI: "\<lbrakk>x\<in>A; x \<notin>C; <C,y> \<in> fold_set(A, B,f,e); f(x,y):B\<rbrakk>
+ \<Longrightarrow> <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
type_intros Fin.intros
definition
fold :: "[i, [i,i]=>i, i, i] => i" (\<open>fold[_]'(_,_,_')\<close>) where
- "fold[B](f,e, A) == THE x. <A, x>\<in>fold_set(A, B, f,e)"
+ "fold[B](f,e, A) \<equiv> THE x. <A, x>\<in>fold_set(A, B, f,e)"
definition
setsum :: "[i=>i, i] => i" where
- "setsum(g, C) == if Finite(C) then
+ "setsum(g, C) \<equiv> if Finite(C) then
fold[int](%x y. g(x) $+ y, #0, C) else #0"
(** foldSet **)
@@ -34,38 +34,38 @@
(* add-hoc lemmas *)
-lemma cons_lemma1: "[| x\<notin>C; x\<notin>B |] ==> cons(x,B)=cons(x,C) \<longleftrightarrow> B = C"
+lemma cons_lemma1: "\<lbrakk>x\<notin>C; x\<notin>B\<rbrakk> \<Longrightarrow> cons(x,B)=cons(x,C) \<longleftrightarrow> B = C"
by (auto elim: equalityE)
-lemma cons_lemma2: "[| cons(x, B)=cons(y, C); x\<noteq>y; x\<notin>B; y\<notin>C |]
- ==> B - {y} = C-{x} & x\<in>C & y\<in>B"
+lemma cons_lemma2: "\<lbrakk>cons(x, B)=cons(y, C); x\<noteq>y; x\<notin>B; y\<notin>C\<rbrakk>
+ \<Longrightarrow> B - {y} = C-{x} & x\<in>C & y\<in>B"
apply (auto elim: equalityE)
done
(* fold_set monotonicity *)
lemma fold_set_mono_lemma:
"<C, x> \<in> fold_set(A, B, f, e)
- ==> \<forall>D. A<=D \<longrightarrow> <C, x> \<in> fold_set(D, B, f, e)"
+ \<Longrightarrow> \<forall>D. A<=D \<longrightarrow> <C, x> \<in> fold_set(D, B, f, e)"
apply (erule fold_set.induct)
apply (auto intro: fold_set.intros)
done
-lemma fold_set_mono: " C<=A ==> fold_set(C, B, f, e) \<subseteq> fold_set(A, B, f, e)"
+lemma fold_set_mono: " C<=A \<Longrightarrow> fold_set(C, B, f, e) \<subseteq> fold_set(A, B, f, e)"
apply clarify
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
apply (auto dest: fold_set_mono_lemma)
done
lemma fold_set_lemma:
- "<C, x>\<in>fold_set(A, B, f, e) ==> <C, x>\<in>fold_set(C, B, f, e) & C<=A"
+ "<C, x>\<in>fold_set(A, B, f, e) \<Longrightarrow> <C, x>\<in>fold_set(C, B, f, e) & C<=A"
apply (erule fold_set.induct)
apply (auto intro!: fold_set.intros intro: fold_set_mono [THEN subsetD])
done
(* Proving that fold_set is deterministic *)
lemma Diff1_fold_set:
- "[| <C-{x},y> \<in> fold_set(A, B, f,e); x\<in>C; x\<in>A; f(x, y):B |]
- ==> <C, f(x, y)> \<in> fold_set(A, B, f, e)"
+ "\<lbrakk><C-{x},y> \<in> fold_set(A, B, f,e); x\<in>C; x\<in>A; f(x, y):B\<rbrakk>
+ \<Longrightarrow> <C, f(x, y)> \<in> fold_set(A, B, f, e)"
apply (frule fold_set.dom_subset [THEN subsetD])
apply (erule cons_Diff [THEN subst], rule fold_set.intros, auto)
done
@@ -73,25 +73,25 @@
locale fold_typing =
fixes A and B and e and f
- assumes ftype [intro,simp]: "[|x \<in> A; y \<in> B|] ==> f(x,y) \<in> B"
+ assumes ftype [intro,simp]: "\<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> f(x,y) \<in> B"
and etype [intro,simp]: "e \<in> B"
- and fcomm: "[|x \<in> A; y \<in> A; z \<in> B|] ==> f(x, f(y, z))=f(y, f(x, z))"
+ and fcomm: "\<lbrakk>x \<in> A; y \<in> A; z \<in> B\<rbrakk> \<Longrightarrow> f(x, f(y, z))=f(y, f(x, z))"
lemma (in fold_typing) Fin_imp_fold_set:
- "C\<in>Fin(A) ==> (\<exists>x. <C, x> \<in> fold_set(A, B, f,e))"
+ "C\<in>Fin(A) \<Longrightarrow> (\<exists>x. <C, x> \<in> fold_set(A, B, f,e))"
apply (erule Fin_induct)
apply (auto dest: fold_set.dom_subset [THEN subsetD]
intro: fold_set.intros etype ftype)
done
lemma Diff_sing_imp:
- "[|C - {b} = D - {a}; a \<noteq> b; b \<in> C|] ==> C = cons(b,D) - {a}"
+ "\<lbrakk>C - {b} = D - {a}; a \<noteq> b; b \<in> C\<rbrakk> \<Longrightarrow> C = cons(b,D) - {a}"
by (blast elim: equalityE)
lemma (in fold_typing) fold_set_determ_lemma [rule_format]:
"n\<in>nat
- ==> \<forall>C. |C|<n \<longrightarrow>
+ \<Longrightarrow> \<forall>C. |C|<n \<longrightarrow>
(\<forall>x. <C, x> \<in> fold_set(A, B, f,e)\<longrightarrow>
(\<forall>y. <C, y> \<in> fold_set(A, B, f,e) \<longrightarrow> y=x))"
apply (erule nat_induct)
@@ -131,8 +131,8 @@
done
lemma (in fold_typing) fold_set_determ:
- "[| <C, x>\<in>fold_set(A, B, f, e);
- <C, y>\<in>fold_set(A, B, f, e)|] ==> y=x"
+ "\<lbrakk><C, x>\<in>fold_set(A, B, f, e);
+ <C, y>\<in>fold_set(A, B, f, e)\<rbrakk> \<Longrightarrow> y=x"
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
apply (drule Fin_into_Finite)
apply (unfold Finite_def, clarify)
@@ -143,7 +143,7 @@
(** The fold function **)
lemma (in fold_typing) fold_equality:
- "<C,y> \<in> fold_set(A,B,f,e) ==> fold[B](f,e,C) = y"
+ "<C,y> \<in> fold_set(A,B,f,e) \<Longrightarrow> fold[B](f,e,C) = y"
apply (unfold fold_def)
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
apply (rule the_equality)
@@ -154,15 +154,15 @@
apply (auto dest: fold_set_lemma intro: ftype etype fcomm)
done
-lemma fold_0 [simp]: "e \<in> B ==> fold[B](f,e,0) = e"
+lemma fold_0 [simp]: "e \<in> B \<Longrightarrow> fold[B](f,e,0) = e"
apply (unfold fold_def)
apply (blast elim!: empty_fold_setE intro: fold_set.intros)
done
text\<open>This result is the right-to-left direction of the subsequent result\<close>
lemma (in fold_typing) fold_set_imp_cons:
- "[| <C, y> \<in> fold_set(C, B, f, e); C \<in> Fin(A); c \<in> A; c\<notin>C |]
- ==> <cons(c, C), f(c,y)> \<in> fold_set(cons(c, C), B, f, e)"
+ "\<lbrakk><C, y> \<in> fold_set(C, B, f, e); C \<in> Fin(A); c \<in> A; c\<notin>C\<rbrakk>
+ \<Longrightarrow> <cons(c, C), f(c,y)> \<in> fold_set(cons(c, C), B, f, e)"
apply (frule FinD [THEN fold_set_mono, THEN subsetD])
apply assumption
apply (frule fold_set.dom_subset [of A, THEN subsetD])
@@ -170,8 +170,8 @@
done
lemma (in fold_typing) fold_cons_lemma [rule_format]:
-"[| C \<in> Fin(A); c \<in> A; c\<notin>C |]
- ==> <cons(c, C), v> \<in> fold_set(cons(c, C), B, f, e) \<longleftrightarrow>
+"\<lbrakk>C \<in> Fin(A); c \<in> A; c\<notin>C\<rbrakk>
+ \<Longrightarrow> <cons(c, C), v> \<in> fold_set(cons(c, C), B, f, e) \<longleftrightarrow>
(\<exists>y. <C, y> \<in> fold_set(C, B, f, e) & v = f(c, y))"
apply auto
prefer 2 apply (blast intro: fold_set_imp_cons)
@@ -191,8 +191,8 @@
done
lemma (in fold_typing) fold_cons:
- "[| C\<in>Fin(A); c\<in>A; c\<notin>C|]
- ==> fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))"
+ "\<lbrakk>C\<in>Fin(A); c\<in>A; c\<notin>C\<rbrakk>
+ \<Longrightarrow> fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))"
apply (unfold fold_def)
apply (simp add: fold_cons_lemma)
apply (rule the_equality, auto)
@@ -203,14 +203,14 @@
done
lemma (in fold_typing) fold_type [simp,TC]:
- "C\<in>Fin(A) ==> fold[B](f,e,C):B"
+ "C\<in>Fin(A) \<Longrightarrow> fold[B](f,e,C):B"
apply (erule Fin_induct)
apply (simp_all add: fold_cons ftype etype)
done
lemma (in fold_typing) fold_commute [rule_format]:
- "[| C\<in>Fin(A); c\<in>A |]
- ==> (\<forall>y\<in>B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))"
+ "\<lbrakk>C\<in>Fin(A); c\<in>A\<rbrakk>
+ \<Longrightarrow> (\<forall>y\<in>B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))"
apply (erule Fin_induct)
apply (simp_all add: fold_typing.fold_cons [of A B _ f]
fold_typing.fold_type [of A B _ f]
@@ -218,8 +218,8 @@
done
lemma (in fold_typing) fold_nest_Un_Int:
- "[| C\<in>Fin(A); D\<in>Fin(A) |]
- ==> fold[B](f, fold[B](f, e, D), C) =
+ "\<lbrakk>C\<in>Fin(A); D\<in>Fin(A)\<rbrakk>
+ \<Longrightarrow> fold[B](f, fold[B](f, e, D), C) =
fold[B](f, fold[B](f, e, (C \<inter> D)), C \<union> D)"
apply (erule Fin_induct, auto)
apply (simp add: Un_cons Int_cons_left fold_type fold_commute
@@ -228,11 +228,11 @@
done
lemma (in fold_typing) fold_nest_Un_disjoint:
- "[| C\<in>Fin(A); D\<in>Fin(A); C \<inter> D = 0 |]
- ==> fold[B](f,e,C \<union> D) = fold[B](f, fold[B](f,e,D), C)"
+ "\<lbrakk>C\<in>Fin(A); D\<in>Fin(A); C \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> fold[B](f,e,C \<union> D) = fold[B](f, fold[B](f,e,D), C)"
by (simp add: fold_nest_Un_Int)
-lemma Finite_cons_lemma: "Finite(C) ==> C\<in>Fin(cons(c, C))"
+lemma Finite_cons_lemma: "Finite(C) \<Longrightarrow> C\<in>Fin(cons(c, C))"
apply (drule Finite_into_Fin)
apply (blast intro: Fin_mono [THEN subsetD])
done
@@ -243,7 +243,7 @@
by (simp add: setsum_def)
lemma setsum_cons [simp]:
- "Finite(C) ==>
+ "Finite(C) \<Longrightarrow>
setsum(g, cons(c,C)) =
(if c \<in> C then setsum(g,C) else g(c) $+ setsum(g,C))"
apply (auto simp add: setsum_def Finite_cons cons_absorb)
@@ -259,8 +259,8 @@
(*The reversed orientation looks more natural, but LOOPS as a simprule!*)
lemma setsum_Un_Int:
- "[| Finite(C); Finite(D) |]
- ==> setsum(g, C \<union> D) $+ setsum(g, C \<inter> D)
+ "\<lbrakk>Finite(C); Finite(D)\<rbrakk>
+ \<Longrightarrow> setsum(g, C \<union> D) $+ setsum(g, C \<inter> D)
= setsum(g, C) $+ setsum(g, D)"
apply (erule Finite_induct)
apply (simp_all add: Int_cons_right cons_absorb Un_cons Int_commute Finite_Un
@@ -274,21 +274,21 @@
done
lemma setsum_Un_disjoint:
- "[| Finite(C); Finite(D); C \<inter> D = 0 |]
- ==> setsum(g, C \<union> D) = setsum(g, C) $+ setsum(g,D)"
+ "\<lbrakk>Finite(C); Finite(D); C \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> setsum(g, C \<union> D) = setsum(g, C) $+ setsum(g,D)"
apply (subst setsum_Un_Int [symmetric])
apply (subgoal_tac [3] "Finite (C \<union> D) ")
apply (auto intro: Finite_Un)
done
lemma Finite_RepFun [rule_format (no_asm)]:
- "Finite(I) ==> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow> Finite(RepFun(I, C))"
+ "Finite(I) \<Longrightarrow> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow> Finite(RepFun(I, C))"
apply (erule Finite_induct, auto)
done
lemma setsum_UN_disjoint [rule_format (no_asm)]:
"Finite(I)
- ==> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow>
+ \<Longrightarrow> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow>
(\<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow>
setsum(f, \<Union>i\<in>I. C(i)) = setsum (%i. setsum(f, C(i)), I)"
apply (erule Finite_induct, auto)
@@ -310,15 +310,15 @@
lemma fold_set_cong:
- "[| A=A'; B=B'; e=e'; (\<forall>x\<in>A'. \<forall>y\<in>B'. f(x,y) = f'(x,y)) |]
- ==> fold_set(A,B,f,e) = fold_set(A',B',f',e')"
+ "\<lbrakk>A=A'; B=B'; e=e'; (\<forall>x\<in>A'. \<forall>y\<in>B'. f(x,y) = f'(x,y))\<rbrakk>
+ \<Longrightarrow> fold_set(A,B,f,e) = fold_set(A',B',f',e')"
apply (simp add: fold_set_def)
apply (intro refl iff_refl lfp_cong Collect_cong disj_cong ex_cong, auto)
done
lemma fold_cong:
-"[| B=B'; A=A'; e=e';
- !!x y. [|x\<in>A'; y\<in>B'|] ==> f(x,y) = f'(x,y) |] ==>
+"\<lbrakk>B=B'; A=A'; e=e';
+ \<And>x y. \<lbrakk>x\<in>A'; y\<in>B'\<rbrakk> \<Longrightarrow> f(x,y) = f'(x,y)\<rbrakk> \<Longrightarrow>
fold[B](f,e,A) = fold[B'](f', e', A')"
apply (simp add: fold_def)
apply (subst fold_set_cong)
@@ -326,27 +326,27 @@
done
lemma setsum_cong:
- "[| A=B; !!x. x\<in>B ==> f(x) = g(x) |] ==>
+ "\<lbrakk>A=B; \<And>x. x\<in>B \<Longrightarrow> f(x) = g(x)\<rbrakk> \<Longrightarrow>
setsum(f, A) = setsum(g, B)"
by (simp add: setsum_def cong add: fold_cong)
lemma setsum_Un:
- "[| Finite(A); Finite(B) |]
- ==> setsum(f, A \<union> B) =
+ "\<lbrakk>Finite(A); Finite(B)\<rbrakk>
+ \<Longrightarrow> setsum(f, A \<union> B) =
setsum(f, A) $+ setsum(f, B) $- setsum(f, A \<inter> B)"
apply (subst setsum_Un_Int [symmetric], auto)
done
lemma setsum_zneg_or_0 [rule_format (no_asm)]:
- "Finite(A) ==> (\<forall>x\<in>A. g(x) $\<le> #0) \<longrightarrow> setsum(g, A) $\<le> #0"
+ "Finite(A) \<Longrightarrow> (\<forall>x\<in>A. g(x) $\<le> #0) \<longrightarrow> setsum(g, A) $\<le> #0"
apply (erule Finite_induct)
apply (auto intro: zneg_or_0_add_zneg_or_0_imp_zneg_or_0)
done
lemma setsum_succD_lemma [rule_format]:
"Finite(A)
- ==> \<forall>n\<in>nat. setsum(f,A) = $# succ(n) \<longrightarrow> (\<exists>a\<in>A. #0 $< f(a))"
+ \<Longrightarrow> \<forall>n\<in>nat. setsum(f,A) = $# succ(n) \<longrightarrow> (\<exists>a\<in>A. #0 $< f(a))"
apply (erule Finite_induct)
apply (auto simp del: int_of_0 int_of_succ simp add: not_zless_iff_zle int_of_0 [symmetric])
apply (subgoal_tac "setsum (f, B) $\<le> #0")
@@ -360,7 +360,7 @@
done
lemma setsum_succD:
- "[| setsum(f, A) = $# succ(n); n\<in>nat |]==> \<exists>a\<in>A. #0 $< f(a)"
+ "\<lbrakk>setsum(f, A) = $# succ(n); n\<in>nat\<rbrakk>\<Longrightarrow> \<exists>a\<in>A. #0 $< f(a)"
apply (case_tac "Finite (A) ")
apply (blast intro: setsum_succD_lemma)
apply (unfold setsum_def)
@@ -368,27 +368,27 @@
done
lemma g_zpos_imp_setsum_zpos [rule_format]:
- "Finite(A) ==> (\<forall>x\<in>A. #0 $\<le> g(x)) \<longrightarrow> #0 $\<le> setsum(g, A)"
+ "Finite(A) \<Longrightarrow> (\<forall>x\<in>A. #0 $\<le> g(x)) \<longrightarrow> #0 $\<le> setsum(g, A)"
apply (erule Finite_induct)
apply (simp (no_asm))
apply (auto intro: zpos_add_zpos_imp_zpos)
done
lemma g_zpos_imp_setsum_zpos2 [rule_format]:
- "[| Finite(A); \<forall>x. #0 $\<le> g(x) |] ==> #0 $\<le> setsum(g, A)"
+ "\<lbrakk>Finite(A); \<forall>x. #0 $\<le> g(x)\<rbrakk> \<Longrightarrow> #0 $\<le> setsum(g, A)"
apply (erule Finite_induct)
apply (auto intro: zpos_add_zpos_imp_zpos)
done
lemma g_zspos_imp_setsum_zspos [rule_format]:
"Finite(A)
- ==> (\<forall>x\<in>A. #0 $< g(x)) \<longrightarrow> A \<noteq> 0 \<longrightarrow> (#0 $< setsum(g, A))"
+ \<Longrightarrow> (\<forall>x\<in>A. #0 $< g(x)) \<longrightarrow> A \<noteq> 0 \<longrightarrow> (#0 $< setsum(g, A))"
apply (erule Finite_induct)
apply (auto intro: zspos_add_zspos_imp_zspos)
done
lemma setsum_Diff [rule_format]:
- "Finite(A) ==> \<forall>a. M(a) = #0 \<longrightarrow> setsum(M, A) = setsum(M, A-{a})"
+ "Finite(A) \<Longrightarrow> \<forall>a. M(a) = #0 \<longrightarrow> setsum(M, A) = setsum(M, A-{a})"
apply (erule Finite_induct)
apply (simp_all add: Diff_cons_eq Finite_Diff)
done
--- a/src/ZF/Induct/ListN.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/ListN.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,11 +17,11 @@
domains "listn(A)" \<subseteq> "nat \<times> list(A)"
intros
NilI: "<0,Nil> \<in> listn(A)"
- ConsI: "[| a \<in> A; <n,l> \<in> listn(A) |] ==> <succ(n), Cons(a,l)> \<in> listn(A)"
+ ConsI: "\<lbrakk>a \<in> A; <n,l> \<in> listn(A)\<rbrakk> \<Longrightarrow> <succ(n), Cons(a,l)> \<in> listn(A)"
type_intros nat_typechecks list.intros
-lemma list_into_listn: "l \<in> list(A) ==> <length(l),l> \<in> listn(A)"
+lemma list_into_listn: "l \<in> list(A) \<Longrightarrow> <length(l),l> \<in> listn(A)"
by (induct set: list) (simp_all add: listn.intros)
lemma listn_iff: "<n,l> \<in> listn(A) \<longleftrightarrow> l \<in> list(A) & length(l)=n"
@@ -36,7 +36,7 @@
apply (simp add: listn_iff separation image_singleton_iff)
done
-lemma listn_mono: "A \<subseteq> B ==> listn(A) \<subseteq> listn(B)"
+lemma listn_mono: "A \<subseteq> B \<Longrightarrow> listn(A) \<subseteq> listn(B)"
apply (unfold listn.defs)
apply (rule lfp_mono)
apply (rule listn.bnd_mono)+
@@ -44,7 +44,7 @@
done
lemma listn_append:
- "[| <n,l> \<in> listn(A); <n',l'> \<in> listn(A) |] ==> <n#+n', l@l'> \<in> listn(A)"
+ "\<lbrakk><n,l> \<in> listn(A); <n',l'> \<in> listn(A)\<rbrakk> \<Longrightarrow> <n#+n', l@l'> \<in> listn(A)"
apply (erule listn.induct)
apply (frule listn.dom_subset [THEN subsetD])
apply (simp_all add: listn.intros)
--- a/src/ZF/Induct/Multiset.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Multiset.thy Tue Sep 27 16:51:35 2022 +0100
@@ -14,64 +14,64 @@
abbreviation (input)
\<comment> \<open>Short cut for multiset space\<close>
Mult :: "i=>i" where
- "Mult(A) == A -||> nat-{0}"
+ "Mult(A) \<equiv> A -||> nat-{0}"
definition
(* This is the original "restrict" from ZF.thy.
Restricts the function f to the domain A
FIXME: adapt Multiset to the new "restrict". *)
funrestrict :: "[i,i] => i" where
- "funrestrict(f,A) == \<lambda>x \<in> A. f`x"
+ "funrestrict(f,A) \<equiv> \<lambda>x \<in> A. f`x"
definition
(* M is a multiset *)
multiset :: "i => o" where
- "multiset(M) == \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
+ "multiset(M) \<equiv> \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
definition
mset_of :: "i=>i" where
- "mset_of(M) == domain(M)"
+ "mset_of(M) \<equiv> domain(M)"
definition
munion :: "[i, i] => i" (infixl \<open>+#\<close> 65) where
- "M +# N == \<lambda>x \<in> mset_of(M) \<union> mset_of(N).
+ "M +# N \<equiv> \<lambda>x \<in> mset_of(M) \<union> mset_of(N).
if x \<in> mset_of(M) \<inter> mset_of(N) then (M`x) #+ (N`x)
else (if x \<in> mset_of(M) then M`x else N`x)"
definition
(*convert a function to a multiset by eliminating 0*)
normalize :: "i => i" where
- "normalize(f) ==
+ "normalize(f) \<equiv>
if (\<exists>A. f \<in> A -> nat & Finite(A)) then
funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
else 0"
definition
mdiff :: "[i, i] => i" (infixl \<open>-#\<close> 65) where
- "M -# N == normalize(\<lambda>x \<in> mset_of(M).
+ "M -# N \<equiv> normalize(\<lambda>x \<in> mset_of(M).
if x \<in> mset_of(N) then M`x #- N`x else M`x)"
definition
(* set of elements of a multiset *)
msingle :: "i => i" (\<open>{#_#}\<close>) where
- "{#a#} == {<a, 1>}"
+ "{#a#} \<equiv> {<a, 1>}"
definition
MCollect :: "[i, i=>o] => i" (*comprehension*) where
- "MCollect(M, P) == funrestrict(M, {x \<in> mset_of(M). P(x)})"
+ "MCollect(M, P) \<equiv> funrestrict(M, {x \<in> mset_of(M). P(x)})"
definition
(* Counts the number of occurrences of an element in a multiset *)
mcount :: "[i, i] => i" where
- "mcount(M, a) == if a \<in> mset_of(M) then M`a else 0"
+ "mcount(M, a) \<equiv> if a \<in> mset_of(M) then M`a else 0"
definition
msize :: "i => i" where
- "msize(M) == setsum(%a. $# mcount(M,a), mset_of(M))"
+ "msize(M) \<equiv> setsum(%a. $# mcount(M,a), mset_of(M))"
abbreviation
melem :: "[i,i] => o" (\<open>(_/ :# _)\<close> [50, 51] 50) where
- "a :# M == a \<in> mset_of(M)"
+ "a :# M \<equiv> a \<in> mset_of(M)"
syntax
"_MColl" :: "[pttrn, i, o] => i" (\<open>(1{# _ \<in> _./ _#})\<close>)
@@ -84,43 +84,43 @@
(* multirel1 has to be a set (not a predicate) so that we can form
its transitive closure and reason about wf(.) and acc(.) *)
multirel1 :: "[i,i]=>i" where
- "multirel1(A, r) ==
+ "multirel1(A, r) \<equiv>
{<M, N> \<in> Mult(A)*Mult(A).
\<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).
N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
definition
multirel :: "[i, i] => i" where
- "multirel(A, r) == multirel1(A, r)^+"
+ "multirel(A, r) \<equiv> multirel1(A, r)^+"
(* ordinal multiset orderings *)
definition
omultiset :: "i => o" where
- "omultiset(M) == \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"
+ "omultiset(M) \<equiv> \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"
definition
mless :: "[i, i] => o" (infixl \<open><#\<close> 50) where
- "M <# N == \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
+ "M <# N \<equiv> \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
definition
mle :: "[i, i] => o" (infixl \<open><#=\<close> 50) where
- "M <#= N == (omultiset(M) & M = N) | M <# N"
+ "M <#= N \<equiv> (omultiset(M) & M = N) | M <# N"
subsection\<open>Properties of the original "restrict" from ZF.thy\<close>
-lemma funrestrict_subset: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<subseteq> f"
+lemma funrestrict_subset: "\<lbrakk>f \<in> Pi(C,B); A\<subseteq>C\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<subseteq> f"
by (auto simp add: funrestrict_def lam_def intro: apply_Pair)
lemma funrestrict_type:
- "[| !!x. x \<in> A ==> f`x \<in> B(x) |] ==> funrestrict(f,A) \<in> Pi(A,B)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f`x \<in> B(x)\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<in> Pi(A,B)"
by (simp add: funrestrict_def lam_type)
-lemma funrestrict_type2: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<in> Pi(A,B)"
+lemma funrestrict_type2: "\<lbrakk>f \<in> Pi(C,B); A\<subseteq>C\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<in> Pi(A,B)"
by (blast intro: apply_type funrestrict_type)
-lemma funrestrict [simp]: "a \<in> A ==> funrestrict(f,A) ` a = f`a"
+lemma funrestrict [simp]: "a \<in> A \<Longrightarrow> funrestrict(f,A) ` a = f`a"
by (simp add: funrestrict_def)
lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0"
@@ -130,7 +130,7 @@
by (auto simp add: funrestrict_def lam_def)
lemma fun_cons_funrestrict_eq:
- "f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, funrestrict(f, b))"
+ "f \<in> cons(a, b) -> B \<Longrightarrow> f = cons(<a, f ` a>, funrestrict(f, b))"
apply (rule equalityI)
prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD])
apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def)
@@ -151,7 +151,7 @@
done
(** The multiset space **)
-lemma multiset_into_Mult: "[| multiset(M); mset_of(M)\<subseteq>A |] ==> M \<in> Mult(A)"
+lemma multiset_into_Mult: "\<lbrakk>multiset(M); mset_of(M)\<subseteq>A\<rbrakk> \<Longrightarrow> M \<in> Mult(A)"
apply (simp add: multiset_def)
apply (auto simp add: multiset_fun_iff mset_of_def)
apply (rule_tac B1 = "nat-{0}" in FiniteFun_mono [THEN subsetD], simp_all)
@@ -159,7 +159,7 @@
apply (simp_all (no_asm_simp) add: multiset_fun_iff)
done
-lemma Mult_into_multiset: "M \<in> Mult(A) ==> multiset(M) & mset_of(M)\<subseteq>A"
+lemma Mult_into_multiset: "M \<in> Mult(A) \<Longrightarrow> multiset(M) & mset_of(M)\<subseteq>A"
apply (simp add: multiset_def mset_of_def)
apply (frule FiniteFun_is_fun)
apply (drule FiniteFun_domain_Fin)
@@ -185,13 +185,13 @@
text\<open>The \<^term>\<open>mset_of\<close> operator\<close>
-lemma multiset_set_of_Finite [simp]: "multiset(M) ==> Finite(mset_of(M))"
+lemma multiset_set_of_Finite [simp]: "multiset(M) \<Longrightarrow> Finite(mset_of(M))"
by (simp add: multiset_def mset_of_def, auto)
lemma mset_of_0 [iff]: "mset_of(0) = 0"
by (simp add: mset_of_def)
-lemma mset_is_0_iff: "multiset(M) ==> mset_of(M)=0 \<longleftrightarrow> M=0"
+lemma mset_is_0_iff: "multiset(M) \<Longrightarrow> mset_of(M)=0 \<longleftrightarrow> M=0"
by (auto simp add: multiset_def mset_of_def)
lemma mset_of_single [iff]: "mset_of({#a#}) = {a}"
@@ -200,7 +200,7 @@
lemma mset_of_union [iff]: "mset_of(M +# N) = mset_of(M) \<union> mset_of(N)"
by (simp add: mset_of_def munion_def)
-lemma mset_of_diff [simp]: "mset_of(M)\<subseteq>A ==> mset_of(M -# N) \<subseteq> A"
+lemma mset_of_diff [simp]: "mset_of(M)\<subseteq>A \<Longrightarrow> mset_of(M -# N) \<subseteq> A"
by (auto simp add: mdiff_def multiset_def normalize_def mset_of_def)
(* msingle *)
@@ -230,7 +230,7 @@
apply (auto intro!: lam_type simp add: Collect_Finite)
done
-lemma normalize_multiset [simp]: "multiset(M) ==> normalize(M) = M"
+lemma normalize_multiset [simp]: "multiset(M) \<Longrightarrow> normalize(M) = M"
by (auto simp add: multiset_def normalize_def mset_of_def funrestrict_def multiset_fun_iff)
lemma multiset_normalize [simp]: "multiset(normalize(f))"
@@ -244,7 +244,7 @@
(* union *)
-lemma munion_multiset [simp]: "[| multiset(M); multiset(N) |] ==> multiset(M +# N)"
+lemma munion_multiset [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk> \<Longrightarrow> multiset(M +# N)"
apply (unfold multiset_def munion_def mset_of_def, auto)
apply (rule_tac x = "A \<union> Aa" in exI)
apply (auto intro!: lam_type intro: Finite_Un simp add: multiset_fun_iff zero_less_add)
@@ -259,7 +259,7 @@
(* Union *)
-lemma munion_0 [simp]: "multiset(M) ==> M +# 0 = M & 0 +# M = M"
+lemma munion_0 [simp]: "multiset(M) \<Longrightarrow> M +# 0 = M & 0 +# M = M"
apply (simp add: multiset_def)
apply (auto simp add: munion_def mset_of_def)
done
@@ -287,10 +287,10 @@
lemma mdiff_0 [simp]: "0 -# M = 0"
by (simp add: mdiff_def normalize_def)
-lemma mdiff_0_right [simp]: "multiset(M) ==> M -# 0 = M"
+lemma mdiff_0_right [simp]: "multiset(M) \<Longrightarrow> M -# 0 = M"
by (auto simp add: multiset_def mdiff_def normalize_def multiset_fun_iff mset_of_def funrestrict_def)
-lemma mdiff_union_inverse2 [simp]: "multiset(M) ==> M +# {#a#} -# {#a#} = M"
+lemma mdiff_union_inverse2 [simp]: "multiset(M) \<Longrightarrow> M +# {#a#} -# {#a#} = M"
apply (unfold multiset_def munion_def mdiff_def msingle_def normalize_def mset_of_def)
apply (auto cong add: if_cong simp add: ltD multiset_fun_iff funrestrict_def subset_Un_iff2 [THEN iffD1])
prefer 2 apply (force intro!: lam_type)
@@ -303,7 +303,7 @@
(** Count of elements **)
-lemma mcount_type [simp,TC]: "multiset(M) ==> mcount(M, a) \<in> nat"
+lemma mcount_type [simp,TC]: "multiset(M) \<Longrightarrow> mcount(M, a) \<in> nat"
by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff)
lemma mcount_0 [simp]: "mcount(0, a) = 0"
@@ -312,13 +312,13 @@
lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)"
by (simp add: mcount_def mset_of_def msingle_def)
-lemma mcount_union [simp]: "[| multiset(M); multiset(N) |]
- ==> mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)"
+lemma mcount_union [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
+ \<Longrightarrow> mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)"
apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def)
done
lemma mcount_diff [simp]:
- "multiset(M) ==> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)"
+ "multiset(M) \<Longrightarrow> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)"
apply (simp add: multiset_def)
apply (auto dest!: not_lt_imp_le
simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def)
@@ -326,7 +326,7 @@
apply (force intro!: lam_type)
done
-lemma mcount_elem: "[| multiset(M); a \<in> mset_of(M) |] ==> 0 < mcount(M, a)"
+lemma mcount_elem: "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> 0 < mcount(M, a)"
apply (simp add: multiset_def, clarify)
apply (simp add: mcount_def mset_of_def)
apply (simp add: multiset_fun_iff)
@@ -343,23 +343,23 @@
lemma msize_type [simp,TC]: "msize(M) \<in> int"
by (simp add: msize_def)
-lemma msize_zpositive: "multiset(M)==> #0 $\<le> msize(M)"
+lemma msize_zpositive: "multiset(M)\<Longrightarrow> #0 $\<le> msize(M)"
by (auto simp add: msize_def intro: g_zpos_imp_setsum_zpos)
-lemma msize_int_of_nat: "multiset(M) ==> \<exists>n \<in> nat. msize(M)= $# n"
+lemma msize_int_of_nat: "multiset(M) \<Longrightarrow> \<exists>n \<in> nat. msize(M)= $# n"
apply (rule not_zneg_int_of)
apply (simp_all (no_asm_simp) add: msize_type [THEN znegative_iff_zless_0] not_zless_iff_zle msize_zpositive)
done
lemma not_empty_multiset_imp_exist:
- "[| M\<noteq>0; multiset(M) |] ==> \<exists>a \<in> mset_of(M). 0 < mcount(M, a)"
+ "\<lbrakk>M\<noteq>0; multiset(M)\<rbrakk> \<Longrightarrow> \<exists>a \<in> mset_of(M). 0 < mcount(M, a)"
apply (simp add: multiset_def)
apply (erule not_emptyE)
apply (auto simp add: mset_of_def mcount_def multiset_fun_iff)
apply (blast dest!: fun_is_rel)
done
-lemma msize_eq_0_iff: "multiset(M) ==> msize(M)=#0 \<longleftrightarrow> M=0"
+lemma msize_eq_0_iff: "multiset(M) \<Longrightarrow> msize(M)=#0 \<longleftrightarrow> M=0"
apply (simp add: msize_def, auto)
apply (rule_tac P = "setsum (u,v) \<noteq> #0" for u v in swap)
apply blast
@@ -376,7 +376,7 @@
done
lemma setsum_mcount_Int:
- "Finite(A) ==> setsum(%a. $# mcount(N, a), A \<inter> mset_of(N))
+ "Finite(A) \<Longrightarrow> setsum(%a. $# mcount(N, a), A \<inter> mset_of(N))
= setsum(%a. $# mcount(N, a), A)"
apply (induct rule: Finite_induct)
apply auto
@@ -386,13 +386,13 @@
done
lemma msize_union [simp]:
- "[| multiset(M); multiset(N) |] ==> msize(M +# N) = msize(M) $+ msize(N)"
+ "\<lbrakk>multiset(M); multiset(N)\<rbrakk> \<Longrightarrow> msize(M +# N) = msize(M) $+ msize(N)"
apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int)
apply (subst Int_commute)
apply (simp add: setsum_mcount_Int)
done
-lemma msize_eq_succ_imp_elem: "[|msize(M)= $# succ(n); n \<in> nat|] ==> \<exists>a. a \<in> mset_of(M)"
+lemma msize_eq_succ_imp_elem: "\<lbrakk>msize(M)= $# succ(n); n \<in> nat\<rbrakk> \<Longrightarrow> \<exists>a. a \<in> mset_of(M)"
apply (unfold msize_def)
apply (blast dest: setsum_succD)
done
@@ -400,8 +400,8 @@
(** Equality of multisets **)
lemma equality_lemma:
- "[| multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a) |]
- ==> mset_of(M)=mset_of(N)"
+ "\<lbrakk>multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a)\<rbrakk>
+ \<Longrightarrow> mset_of(M)=mset_of(N)"
apply (simp add: multiset_def)
apply (rule sym, rule equalityI)
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)
@@ -410,7 +410,7 @@
done
lemma multiset_equality:
- "[| multiset(M); multiset(N) |]==> M=N\<longleftrightarrow>(\<forall>a. mcount(M, a)=mcount(N, a))"
+ "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow> M=N\<longleftrightarrow>(\<forall>a. mcount(M, a)=mcount(N, a))"
apply auto
apply (subgoal_tac "mset_of (M) = mset_of (N) ")
prefer 2 apply (blast intro: equality_lemma)
@@ -424,28 +424,28 @@
(** More algebraic properties of multisets **)
-lemma munion_eq_0_iff [simp]: "[|multiset(M); multiset(N)|]==>(M +# N =0) \<longleftrightarrow> (M=0 & N=0)"
+lemma munion_eq_0_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(M +# N =0) \<longleftrightarrow> (M=0 & N=0)"
by (auto simp add: multiset_equality)
-lemma empty_eq_munion_iff [simp]: "[|multiset(M); multiset(N)|]==>(0=M +# N) \<longleftrightarrow> (M=0 & N=0)"
+lemma empty_eq_munion_iff [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow>(0=M +# N) \<longleftrightarrow> (M=0 & N=0)"
apply (rule iffI, drule sym)
apply (simp_all add: multiset_equality)
done
lemma munion_right_cancel [simp]:
- "[| multiset(M); multiset(N); multiset(K) |]==>(M +# K = N +# K)\<longleftrightarrow>(M=N)"
+ "\<lbrakk>multiset(M); multiset(N); multiset(K)\<rbrakk>\<Longrightarrow>(M +# K = N +# K)\<longleftrightarrow>(M=N)"
by (auto simp add: multiset_equality)
lemma munion_left_cancel [simp]:
- "[|multiset(K); multiset(M); multiset(N)|] ==>(K +# M = K +# N) \<longleftrightarrow> (M = N)"
+ "\<lbrakk>multiset(K); multiset(M); multiset(N)\<rbrakk> \<Longrightarrow>(K +# M = K +# N) \<longleftrightarrow> (M = N)"
by (auto simp add: multiset_equality)
-lemma nat_add_eq_1_cases: "[| m \<in> nat; n \<in> nat |] ==> (m #+ n = 1) \<longleftrightarrow> (m=1 & n=0) | (m=0 & n=1)"
+lemma nat_add_eq_1_cases: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> (m #+ n = 1) \<longleftrightarrow> (m=1 & n=0) | (m=0 & n=1)"
by (induct_tac n) auto
lemma munion_is_single:
- "[|multiset(M); multiset(N)|]
- ==> (M +# N = {#a#}) \<longleftrightarrow> (M={#a#} & N=0) | (M = 0 & N = {#a#})"
+ "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
+ \<Longrightarrow> (M +# N = {#a#}) \<longleftrightarrow> (M={#a#} & N=0) | (M = 0 & N = {#a#})"
apply (simp (no_asm_simp) add: multiset_equality)
apply safe
apply simp_all
@@ -464,8 +464,8 @@
apply (simp_all add: nat_add_eq_1_cases)
done
-lemma msingle_is_union: "[| multiset(M); multiset(N) |]
- ==> ({#a#} = M +# N) \<longleftrightarrow> ({#a#} = M & N=0 | M = 0 & {#a#} = N)"
+lemma msingle_is_union: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
+ \<Longrightarrow> ({#a#} = M +# N) \<longleftrightarrow> ({#a#} = M & N=0 | M = 0 & {#a#} = N)"
apply (subgoal_tac " ({#a#} = M +# N) \<longleftrightarrow> (M +# N = {#a#}) ")
apply (simp (no_asm_simp) add: munion_is_single)
apply blast
@@ -476,7 +476,7 @@
lemma setsum_decr:
"Finite(A)
- ==> (\<forall>M. multiset(M) \<longrightarrow>
+ \<Longrightarrow> (\<forall>M. multiset(M) \<longrightarrow>
(\<forall>a \<in> mset_of(M). setsum(%z. $# mcount(M(a:=M`a #- 1), z), A) =
(if a \<in> A then setsum(%z. $# mcount(M, z), A) $- #1
else setsum(%z. $# mcount(M, z), A))))"
@@ -492,7 +492,7 @@
lemma setsum_decr2:
"Finite(A)
- ==> \<forall>M. multiset(M) \<longrightarrow> (\<forall>a \<in> mset_of(M).
+ \<Longrightarrow> \<forall>M. multiset(M) \<longrightarrow> (\<forall>a \<in> mset_of(M).
setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) =
(if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
else setsum(%x. $# mcount(M, x), A)))"
@@ -501,8 +501,8 @@
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)
done
-lemma setsum_decr3: "[| Finite(A); multiset(M); a \<in> mset_of(M) |]
- ==> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
+lemma setsum_decr3: "\<lbrakk>Finite(A); multiset(M); a \<in> mset_of(M)\<rbrakk>
+ \<Longrightarrow> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
(if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
else setsum(%x. $# mcount(M, x), A))"
apply (subgoal_tac "setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ")
@@ -512,10 +512,10 @@
apply (simp add: mcount_def mset_of_def)
done
-lemma nat_le_1_cases: "n \<in> nat ==> n \<le> 1 \<longleftrightarrow> (n=0 | n=1)"
+lemma nat_le_1_cases: "n \<in> nat \<Longrightarrow> n \<le> 1 \<longleftrightarrow> (n=0 | n=1)"
by (auto elim: natE)
-lemma succ_pred_eq_self: "[| 0<n; n \<in> nat |] ==> succ(n #- 1) = n"
+lemma succ_pred_eq_self: "\<lbrakk>0<n; n \<in> nat\<rbrakk> \<Longrightarrow> succ(n #- 1) = n"
apply (subgoal_tac "1 \<le> n")
apply (drule add_diff_inverse2, auto)
done
@@ -530,11 +530,11 @@
done
lemma multiset_induct_aux:
- assumes prem1: "!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M))"
- and prem2: "!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))"
+ assumes prem1: "\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(<a, 1>, M))"
+ and prem2: "\<And>M b. \<lbrakk>multiset(M); b \<in> mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(M(b:= M`b #+ 1))"
shows
- "[| n \<in> nat; P(0) |]
- ==> (\<forall>M. multiset(M)\<longrightarrow>
+ "\<lbrakk>n \<in> nat; P(0)\<rbrakk>
+ \<Longrightarrow> (\<forall>M. multiset(M)\<longrightarrow>
(setsum(%x. $# mcount(M, x), {x \<in> mset_of(M). 0 < M`x}) = $# n) \<longrightarrow> P(M))"
apply (erule nat_induct, clarify)
apply (frule msize_eq_0_iff)
@@ -600,10 +600,10 @@
done
lemma multiset_induct2:
- "[| multiset(M); P(0);
- (!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M)));
- (!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))) |]
- ==> P(M)"
+ "\<lbrakk>multiset(M); P(0);
+ (\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(<a, 1>, M)));
+ (\<And>M b. \<lbrakk>multiset(M); b \<in> mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(M(b:= M`b #+ 1)))\<rbrakk>
+ \<Longrightarrow> P(M)"
apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n")
apply (rule_tac [2] not_zneg_int_of)
apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle)
@@ -614,7 +614,7 @@
done
lemma munion_single_case1:
- "[| multiset(M); a \<notin>mset_of(M) |] ==> M +# {#a#} = cons(<a, 1>, M)"
+ "\<lbrakk>multiset(M); a \<notin>mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = cons(<a, 1>, M)"
apply (simp add: multiset_def msingle_def)
apply (auto simp add: munion_def)
apply (unfold mset_of_def, simp)
@@ -625,7 +625,7 @@
done
lemma munion_single_case2:
- "[| multiset(M); a \<in> mset_of(M) |] ==> M +# {#a#} = M(a:=M`a #+ 1)"
+ "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = M(a:=M`a #+ 1)"
apply (simp add: multiset_def)
apply (auto simp add: munion_def multiset_fun_iff msingle_def)
apply (unfold mset_of_def, simp)
@@ -639,7 +639,7 @@
lemma multiset_induct:
assumes M: "multiset(M)"
and P0: "P(0)"
- and step: "!!M a. [| multiset(M); P(M) |] ==> P(M +# {#a#})"
+ and step: "\<And>M a. \<lbrakk>multiset(M); P(M)\<rbrakk> \<Longrightarrow> P(M +# {#a#})"
shows "P(M)"
apply (rule multiset_induct2 [OF M])
apply (simp_all add: P0)
@@ -651,7 +651,7 @@
(** MCollect **)
lemma MCollect_multiset [simp]:
- "multiset(M) ==> multiset({# x \<in> M. P(x)#})"
+ "multiset(M) \<Longrightarrow> multiset({# x \<in> M. P(x)#})"
apply (simp add: MCollect_def multiset_def mset_of_def, clarify)
apply (rule_tac x = "{x \<in> A. P (x) }" in exI)
apply (auto dest: CollectD1 [THEN [2] apply_type]
@@ -659,7 +659,7 @@
done
lemma mset_of_MCollect [simp]:
- "multiset(M) ==> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)"
+ "multiset(M) \<Longrightarrow> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)"
by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def)
lemma MCollect_mem_iff [iff]:
@@ -670,17 +670,17 @@
"mcount({# x \<in> M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)"
by (simp add: mcount_def MCollect_def mset_of_def)
-lemma multiset_partition: "multiset(M) ==> M = {# x \<in> M. P(x) #} +# {# x \<in> M. ~ P(x) #}"
+lemma multiset_partition: "multiset(M) \<Longrightarrow> M = {# x \<in> M. P(x) #} +# {# x \<in> M. \<not> P(x) #}"
by (simp add: multiset_equality)
lemma natify_elem_is_self [simp]:
- "[| multiset(M); a \<in> mset_of(M) |] ==> natify(M`a) = M`a"
+ "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> natify(M`a) = M`a"
by (auto simp add: multiset_def mset_of_def multiset_fun_iff)
(* and more algebraic laws on multisets *)
-lemma munion_eq_conv_diff: "[| multiset(M); multiset(N) |]
- ==> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow> (M = N & a = b |
+lemma munion_eq_conv_diff: "\<lbrakk>multiset(M); multiset(N)\<rbrakk>
+ \<Longrightarrow> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow> (M = N & a = b |
M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})"
apply (simp del: mcount_single add: multiset_equality)
apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE)
@@ -694,7 +694,7 @@
done
lemma melem_diff_single:
-"multiset(M) ==>
+"multiset(M) \<Longrightarrow>
k \<in> mset_of(M -# {#a#}) \<longleftrightarrow> (k=a & 1 < mcount(M,a)) | (k\<noteq> a & k \<in> mset_of(M))"
apply (simp add: multiset_def)
apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def)
@@ -706,8 +706,8 @@
done
lemma munion_eq_conv_exist:
-"[| M \<in> Mult(A); N \<in> Mult(A) |]
- ==> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow>
+"\<lbrakk>M \<in> Mult(A); N \<in> Mult(A)\<rbrakk>
+ \<Longrightarrow> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow>
(M=N & a=b | (\<exists>K \<in> Mult(A). M= K +# {#b#} & N=K +# {#a#}))"
by (auto simp add: Mult_iff_multiset melem_diff_single munion_eq_conv_diff)
@@ -735,7 +735,7 @@
text\<open>Monotonicity of \<^term>\<open>multirel1\<close>\<close>
-lemma multirel1_mono1: "A\<subseteq>B ==> multirel1(A, r)\<subseteq>multirel1(B, r)"
+lemma multirel1_mono1: "A\<subseteq>B \<Longrightarrow> multirel1(A, r)\<subseteq>multirel1(B, r)"
apply (auto simp add: multirel1_def)
apply (auto simp add: Un_subset_iff Mult_iff_multiset)
apply (rule_tac x = a in bexI)
@@ -744,7 +744,7 @@
apply (auto simp add: Mult_iff_multiset)
done
-lemma multirel1_mono2: "r\<subseteq>s ==> multirel1(A,r)\<subseteq>multirel1(A, s)"
+lemma multirel1_mono2: "r\<subseteq>s \<Longrightarrow> multirel1(A,r)\<subseteq>multirel1(A, s)"
apply (simp add: multirel1_def, auto)
apply (rule_tac x = a in bexI)
apply (rule_tac x = M0 in bexI)
@@ -754,7 +754,7 @@
done
lemma multirel1_mono:
- "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel1(A, r) \<subseteq> multirel1(B, s)"
+ "\<lbrakk>A\<subseteq>B; r\<subseteq>s\<rbrakk> \<Longrightarrow> multirel1(A, r) \<subseteq> multirel1(B, s)"
apply (rule subset_trans)
apply (rule multirel1_mono1)
apply (rule_tac [2] multirel1_mono2, auto)
@@ -765,7 +765,7 @@
lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)"
by (auto simp add: multirel1_def Mult_iff_multiset)
-lemma less_munion: "[| <N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A) |] ==>
+lemma less_munion: "\<lbrakk><N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A)\<rbrakk> \<Longrightarrow>
(\<exists>M. <M, M0> \<in> multirel1(A, r) & N = M +# {#a#}) |
(\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)"
apply (frule multirel1_type [THEN subsetD])
@@ -776,7 +776,7 @@
apply (auto simp add: munion_commute)
done
-lemma multirel1_base: "[| M \<in> Mult(A); a \<in> A |] ==> <M, M +# {#a#}> \<in> multirel1(A, r)"
+lemma multirel1_base: "\<lbrakk>M \<in> Mult(A); a \<in> A\<rbrakk> \<Longrightarrow> <M, M +# {#a#}> \<in> multirel1(A, r)"
apply (auto simp add: multirel1_iff)
apply (simp add: Mult_iff_multiset)
apply (rule_tac x = a in exI, clarify)
@@ -787,11 +787,11 @@
lemma acc_0: "acc(0)=0"
by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD])
-lemma lemma1: "[| \<forall>b \<in> A. <b,a> \<in> r \<longrightarrow>
+lemma lemma1: "\<lbrakk>\<forall>b \<in> A. <b,a> \<in> r \<longrightarrow>
(\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r)));
M0 \<in> acc(multirel1(A, r)); a \<in> A;
- \<forall>M. <M,M0> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r)) |]
- ==> M0 +# {#a#} \<in> acc(multirel1(A, r))"
+ \<forall>M. <M,M0> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))\<rbrakk>
+ \<Longrightarrow> M0 +# {#a#} \<in> acc(multirel1(A, r))"
apply (subgoal_tac "M0 \<in> Mult(A) ")
prefer 2
apply (erule acc.cases)
@@ -819,20 +819,20 @@
apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset)
done
-lemma lemma2: "[| \<forall>b \<in> A. <b,a> \<in> r
+lemma lemma2: "\<lbrakk>\<forall>b \<in> A. <b,a> \<in> r
\<longrightarrow> (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r)));
- M \<in> acc(multirel1(A, r)); a \<in> A|] ==> M +# {#a#} \<in> acc(multirel1(A, r))"
+ M \<in> acc(multirel1(A, r)); a \<in> A\<rbrakk> \<Longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))"
apply (erule acc_induct)
apply (blast intro: lemma1)
done
-lemma lemma3: "[| wf[A](r); a \<in> A |]
- ==> \<forall>M \<in> acc(multirel1(A, r)). M +# {#a#} \<in> acc(multirel1(A, r))"
+lemma lemma3: "\<lbrakk>wf[A](r); a \<in> A\<rbrakk>
+ \<Longrightarrow> \<forall>M \<in> acc(multirel1(A, r)). M +# {#a#} \<in> acc(multirel1(A, r))"
apply (erule_tac a = a in wf_on_induct, blast)
apply (blast intro: lemma2)
done
-lemma lemma4: "multiset(M) ==> mset_of(M)\<subseteq>A \<longrightarrow>
+lemma lemma4: "multiset(M) \<Longrightarrow> mset_of(M)\<subseteq>A \<longrightarrow>
wf[A](r) \<longrightarrow> M \<in> field(multirel1(A, r)) \<longrightarrow> M \<in> acc(multirel1(A, r))"
apply (erule multiset_induct)
(* proving the base case *)
@@ -852,14 +852,14 @@
apply (auto simp add: Mult_iff_multiset)
done
-lemma all_accessible: "[| wf[A](r); M \<in> Mult(A); A \<noteq> 0|] ==> M \<in> acc(multirel1(A, r))"
+lemma all_accessible: "\<lbrakk>wf[A](r); M \<in> Mult(A); A \<noteq> 0\<rbrakk> \<Longrightarrow> M \<in> acc(multirel1(A, r))"
apply (erule not_emptyE)
apply (rule lemma4 [THEN mp, THEN mp, THEN mp])
apply (rule_tac [4] multirel1_base [THEN fieldI1])
apply (auto simp add: Mult_iff_multiset)
done
-lemma wf_on_multirel1: "wf[A](r) ==> wf[A-||>nat-{0}](multirel1(A, r))"
+lemma wf_on_multirel1: "wf[A](r) \<Longrightarrow> wf[A-||>nat-{0}](multirel1(A, r))"
apply (case_tac "A=0")
apply (simp (no_asm_simp))
apply (rule wf_imp_wf_on)
@@ -870,7 +870,7 @@
apply (blast intro: all_accessible)
done
-lemma wf_multirel1: "wf(r) ==>wf(multirel1(field(r), r))"
+lemma wf_multirel1: "wf(r) \<Longrightarrow>wf(multirel1(field(r), r))"
apply (simp (no_asm_use) add: wf_iff_wf_on_field)
apply (drule wf_on_multirel1)
apply (rule_tac A = "field (r) -||> nat - {0}" in wf_on_subset_A)
@@ -889,7 +889,7 @@
(* Monotonicity of multirel *)
lemma multirel_mono:
- "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel(A, r)\<subseteq>multirel(B,s)"
+ "\<lbrakk>A\<subseteq>B; r\<subseteq>s\<rbrakk> \<Longrightarrow> multirel(A, r)\<subseteq>multirel(B,s)"
apply (simp add: multirel_def)
apply (rule trancl_mono)
apply (rule multirel1_mono, auto)
@@ -897,24 +897,24 @@
(* Equivalence of multirel with the usual (closure-free) definition *)
-lemma add_diff_eq: "k \<in> nat ==> 0 < k \<longrightarrow> n #+ k #- 1 = n #+ (k #- 1)"
+lemma add_diff_eq: "k \<in> nat \<Longrightarrow> 0 < k \<longrightarrow> n #+ k #- 1 = n #+ (k #- 1)"
by (erule nat_induct, auto)
-lemma mdiff_union_single_conv: "[|a \<in> mset_of(J); multiset(I); multiset(J) |]
- ==> I +# J -# {#a#} = I +# (J-# {#a#})"
+lemma mdiff_union_single_conv: "\<lbrakk>a \<in> mset_of(J); multiset(I); multiset(J)\<rbrakk>
+ \<Longrightarrow> I +# J -# {#a#} = I +# (J-# {#a#})"
apply (simp (no_asm_simp) add: multiset_equality)
apply (case_tac "a \<notin> mset_of (I) ")
apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff)
apply (auto dest: domain_type simp add: add_diff_eq)
done
-lemma diff_add_commute: "[| n \<le> m; m \<in> nat; n \<in> nat; k \<in> nat |] ==> m #- n #+ k = m #+ k #- n"
+lemma diff_add_commute: "\<lbrakk>n \<le> m; m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> m #- n #+ k = m #+ k #- n"
by (auto simp add: le_iff less_iff_succ_add)
(* One direction *)
lemma multirel_implies_one_step:
-"<M,N> \<in> multirel(A, r) ==>
+"<M,N> \<in> multirel(A, r) \<Longrightarrow>
trans[A](r) \<longrightarrow>
(\<exists>I J K.
I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &
@@ -964,12 +964,12 @@
apply (drule_tac [2] sym, auto)
done
-lemma melem_imp_eq_diff_union [simp]: "[| a \<in> mset_of(M); multiset(M) |] ==> M -# {#a#} +# {#a#} = M"
+lemma melem_imp_eq_diff_union [simp]: "\<lbrakk>a \<in> mset_of(M); multiset(M)\<rbrakk> \<Longrightarrow> M -# {#a#} +# {#a#} = M"
by (simp add: multiset_equality mcount_elem [THEN succ_pred_eq_self])
lemma msize_eq_succ_imp_eq_union:
- "[| msize(M)=$# succ(n); M \<in> Mult(A); n \<in> nat |]
- ==> \<exists>a N. M = N +# {#a#} & N \<in> Mult(A) & a \<in> A"
+ "\<lbrakk>msize(M)=$# succ(n); M \<in> Mult(A); n \<in> nat\<rbrakk>
+ \<Longrightarrow> \<exists>a N. M = N +# {#a#} & N \<in> Mult(A) & a \<in> A"
apply (drule msize_eq_succ_imp_elem, auto)
apply (rule_tac x = a in exI)
apply (rule_tac x = "M -# {#a#}" in exI)
@@ -981,7 +981,7 @@
(* The second direction *)
lemma one_step_implies_multirel_lemma [rule_format (no_asm)]:
-"n \<in> nat ==>
+"n \<in> nat \<Longrightarrow>
(\<forall>I J K.
I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &
(msize(J) = $# n & J \<noteq>0 & (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k, j> \<in> r))
@@ -1024,9 +1024,9 @@
done
lemma one_step_implies_multirel:
- "[| J \<noteq> 0; \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r;
- I \<in> Mult(A); J \<in> Mult(A); K \<in> Mult(A) |]
- ==> <I+#K, I+#J> \<in> multirel(A, r)"
+ "\<lbrakk>J \<noteq> 0; \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r;
+ I \<in> Mult(A); J \<in> Mult(A); K \<in> Mult(A)\<rbrakk>
+ \<Longrightarrow> <I+#K, I+#J> \<in> multirel(A, r)"
apply (subgoal_tac "multiset (J) ")
prefer 2 apply (simp add: Mult_iff_multiset)
apply (frule_tac M = J in msize_int_of_nat)
@@ -1038,7 +1038,7 @@
(*irreflexivity*)
lemma multirel_irrefl_lemma:
- "Finite(A) ==> part_ord(A, r) \<longrightarrow> (\<forall>x \<in> A. \<exists>y \<in> A. <x,y> \<in> r) \<longrightarrow>A=0"
+ "Finite(A) \<Longrightarrow> part_ord(A, r) \<longrightarrow> (\<forall>x \<in> A. \<exists>y \<in> A. <x,y> \<in> r) \<longrightarrow>A=0"
apply (erule Finite_induct)
apply (auto dest: subset_consI [THEN [2] part_ord_subset])
apply (auto simp add: part_ord_def irrefl_def)
@@ -1047,7 +1047,7 @@
done
lemma irrefl_on_multirel:
- "part_ord(A, r) ==> irrefl(Mult(A), multirel(A, r))"
+ "part_ord(A, r) \<Longrightarrow> irrefl(Mult(A), multirel(A, r))"
apply (simp add: irrefl_def)
apply (subgoal_tac "trans[A](r) ")
prefer 2 apply (simp add: part_ord_def, clarify)
@@ -1066,7 +1066,7 @@
done
lemma multirel_trans:
- "[| <M, N> \<in> multirel(A, r); <N, K> \<in> multirel(A, r) |] ==> <M, K> \<in> multirel(A,r)"
+ "\<lbrakk><M, N> \<in> multirel(A, r); <N, K> \<in> multirel(A, r)\<rbrakk> \<Longrightarrow> <M, K> \<in> multirel(A,r)"
apply (simp add: multirel_def)
apply (blast intro: trancl_trans)
done
@@ -1076,7 +1076,7 @@
apply (rule trans_trancl)
done
-lemma part_ord_multirel: "part_ord(A,r) ==> part_ord(Mult(A), multirel(A, r))"
+lemma part_ord_multirel: "part_ord(A,r) \<Longrightarrow> part_ord(Mult(A), multirel(A, r))"
apply (simp (no_asm) add: part_ord_def)
apply (blast intro: irrefl_on_multirel trans_on_multirel)
done
@@ -1084,7 +1084,7 @@
(** Monotonicity of multiset union **)
lemma munion_multirel1_mono:
-"[|<M,N> \<in> multirel1(A, r); K \<in> Mult(A) |] ==> <K +# M, K +# N> \<in> multirel1(A, r)"
+"\<lbrakk><M,N> \<in> multirel1(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <K +# M, K +# N> \<in> multirel1(A, r)"
apply (frule multirel1_type [THEN subsetD])
apply (auto simp add: multirel1_iff Mult_iff_multiset)
apply (rule_tac x = a in exI)
@@ -1096,7 +1096,7 @@
done
lemma munion_multirel_mono2:
- "[| <M, N> \<in> multirel(A, r); K \<in> Mult(A) |]==><K +# M, K +# N> \<in> multirel(A, r)"
+ "\<lbrakk><M, N> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk>\<Longrightarrow><K +# M, K +# N> \<in> multirel(A, r)"
apply (frule multirel_type [THEN subsetD])
apply (simp (no_asm_use) add: multirel_def)
apply clarify
@@ -1115,7 +1115,7 @@
done
lemma munion_multirel_mono1:
- "[|<M, N> \<in> multirel(A, r); K \<in> Mult(A)|] ==> <M +# K, N +# K> \<in> multirel(A, r)"
+ "\<lbrakk><M, N> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <M +# K, N +# K> \<in> multirel(A, r)"
apply (frule multirel_type [THEN subsetD])
apply (rule_tac P = "%x. <x,u> \<in> multirel(A, r)" for u in munion_commute [THEN subst])
apply (subst munion_commute [of N])
@@ -1124,8 +1124,8 @@
done
lemma munion_multirel_mono:
- "[|<M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)|]
- ==> <M +# N, K +# L> \<in> multirel(A, r)"
+ "\<lbrakk><M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)\<rbrakk>
+ \<Longrightarrow> <M +# N, K +# L> \<in> multirel(A, r)"
apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ")
prefer 2 apply (blast dest: multirel_type [THEN subsetD])
apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2)
@@ -1134,29 +1134,29 @@
subsection\<open>Ordinal Multisets\<close>
-(* A \<subseteq> B ==> field(Memrel(A)) \<subseteq> field(Memrel(B)) *)
+(* A \<subseteq> B \<Longrightarrow> field(Memrel(A)) \<subseteq> field(Memrel(B)) *)
lemmas field_Memrel_mono = Memrel_mono [THEN field_mono]
(*
-[| Aa \<subseteq> Ba; A \<subseteq> B |] ==>
+\<lbrakk>Aa \<subseteq> Ba; A \<subseteq> B\<rbrakk> \<Longrightarrow>
multirel(field(Memrel(Aa)), Memrel(A))\<subseteq> multirel(field(Memrel(Ba)), Memrel(B))
*)
lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono]
-lemma omultiset_is_multiset [simp]: "omultiset(M) ==> multiset(M)"
+lemma omultiset_is_multiset [simp]: "omultiset(M) \<Longrightarrow> multiset(M)"
apply (simp add: omultiset_def)
apply (auto simp add: Mult_iff_multiset)
done
-lemma munion_omultiset [simp]: "[| omultiset(M); omultiset(N) |] ==> omultiset(M +# N)"
+lemma munion_omultiset [simp]: "\<lbrakk>omultiset(M); omultiset(N)\<rbrakk> \<Longrightarrow> omultiset(M +# N)"
apply (simp add: omultiset_def, clarify)
apply (rule_tac x = "i \<union> ia" in exI)
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)
apply (blast intro: field_Memrel_mono)
done
-lemma mdiff_omultiset [simp]: "omultiset(M) ==> omultiset(M -# N)"
+lemma mdiff_omultiset [simp]: "omultiset(M) \<Longrightarrow> omultiset(M -# N)"
apply (simp add: omultiset_def, clarify)
apply (simp add: Mult_iff_multiset)
apply (rule_tac x = i in exI)
@@ -1165,7 +1165,7 @@
(** Proving that Memrel is a partial order **)
-lemma irrefl_Memrel: "Ord(i) ==> irrefl(field(Memrel(i)), Memrel(i))"
+lemma irrefl_Memrel: "Ord(i) \<Longrightarrow> irrefl(field(Memrel(i)), Memrel(i))"
apply (rule irreflI, clarify)
apply (subgoal_tac "Ord (x) ")
prefer 2 apply (blast intro: Ord_in_Ord)
@@ -1175,14 +1175,14 @@
lemma trans_iff_trans_on: "trans(r) \<longleftrightarrow> trans[field(r)](r)"
by (simp add: trans_on_def trans_def, auto)
-lemma part_ord_Memrel: "Ord(i) ==>part_ord(field(Memrel(i)), Memrel(i))"
+lemma part_ord_Memrel: "Ord(i) \<Longrightarrow>part_ord(field(Memrel(i)), Memrel(i))"
apply (simp add: part_ord_def)
apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym])
apply (blast intro: trans_Memrel irrefl_Memrel)
done
(*
- Ord(i) ==>
+ Ord(i) \<Longrightarrow>
part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i)))
*)
@@ -1190,18 +1190,18 @@
(*irreflexivity*)
-lemma mless_not_refl: "~(M <# M)"
+lemma mless_not_refl: "\<not>(M <# M)"
apply (simp add: mless_def, clarify)
apply (frule multirel_type [THEN subsetD])
apply (drule part_ord_mless)
apply (simp add: part_ord_def irrefl_def)
done
-(* N<N ==> R *)
+(* N<N \<Longrightarrow> R *)
lemmas mless_irrefl = mless_not_refl [THEN notE, elim!]
(*transitivity*)
-lemma mless_trans: "[| K <# M; M <# N |] ==> K <# N"
+lemma mless_trans: "\<lbrakk>K <# M; M <# N\<rbrakk> \<Longrightarrow> K <# N"
apply (simp add: mless_def, clarify)
apply (rule_tac x = "i \<union> ia" in exI)
apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD]
@@ -1210,27 +1210,27 @@
done
(*asymmetry*)
-lemma mless_not_sym: "M <# N ==> ~ N <# M"
+lemma mless_not_sym: "M <# N \<Longrightarrow> \<not> N <# M"
apply clarify
apply (rule mless_not_refl [THEN notE])
apply (erule mless_trans, assumption)
done
-lemma mless_asym: "[| M <# N; ~P ==> N <# M |] ==> P"
+lemma mless_asym: "\<lbrakk>M <# N; \<not>P \<Longrightarrow> N <# M\<rbrakk> \<Longrightarrow> P"
by (blast dest: mless_not_sym)
-lemma mle_refl [simp]: "omultiset(M) ==> M <#= M"
+lemma mle_refl [simp]: "omultiset(M) \<Longrightarrow> M <#= M"
by (simp add: mle_def)
(*anti-symmetry*)
lemma mle_antisym:
- "[| M <#= N; N <#= M |] ==> M = N"
+ "\<lbrakk>M <#= N; N <#= M\<rbrakk> \<Longrightarrow> M = N"
apply (simp add: mle_def)
apply (blast dest: mless_not_sym)
done
(*transitivity*)
-lemma mle_trans: "[| K <#= M; M <#= N |] ==> K <#= N"
+lemma mle_trans: "\<lbrakk>K <#= M; M <#= N\<rbrakk> \<Longrightarrow> K <#= N"
apply (simp add: mle_def)
apply (blast intro: mless_trans)
done
@@ -1240,7 +1240,7 @@
(** Monotonicity of mless **)
-lemma munion_less_mono2: "[| M <# N; omultiset(K) |] ==> K +# M <# K +# N"
+lemma munion_less_mono2: "\<lbrakk>M <# N; omultiset(K)\<rbrakk> \<Longrightarrow> K +# M <# K +# N"
apply (simp add: mless_def omultiset_def, clarify)
apply (rule_tac x = "i \<union> ia" in exI)
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)
@@ -1250,13 +1250,13 @@
apply (blast intro: field_Memrel_mono [THEN subsetD])
done
-lemma munion_less_mono1: "[| M <# N; omultiset(K) |] ==> M +# K <# N +# K"
+lemma munion_less_mono1: "\<lbrakk>M <# N; omultiset(K)\<rbrakk> \<Longrightarrow> M +# K <# N +# K"
by (force dest: munion_less_mono2 simp add: munion_commute)
-lemma mless_imp_omultiset: "M <# N ==> omultiset(M) & omultiset(N)"
+lemma mless_imp_omultiset: "M <# N \<Longrightarrow> omultiset(M) & omultiset(N)"
by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD])
-lemma munion_less_mono: "[| M <# K; N <# L |] ==> M +# N <# K +# L"
+lemma munion_less_mono: "\<lbrakk>M <# K; N <# L\<rbrakk> \<Longrightarrow> M +# N <# K +# L"
apply (frule_tac M = M in mless_imp_omultiset)
apply (frule_tac M = N in mless_imp_omultiset)
apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans)
@@ -1264,10 +1264,10 @@
(* <#= *)
-lemma mle_imp_omultiset: "M <#= N ==> omultiset(M) & omultiset(N)"
+lemma mle_imp_omultiset: "M <#= N \<Longrightarrow> omultiset(M) & omultiset(N)"
by (auto simp add: mle_def mless_imp_omultiset)
-lemma mle_mono: "[| M <#= K; N <#= L |] ==> M +# N <#= K +# L"
+lemma mle_mono: "\<lbrakk>M <#= K; N <#= L\<rbrakk> \<Longrightarrow> M +# N <#= K +# L"
apply (frule_tac M = M in mle_imp_omultiset)
apply (frule_tac M = N in mle_imp_omultiset)
apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono)
@@ -1276,7 +1276,7 @@
lemma omultiset_0 [iff]: "omultiset(0)"
by (auto simp add: omultiset_def Mult_iff_multiset)
-lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M"
+lemma empty_leI [simp]: "omultiset(M) \<Longrightarrow> 0 <#= M"
apply (simp add: mle_def mless_def)
apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ")
prefer 2 apply (simp add: omultiset_def)
@@ -1286,7 +1286,7 @@
apply (auto simp add: Mult_iff_multiset)
done
-lemma munion_upper1: "[| omultiset(M); omultiset(N) |] ==> M <#= M +# N"
+lemma munion_upper1: "\<lbrakk>omultiset(M); omultiset(N)\<rbrakk> \<Longrightarrow> M <#= M +# N"
apply (subgoal_tac "M +# 0 <#= M +# N")
apply (rule_tac [2] mle_mono, auto)
done
--- a/src/ZF/Induct/Mutil.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Mutil.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,21 +19,21 @@
inductive
domains "domino" \<subseteq> "Pow(nat \<times> nat)"
intros
- horiz: "[| i \<in> nat; j \<in> nat |] ==> {<i,j>, <i,succ(j)>} \<in> domino"
- vertl: "[| i \<in> nat; j \<in> nat |] ==> {<i,j>, <succ(i),j>} \<in> domino"
+ horiz: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> {<i,j>, <i,succ(j)>} \<in> domino"
+ vertl: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> {<i,j>, <succ(i),j>} \<in> domino"
type_intros empty_subsetI cons_subsetI PowI SigmaI nat_succI
inductive
domains "tiling(A)" \<subseteq> "Pow(\<Union>(A))"
intros
empty: "0 \<in> tiling(A)"
- Un: "[| a \<in> A; t \<in> tiling(A); a \<inter> t = 0 |] ==> a \<union> t \<in> tiling(A)"
+ Un: "\<lbrakk>a \<in> A; t \<in> tiling(A); a \<inter> t = 0\<rbrakk> \<Longrightarrow> a \<union> t \<in> tiling(A)"
type_intros empty_subsetI Union_upper Un_least PowI
type_elims PowD [elim_format]
definition
evnodd :: "[i, i] => i" where
- "evnodd(A,b) == {z \<in> A. \<exists>i j. z = <i,j> \<and> (i #+ j) mod 2 = b}"
+ "evnodd(A,b) \<equiv> {z \<in> A. \<exists>i j. z = <i,j> \<and> (i #+ j) mod 2 = b}"
subsection \<open>Basic properties of evnodd\<close>
@@ -44,7 +44,7 @@
lemma evnodd_subset: "evnodd(A, b) \<subseteq> A"
by (unfold evnodd_def) blast
-lemma Finite_evnodd: "Finite(X) ==> Finite(evnodd(X,b))"
+lemma Finite_evnodd: "Finite(X) \<Longrightarrow> Finite(evnodd(X,b))"
by (rule lepoll_Finite, rule subset_imp_lepoll, rule evnodd_subset)
lemma evnodd_Un: "evnodd(A \<union> B, b) = evnodd(A,b) \<union> evnodd(B,b)"
@@ -64,11 +64,11 @@
subsection \<open>Dominoes\<close>
-lemma domino_Finite: "d \<in> domino ==> Finite(d)"
+lemma domino_Finite: "d \<in> domino \<Longrightarrow> Finite(d)"
by (blast intro!: Finite_cons Finite_0 elim: domino.cases)
lemma domino_singleton:
- "[| d \<in> domino; b<2 |] ==> \<exists>i' j'. evnodd(d,b) = {<i',j'>}"
+ "\<lbrakk>d \<in> domino; b<2\<rbrakk> \<Longrightarrow> \<exists>i' j'. evnodd(d,b) = {<i',j'>}"
apply (erule domino.cases)
apply (rule_tac [2] k1 = "i#+j" in mod2_cases [THEN disjE])
apply (rule_tac k1 = "i#+j" in mod2_cases [THEN disjE])
@@ -83,20 +83,20 @@
text \<open>The union of two disjoint tilings is a tiling\<close>
lemma tiling_UnI:
- "t \<in> tiling(A) ==> u \<in> tiling(A) ==> t \<inter> u = 0 ==> t \<union> u \<in> tiling(A)"
+ "t \<in> tiling(A) \<Longrightarrow> u \<in> tiling(A) \<Longrightarrow> t \<inter> u = 0 \<Longrightarrow> t \<union> u \<in> tiling(A)"
apply (induct set: tiling)
apply (simp add: tiling.intros)
apply (simp add: Un_assoc subset_empty_iff [THEN iff_sym])
apply (blast intro: tiling.intros)
done
-lemma tiling_domino_Finite: "t \<in> tiling(domino) ==> Finite(t)"
+lemma tiling_domino_Finite: "t \<in> tiling(domino) \<Longrightarrow> Finite(t)"
apply (induct set: tiling)
apply (rule Finite_0)
apply (blast intro!: Finite_Un intro: domino_Finite)
done
-lemma tiling_domino_0_1: "t \<in> tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|"
+lemma tiling_domino_0_1: "t \<in> tiling(domino) \<Longrightarrow> |evnodd(t,0)| = |evnodd(t,1)|"
apply (induct set: tiling)
apply (simp add: evnodd_def)
apply (rule_tac b1 = 0 in domino_singleton [THEN exE])
@@ -115,7 +115,7 @@
done
lemma dominoes_tile_row:
- "[| i \<in> nat; n \<in> nat |] ==> {i} * (n #+ n) \<in> tiling(domino)"
+ "\<lbrakk>i \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> {i} * (n #+ n) \<in> tiling(domino)"
apply (induct_tac n)
apply (simp add: tiling.intros)
apply (simp add: Un_assoc [symmetric] Sigma_succ2)
@@ -131,20 +131,20 @@
done
lemma dominoes_tile_matrix:
- "[| m \<in> nat; n \<in> nat |] ==> m * (n #+ n) \<in> tiling(domino)"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> m * (n #+ n) \<in> tiling(domino)"
apply (induct_tac m)
apply (simp add: tiling.intros)
apply (simp add: Sigma_succ1)
apply (blast intro: tiling_UnI dominoes_tile_row elim: mem_irrefl)
done
-lemma eq_lt_E: "[| x=y; x<y |] ==> P"
+lemma eq_lt_E: "\<lbrakk>x=y; x<y\<rbrakk> \<Longrightarrow> P"
by auto
-theorem mutil_not_tiling: "[| m \<in> nat; n \<in> nat;
+theorem mutil_not_tiling: "\<lbrakk>m \<in> nat; n \<in> nat;
t = (succ(m)#+succ(m))*(succ(n)#+succ(n));
- t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>} |]
- ==> t' \<notin> tiling(domino)"
+ t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>}\<rbrakk>
+ \<Longrightarrow> t' \<notin> tiling(domino)"
apply (rule notI)
apply (drule tiling_domino_0_1)
apply (erule_tac x = "|A|" for A in eq_lt_E)
--- a/src/ZF/Induct/Ntree.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Ntree.thy Tue Sep 27 16:51:35 2022 +0100
@@ -32,12 +32,12 @@
definition
ntree_rec :: "[[i, i, i] => i, i] => i" where
- "ntree_rec(b) ==
+ "ntree_rec(b) \<equiv>
Vrecursor(\<lambda>pr. ntree_case(\<lambda>x h. b(x, h, \<lambda>i \<in> domain(h). pr`(h`i))))"
definition
ntree_copy :: "i => i" where
- "ntree_copy(z) == ntree_rec(\<lambda>x h r. Branch(x,r), z)"
+ "ntree_copy(z) \<equiv> ntree_rec(\<lambda>x h r. Branch(x,r), z)"
text \<open>
@@ -50,8 +50,8 @@
lemma ntree_induct [consumes 1, case_names Branch, induct set: ntree]:
assumes t: "t \<in> ntree(A)"
- and step: "!!x n h. [| x \<in> A; n \<in> nat; h \<in> n -> ntree(A); \<forall>i \<in> n. P(h`i)
- |] ==> P(Branch(x,h))"
+ and step: "\<And>x n h. \<lbrakk>x \<in> A; n \<in> nat; h \<in> n -> ntree(A); \<forall>i \<in> n. P(h`i)
+\<rbrakk> \<Longrightarrow> P(Branch(x,h))"
shows "P(t)"
\<comment> \<open>A nicer induction rule than the standard one.\<close>
using t
@@ -66,7 +66,7 @@
assumes t: "t \<in> ntree(A)"
and f: "f \<in> ntree(A)->B"
and g: "g \<in> ntree(A)->B"
- and step: "!!x n h. [| x \<in> A; n \<in> nat; h \<in> n -> ntree(A); f O h = g O h |] ==>
+ and step: "\<And>x n h. \<lbrakk>x \<in> A; n \<in> nat; h \<in> n -> ntree(A); f O h = g O h\<rbrakk> \<Longrightarrow>
f ` Branch(x,h) = g ` Branch(x,h)"
shows "f`t=g`t"
\<comment> \<open>Induction on \<^term>\<open>ntree(A)\<close> to prove an equation\<close>
@@ -84,7 +84,7 @@
type definitions.
\<close>
-lemma ntree_mono: "A \<subseteq> B ==> ntree(A) \<subseteq> ntree(B)"
+lemma ntree_mono: "A \<subseteq> B \<Longrightarrow> ntree(A) \<subseteq> ntree(B)"
apply (unfold ntree.defs)
apply (rule lfp_mono)
apply (rule ntree.bnd_mono)+
@@ -99,7 +99,7 @@
apply (blast intro: Pair_in_univ nat_fun_univ [THEN subsetD])
done
-lemma ntree_subset_univ: "A \<subseteq> univ(B) ==> ntree(A) \<subseteq> univ(B)"
+lemma ntree_subset_univ: "A \<subseteq> univ(B) \<Longrightarrow> ntree(A) \<subseteq> univ(B)"
by (rule subset_trans [OF ntree_mono ntree_univ])
@@ -108,18 +108,18 @@
\<close>
lemma ntree_rec_Branch:
- "function(h) ==>
+ "function(h) \<Longrightarrow>
ntree_rec(b, Branch(x,h)) = b(x, h, \<lambda>i \<in> domain(h). ntree_rec(b, h`i))"
apply (rule ntree_rec_def [THEN def_Vrecursor, THEN trans])
apply (simp add: ntree.con_defs rank_pair2 [THEN [2] lt_trans] rank_apply)
done
lemma ntree_copy_Branch [simp]:
- "function(h) ==>
+ "function(h) \<Longrightarrow>
ntree_copy (Branch(x, h)) = Branch(x, \<lambda>i \<in> domain(h). ntree_copy (h`i))"
by (simp add: ntree_copy_def ntree_rec_Branch)
-lemma ntree_copy_is_ident: "z \<in> ntree(A) ==> ntree_copy(z) = z"
+lemma ntree_copy_is_ident: "z \<in> ntree(A) \<Longrightarrow> ntree_copy(z) = z"
by (induct z set: ntree)
(auto simp add: domain_of_fun Pi_Collect_iff fun_is_function)
@@ -134,9 +134,9 @@
lemma maptree_induct [consumes 1, induct set: maptree]:
assumes t: "t \<in> maptree(A)"
- and step: "!!x n h. [| x \<in> A; h \<in> maptree(A) -||> maptree(A);
+ and step: "\<And>x n h. \<lbrakk>x \<in> A; h \<in> maptree(A) -||> maptree(A);
\<forall>y \<in> field(h). P(y)
- |] ==> P(Sons(x,h))"
+\<rbrakk> \<Longrightarrow> P(Sons(x,h))"
shows "P(t)"
\<comment> \<open>A nicer induction rule than the standard one.\<close>
using t
@@ -159,8 +159,8 @@
lemma maptree2_induct [consumes 1, induct set: maptree2]:
assumes t: "t \<in> maptree2(A, B)"
- and step: "!!x n h. [| x \<in> A; h \<in> B -||> maptree2(A,B); \<forall>y \<in> range(h). P(y)
- |] ==> P(Sons2(x,h))"
+ and step: "\<And>x n h. \<lbrakk>x \<in> A; h \<in> B -||> maptree2(A,B); \<forall>y \<in> range(h). P(y)
+\<rbrakk> \<Longrightarrow> P(Sons2(x,h))"
shows "P(t)"
using t
apply induct
--- a/src/ZF/Induct/Primrec.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Primrec.thy Tue Sep 27 16:51:35 2022 +0100
@@ -18,23 +18,23 @@
definition
SC :: "i" where
- "SC == \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. succ(x), l)"
+ "SC \<equiv> \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. succ(x), l)"
definition
CONSTANT :: "i=>i" where
- "CONSTANT(k) == \<lambda>l \<in> list(nat). k"
+ "CONSTANT(k) \<equiv> \<lambda>l \<in> list(nat). k"
definition
PROJ :: "i=>i" where
- "PROJ(i) == \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. x, drop(i,l))"
+ "PROJ(i) \<equiv> \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. x, drop(i,l))"
definition
COMP :: "[i,i]=>i" where
- "COMP(g,fs) == \<lambda>l \<in> list(nat). g ` map(\<lambda>f. f`l, fs)"
+ "COMP(g,fs) \<equiv> \<lambda>l \<in> list(nat). g ` map(\<lambda>f. f`l, fs)"
definition
PREC :: "[i,i]=>i" where
- "PREC(f,g) ==
+ "PREC(f,g) \<equiv>
\<lambda>l \<in> list(nat). list_case(0,
\<lambda>x xs. rec(x, f`xs, \<lambda>y r. g ` Cons(r, Cons(y, xs))), l)"
\<comment> \<open>Note that \<open>g\<close> is applied first to \<^term>\<open>PREC(f,g)`y\<close> and then to \<open>y\<close>!\<close>
@@ -47,31 +47,31 @@
abbreviation
ack :: "[i,i]=>i" where
- "ack(x,y) == ACK(x) ` [y]"
+ "ack(x,y) \<equiv> ACK(x) ` [y]"
text \<open>
\medskip Useful special cases of evaluation.
\<close>
-lemma SC: "[| x \<in> nat; l \<in> list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)"
+lemma SC: "\<lbrakk>x \<in> nat; l \<in> list(nat)\<rbrakk> \<Longrightarrow> SC ` (Cons(x,l)) = succ(x)"
by (simp add: SC_def)
-lemma CONSTANT: "l \<in> list(nat) ==> CONSTANT(k) ` l = k"
+lemma CONSTANT: "l \<in> list(nat) \<Longrightarrow> CONSTANT(k) ` l = k"
by (simp add: CONSTANT_def)
-lemma PROJ_0: "[| x \<in> nat; l \<in> list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x"
+lemma PROJ_0: "\<lbrakk>x \<in> nat; l \<in> list(nat)\<rbrakk> \<Longrightarrow> PROJ(0) ` (Cons(x,l)) = x"
by (simp add: PROJ_def)
-lemma COMP_1: "l \<in> list(nat) ==> COMP(g,[f]) ` l = g` [f`l]"
+lemma COMP_1: "l \<in> list(nat) \<Longrightarrow> COMP(g,[f]) ` l = g` [f`l]"
by (simp add: COMP_def)
-lemma PREC_0: "l \<in> list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l"
+lemma PREC_0: "l \<in> list(nat) \<Longrightarrow> PREC(f,g) ` (Cons(0,l)) = f`l"
by (simp add: PREC_def)
lemma PREC_succ:
- "[| x \<in> nat; l \<in> list(nat) |]
- ==> PREC(f,g) ` (Cons(succ(x),l)) =
+ "\<lbrakk>x \<in> nat; l \<in> list(nat)\<rbrakk>
+ \<Longrightarrow> PREC(f,g) ` (Cons(succ(x),l)) =
g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))"
by (simp add: PREC_def)
@@ -85,10 +85,10 @@
domains prim_rec \<subseteq> "list(nat)->nat"
intros
"SC \<in> prim_rec"
- "k \<in> nat ==> CONSTANT(k) \<in> prim_rec"
- "i \<in> nat ==> PROJ(i) \<in> prim_rec"
- "[| g \<in> prim_rec; fs\<in>list(prim_rec) |] ==> COMP(g,fs) \<in> prim_rec"
- "[| f \<in> prim_rec; g \<in> prim_rec |] ==> PREC(f,g) \<in> prim_rec"
+ "k \<in> nat \<Longrightarrow> CONSTANT(k) \<in> prim_rec"
+ "i \<in> nat \<Longrightarrow> PROJ(i) \<in> prim_rec"
+ "\<lbrakk>g \<in> prim_rec; fs\<in>list(prim_rec)\<rbrakk> \<Longrightarrow> COMP(g,fs) \<in> prim_rec"
+ "\<lbrakk>f \<in> prim_rec; g \<in> prim_rec\<rbrakk> \<Longrightarrow> PREC(f,g) \<in> prim_rec"
monos list_mono
con_defs SC_def CONSTANT_def PROJ_def COMP_def PREC_def
type_intros nat_typechecks list.intros
@@ -96,7 +96,7 @@
apply_type rec_type
-lemma prim_rec_into_fun [TC]: "c \<in> prim_rec ==> c \<in> list(nat) -> nat"
+lemma prim_rec_into_fun [TC]: "c \<in> prim_rec \<Longrightarrow> c \<in> list(nat) -> nat"
by (erule subsetD [OF prim_rec.dom_subset])
lemmas [TC] = apply_type [OF prim_rec_into_fun]
@@ -105,16 +105,16 @@
declare nat_into_Ord [TC]
declare rec_type [TC]
-lemma ACK_in_prim_rec [TC]: "i \<in> nat ==> ACK(i) \<in> prim_rec"
+lemma ACK_in_prim_rec [TC]: "i \<in> nat \<Longrightarrow> ACK(i) \<in> prim_rec"
by (induct set: nat) simp_all
-lemma ack_type [TC]: "[| i \<in> nat; j \<in> nat |] ==> ack(i,j) \<in> nat"
+lemma ack_type [TC]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> ack(i,j) \<in> nat"
by auto
subsection \<open>Ackermann's function cases\<close>
-lemma ack_0: "j \<in> nat ==> ack(0,j) = succ(j)"
+lemma ack_0: "j \<in> nat \<Longrightarrow> ack(0,j) = succ(j)"
\<comment> \<open>PROPERTY A 1\<close>
by (simp add: SC)
@@ -123,7 +123,7 @@
by (simp add: CONSTANT PREC_0)
lemma ack_succ_succ:
- "[| i\<in>nat; j\<in>nat |] ==> ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))"
+ "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))"
\<comment> \<open>PROPERTY A 3\<close>
by (simp add: CONSTANT PREC_succ COMP_1 PROJ_0)
@@ -131,7 +131,7 @@
and [simp del] = ACK.simps
-lemma lt_ack2: "i \<in> nat ==> j \<in> nat ==> j < ack(i,j)"
+lemma lt_ack2: "i \<in> nat \<Longrightarrow> j \<in> nat \<Longrightarrow> j < ack(i,j)"
\<comment> \<open>PROPERTY A 4\<close>
apply (induct i arbitrary: j set: nat)
apply simp
@@ -141,11 +141,11 @@
apply auto
done
-lemma ack_lt_ack_succ2: "[|i\<in>nat; j\<in>nat|] ==> ack(i,j) < ack(i, succ(j))"
+lemma ack_lt_ack_succ2: "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> ack(i,j) < ack(i, succ(j))"
\<comment> \<open>PROPERTY A 5-, the single-step lemma\<close>
by (induct set: nat) (simp_all add: lt_ack2)
-lemma ack_lt_mono2: "[| j<k; i \<in> nat; k \<in> nat |] ==> ack(i,j) < ack(i,k)"
+lemma ack_lt_mono2: "\<lbrakk>j<k; i \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> ack(i,j) < ack(i,k)"
\<comment> \<open>PROPERTY A 5, monotonicity for \<open><\<close>\<close>
apply (frule lt_nat_in_nat, assumption)
apply (erule succ_lt_induct)
@@ -154,14 +154,14 @@
apply (auto intro: ack_lt_ack_succ2)
done
-lemma ack_le_mono2: "[|j\<le>k; i\<in>nat; k\<in>nat|] ==> ack(i,j) \<le> ack(i,k)"
+lemma ack_le_mono2: "\<lbrakk>j\<le>k; i\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> ack(i,j) \<le> ack(i,k)"
\<comment> \<open>PROPERTY A 5', monotonicity for \<open>\<le>\<close>\<close>
apply (rule_tac f = "\<lambda>j. ack (i,j) " in Ord_lt_mono_imp_le_mono)
apply (assumption | rule ack_lt_mono2 ack_type [THEN nat_into_Ord])+
done
lemma ack2_le_ack1:
- "[| i\<in>nat; j\<in>nat |] ==> ack(i, succ(j)) \<le> ack(succ(i), j)"
+ "\<lbrakk>i\<in>nat; j\<in>nat\<rbrakk> \<Longrightarrow> ack(i, succ(j)) \<le> ack(succ(i), j)"
\<comment> \<open>PROPERTY A 6\<close>
apply (induct_tac j)
apply simp_all
@@ -170,14 +170,14 @@
apply auto
done
-lemma ack_lt_ack_succ1: "[| i \<in> nat; j \<in> nat |] ==> ack(i,j) < ack(succ(i),j)"
+lemma ack_lt_ack_succ1: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> ack(i,j) < ack(succ(i),j)"
\<comment> \<open>PROPERTY A 7-, the single-step lemma\<close>
apply (rule ack_lt_mono2 [THEN lt_trans2])
apply (rule_tac [4] ack2_le_ack1)
apply auto
done
-lemma ack_lt_mono1: "[| i<j; j \<in> nat; k \<in> nat |] ==> ack(i,k) < ack(j,k)"
+lemma ack_lt_mono1: "\<lbrakk>i<j; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> ack(i,k) < ack(j,k)"
\<comment> \<open>PROPERTY A 7, monotonicity for \<open><\<close>\<close>
apply (frule lt_nat_in_nat, assumption)
apply (erule succ_lt_induct)
@@ -186,23 +186,23 @@
apply (auto intro: ack_lt_ack_succ1)
done
-lemma ack_le_mono1: "[| i\<le>j; j \<in> nat; k \<in> nat |] ==> ack(i,k) \<le> ack(j,k)"
+lemma ack_le_mono1: "\<lbrakk>i\<le>j; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> ack(i,k) \<le> ack(j,k)"
\<comment> \<open>PROPERTY A 7', monotonicity for \<open>\<le>\<close>\<close>
apply (rule_tac f = "\<lambda>j. ack (j,k) " in Ord_lt_mono_imp_le_mono)
apply (assumption | rule ack_lt_mono1 ack_type [THEN nat_into_Ord])+
done
-lemma ack_1: "j \<in> nat ==> ack(1,j) = succ(succ(j))"
+lemma ack_1: "j \<in> nat \<Longrightarrow> ack(1,j) = succ(succ(j))"
\<comment> \<open>PROPERTY A 8\<close>
by (induct set: nat) simp_all
-lemma ack_2: "j \<in> nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))"
+lemma ack_2: "j \<in> nat \<Longrightarrow> ack(succ(1),j) = succ(succ(succ(j#+j)))"
\<comment> \<open>PROPERTY A 9\<close>
by (induct set: nat) (simp_all add: ack_1)
lemma ack_nest_bound:
- "[| i1 \<in> nat; i2 \<in> nat; j \<in> nat |]
- ==> ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)"
+ "\<lbrakk>i1 \<in> nat; i2 \<in> nat; j \<in> nat\<rbrakk>
+ \<Longrightarrow> ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)"
\<comment> \<open>PROPERTY A 10\<close>
apply (rule lt_trans2 [OF _ ack2_le_ack1])
apply simp
@@ -212,8 +212,8 @@
done
lemma ack_add_bound:
- "[| i1 \<in> nat; i2 \<in> nat; j \<in> nat |]
- ==> ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)"
+ "\<lbrakk>i1 \<in> nat; i2 \<in> nat; j \<in> nat\<rbrakk>
+ \<Longrightarrow> ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)"
\<comment> \<open>PROPERTY A 11\<close>
apply (rule_tac j = "ack (succ (1), ack (i1 #+ i2, j))" in lt_trans)
apply (simp add: ack_2)
@@ -223,8 +223,8 @@
done
lemma ack_add_bound2:
- "[| i < ack(k,j); j \<in> nat; k \<in> nat |]
- ==> i#+j < ack(succ(succ(succ(succ(k)))), j)"
+ "\<lbrakk>i < ack(k,j); j \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> i#+j < ack(succ(succ(succ(succ(k)))), j)"
\<comment> \<open>PROPERTY A 12.\<close>
\<comment> \<open>Article uses existential quantifier but the ALF proof used \<^term>\<open>k#+#4\<close>.\<close>
\<comment> \<open>Quantified version must be nested \<open>\<exists>k'. \<forall>i,j \<dots>\<close>.\<close>
@@ -239,14 +239,14 @@
declare list_add_type [simp]
-lemma SC_case: "l \<in> list(nat) ==> SC ` l < ack(1, list_add(l))"
+lemma SC_case: "l \<in> list(nat) \<Longrightarrow> SC ` l < ack(1, list_add(l))"
apply (unfold SC_def)
apply (erule list.cases)
apply (simp add: succ_iff)
apply (simp add: ack_1 add_le_self)
done
-lemma lt_ack1: "[| i \<in> nat; j \<in> nat |] ==> i < ack(i,j)"
+lemma lt_ack1: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> i < ack(i,j)"
\<comment> \<open>PROPERTY A 4'? Extra lemma needed for \<open>CONSTANT\<close> case, constant functions.\<close>
apply (induct_tac i)
apply (simp add: nat_0_le)
@@ -255,11 +255,11 @@
done
lemma CONSTANT_case:
- "[| l \<in> list(nat); k \<in> nat |] ==> CONSTANT(k) ` l < ack(k, list_add(l))"
+ "\<lbrakk>l \<in> list(nat); k \<in> nat\<rbrakk> \<Longrightarrow> CONSTANT(k) ` l < ack(k, list_add(l))"
by (simp add: CONSTANT_def lt_ack1)
lemma PROJ_case [rule_format]:
- "l \<in> list(nat) ==> \<forall>i \<in> nat. PROJ(i) ` l < ack(0, list_add(l))"
+ "l \<in> list(nat) \<Longrightarrow> \<forall>i \<in> nat. PROJ(i) ` l < ack(0, list_add(l))"
apply (unfold PROJ_def)
apply simp
apply (erule list.induct)
@@ -280,7 +280,7 @@
lemma COMP_map_lemma:
"fs \<in> list({f \<in> prim_rec. \<exists>kf \<in> nat. \<forall>l \<in> list(nat). f`l < ack(kf, list_add(l))})
- ==> \<exists>k \<in> nat. \<forall>l \<in> list(nat).
+ \<Longrightarrow> \<exists>k \<in> nat. \<forall>l \<in> list(nat).
list_add(map(\<lambda>f. f ` l, fs)) < ack(k, list_add(l))"
apply (induct set: list)
apply (rule_tac x = 0 in bexI)
@@ -294,12 +294,12 @@
done
lemma COMP_case:
- "[| kg\<in>nat;
+ "\<lbrakk>kg\<in>nat;
\<forall>l \<in> list(nat). g`l < ack(kg, list_add(l));
fs \<in> list({f \<in> prim_rec .
\<exists>kf \<in> nat. \<forall>l \<in> list(nat).
- f`l < ack(kf, list_add(l))}) |]
- ==> \<exists>k \<in> nat. \<forall>l \<in> list(nat). COMP(g,fs)`l < ack(k, list_add(l))"
+ f`l < ack(kf, list_add(l))})\<rbrakk>
+ \<Longrightarrow> \<exists>k \<in> nat. \<forall>l \<in> list(nat). COMP(g,fs)`l < ack(k, list_add(l))"
apply (simp add: COMP_def)
apply (frule list_CollectD)
apply (erule COMP_map_lemma [THEN bexE])
@@ -316,12 +316,12 @@
\<close>
lemma PREC_case_lemma:
- "[| \<forall>l \<in> list(nat). f`l #+ list_add(l) < ack(kf, list_add(l));
+ "\<lbrakk>\<forall>l \<in> list(nat). f`l #+ list_add(l) < ack(kf, list_add(l));
\<forall>l \<in> list(nat). g`l #+ list_add(l) < ack(kg, list_add(l));
f \<in> prim_rec; kf\<in>nat;
g \<in> prim_rec; kg\<in>nat;
- l \<in> list(nat) |]
- ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))"
+ l \<in> list(nat)\<rbrakk>
+ \<Longrightarrow> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))"
apply (unfold PREC_def)
apply (erule list.cases)
apply (simp add: lt_trans [OF nat_le_refl lt_ack2])
@@ -347,11 +347,11 @@
done
lemma PREC_case:
- "[| f \<in> prim_rec; kf\<in>nat;
+ "\<lbrakk>f \<in> prim_rec; kf\<in>nat;
g \<in> prim_rec; kg\<in>nat;
\<forall>l \<in> list(nat). f`l < ack(kf, list_add(l));
- \<forall>l \<in> list(nat). g`l < ack(kg, list_add(l)) |]
- ==> \<exists>k \<in> nat. \<forall>l \<in> list(nat). PREC(f,g)`l< ack(k, list_add(l))"
+ \<forall>l \<in> list(nat). g`l < ack(kg, list_add(l))\<rbrakk>
+ \<Longrightarrow> \<exists>k \<in> nat. \<forall>l \<in> list(nat). PREC(f,g)`l< ack(k, list_add(l))"
apply (rule ballI [THEN bexI])
apply (rule lt_trans1 [OF add_le_self PREC_case_lemma])
apply typecheck
@@ -359,7 +359,7 @@
done
lemma ack_bounds_prim_rec:
- "f \<in> prim_rec ==> \<exists>k \<in> nat. \<forall>l \<in> list(nat). f`l < ack(k, list_add(l))"
+ "f \<in> prim_rec \<Longrightarrow> \<exists>k \<in> nat. \<forall>l \<in> list(nat). f`l < ack(k, list_add(l))"
apply (induct set: prim_rec)
apply (auto intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case)
done
--- a/src/ZF/Induct/PropLog.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/PropLog.thy Tue Sep 27 16:51:35 2022 +0100
@@ -36,17 +36,17 @@
abbreviation
thms_syntax :: "[i,i] => o" (infixl \<open>|-\<close> 50)
- where "H |- p == p \<in> thms(H)"
+ where "H |- p \<equiv> p \<in> thms(H)"
inductive
domains "thms(H)" \<subseteq> "propn"
intros
- H: "[| p \<in> H; p \<in> propn |] ==> H |- p"
- K: "[| p \<in> propn; q \<in> propn |] ==> H |- p=>q=>p"
- S: "[| p \<in> propn; q \<in> propn; r \<in> propn |]
- ==> H |- (p=>q=>r) => (p=>q) => p=>r"
- DN: "p \<in> propn ==> H |- ((p=>Fls) => Fls) => p"
- MP: "[| H |- p=>q; H |- p; p \<in> propn; q \<in> propn |] ==> H |- q"
+ H: "\<lbrakk>p \<in> H; p \<in> propn\<rbrakk> \<Longrightarrow> H |- p"
+ K: "\<lbrakk>p \<in> propn; q \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q=>p"
+ S: "\<lbrakk>p \<in> propn; q \<in> propn; r \<in> propn\<rbrakk>
+ \<Longrightarrow> H |- (p=>q=>r) => (p=>q) => p=>r"
+ DN: "p \<in> propn \<Longrightarrow> H |- ((p=>Fls) => Fls) => p"
+ MP: "\<lbrakk>H |- p=>q; H |- p; p \<in> propn; q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
type_intros "propn.intros"
declare propn.intros [simp]
@@ -65,7 +65,7 @@
definition
is_true :: "[i,i] => o" where
- "is_true(p,t) == is_true_fun(p,t) = 1"
+ "is_true(p,t) \<equiv> is_true_fun(p,t) = 1"
\<comment> \<open>this definition is required since predicates can't be recursive\<close>
lemma is_true_Fls [simp]: "is_true(Fls,t) \<longleftrightarrow> False"
@@ -87,7 +87,7 @@
definition
logcon :: "[i,i] => o" (infixl \<open>|=\<close> 50) where
- "H |= p == \<forall>t. (\<forall>q \<in> H. is_true(q,t)) \<longrightarrow> is_true(p,t)"
+ "H |= p \<equiv> \<forall>t. (\<forall>q \<in> H. is_true(q,t)) \<longrightarrow> is_true(p,t)"
text \<open>
@@ -106,7 +106,7 @@
subsection \<open>Proof theory of propositional logic\<close>
-lemma thms_mono: "G \<subseteq> H ==> thms(G) \<subseteq> thms(H)"
+lemma thms_mono: "G \<subseteq> H \<Longrightarrow> thms(G) \<subseteq> thms(H)"
apply (unfold thms.defs)
apply (rule lfp_mono)
apply (rule thms.bnd_mono)+
@@ -117,13 +117,13 @@
inductive_cases ImpE: "p=>q \<in> propn"
-lemma thms_MP: "[| H |- p=>q; H |- p |] ==> H |- q"
+lemma thms_MP: "\<lbrakk>H |- p=>q; H |- p\<rbrakk> \<Longrightarrow> H |- q"
\<comment> \<open>Stronger Modus Ponens rule: no typechecking!\<close>
apply (rule thms.MP)
apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+
done
-lemma thms_I: "p \<in> propn ==> H |- p=>p"
+lemma thms_I: "p \<in> propn \<Longrightarrow> H |- p=>p"
\<comment> \<open>Rule is called \<open>I\<close> for Identity Combinator, not for Introduction.\<close>
apply (rule thms.S [THEN thms_MP, THEN thms_MP])
apply (rule_tac [5] thms.K)
@@ -134,23 +134,23 @@
subsubsection \<open>Weakening, left and right\<close>
-lemma weaken_left: "[| G \<subseteq> H; G|-p |] ==> H|-p"
+lemma weaken_left: "\<lbrakk>G \<subseteq> H; G|-p\<rbrakk> \<Longrightarrow> H|-p"
\<comment> \<open>Order of premises is convenient with \<open>THEN\<close>\<close>
by (erule thms_mono [THEN subsetD])
-lemma weaken_left_cons: "H |- p ==> cons(a,H) |- p"
+lemma weaken_left_cons: "H |- p \<Longrightarrow> cons(a,H) |- p"
by (erule subset_consI [THEN weaken_left])
lemmas weaken_left_Un1 = Un_upper1 [THEN weaken_left]
lemmas weaken_left_Un2 = Un_upper2 [THEN weaken_left]
-lemma weaken_right: "[| H |- q; p \<in> propn |] ==> H |- p=>q"
+lemma weaken_right: "\<lbrakk>H |- q; p \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q"
by (simp_all add: thms.K [THEN thms_MP] thms_in_pl)
subsubsection \<open>The deduction theorem\<close>
-theorem deduction: "[| cons(p,H) |- q; p \<in> propn |] ==> H |- p=>q"
+theorem deduction: "\<lbrakk>cons(p,H) |- q; p \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q"
apply (erule thms.induct)
apply (blast intro: thms_I thms.H [THEN weaken_right])
apply (blast intro: thms.K [THEN weaken_right])
@@ -162,24 +162,24 @@
subsubsection \<open>The cut rule\<close>
-lemma cut: "[| H|-p; cons(p,H) |- q |] ==> H |- q"
+lemma cut: "\<lbrakk>H|-p; cons(p,H) |- q\<rbrakk> \<Longrightarrow> H |- q"
apply (rule deduction [THEN thms_MP])
apply (simp_all add: thms_in_pl)
done
-lemma thms_FlsE: "[| H |- Fls; p \<in> propn |] ==> H |- p"
+lemma thms_FlsE: "\<lbrakk>H |- Fls; p \<in> propn\<rbrakk> \<Longrightarrow> H |- p"
apply (rule thms.DN [THEN thms_MP])
apply (rule_tac [2] weaken_right)
apply (simp_all add: propn.intros)
done
-lemma thms_notE: "[| H |- p=>Fls; H |- p; q \<in> propn |] ==> H |- q"
+lemma thms_notE: "\<lbrakk>H |- p=>Fls; H |- p; q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
by (erule thms_MP [THEN thms_FlsE])
subsubsection \<open>Soundness of the rules wrt truth-table semantics\<close>
-theorem soundness: "H |- p ==> H |= p"
+theorem soundness: "H |- p \<Longrightarrow> H |= p"
apply (unfold logcon_def)
apply (induct set: thms)
apply auto
@@ -190,14 +190,14 @@
subsubsection \<open>Towards the completeness proof\<close>
-lemma Fls_Imp: "[| H |- p=>Fls; q \<in> propn |] ==> H |- p=>q"
+lemma Fls_Imp: "\<lbrakk>H |- p=>Fls; q \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q"
apply (frule thms_in_pl)
apply (rule deduction)
apply (rule weaken_left_cons [THEN thms_notE])
apply (blast intro: thms.H elim: ImpE)+
done
-lemma Imp_Fls: "[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"
+lemma Imp_Fls: "\<lbrakk>H |- p; H |- q=>Fls\<rbrakk> \<Longrightarrow> H |- (p=>q)=>Fls"
apply (frule thms_in_pl)
apply (frule thms_in_pl [of concl: "q=>Fls"])
apply (rule deduction)
@@ -207,7 +207,7 @@
done
lemma hyps_thms_if:
- "p \<in> propn ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"
+ "p \<in> propn \<Longrightarrow> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"
\<comment> \<open>Typical example of strengthening the induction statement.\<close>
apply simp
apply (induct_tac p)
@@ -216,7 +216,7 @@
apply (blast intro: weaken_left_Un1 weaken_left_Un2 weaken_right Imp_Fls)+
done
-lemma logcon_thms_p: "[| p \<in> propn; 0 |= p |] ==> hyps(p,t) |- p"
+lemma logcon_thms_p: "\<lbrakk>p \<in> propn; 0 |= p\<rbrakk> \<Longrightarrow> hyps(p,t) |- p"
\<comment> \<open>Key lemma for completeness; yields a set of assumptions satisfying \<open>p\<close>\<close>
apply (drule hyps_thms_if)
apply (simp add: logcon_def)
@@ -234,14 +234,14 @@
\<close>
lemma thms_excluded_middle:
- "[| p \<in> propn; q \<in> propn |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"
+ "\<lbrakk>p \<in> propn; q \<in> propn\<rbrakk> \<Longrightarrow> H |- (p=>q) => ((p=>Fls)=>q) => q"
apply (rule deduction [THEN deduction])
apply (rule thms.DN [THEN thms_MP])
apply (best intro!: propn_SIs intro: propn_Is)+
done
lemma thms_excluded_middle_rule:
- "[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p \<in> propn |] ==> H |- q"
+ "\<lbrakk>cons(p,H) |- q; cons(p=>Fls,H) |- q; p \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
\<comment> \<open>Hard to prove directly because it requires cuts\<close>
apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP])
apply (blast intro!: propn_SIs intro: propn_Is)+
@@ -255,7 +255,7 @@
\<close>
lemma hyps_Diff:
- "p \<in> propn ==> hyps(p, t-{v}) \<subseteq> cons(#v=>Fls, hyps(p,t)-{#v})"
+ "p \<in> propn \<Longrightarrow> hyps(p, t-{v}) \<subseteq> cons(#v=>Fls, hyps(p,t)-{#v})"
by (induct set: propn) auto
text \<open>
@@ -264,7 +264,7 @@
\<close>
lemma hyps_cons:
- "p \<in> propn ==> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v=>Fls})"
+ "p \<in> propn \<Longrightarrow> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v=>Fls})"
by (induct set: propn) auto
text \<open>Two lemmas for use with \<open>weaken_left\<close>\<close>
@@ -281,7 +281,7 @@
\<^prop>\<open>hyps(p,t) \<in> Fin(hyps(p,0) \<union> hyps(p,nat))\<close>.
\<close>
-lemma hyps_finite: "p \<in> propn ==> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v=>Fls})"
+lemma hyps_finite: "p \<in> propn \<Longrightarrow> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v=>Fls})"
by (induct set: propn) auto
lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left]
@@ -292,7 +292,7 @@
\<close>
lemma completeness_0_lemma [rule_format]:
- "[| p \<in> propn; 0 |= p |] ==> \<forall>t. hyps(p,t) - hyps(p,t0) |- p"
+ "\<lbrakk>p \<in> propn; 0 |= p\<rbrakk> \<Longrightarrow> \<forall>t. hyps(p,t) - hyps(p,t0) |- p"
apply (frule hyps_finite)
apply (erule Fin_induct)
apply (simp add: logcon_thms_p Diff_0)
@@ -315,18 +315,18 @@
subsubsection \<open>Completeness theorem\<close>
-lemma completeness_0: "[| p \<in> propn; 0 |= p |] ==> 0 |- p"
+lemma completeness_0: "\<lbrakk>p \<in> propn; 0 |= p\<rbrakk> \<Longrightarrow> 0 |- p"
\<comment> \<open>The base case for completeness\<close>
apply (rule Diff_cancel [THEN subst])
apply (blast intro: completeness_0_lemma)
done
-lemma logcon_Imp: "[| cons(p,H) |= q |] ==> H |= p=>q"
+lemma logcon_Imp: "\<lbrakk>cons(p,H) |= q\<rbrakk> \<Longrightarrow> H |= p=>q"
\<comment> \<open>A semantic analogue of the Deduction Theorem\<close>
by (simp add: logcon_def)
lemma completeness:
- "H \<in> Fin(propn) ==> p \<in> propn \<Longrightarrow> H |= p \<Longrightarrow> H |- p"
+ "H \<in> Fin(propn) \<Longrightarrow> p \<in> propn \<Longrightarrow> H |= p \<Longrightarrow> H |- p"
apply (induct arbitrary: p set: Fin)
apply (safe intro!: completeness_0)
apply (rule weaken_left_cons [THEN thms_MP])
@@ -334,7 +334,7 @@
apply (blast intro: propn_Is)
done
-theorem thms_iff: "H \<in> Fin(propn) ==> H |- p \<longleftrightarrow> H |= p \<and> p \<in> propn"
+theorem thms_iff: "H \<in> Fin(propn) \<Longrightarrow> H |- p \<longleftrightarrow> H |= p \<and> p \<in> propn"
by (blast intro: soundness completeness thms_in_pl)
end
--- a/src/ZF/Induct/Rmap.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Rmap.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,12 +15,12 @@
intros
NilI: "<Nil,Nil> \<in> rmap(r)"
- ConsI: "[| <x,y>: r; <xs,ys> \<in> rmap(r) |]
- ==> <Cons(x,xs), Cons(y,ys)> \<in> rmap(r)"
+ ConsI: "\<lbrakk><x,y>: r; <xs,ys> \<in> rmap(r)\<rbrakk>
+ \<Longrightarrow> <Cons(x,xs), Cons(y,ys)> \<in> rmap(r)"
type_intros domainI rangeI list.intros
-lemma rmap_mono: "r \<subseteq> s ==> rmap(r) \<subseteq> rmap(s)"
+lemma rmap_mono: "r \<subseteq> s \<Longrightarrow> rmap(r) \<subseteq> rmap(s)"
apply (unfold rmap.defs)
apply (rule lfp_mono)
apply (rule rmap.bnd_mono)+
@@ -33,19 +33,19 @@
declare rmap.intros [intro]
-lemma rmap_rel_type: "r \<subseteq> A \<times> B ==> rmap(r) \<subseteq> list(A) \<times> list(B)"
+lemma rmap_rel_type: "r \<subseteq> A \<times> B \<Longrightarrow> rmap(r) \<subseteq> list(A) \<times> list(B)"
apply (rule rmap.dom_subset [THEN subset_trans])
apply (assumption |
rule domain_rel_subset range_rel_subset Sigma_mono list_mono)+
done
-lemma rmap_total: "A \<subseteq> domain(r) ==> list(A) \<subseteq> domain(rmap(r))"
+lemma rmap_total: "A \<subseteq> domain(r) \<Longrightarrow> list(A) \<subseteq> domain(rmap(r))"
apply (rule subsetI)
apply (erule list.induct)
apply blast+
done
-lemma rmap_functional: "function(r) ==> function(rmap(r))"
+lemma rmap_functional: "function(r) \<Longrightarrow> function(rmap(r))"
apply (unfold function_def)
apply (rule impI [THEN allI, THEN allI])
apply (erule rmap.induct)
@@ -57,14 +57,14 @@
as expected.
\<close>
-lemma rmap_fun_type: "f \<in> A->B ==> rmap(f): list(A)->list(B)"
+lemma rmap_fun_type: "f \<in> A->B \<Longrightarrow> rmap(f): list(A)->list(B)"
by (simp add: Pi_iff rmap_rel_type rmap_functional rmap_total)
lemma rmap_Nil: "rmap(f)`Nil = Nil"
by (unfold apply_def) blast
-lemma rmap_Cons: "[| f \<in> A->B; x \<in> A; xs: list(A) |]
- ==> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)"
+lemma rmap_Cons: "\<lbrakk>f \<in> A->B; x \<in> A; xs: list(A)\<rbrakk>
+ \<Longrightarrow> rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)"
by (blast intro: apply_equality apply_Pair rmap_fun_type rmap.intros)
end
--- a/src/ZF/Induct/Term.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Term.thy Tue Sep 27 16:51:35 2022 +0100
@@ -22,39 +22,39 @@
definition
term_rec :: "[i, [i, i, i] => i] => i" where
- "term_rec(t,d) ==
+ "term_rec(t,d) \<equiv>
Vrec(t, \<lambda>t g. term_case(\<lambda>x zs. d(x, zs, map(\<lambda>z. g`z, zs)), t))"
definition
term_map :: "[i => i, i] => i" where
- "term_map(f,t) == term_rec(t, \<lambda>x zs rs. Apply(f(x), rs))"
+ "term_map(f,t) \<equiv> term_rec(t, \<lambda>x zs rs. Apply(f(x), rs))"
definition
term_size :: "i => i" where
- "term_size(t) == term_rec(t, \<lambda>x zs rs. succ(list_add(rs)))"
+ "term_size(t) \<equiv> term_rec(t, \<lambda>x zs rs. succ(list_add(rs)))"
definition
reflect :: "i => i" where
- "reflect(t) == term_rec(t, \<lambda>x zs rs. Apply(x, rev(rs)))"
+ "reflect(t) \<equiv> term_rec(t, \<lambda>x zs rs. Apply(x, rev(rs)))"
definition
preorder :: "i => i" where
- "preorder(t) == term_rec(t, \<lambda>x zs rs. Cons(x, flat(rs)))"
+ "preorder(t) \<equiv> term_rec(t, \<lambda>x zs rs. Cons(x, flat(rs)))"
definition
postorder :: "i => i" where
- "postorder(t) == term_rec(t, \<lambda>x zs rs. flat(rs) @ [x])"
+ "postorder(t) \<equiv> term_rec(t, \<lambda>x zs rs. flat(rs) @ [x])"
lemma term_unfold: "term(A) = A * list(term(A))"
by (fast intro!: term.intros [unfolded term.con_defs]
elim: term.cases [unfolded term.con_defs])
lemma term_induct2:
- "[| t \<in> term(A);
- !!x. [| x \<in> A |] ==> P(Apply(x,Nil));
- !!x z zs. [| x \<in> A; z \<in> term(A); zs: list(term(A)); P(Apply(x,zs))
- |] ==> P(Apply(x, Cons(z,zs)))
- |] ==> P(t)"
+ "\<lbrakk>t \<in> term(A);
+ \<And>x. \<lbrakk>x \<in> A\<rbrakk> \<Longrightarrow> P(Apply(x,Nil));
+ \<And>x z zs. \<lbrakk>x \<in> A; z \<in> term(A); zs: list(term(A)); P(Apply(x,zs))
+\<rbrakk> \<Longrightarrow> P(Apply(x, Cons(z,zs)))
+\<rbrakk> \<Longrightarrow> P(t)"
\<comment> \<open>Induction on \<^term>\<open>term(A)\<close> followed by induction on \<^term>\<open>list\<close>.\<close>
apply (induct_tac t)
apply (erule list.induct)
@@ -62,10 +62,10 @@
done
lemma term_induct_eqn [consumes 1, case_names Apply]:
- "[| t \<in> term(A);
- !!x zs. [| x \<in> A; zs: list(term(A)); map(f,zs) = map(g,zs) |] ==>
+ "\<lbrakk>t \<in> term(A);
+ \<And>x zs. \<lbrakk>x \<in> A; zs: list(term(A)); map(f,zs) = map(g,zs)\<rbrakk> \<Longrightarrow>
f(Apply(x,zs)) = g(Apply(x,zs))
- |] ==> f(t) = g(t)"
+\<rbrakk> \<Longrightarrow> f(t) = g(t)"
\<comment> \<open>Induction on \<^term>\<open>term(A)\<close> to prove an equation.\<close>
apply (induct_tac t)
apply (auto dest: map_list_Collect list_CollectD)
@@ -76,7 +76,7 @@
type definitions.
\<close>
-lemma term_mono: "A \<subseteq> B ==> term(A) \<subseteq> term(B)"
+lemma term_mono: "A \<subseteq> B \<Longrightarrow> term(A) \<subseteq> term(B)"
apply (unfold term.defs)
apply (rule lfp_mono)
apply (rule term.bnd_mono)+
@@ -92,21 +92,21 @@
apply (assumption | rule Pair_in_univ list_univ [THEN subsetD])+
done
-lemma term_subset_univ: "A \<subseteq> univ(B) ==> term(A) \<subseteq> univ(B)"
+lemma term_subset_univ: "A \<subseteq> univ(B) \<Longrightarrow> term(A) \<subseteq> univ(B)"
apply (rule subset_trans)
apply (erule term_mono)
apply (rule term_univ)
done
-lemma term_into_univ: "[| t \<in> term(A); A \<subseteq> univ(B) |] ==> t \<in> univ(B)"
+lemma term_into_univ: "\<lbrakk>t \<in> term(A); A \<subseteq> univ(B)\<rbrakk> \<Longrightarrow> t \<in> univ(B)"
by (rule term_subset_univ [THEN subsetD])
text \<open>
\medskip \<open>term_rec\<close> -- by \<open>Vset\<close> recursion.
\<close>
-lemma map_lemma: "[| l \<in> list(A); Ord(i); rank(l)<i |]
- ==> map(\<lambda>z. (\<lambda>x \<in> Vset(i).h(x)) ` z, l) = map(h,l)"
+lemma map_lemma: "\<lbrakk>l \<in> list(A); Ord(i); rank(l)<i\<rbrakk>
+ \<Longrightarrow> map(\<lambda>z. (\<lambda>x \<in> Vset(i).h(x)) ` z, l) = map(h,l)"
\<comment> \<open>\<^term>\<open>map\<close> works correctly on the underlying list of terms.\<close>
apply (induct set: list)
apply simp
@@ -116,7 +116,7 @@
apply (blast dest: rank_rls [THEN lt_trans])
done
-lemma term_rec [simp]: "ts \<in> list(A) ==>
+lemma term_rec [simp]: "ts \<in> list(A) \<Longrightarrow>
term_rec(Apply(a,ts), d) = d(a, ts, map (\<lambda>z. term_rec(z,d), ts))"
\<comment> \<open>Typing premise is necessary to invoke \<open>map_lemma\<close>.\<close>
apply (rule term_rec_def [THEN def_Vrec, THEN trans])
@@ -126,9 +126,9 @@
lemma term_rec_type:
assumes t: "t \<in> term(A)"
- and a: "!!x zs r. [| x \<in> A; zs: list(term(A));
- r \<in> list(\<Union>t \<in> term(A). C(t)) |]
- ==> d(x, zs, r): C(Apply(x,zs))"
+ and a: "\<And>x zs r. \<lbrakk>x \<in> A; zs: list(term(A));
+ r \<in> list(\<Union>t \<in> term(A). C(t))\<rbrakk>
+ \<Longrightarrow> d(x, zs, r): C(Apply(x,zs))"
shows "term_rec(t,d) \<in> C(t)"
\<comment> \<open>Slightly odd typing condition on \<open>r\<close> in the second premise!\<close>
using t
@@ -141,17 +141,17 @@
done
lemma def_term_rec:
- "[| !!t. j(t)==term_rec(t,d); ts: list(A) |] ==>
+ "\<lbrakk>\<And>t. j(t)\<equiv>term_rec(t,d); ts: list(A)\<rbrakk> \<Longrightarrow>
j(Apply(a,ts)) = d(a, ts, map(\<lambda>Z. j(Z), ts))"
apply (simp only:)
apply (erule term_rec)
done
lemma term_rec_simple_type [TC]:
- "[| t \<in> term(A);
- !!x zs r. [| x \<in> A; zs: list(term(A)); r \<in> list(C) |]
- ==> d(x, zs, r): C
- |] ==> term_rec(t,d) \<in> C"
+ "\<lbrakk>t \<in> term(A);
+ \<And>x zs r. \<lbrakk>x \<in> A; zs: list(term(A)); r \<in> list(C)\<rbrakk>
+ \<Longrightarrow> d(x, zs, r): C
+\<rbrakk> \<Longrightarrow> term_rec(t,d) \<in> C"
apply (erule term_rec_type)
apply (drule subset_refl [THEN UN_least, THEN list_mono, THEN subsetD])
apply simp
@@ -163,19 +163,19 @@
\<close>
lemma term_map [simp]:
- "ts \<in> list(A) ==>
+ "ts \<in> list(A) \<Longrightarrow>
term_map(f, Apply(a, ts)) = Apply(f(a), map(term_map(f), ts))"
by (rule term_map_def [THEN def_term_rec])
lemma term_map_type [TC]:
- "[| t \<in> term(A); !!x. x \<in> A ==> f(x): B |] ==> term_map(f,t) \<in> term(B)"
+ "\<lbrakk>t \<in> term(A); \<And>x. x \<in> A \<Longrightarrow> f(x): B\<rbrakk> \<Longrightarrow> term_map(f,t) \<in> term(B)"
apply (unfold term_map_def)
apply (erule term_rec_simple_type)
apply fast
done
lemma term_map_type2 [TC]:
- "t \<in> term(A) ==> term_map(f,t) \<in> term({f(u). u \<in> A})"
+ "t \<in> term(A) \<Longrightarrow> term_map(f,t) \<in> term({f(u). u \<in> A})"
apply (erule term_map_type)
apply (erule RepFunI)
done
@@ -185,10 +185,10 @@
\<close>
lemma term_size [simp]:
- "ts \<in> list(A) ==> term_size(Apply(a, ts)) = succ(list_add(map(term_size, ts)))"
+ "ts \<in> list(A) \<Longrightarrow> term_size(Apply(a, ts)) = succ(list_add(map(term_size, ts)))"
by (rule term_size_def [THEN def_term_rec])
-lemma term_size_type [TC]: "t \<in> term(A) ==> term_size(t) \<in> nat"
+lemma term_size_type [TC]: "t \<in> term(A) \<Longrightarrow> term_size(t) \<in> nat"
by (auto simp add: term_size_def)
@@ -197,10 +197,10 @@
\<close>
lemma reflect [simp]:
- "ts \<in> list(A) ==> reflect(Apply(a, ts)) = Apply(a, rev(map(reflect, ts)))"
+ "ts \<in> list(A) \<Longrightarrow> reflect(Apply(a, ts)) = Apply(a, rev(map(reflect, ts)))"
by (rule reflect_def [THEN def_term_rec])
-lemma reflect_type [TC]: "t \<in> term(A) ==> reflect(t) \<in> term(A)"
+lemma reflect_type [TC]: "t \<in> term(A) \<Longrightarrow> reflect(t) \<in> term(A)"
by (auto simp add: reflect_def)
@@ -209,10 +209,10 @@
\<close>
lemma preorder [simp]:
- "ts \<in> list(A) ==> preorder(Apply(a, ts)) = Cons(a, flat(map(preorder, ts)))"
+ "ts \<in> list(A) \<Longrightarrow> preorder(Apply(a, ts)) = Cons(a, flat(map(preorder, ts)))"
by (rule preorder_def [THEN def_term_rec])
-lemma preorder_type [TC]: "t \<in> term(A) ==> preorder(t) \<in> list(A)"
+lemma preorder_type [TC]: "t \<in> term(A) \<Longrightarrow> preorder(t) \<in> list(A)"
by (simp add: preorder_def)
@@ -221,10 +221,10 @@
\<close>
lemma postorder [simp]:
- "ts \<in> list(A) ==> postorder(Apply(a, ts)) = flat(map(postorder, ts)) @ [a]"
+ "ts \<in> list(A) \<Longrightarrow> postorder(Apply(a, ts)) = flat(map(postorder, ts)) @ [a]"
by (rule postorder_def [THEN def_term_rec])
-lemma postorder_type [TC]: "t \<in> term(A) ==> postorder(t) \<in> list(A)"
+lemma postorder_type [TC]: "t \<in> term(A) \<Longrightarrow> postorder(t) \<in> list(A)"
by (simp add: postorder_def)
@@ -234,15 +234,15 @@
declare map_compose [simp]
-lemma term_map_ident: "t \<in> term(A) ==> term_map(\<lambda>u. u, t) = t"
+lemma term_map_ident: "t \<in> term(A) \<Longrightarrow> term_map(\<lambda>u. u, t) = t"
by (induct rule: term_induct_eqn) simp
lemma term_map_compose:
- "t \<in> term(A) ==> term_map(f, term_map(g,t)) = term_map(\<lambda>u. f(g(u)), t)"
+ "t \<in> term(A) \<Longrightarrow> term_map(f, term_map(g,t)) = term_map(\<lambda>u. f(g(u)), t)"
by (induct rule: term_induct_eqn) simp
lemma term_map_reflect:
- "t \<in> term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))"
+ "t \<in> term(A) \<Longrightarrow> term_map(f, reflect(t)) = reflect(term_map(f,t))"
by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric])
@@ -250,13 +250,13 @@
\medskip Theorems about \<open>term_size\<close>.
\<close>
-lemma term_size_term_map: "t \<in> term(A) ==> term_size(term_map(f,t)) = term_size(t)"
+lemma term_size_term_map: "t \<in> term(A) \<Longrightarrow> term_size(term_map(f,t)) = term_size(t)"
by (induct rule: term_induct_eqn) simp
-lemma term_size_reflect: "t \<in> term(A) ==> term_size(reflect(t)) = term_size(t)"
+lemma term_size_reflect: "t \<in> term(A) \<Longrightarrow> term_size(reflect(t)) = term_size(t)"
by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric] list_add_rev)
-lemma term_size_length: "t \<in> term(A) ==> term_size(t) = length(preorder(t))"
+lemma term_size_length: "t \<in> term(A) \<Longrightarrow> term_size(t) = length(preorder(t))"
by (induct rule: term_induct_eqn) (simp add: length_flat)
@@ -264,7 +264,7 @@
\medskip Theorems about \<open>reflect\<close>.
\<close>
-lemma reflect_reflect_ident: "t \<in> term(A) ==> reflect(reflect(t)) = t"
+lemma reflect_reflect_ident: "t \<in> term(A) \<Longrightarrow> reflect(reflect(t)) = t"
by (induct rule: term_induct_eqn) (simp add: rev_map_distrib)
@@ -273,11 +273,11 @@
\<close>
lemma preorder_term_map:
- "t \<in> term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))"
+ "t \<in> term(A) \<Longrightarrow> preorder(term_map(f,t)) = map(f, preorder(t))"
by (induct rule: term_induct_eqn) (simp add: map_flat)
lemma preorder_reflect_eq_rev_postorder:
- "t \<in> term(A) ==> preorder(reflect(t)) = rev(postorder(t))"
+ "t \<in> term(A) \<Longrightarrow> preorder(reflect(t)) = rev(postorder(t))"
by (induct rule: term_induct_eqn)
(simp add: rev_app_distrib rev_flat rev_map_distrib [symmetric])
--- a/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 16:51:35 2022 +0100
@@ -26,10 +26,10 @@
declare tree_forest.intros [simp, TC]
-lemma tree_def: "tree(A) == Part(tree_forest(A), Inl)"
+lemma tree_def: "tree(A) \<equiv> Part(tree_forest(A), Inl)"
by (simp only: tree_forest.defs)
-lemma forest_def: "forest(A) == Part(tree_forest(A), Inr)"
+lemma forest_def: "forest(A) \<equiv> Part(tree_forest(A), Inr)"
by (simp only: tree_forest.defs)
@@ -43,7 +43,7 @@
apply (rule Part_subset)
done
-lemma treeI [TC]: "x \<in> tree(A) ==> x \<in> tree_forest(A)"
+lemma treeI [TC]: "x \<in> tree(A) \<Longrightarrow> x \<in> tree_forest(A)"
by (rule tree_subset_TF [THEN subsetD])
lemma forest_subset_TF: "forest(A) \<subseteq> tree_forest(A)"
@@ -51,7 +51,7 @@
apply (rule Part_subset)
done
-lemma treeI' [TC]: "x \<in> forest(A) ==> x \<in> tree_forest(A)"
+lemma treeI' [TC]: "x \<in> forest(A) \<Longrightarrow> x \<in> tree_forest(A)"
by (rule forest_subset_TF [THEN subsetD])
lemma TF_equals_Un: "tree(A) \<union> forest(A) = tree_forest(A)"
@@ -91,22 +91,22 @@
\<close>
lemma TF_rec_type:
- "[| z \<in> tree_forest(A);
- !!x f r. [| x \<in> A; f \<in> forest(A); r \<in> C(f)
- |] ==> b(x,f,r) \<in> C(Tcons(x,f));
+ "\<lbrakk>z \<in> tree_forest(A);
+ \<And>x f r. \<lbrakk>x \<in> A; f \<in> forest(A); r \<in> C(f)
+\<rbrakk> \<Longrightarrow> b(x,f,r) \<in> C(Tcons(x,f));
c \<in> C(Fnil);
- !!t f r1 r2. [| t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> C(f)
- |] ==> d(t,f,r1,r2) \<in> C(Fcons(t,f))
- |] ==> tree_forest_rec(b,c,d,z) \<in> C(z)"
+ \<And>t f r1 r2. \<lbrakk>t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> C(f)
+\<rbrakk> \<Longrightarrow> d(t,f,r1,r2) \<in> C(Fcons(t,f))
+\<rbrakk> \<Longrightarrow> tree_forest_rec(b,c,d,z) \<in> C(z)"
by (induct_tac z) simp_all
lemma tree_forest_rec_type:
- "[| !!x f r. [| x \<in> A; f \<in> forest(A); r \<in> D(f)
- |] ==> b(x,f,r) \<in> C(Tcons(x,f));
+ "\<lbrakk>\<And>x f r. \<lbrakk>x \<in> A; f \<in> forest(A); r \<in> D(f)
+\<rbrakk> \<Longrightarrow> b(x,f,r) \<in> C(Tcons(x,f));
c \<in> D(Fnil);
- !!t f r1 r2. [| t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> D(f)
- |] ==> d(t,f,r1,r2) \<in> D(Fcons(t,f))
- |] ==> (\<forall>t \<in> tree(A). tree_forest_rec(b,c,d,t) \<in> C(t)) \<and>
+ \<And>t f r1 r2. \<lbrakk>t \<in> tree(A); f \<in> forest(A); r1 \<in> C(t); r2 \<in> D(f)
+\<rbrakk> \<Longrightarrow> d(t,f,r1,r2) \<in> D(Fcons(t,f))
+\<rbrakk> \<Longrightarrow> (\<forall>t \<in> tree(A). tree_forest_rec(b,c,d,t) \<in> C(t)) \<and>
(\<forall>f \<in> forest(A). tree_forest_rec(b,c,d,f) \<in> D(f))"
\<comment> \<open>Mutually recursive version.\<close>
apply (unfold Ball_def)
@@ -161,10 +161,10 @@
\<close>
lemma list_of_TF_type [TC]:
- "z \<in> tree_forest(A) ==> list_of_TF(z) \<in> list(tree(A))"
+ "z \<in> tree_forest(A) \<Longrightarrow> list_of_TF(z) \<in> list(tree(A))"
by (induct set: tree_forest) simp_all
-lemma of_list_type [TC]: "l \<in> list(tree(A)) ==> of_list(l) \<in> forest(A)"
+lemma of_list_type [TC]: "l \<in> list(tree(A)) \<Longrightarrow> of_list(l) \<in> forest(A)"
by (induct set: list) simp_all
text \<open>
@@ -172,9 +172,9 @@
\<close>
lemma
- assumes "!!x. x \<in> A ==> h(x): B"
- shows map_tree_type: "t \<in> tree(A) ==> map(h,t) \<in> tree(B)"
- and map_forest_type: "f \<in> forest(A) ==> map(h,f) \<in> forest(B)"
+ assumes "\<And>x. x \<in> A \<Longrightarrow> h(x): B"
+ shows map_tree_type: "t \<in> tree(A) \<Longrightarrow> map(h,t) \<in> tree(B)"
+ and map_forest_type: "f \<in> forest(A) \<Longrightarrow> map(h,f) \<in> forest(B)"
using assms
by (induct rule: tree'induct forest'induct) simp_all
@@ -182,7 +182,7 @@
\medskip \<open>size\<close>.
\<close>
-lemma size_type [TC]: "z \<in> tree_forest(A) ==> size(z) \<in> nat"
+lemma size_type [TC]: "z \<in> tree_forest(A) \<Longrightarrow> size(z) \<in> nat"
by (induct set: tree_forest) simp_all
@@ -190,7 +190,7 @@
\medskip \<open>preorder\<close>.
\<close>
-lemma preorder_type [TC]: "z \<in> tree_forest(A) ==> preorder(z) \<in> list(A)"
+lemma preorder_type [TC]: "z \<in> tree_forest(A) \<Longrightarrow> preorder(z) \<in> list(A)"
by (induct set: tree_forest) simp_all
@@ -199,10 +199,10 @@
\<close>
lemma forest_induct [consumes 1, case_names Fnil Fcons]:
- "[| f \<in> forest(A);
+ "\<lbrakk>f \<in> forest(A);
R(Fnil);
- !!t f. [| t \<in> tree(A); f \<in> forest(A); R(f) |] ==> R(Fcons(t,f))
- |] ==> R(f)"
+ \<And>t f. \<lbrakk>t \<in> tree(A); f \<in> forest(A); R(f)\<rbrakk> \<Longrightarrow> R(Fcons(t,f))
+\<rbrakk> \<Longrightarrow> R(f)"
\<comment> \<open>Essentially the same as list induction.\<close>
apply (erule tree_forest.mutual_induct
[THEN conjunct2, THEN spec, THEN [2] rev_mp])
@@ -211,10 +211,10 @@
apply simp
done
-lemma forest_iso: "f \<in> forest(A) ==> of_list(list_of_TF(f)) = f"
+lemma forest_iso: "f \<in> forest(A) \<Longrightarrow> of_list(list_of_TF(f)) = f"
by (induct rule: forest_induct) simp_all
-lemma tree_list_iso: "ts: list(tree(A)) ==> list_of_TF(of_list(ts)) = ts"
+lemma tree_list_iso: "ts: list(tree(A)) \<Longrightarrow> list_of_TF(of_list(ts)) = ts"
by (induct set: list) simp_all
@@ -222,11 +222,11 @@
\medskip Theorems about \<open>map\<close>.
\<close>
-lemma map_ident: "z \<in> tree_forest(A) ==> map(\<lambda>u. u, z) = z"
+lemma map_ident: "z \<in> tree_forest(A) \<Longrightarrow> map(\<lambda>u. u, z) = z"
by (induct set: tree_forest) simp_all
lemma map_compose:
- "z \<in> tree_forest(A) ==> map(h, map(j,z)) = map(\<lambda>u. h(j(u)), z)"
+ "z \<in> tree_forest(A) \<Longrightarrow> map(h, map(j,z)) = map(\<lambda>u. h(j(u)), z)"
by (induct set: tree_forest) simp_all
@@ -234,10 +234,10 @@
\medskip Theorems about \<open>size\<close>.
\<close>
-lemma size_map: "z \<in> tree_forest(A) ==> size(map(h,z)) = size(z)"
+lemma size_map: "z \<in> tree_forest(A) \<Longrightarrow> size(map(h,z)) = size(z)"
by (induct set: tree_forest) simp_all
-lemma size_length: "z \<in> tree_forest(A) ==> size(z) = length(preorder(z))"
+lemma size_length: "z \<in> tree_forest(A) \<Longrightarrow> size(z) = length(preorder(z))"
by (induct set: tree_forest) (simp_all add: length_app)
text \<open>
@@ -245,7 +245,7 @@
\<close>
lemma preorder_map:
- "z \<in> tree_forest(A) ==> preorder(map(h,z)) = List.map(h, preorder(z))"
+ "z \<in> tree_forest(A) \<Longrightarrow> preorder(map(h,z)) = List.map(h, preorder(z))"
by (induct set: tree_forest) (simp_all add: map_app_distrib)
end
--- a/src/ZF/Inductive.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Inductive.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,13 +19,13 @@
"elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command
begin
-lemma def_swap_iff: "a == b ==> a = c \<longleftrightarrow> c = b"
+lemma def_swap_iff: "a \<equiv> b \<Longrightarrow> a = c \<longleftrightarrow> c = b"
by blast
-lemma def_trans: "f == g ==> g(a) = b ==> f(a) = b"
+lemma def_trans: "f \<equiv> g \<Longrightarrow> g(a) = b \<Longrightarrow> f(a) = b"
by simp
-lemma refl_thin: "!!P. a = a ==> P ==> P" .
+lemma refl_thin: "\<And>P. a = a \<Longrightarrow> P \<Longrightarrow> P" .
ML_file \<open>ind_syntax.ML\<close>
ML_file \<open>Tools/ind_cases.ML\<close>
--- a/src/ZF/InfDatatype.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/InfDatatype.thy Tue Sep 27 16:51:35 2022 +0100
@@ -45,13 +45,13 @@
qed
lemma subset_Vcsucc:
- "[| D \<subseteq> Vfrom(A,csucc(K)); |D| \<le> K; InfCard(K) |]
- ==> \<exists>j. D \<subseteq> Vfrom(A,j) & j < csucc(K)"
+ "\<lbrakk>D \<subseteq> Vfrom(A,csucc(K)); |D| \<le> K; InfCard(K)\<rbrakk>
+ \<Longrightarrow> \<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 \<subseteq> Vfrom(A,csucc(K)) |] ==>
+ "\<lbrakk>|D| \<le> K; InfCard(K); D \<subseteq> Vfrom(A,csucc(K))\<rbrakk> \<Longrightarrow>
D -> Vfrom(A,csucc(K)) \<subseteq> Vfrom(A,csucc(K))"
apply (safe dest!: fun_Vcsucc_lemma subset_Vcsucc)
apply (rule Vfrom [THEN ssubst])
@@ -65,9 +65,9 @@
done
lemma fun_in_Vcsucc:
- "[| f: D -> Vfrom(A, csucc(K)); |D| \<le> K; InfCard(K);
- D \<subseteq> Vfrom(A,csucc(K)) |]
- ==> f: Vfrom(A,csucc(K))"
+ "\<lbrakk>f: D -> Vfrom(A, csucc(K)); |D| \<le> K; InfCard(K);
+ D \<subseteq> Vfrom(A,csucc(K))\<rbrakk>
+ \<Longrightarrow> f: Vfrom(A,csucc(K))"
by (blast intro: fun_Vcsucc [THEN subsetD])
text\<open>Remove \<open>\<subseteq>\<close> from the rule above\<close>
@@ -76,7 +76,7 @@
(** Version where K itself is the index set **)
lemma Card_fun_Vcsucc:
- "InfCard(K) ==> K -> Vfrom(A,csucc(K)) \<subseteq> Vfrom(A,csucc(K))"
+ "InfCard(K) \<Longrightarrow> 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
@@ -84,10 +84,10 @@
done
lemma Card_fun_in_Vcsucc:
- "[| f: K -> Vfrom(A, csucc(K)); InfCard(K) |] ==> f: Vfrom(A,csucc(K))"
+ "\<lbrakk>f: K -> Vfrom(A, csucc(K)); InfCard(K)\<rbrakk> \<Longrightarrow> f: Vfrom(A,csucc(K))"
by (blast intro: Card_fun_Vcsucc [THEN subsetD])
-lemma Limit_csucc: "InfCard(K) ==> Limit(csucc(K))"
+lemma Limit_csucc: "InfCard(K) \<Longrightarrow> Limit(csucc(K))"
by (erule InfCard_csucc [THEN InfCard_is_Limit])
lemmas Pair_in_Vcsucc = Pair_in_VLimit [OF _ _ Limit_csucc]
--- a/src/ZF/Int.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Int.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,50 +9,50 @@
definition
intrel :: i where
- "intrel == {p \<in> (nat*nat)*(nat*nat).
+ "intrel \<equiv> {p \<in> (nat*nat)*(nat*nat).
\<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
definition
int :: i where
- "int == (nat*nat)//intrel"
+ "int \<equiv> (nat*nat)//intrel"
definition
int_of :: "i=>i" \<comment> \<open>coercion from nat to int\<close> (\<open>$# _\<close> [80] 80) where
- "$# m == intrel `` {<natify(m), 0>}"
+ "$# m \<equiv> intrel `` {<natify(m), 0>}"
definition
intify :: "i=>i" \<comment> \<open>coercion from ANYTHING to int\<close> where
- "intify(m) == if m \<in> int then m else $#0"
+ "intify(m) \<equiv> if m \<in> int then m else $#0"
definition
raw_zminus :: "i=>i" where
- "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}"
+ "raw_zminus(z) \<equiv> \<Union><x,y>\<in>z. intrel``{<y,x>}"
definition
zminus :: "i=>i" (\<open>$- _\<close> [80] 80) where
- "$- z == raw_zminus (intify(z))"
+ "$- z \<equiv> raw_zminus (intify(z))"
definition
znegative :: "i=>o" where
- "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
+ "znegative(z) \<equiv> \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z"
definition
iszero :: "i=>o" where
- "iszero(z) == z = $# 0"
+ "iszero(z) \<equiv> z = $# 0"
definition
raw_nat_of :: "i=>i" where
- "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)"
+ "raw_nat_of(z) \<equiv> natify (\<Union><x,y>\<in>z. x#-y)"
definition
nat_of :: "i=>i" where
- "nat_of(z) == raw_nat_of (intify(z))"
+ "nat_of(z) \<equiv> raw_nat_of (intify(z))"
definition
zmagnitude :: "i=>i" where
\<comment> \<open>could be replaced by an absolute value function from int to int?\<close>
- "zmagnitude(z) ==
- THE m. m\<in>nat & ((~ znegative(z) & z = $# m) |
+ "zmagnitude(z) \<equiv>
+ THE m. m\<in>nat & ((\<not> znegative(z) & z = $# m) |
(znegative(z) & $- z = $# m))"
definition
@@ -60,35 +60,35 @@
(*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
Perhaps a "curried" or even polymorphic congruent predicate would be
better.*)
- "raw_zmult(z1,z2) ==
+ "raw_zmult(z1,z2) \<equiv>
\<Union>p1\<in>z1. \<Union>p2\<in>z2. split(%x1 y1. split(%x2 y2.
intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
definition
zmult :: "[i,i]=>i" (infixl \<open>$*\<close> 70) where
- "z1 $* z2 == raw_zmult (intify(z1),intify(z2))"
+ "z1 $* z2 \<equiv> raw_zmult (intify(z1),intify(z2))"
definition
raw_zadd :: "[i,i]=>i" where
- "raw_zadd (z1, z2) ==
+ "raw_zadd (z1, z2) \<equiv>
\<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2
in intrel``{<x1#+x2, y1#+y2>}"
definition
zadd :: "[i,i]=>i" (infixl \<open>$+\<close> 65) where
- "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))"
+ "z1 $+ z2 \<equiv> raw_zadd (intify(z1),intify(z2))"
definition
zdiff :: "[i,i]=>i" (infixl \<open>$-\<close> 65) where
- "z1 $- z2 == z1 $+ zminus(z2)"
+ "z1 $- z2 \<equiv> z1 $+ zminus(z2)"
definition
zless :: "[i,i]=>o" (infixl \<open>$<\<close> 50) where
- "z1 $< z2 == znegative(z1 $- z2)"
+ "z1 $< z2 \<equiv> znegative(z1 $- z2)"
definition
zle :: "[i,i]=>o" (infixl \<open>$\<le>\<close> 50) where
- "z1 $\<le> z2 == z1 $< z2 | intify(z1)=intify(z2)"
+ "z1 $\<le> z2 \<equiv> z1 $< z2 | intify(z1)=intify(z2)"
declare quotientE [elim!]
@@ -103,19 +103,19 @@
by (simp add: intrel_def)
lemma intrelI [intro!]:
- "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]
- ==> <<x1,y1>,<x2,y2>>: intrel"
+ "\<lbrakk>x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
+ \<Longrightarrow> <<x1,y1>,<x2,y2>>: intrel"
by (simp add: intrel_def)
lemma intrelE [elim!]:
- "[| p \<in> intrel;
- !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1;
- x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] ==> Q |]
- ==> Q"
+ "\<lbrakk>p \<in> intrel;
+ \<And>x1 y1 x2 y2. \<lbrakk>p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1;
+ x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk> \<Longrightarrow> Q\<rbrakk>
+ \<Longrightarrow> Q"
by (simp add: intrel_def, blast)
lemma int_trans_lemma:
- "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1"
+ "\<lbrakk>x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2\<rbrakk> \<Longrightarrow> x1 #+ y3 = x3 #+ y1"
apply (rule sym)
apply (erule add_left_cancel)+
apply (simp_all (no_asm_simp))
@@ -126,7 +126,7 @@
apply (fast elim!: sym int_trans_lemma)
done
-lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} \<in> int"
+lemma image_intrel_int: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> intrel `` {<m,n>} \<in> int"
by (simp add: int_def)
declare equiv_intrel [THEN eq_equiv_class_iff, simp]
@@ -142,7 +142,7 @@
lemma int_of_eq [iff]: "($# m = $# n) \<longleftrightarrow> natify(m)=natify(n)"
by (simp add: int_of_def)
-lemma int_of_inject: "[| $#m = $#n; m\<in>nat; n\<in>nat |] ==> m=n"
+lemma int_of_inject: "\<lbrakk>$#m = $#n; m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> m=n"
by (drule int_of_eq [THEN iffD1], auto)
@@ -151,7 +151,7 @@
lemma intify_in_int [iff,TC]: "intify(x) \<in> int"
by (simp add: intify_def)
-lemma intify_ident [simp]: "n \<in> int ==> intify(n) = n"
+lemma intify_ident [simp]: "n \<in> int \<Longrightarrow> intify(n) = n"
by (simp add: intify_def)
@@ -211,7 +211,7 @@
lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel"
by (auto simp add: congruent_def add_ac)
-lemma raw_zminus_type: "z \<in> int ==> raw_zminus(z) \<in> int"
+lemma raw_zminus_type: "z \<in> int \<Longrightarrow> 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
@@ -220,31 +220,31 @@
by (simp add: zminus_def raw_zminus_type)
lemma raw_zminus_inject:
- "[| raw_zminus(z) = raw_zminus(w); z \<in> int; w \<in> int |] ==> z=w"
+ "\<lbrakk>raw_zminus(z) = raw_zminus(w); z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> z=w"
apply (simp add: int_def raw_zminus_def)
apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe)
apply (auto dest: eq_intrelD simp add: add_ac)
done
-lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)"
+lemma zminus_inject_intify [dest!]: "$-z = $-w \<Longrightarrow> intify(z) = intify(w)"
apply (simp add: zminus_def)
apply (blast dest!: raw_zminus_inject)
done
-lemma zminus_inject: "[| $-z = $-w; z \<in> int; w \<in> int |] ==> z=w"
+lemma zminus_inject: "\<lbrakk>$-z = $-w; z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> z=w"
by auto
lemma raw_zminus:
- "[| x\<in>nat; y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}"
+ "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}"
apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent])
done
lemma zminus:
- "[| x\<in>nat; y\<in>nat |]
- ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
+ "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk>
+ \<Longrightarrow> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
by (simp add: zminus_def raw_zminus image_intrel_int)
-lemma raw_zminus_zminus: "z \<in> int ==> raw_zminus (raw_zminus(z)) = z"
+lemma raw_zminus_zminus: "z \<in> int \<Longrightarrow> raw_zminus (raw_zminus(z)) = z"
by (auto simp add: int_def raw_zminus)
lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)"
@@ -253,13 +253,13 @@
lemma zminus_int0 [simp]: "$- ($#0) = $#0"
by (simp add: int_of_def zminus)
-lemma zminus_zminus: "z \<in> int ==> $- ($- z) = z"
+lemma zminus_zminus: "z \<in> int \<Longrightarrow> $- ($- z) = z"
by simp
subsection\<open>\<^term>\<open>znegative\<close>: the test for negative integers\<close>
-lemma znegative: "[| x\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) \<longleftrightarrow> x<y"
+lemma znegative: "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> znegative(intrel``{<x,y>}) \<longleftrightarrow> x<y"
apply (cases "x<y")
apply (auto simp add: znegative_def not_lt_iff_le)
apply (subgoal_tac "y #+ x2 < x #+ y2", force)
@@ -267,13 +267,13 @@
done
(*No natural number is negative!*)
-lemma not_znegative_int_of [iff]: "~ znegative($# n)"
+lemma not_znegative_int_of [iff]: "\<not> znegative($# n)"
by (simp add: znegative int_of_def)
lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))"
by (simp add: znegative int_of_def zminus natify_succ)
-lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0"
+lemma not_znegative_imp_zero: "\<not> znegative($- $# n) \<Longrightarrow> natify(n)=0"
by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym])
@@ -286,7 +286,7 @@
by (auto simp add: congruent_def split: nat_diff_split)
lemma raw_nat_of:
- "[| x\<in>nat; y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = x#-y"
+ "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> raw_nat_of(intrel``{<x,y>}) = x#-y"
by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent])
lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)"
@@ -306,7 +306,7 @@
lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)"
by (auto simp add: zmagnitude_def int_of_eq)
-lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n"
+lemma natify_int_of_eq: "natify(x)=n \<Longrightarrow> $#x = $# n"
apply (drule sym)
apply (simp (no_asm_simp) add: int_of_eq)
done
@@ -324,7 +324,7 @@
done
lemma not_zneg_int_of:
- "[| z \<in> int; ~ znegative(z) |] ==> \<exists>n\<in>nat. z = $# n"
+ "\<lbrakk>z \<in> int; \<not> znegative(z)\<rbrakk> \<Longrightarrow> \<exists>n\<in>nat. z = $# n"
apply (auto simp add: int_def znegative int_of_def not_lt_iff_le)
apply (rename_tac x y)
apply (rule_tac x="x#-y" in bexI)
@@ -332,38 +332,38 @@
done
lemma not_zneg_mag [simp]:
- "[| z \<in> int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z"
+ "\<lbrakk>z \<in> int; \<not> znegative(z)\<rbrakk> \<Longrightarrow> $# (zmagnitude(z)) = z"
by (drule not_zneg_int_of, auto)
lemma zneg_int_of:
- "[| znegative(z); z \<in> int |] ==> \<exists>n\<in>nat. z = $- ($# succ(n))"
+ "\<lbrakk>znegative(z); z \<in> int\<rbrakk> \<Longrightarrow> \<exists>n\<in>nat. z = $- ($# succ(n))"
by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add)
lemma zneg_mag [simp]:
- "[| znegative(z); z \<in> int |] ==> $# (zmagnitude(z)) = $- z"
+ "\<lbrakk>znegative(z); z \<in> int\<rbrakk> \<Longrightarrow> $# (zmagnitude(z)) = $- z"
by (drule zneg_int_of, auto)
-lemma int_cases: "z \<in> int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))"
+lemma int_cases: "z \<in> int \<Longrightarrow> \<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)
done
lemma not_zneg_raw_nat_of:
- "[| ~ znegative(z); z \<in> int |] ==> $# (raw_nat_of(z)) = z"
+ "\<lbrakk>\<not> znegative(z); z \<in> int\<rbrakk> \<Longrightarrow> $# (raw_nat_of(z)) = z"
apply (drule not_zneg_int_of)
apply (auto simp add: raw_nat_of_type raw_nat_of_int_of)
done
lemma not_zneg_nat_of_intify:
- "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)"
+ "\<not> znegative(intify(z)) \<Longrightarrow> $# (nat_of(z)) = intify(z)"
by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of)
-lemma not_zneg_nat_of: "[| ~ znegative(z); z \<in> int |] ==> $# (nat_of(z)) = z"
+lemma not_zneg_nat_of: "\<lbrakk>\<not> znegative(z); z \<in> int\<rbrakk> \<Longrightarrow> $# (nat_of(z)) = z"
apply (simp (no_asm_simp) add: not_zneg_nat_of_intify)
done
-lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0"
+lemma zneg_nat_of [simp]: "znegative(intify(z)) \<Longrightarrow> nat_of(z) = 0"
apply (subgoal_tac "intify(z) \<in> int")
apply (simp add: int_def)
apply (auto simp add: znegative nat_of_def raw_nat_of
@@ -389,7 +389,7 @@
apply (simp (no_asm_simp) add: add_assoc [symmetric])
done
-lemma raw_zadd_type: "[| z \<in> int; w \<in> int |] ==> raw_zadd(z,w) \<in> int"
+lemma raw_zadd_type: "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> 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)
@@ -399,8 +399,8 @@
by (simp add: zadd_def raw_zadd_type)
lemma raw_zadd:
- "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]
- ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
+ "\<lbrakk>x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
+ \<Longrightarrow> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
intrel `` {<x1#+x2, y1#+y2>}"
apply (simp add: raw_zadd_def
UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2])
@@ -408,37 +408,37 @@
done
lemma zadd:
- "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]
- ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =
+ "\<lbrakk>x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
+ \<Longrightarrow> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =
intrel `` {<x1#+x2, y1#+y2>}"
by (simp add: zadd_def raw_zadd image_intrel_int)
-lemma raw_zadd_int0: "z \<in> int ==> raw_zadd ($#0,z) = z"
+lemma raw_zadd_int0: "z \<in> int \<Longrightarrow> 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)"
by (simp add: zadd_def raw_zadd_int0)
-lemma zadd_int0: "z \<in> int ==> $#0 $+ z = z"
+lemma zadd_int0: "z \<in> int \<Longrightarrow> $#0 $+ z = z"
by simp
lemma raw_zminus_zadd_distrib:
- "[| z \<in> int; w \<in> int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)"
+ "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> $- raw_zadd(z,w) = raw_zadd($- z, $- w)"
by (auto simp add: zminus raw_zadd int_def)
lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w"
by (simp add: zadd_def raw_zminus_zadd_distrib)
lemma raw_zadd_commute:
- "[| z \<in> int; w \<in> int |] ==> raw_zadd(z,w) = raw_zadd(w,z)"
+ "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> raw_zadd(z,w) = raw_zadd(w,z)"
by (auto simp add: raw_zadd add_ac int_def)
lemma zadd_commute: "z $+ w = w $+ z"
by (simp add: zadd_def raw_zadd_commute)
lemma raw_zadd_assoc:
- "[| z1: int; z2: int; z3: int |]
- ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))"
+ "\<lbrakk>z1: int; z2: int; z3: int\<rbrakk>
+ \<Longrightarrow> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))"
by (auto simp add: int_def raw_zadd add_assoc)
lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"
@@ -460,13 +460,13 @@
by (simp add: int_of_add [symmetric] natify_succ)
lemma int_of_diff:
- "[| m\<in>nat; n \<le> m |] ==> $# (m #- n) = ($#m) $- ($#n)"
+ "\<lbrakk>m\<in>nat; n \<le> m\<rbrakk> \<Longrightarrow> $# (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 \<in> int ==> raw_zadd (z, $- z) = $#0"
+lemma raw_zadd_zminus_inverse: "z \<in> int \<Longrightarrow> 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"
@@ -481,7 +481,7 @@
lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)"
by (rule trans [OF zadd_commute zadd_int0_intify])
-lemma zadd_int0_right: "z \<in> int ==> z $+ $#0 = z"
+lemma zadd_int0_right: "z \<in> int \<Longrightarrow> z $+ $#0 = z"
by simp
@@ -502,7 +502,7 @@
done
-lemma raw_zmult_type: "[| z \<in> int; w \<in> int |] ==> raw_zmult(z,w) \<in> int"
+lemma raw_zmult_type: "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> 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)
@@ -512,42 +512,42 @@
by (simp add: zmult_def raw_zmult_type)
lemma raw_zmult:
- "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]
- ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
+ "\<lbrakk>x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
+ \<Longrightarrow> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
by (simp add: raw_zmult_def
UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2])
lemma zmult:
- "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |]
- ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =
+ "\<lbrakk>x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
+ \<Longrightarrow> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =
intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
by (simp add: zmult_def raw_zmult image_intrel_int)
-lemma raw_zmult_int0: "z \<in> int ==> raw_zmult ($#0,z) = $#0"
+lemma raw_zmult_int0: "z \<in> int \<Longrightarrow> 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 \<in> int ==> raw_zmult ($#1,z) = z"
+lemma raw_zmult_int1: "z \<in> int \<Longrightarrow> 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 \<in> int ==> $#1 $* z = z"
+lemma zmult_int1: "z \<in> int \<Longrightarrow> $#1 $* z = z"
by simp
lemma raw_zmult_commute:
- "[| z \<in> int; w \<in> int |] ==> raw_zmult(z,w) = raw_zmult(w,z)"
+ "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> raw_zmult(z,w) = raw_zmult(w,z)"
by (auto simp add: int_def raw_zmult add_ac mult_ac)
lemma zmult_commute: "z $* w = w $* z"
by (simp add: zmult_def raw_zmult_commute)
lemma raw_zmult_zminus:
- "[| z \<in> int; w \<in> int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)"
+ "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> raw_zmult($- z, w) = $- raw_zmult(z, w)"
by (auto simp add: int_def zminus raw_zmult add_ac)
lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)"
@@ -559,8 +559,8 @@
by (simp add: zmult_commute [of w])
lemma raw_zmult_assoc:
- "[| z1: int; z2: int; z3: int |]
- ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))"
+ "\<lbrakk>z1: int; z2: int; z3: int\<rbrakk>
+ \<Longrightarrow> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))"
by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac)
lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)"
@@ -576,8 +576,8 @@
lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute
lemma raw_zadd_zmult_distrib:
- "[| z1: int; z2: int; w \<in> int |]
- ==> raw_zmult(raw_zadd(z1,z2), w) =
+ "\<lbrakk>z1: int; z2: int; w \<in> int\<rbrakk>
+ \<Longrightarrow> raw_zmult(raw_zadd(z1,z2), w) =
raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))"
by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac)
@@ -620,7 +620,7 @@
(*"Less than" is a linear ordering*)
lemma zless_linear_lemma:
- "[| z \<in> int; w \<in> int |] ==> z$<w | z=w | w$<z"
+ "\<lbrakk>z \<in> int; w \<in> int\<rbrakk> \<Longrightarrow> z$<w | z=w | w$<z"
apply (simp add: int_def zless_def znegative_def zdiff_def, auto)
apply (simp add: zadd zminus image_iff Bex_def)
apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt)
@@ -632,15 +632,15 @@
apply auto
done
-lemma zless_not_refl [iff]: "~ (z$<z)"
+lemma zless_not_refl [iff]: "\<not> (z$<z)"
by (auto simp add: zless_def znegative_def int_of_def zdiff_def)
-lemma neq_iff_zless: "[| x \<in> int; y \<in> int |] ==> (x \<noteq> y) \<longleftrightarrow> (x $< y | y $< x)"
+lemma neq_iff_zless: "\<lbrakk>x \<in> int; y \<in> int\<rbrakk> \<Longrightarrow> (x \<noteq> y) \<longleftrightarrow> (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) \<noteq> intify(z)"
+lemma zless_imp_intify_neq: "w $< z \<Longrightarrow> intify(w) \<noteq> intify(z)"
apply auto
-apply (subgoal_tac "~ (intify (w) $< intify (z))")
+apply (subgoal_tac "\<not> (intify (w) $< intify (z))")
apply (erule_tac [2] ssubst)
apply (simp (no_asm_use))
apply auto
@@ -648,7 +648,7 @@
(*This lemma allows direct proofs of other <-properties*)
lemma zless_imp_succ_zadd_lemma:
- "[| w $< z; w \<in> int; z \<in> int |] ==> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))"
+ "\<lbrakk>w $< z; w \<in> int; z \<in> int\<rbrakk> \<Longrightarrow> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))"
apply (simp add: zless_def znegative_def zdiff_def int_def)
apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def)
apply (rule_tac x = k in bexI)
@@ -656,14 +656,14 @@
done
lemma zless_imp_succ_zadd:
- "w $< z ==> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
+ "w $< z \<Longrightarrow> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))"
apply (subgoal_tac "intify (w) $< intify (z) ")
apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma)
apply auto
done
lemma zless_succ_zadd_lemma:
- "w \<in> int ==> w $< w $+ $# succ(n)"
+ "w \<in> int \<Longrightarrow> 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)
@@ -680,13 +680,13 @@
apply (cut_tac w = w and n = n in zless_succ_zadd, auto)
done
-lemma zless_int_of [simp]: "[| m\<in>nat; n\<in>nat |] ==> ($#m $< $#n) \<longleftrightarrow> (m<n)"
+lemma zless_int_of [simp]: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> ($#m $< $#n) \<longleftrightarrow> (m<n)"
apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric])
apply (blast intro: sym)
done
lemma zless_trans_lemma:
- "[| x $< y; y $< z; x \<in> int; y \<in> int; z \<in> int |] ==> x $< z"
+ "\<lbrakk>x $< y; y $< z; x \<in> int; y \<in> int; z \<in> int\<rbrakk> \<Longrightarrow> 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)
@@ -699,19 +699,19 @@
apply auto
done
-lemma zless_trans [trans]: "[| x $< y; y $< z |] ==> x $< z"
+lemma zless_trans [trans]: "\<lbrakk>x $< y; y $< z\<rbrakk> \<Longrightarrow> x $< z"
apply (subgoal_tac "intify (x) $< intify (z) ")
apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma)
apply auto
done
-lemma zless_not_sym: "z $< w ==> ~ (w $< z)"
+lemma zless_not_sym: "z $< w \<Longrightarrow> \<not> (w $< z)"
by (blast dest: zless_trans)
-(* [| z $< w; ~ P ==> w $< z |] ==> P *)
+(* \<lbrakk>z $< w; \<not> P \<Longrightarrow> w $< z\<rbrakk> \<Longrightarrow> P *)
lemmas zless_asym = zless_not_sym [THEN swap]
-lemma zless_imp_zle: "z $< w ==> z $\<le> w"
+lemma zless_imp_zle: "z $< w \<Longrightarrow> z $\<le> w"
by (simp add: zle_def)
lemma zle_linear: "z $\<le> w | w $\<le> z"
@@ -725,48 +725,48 @@
lemma zle_refl: "z $\<le> z"
by (simp add: zle_def)
-lemma zle_eq_refl: "x=y ==> x $\<le> y"
+lemma zle_eq_refl: "x=y \<Longrightarrow> x $\<le> y"
by (simp add: zle_refl)
-lemma zle_anti_sym_intify: "[| x $\<le> y; y $\<le> x |] ==> intify(x) = intify(y)"
+lemma zle_anti_sym_intify: "\<lbrakk>x $\<le> y; y $\<le> x\<rbrakk> \<Longrightarrow> intify(x) = intify(y)"
apply (simp add: zle_def, auto)
apply (blast dest: zless_trans)
done
-lemma zle_anti_sym: "[| x $\<le> y; y $\<le> x; x \<in> int; y \<in> int |] ==> x=y"
+lemma zle_anti_sym: "\<lbrakk>x $\<le> y; y $\<le> x; x \<in> int; y \<in> int\<rbrakk> \<Longrightarrow> x=y"
by (drule zle_anti_sym_intify, auto)
lemma zle_trans_lemma:
- "[| x \<in> int; y \<in> int; z \<in> int; x $\<le> y; y $\<le> z |] ==> x $\<le> z"
+ "\<lbrakk>x \<in> int; y \<in> int; z \<in> int; x $\<le> y; y $\<le> z\<rbrakk> \<Longrightarrow> x $\<le> z"
apply (simp add: zle_def, auto)
apply (blast intro: zless_trans)
done
-lemma zle_trans [trans]: "[| x $\<le> y; y $\<le> z |] ==> x $\<le> z"
+lemma zle_trans [trans]: "\<lbrakk>x $\<le> y; y $\<le> z\<rbrakk> \<Longrightarrow> x $\<le> z"
apply (subgoal_tac "intify (x) $\<le> intify (z) ")
apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma)
apply auto
done
-lemma zle_zless_trans [trans]: "[| i $\<le> j; j $< k |] ==> i $< k"
+lemma zle_zless_trans [trans]: "\<lbrakk>i $\<le> j; j $< k\<rbrakk> \<Longrightarrow> i $< k"
apply (auto simp add: zle_def)
apply (blast intro: zless_trans)
apply (simp add: zless_def zdiff_def zadd_def)
done
-lemma zless_zle_trans [trans]: "[| i $< j; j $\<le> k |] ==> i $< k"
+lemma zless_zle_trans [trans]: "\<lbrakk>i $< j; j $\<le> k\<rbrakk> \<Longrightarrow> i $< k"
apply (auto simp add: zle_def)
apply (blast intro: zless_trans)
apply (simp add: zless_def zdiff_def zminus_def)
done
-lemma not_zless_iff_zle: "~ (z $< w) \<longleftrightarrow> (w $\<le> z)"
+lemma not_zless_iff_zle: "\<not> (z $< w) \<longleftrightarrow> (w $\<le> z)"
apply (cut_tac z = z and w = w in zless_linear)
apply (auto dest: zless_trans simp add: zle_def)
apply (auto dest!: zless_imp_intify_neq)
done
-lemma not_zle_iff_zless: "~ (z $\<le> w) \<longleftrightarrow> (w $< z)"
+lemma not_zle_iff_zless: "\<not> (z $\<le> w) \<longleftrightarrow> (w $< z)"
by (simp add: not_zless_iff_zle [THEN iff_sym])
@@ -784,21 +784,21 @@
lemma zless_zdiff_iff: "(x $< z$-y) \<longleftrightarrow> (x $+ y $< z)"
by (simp add: zless_def zdiff_def zadd_ac)
-lemma zdiff_eq_iff: "[| x \<in> int; z \<in> int |] ==> (x$-y = z) \<longleftrightarrow> (x = z $+ y)"
+lemma zdiff_eq_iff: "\<lbrakk>x \<in> int; z \<in> int\<rbrakk> \<Longrightarrow> (x$-y = z) \<longleftrightarrow> (x = z $+ y)"
by (auto simp add: zdiff_def zadd_assoc)
-lemma eq_zdiff_iff: "[| x \<in> int; z \<in> int |] ==> (x = z$-y) \<longleftrightarrow> (x $+ y = z)"
+lemma eq_zdiff_iff: "\<lbrakk>x \<in> int; z \<in> int\<rbrakk> \<Longrightarrow> (x = z$-y) \<longleftrightarrow> (x $+ y = z)"
by (auto simp add: zdiff_def zadd_assoc)
lemma zdiff_zle_iff_lemma:
- "[| x \<in> int; z \<in> int |] ==> (x$-y $\<le> z) \<longleftrightarrow> (x $\<le> z $+ y)"
+ "\<lbrakk>x \<in> int; z \<in> int\<rbrakk> \<Longrightarrow> (x$-y $\<le> z) \<longleftrightarrow> (x $\<le> z $+ y)"
by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff)
lemma zdiff_zle_iff: "(x$-y $\<le> z) \<longleftrightarrow> (x $\<le> z $+ y)"
by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp)
lemma zle_zdiff_iff_lemma:
- "[| x \<in> int; z \<in> int |] ==>(x $\<le> z$-y) \<longleftrightarrow> (x $+ y $\<le> z)"
+ "\<lbrakk>x \<in> int; z \<in> int\<rbrakk> \<Longrightarrow>(x $\<le> z$-y) \<longleftrightarrow> (x $+ y $\<le> z)"
apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff)
apply (auto simp add: zdiff_def zadd_assoc)
done
@@ -820,7 +820,7 @@
of the CancelNumerals Simprocs\<close>
lemma zadd_left_cancel:
- "[| w \<in> int; w': int |] ==> (z $+ w' = z $+ w) \<longleftrightarrow> (w' = w)"
+ "\<lbrakk>w \<in> int; w': int\<rbrakk> \<Longrightarrow> (z $+ w' = z $+ w) \<longleftrightarrow> (w' = w)"
apply safe
apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
apply (simp add: zadd_ac)
@@ -833,7 +833,7 @@
done
lemma zadd_right_cancel:
- "[| w \<in> int; w': int |] ==> (w' $+ z = w $+ z) \<longleftrightarrow> (w' = w)"
+ "\<lbrakk>w \<in> int; w': int\<rbrakk> \<Longrightarrow> (w' $+ z = w $+ z) \<longleftrightarrow> (w' = w)"
apply safe
apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
apply (simp add: zadd_ac)
@@ -858,22 +858,22 @@
by (simp add: zadd_commute [of z] zadd_right_cancel_zle)
-(*"v $\<le> w ==> v$+z $\<le> w$+z"*)
+(*"v $\<le> w \<Longrightarrow> v$+z $\<le> w$+z"*)
lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2]
-(*"v $\<le> w ==> z$+v $\<le> z$+w"*)
+(*"v $\<le> w \<Longrightarrow> z$+v $\<le> z$+w"*)
lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2]
-(*"v $\<le> w ==> v$+z $\<le> w$+z"*)
+(*"v $\<le> w \<Longrightarrow> v$+z $\<le> w$+z"*)
lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2]
-(*"v $\<le> w ==> z$+v $\<le> z$+w"*)
+(*"v $\<le> w \<Longrightarrow> z$+v $\<le> z$+w"*)
lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2]
-lemma zadd_zle_mono: "[| w' $\<le> w; z' $\<le> z |] ==> w' $+ z' $\<le> w $+ z"
+lemma zadd_zle_mono: "\<lbrakk>w' $\<le> w; z' $\<le> z\<rbrakk> \<Longrightarrow> w' $+ z' $\<le> w $+ z"
by (erule zadd_zle_mono1 [THEN zle_trans], simp)
-lemma zadd_zless_mono: "[| w' $< w; z' $\<le> z |] ==> w' $+ z' $< w $+ z"
+lemma zadd_zless_mono: "\<lbrakk>w' $< w; z' $\<le> z\<rbrakk> \<Longrightarrow> w' $+ z' $< w $+ z"
by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp)
@@ -887,10 +887,10 @@
subsubsection\<open>More inequality lemmas\<close>
-lemma equation_zminus: "[| x \<in> int; y \<in> int |] ==> (x = $- y) \<longleftrightarrow> (y = $- x)"
+lemma equation_zminus: "\<lbrakk>x \<in> int; y \<in> int\<rbrakk> \<Longrightarrow> (x = $- y) \<longleftrightarrow> (y = $- x)"
by auto
-lemma zminus_equation: "[| x \<in> int; y \<in> int |] ==> ($- x = y) \<longleftrightarrow> ($- y = x)"
+lemma zminus_equation: "\<lbrakk>x \<in> int; y \<in> int\<rbrakk> \<Longrightarrow> ($- x = y) \<longleftrightarrow> ($- y = x)"
by auto
lemma equation_zminus_intify: "(intify(x) = $- y) \<longleftrightarrow> (intify(y) = $- x)"
--- a/src/ZF/IntDiv.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/IntDiv.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,20 +11,20 @@
end
fun negDivAlg (a,b) =
- if 0<=a+b then (~1,a+b)
+ if 0<=a+b then (\<not>1,a+b)
else let val (q,r) = negDivAlg(a, 2*b)
in if 0<=r-b then (2*q+1, r-b) else (2*q, r)
end;
- fun negateSnd (q,r:int) = (q,~r);
+ fun negateSnd (q,r:int) = (q,\<not>r);
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 negateSnd (negDivAlg (\<not>a,\<not>b))
else
if 0<b then negDivAlg (a,b)
- else negateSnd (posDivAlg (~a,~b));
+ else negateSnd (posDivAlg (\<not>a,\<not>b));
*)
section\<open>The Division Operators Div and Mod\<close>
@@ -35,13 +35,13 @@
definition
quorem :: "[i,i] => o" where
- "quorem == %<a,b> <q,r>.
+ "quorem \<equiv> %<a,b> <q,r>.
a = b$*q $+ r &
- (#0$<b & #0$\<le>r & r$<b | ~(#0$<b) & b$<r & r $\<le> #0)"
+ (#0$<b & #0$\<le>r & r$<b | \<not>(#0$<b) & b$<r & r $\<le> #0)"
definition
adjust :: "[i,i] => i" where
- "adjust(b) == %<q,r>. if #0 $\<le> r$-b then <#2$*q $+ #1,r$-b>
+ "adjust(b) \<equiv> %<q,r>. if #0 $\<le> r$-b then <#2$*q $+ #1,r$-b>
else <#2$*q,r>"
@@ -51,7 +51,7 @@
posDivAlg :: "i => i" where
(*for the case a>=0, b>0*)
(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
- "posDivAlg(ab) ==
+ "posDivAlg(ab) \<equiv>
wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
ab,
%<a,b> f. if (a$<b | b$\<le>#0) then <#0,a>
@@ -62,7 +62,7 @@
definition
negDivAlg :: "i => i" where
(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
- "negDivAlg(ab) ==
+ "negDivAlg(ab) \<equiv>
wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
ab,
%<a,b> f. if (#0 $\<le> a$+b | b$\<le>#0) then <#-1,a$+b>
@@ -72,13 +72,13 @@
definition
negateSnd :: "i => i" where
- "negateSnd == %<q,r>. <q, $-r>"
+ "negateSnd \<equiv> %<q,r>. <q, $-r>"
(*The full division algorithm considers all possible signs for a, b
including the special case a=0, b<0, because negDivAlg requires a<0*)
definition
divAlg :: "i => i" where
- "divAlg ==
+ "divAlg \<equiv>
%<a,b>. if #0 $\<le> a then
if #0 $\<le> b then posDivAlg (<a,b>)
else if a=#0 then <#0,#0>
@@ -89,28 +89,28 @@
definition
zdiv :: "[i,i]=>i" (infixl \<open>zdiv\<close> 70) where
- "a zdiv b == fst (divAlg (<intify(a), intify(b)>))"
+ "a zdiv b \<equiv> fst (divAlg (<intify(a), intify(b)>))"
definition
zmod :: "[i,i]=>i" (infixl \<open>zmod\<close> 70) where
- "a zmod b == snd (divAlg (<intify(a), intify(b)>))"
+ "a zmod b \<equiv> snd (divAlg (<intify(a), intify(b)>))"
(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **)
-lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y"
+lemma zspos_add_zspos_imp_zspos: "\<lbrakk>#0 $< x; #0 $< y\<rbrakk> \<Longrightarrow> #0 $< x $+ y"
apply (rule_tac y = "y" in zless_trans)
apply (rule_tac [2] zdiff_zless_iff [THEN iffD1])
apply auto
done
-lemma zpos_add_zpos_imp_zpos: "[| #0 $\<le> x; #0 $\<le> y |] ==> #0 $\<le> x $+ y"
+lemma zpos_add_zpos_imp_zpos: "\<lbrakk>#0 $\<le> x; #0 $\<le> y\<rbrakk> \<Longrightarrow> #0 $\<le> x $+ y"
apply (rule_tac y = "y" in zle_trans)
apply (rule_tac [2] zdiff_zle_iff [THEN iffD1])
apply auto
done
-lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0"
+lemma zneg_add_zneg_imp_zneg: "\<lbrakk>x $< #0; y $< #0\<rbrakk> \<Longrightarrow> x $+ y $< #0"
apply (rule_tac y = "y" in zless_trans)
apply (rule zless_zdiff_iff [THEN iffD1])
apply auto
@@ -118,13 +118,13 @@
(* this theorem is used below *)
lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0:
- "[| x $\<le> #0; y $\<le> #0 |] ==> x $+ y $\<le> #0"
+ "\<lbrakk>x $\<le> #0; y $\<le> #0\<rbrakk> \<Longrightarrow> x $+ y $\<le> #0"
apply (rule_tac y = "y" in zle_trans)
apply (rule zle_zdiff_iff [THEN iffD1])
apply auto
done
-lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)"
+lemma zero_lt_zmagnitude: "\<lbrakk>#0 $< k; k \<in> int\<rbrakk> \<Longrightarrow> 0 < zmagnitude(k)"
apply (drule zero_zless_imp_znegative_zminus)
apply (drule_tac [2] zneg_int_of)
apply (auto simp add: zminus_equation [of k])
@@ -151,7 +151,7 @@
done
lemma zadd_succ_lemma:
- "z \<in> int ==> (w $+ $# succ(m) $\<le> z) \<longleftrightarrow> (w $+ $#m $< z)"
+ "z \<in> int \<Longrightarrow> (w $+ $# succ(m) $\<le> z) \<longleftrightarrow> (w $+ $#m $< z)"
apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff)
apply (auto intro: zle_anti_sym elim: zless_asym
simp add: zless_imp_zle not_zless_iff_zle)
@@ -184,13 +184,13 @@
(*** Monotonicity of Multiplication ***)
-lemma zmult_mono_lemma: "k \<in> nat ==> i $\<le> j ==> i $* $#k $\<le> j $* $#k"
+lemma zmult_mono_lemma: "k \<in> nat \<Longrightarrow> i $\<le> j \<Longrightarrow> i $* $#k $\<le> j $* $#k"
apply (induct_tac "k")
prefer 2 apply (subst int_succ_int_1)
apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono)
done
-lemma zmult_zle_mono1: "[| i $\<le> j; #0 $\<le> k |] ==> i$*k $\<le> j$*k"
+lemma zmult_zle_mono1: "\<lbrakk>i $\<le> j; #0 $\<le> k\<rbrakk> \<Longrightarrow> i$*k $\<le> j$*k"
apply (subgoal_tac "i $* intify (k) $\<le> j $* intify (k) ")
apply (simp (no_asm_use))
apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
@@ -199,25 +199,25 @@
apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym])
done
-lemma zmult_zle_mono1_neg: "[| i $\<le> j; k $\<le> #0 |] ==> j$*k $\<le> i$*k"
+lemma zmult_zle_mono1_neg: "\<lbrakk>i $\<le> j; k $\<le> #0\<rbrakk> \<Longrightarrow> j$*k $\<le> i$*k"
apply (rule zminus_zle_zminus [THEN iffD1])
apply (simp del: zmult_zminus_right
add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus)
done
-lemma zmult_zle_mono2: "[| i $\<le> j; #0 $\<le> k |] ==> k$*i $\<le> k$*j"
+lemma zmult_zle_mono2: "\<lbrakk>i $\<le> j; #0 $\<le> k\<rbrakk> \<Longrightarrow> k$*i $\<le> k$*j"
apply (drule zmult_zle_mono1)
apply (simp_all add: zmult_commute)
done
-lemma zmult_zle_mono2_neg: "[| i $\<le> j; k $\<le> #0 |] ==> k$*j $\<le> k$*i"
+lemma zmult_zle_mono2_neg: "\<lbrakk>i $\<le> j; k $\<le> #0\<rbrakk> \<Longrightarrow> k$*j $\<le> k$*i"
apply (drule zmult_zle_mono1_neg)
apply (simp_all add: zmult_commute)
done
(* $\<le> monotonicity, BOTH arguments*)
lemma zmult_zle_mono:
- "[| i $\<le> j; k $\<le> l; #0 $\<le> j; #0 $\<le> k |] ==> i$*k $\<le> j$*l"
+ "\<lbrakk>i $\<le> j; k $\<le> l; #0 $\<le> j; #0 $\<le> k\<rbrakk> \<Longrightarrow> i$*k $\<le> j$*l"
apply (erule zmult_zle_mono1 [THEN zle_trans])
apply assumption
apply (erule zmult_zle_mono2)
@@ -228,7 +228,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 \<longrightarrow> $#k $* i $< $#k $* j"
+ "\<lbrakk>i$<j; k \<in> nat\<rbrakk> \<Longrightarrow> 0<k \<longrightarrow> $#k $* i $< $#k $* j"
apply (induct_tac "k")
prefer 2
apply (subst int_succ_int_1)
@@ -241,7 +241,7 @@
apply (simp_all (no_asm_simp) add: zle_def)
done
-lemma zmult_zless_mono2: "[| i$<j; #0 $< k |] ==> k$*i $< k$*j"
+lemma zmult_zless_mono2: "\<lbrakk>i$<j; #0 $< k\<rbrakk> \<Longrightarrow> k$*i $< k$*j"
apply (subgoal_tac "intify (k) $* i $< intify (k) $* j")
apply (simp (no_asm_use))
apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst])
@@ -252,27 +252,27 @@
apply (auto simp add: zero_lt_zmagnitude)
done
-lemma zmult_zless_mono1: "[| i$<j; #0 $< k |] ==> i$*k $< j$*k"
+lemma zmult_zless_mono1: "\<lbrakk>i$<j; #0 $< k\<rbrakk> \<Longrightarrow> i$*k $< j$*k"
apply (drule zmult_zless_mono2)
apply (simp_all add: zmult_commute)
done
(* < monotonicity, BOTH arguments*)
lemma zmult_zless_mono:
- "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l"
+ "\<lbrakk>i $< j; k $< l; #0 $< j; #0 $< k\<rbrakk> \<Longrightarrow> i$*k $< j$*l"
apply (erule zmult_zless_mono1 [THEN zless_trans])
apply assumption
apply (erule zmult_zless_mono2)
apply assumption
done
-lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k"
+lemma zmult_zless_mono1_neg: "\<lbrakk>i $< j; k $< #0\<rbrakk> \<Longrightarrow> j$*k $< i$*k"
apply (rule zminus_zless_zminus [THEN iffD1])
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"
+lemma zmult_zless_mono2_neg: "\<lbrakk>i $< j; k $< #0\<rbrakk> \<Longrightarrow> k$*j $< k$*i"
apply (rule zminus_zless_zminus [THEN iffD1])
apply (simp del: zmult_zminus
add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus)
@@ -282,7 +282,7 @@
(** Products of zeroes **)
lemma zmult_eq_lemma:
- "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) \<longleftrightarrow> (m$*n = #0)"
+ "\<lbrakk>m \<in> int; n \<in> int\<rbrakk> \<Longrightarrow> (m = #0 | n = #0) \<longleftrightarrow> (m$*n = #0)"
apply (case_tac "m $< #0")
apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless)
apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+
@@ -297,8 +297,8 @@
but not (yet?) for k*m < n*k. **)
lemma zmult_zless_lemma:
- "[| k \<in> int; m \<in> int; n \<in> int |]
- ==> (m$*k $< n$*k) \<longleftrightarrow> ((#0 $< k & m$<n) | (k $< #0 & n$<m))"
+ "\<lbrakk>k \<in> int; m \<in> int; n \<in> int\<rbrakk>
+ \<Longrightarrow> (m$*k $< n$*k) \<longleftrightarrow> ((#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
@@ -327,12 +327,12 @@
"(k$*m $\<le> k$*n) \<longleftrightarrow> ((#0 $< k \<longrightarrow> m$\<le>n) & (k $< #0 \<longrightarrow> n$\<le>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 \<longleftrightarrow> (m $\<le> n & n $\<le> m)"
+lemma int_eq_iff_zle: "\<lbrakk>m \<in> int; n \<in> int\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m $\<le> n & n $\<le> m)"
apply (blast intro: zle_refl zle_anti_sym)
done
lemma zmult_cancel2_lemma:
- "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) \<longleftrightarrow> (k=#0 | m=n)"
+ "\<lbrakk>k \<in> int; m \<in> int; n \<in> int\<rbrakk> \<Longrightarrow> (m$*k = n$*k) \<longleftrightarrow> (k=#0 | m=n)"
apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m])
apply (auto simp add: zmult_zle_cancel2 neq_iff_zless)
done
@@ -352,8 +352,8 @@
subsection\<open>Uniqueness and monotonicity of quotients and remainders\<close>
lemma unique_quotient_lemma:
- "[| b$*q' $+ r' $\<le> b$*q $+ r; #0 $\<le> r'; #0 $< b; r $< b |]
- ==> q' $\<le> q"
+ "\<lbrakk>b$*q' $+ r' $\<le> b$*q $+ r; #0 $\<le> r'; #0 $< b; r $< b\<rbrakk>
+ \<Longrightarrow> q' $\<le> q"
apply (subgoal_tac "r' $+ b $* (q'$-q) $\<le> r")
prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls)
apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ")
@@ -370,8 +370,8 @@
done
lemma unique_quotient_lemma_neg:
- "[| b$*q' $+ r' $\<le> b$*q $+ r; r $\<le> #0; b $< #0; b $< r' |]
- ==> q $\<le> q'"
+ "\<lbrakk>b$*q' $+ r' $\<le> b$*q $+ r; r $\<le> #0; b $< #0; b $< r'\<rbrakk>
+ \<Longrightarrow> q $\<le> q'"
apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r"
in unique_quotient_lemma)
apply (auto simp del: zminus_zadd_distrib
@@ -380,8 +380,8 @@
lemma unique_quotient:
- "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;
- q \<in> int; q' \<in> int |] ==> q = q'"
+ "\<lbrakk>quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;
+ q \<in> int; q' \<in> int\<rbrakk> \<Longrightarrow> q = q'"
apply (simp add: split_ifs quorem_def neq_iff_zless)
apply safe
apply simp_all
@@ -391,9 +391,9 @@
done
lemma unique_remainder:
- "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b \<noteq> #0;
+ "\<lbrakk>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'"
+ r \<in> int; r' \<in> int\<rbrakk> \<Longrightarrow> r = r'"
apply (subgoal_tac "q = q'")
prefer 2 apply (blast intro: unique_quotient)
apply (simp add: quorem_def)
@@ -411,8 +411,8 @@
lemma posDivAlg_termination:
- "[| #0 $< b; ~ a $< b |]
- ==> nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)"
+ "\<lbrakk>#0 $< b; \<not> a $< b\<rbrakk>
+ \<Longrightarrow> nat_of(a $- #2 $* 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)
done
@@ -420,7 +420,7 @@
lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure]
lemma posDivAlg_eqn:
- "[| #0 $< b; a \<in> int; b \<in> int |] ==>
+ "\<lbrakk>#0 $< b; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow>
posDivAlg(<a,b>) =
(if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))"
apply (rule posDivAlg_unfold [THEN trans])
@@ -430,8 +430,8 @@
lemma posDivAlg_induct_lemma [rule_format]:
assumes prem:
- "!!a b. [| a \<in> int; b \<in> int;
- ~ (a $< b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>) |] ==> P(<a,b>)"
+ "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
+ \<not> (a $< b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>)\<rbrakk> \<Longrightarrow> P(<a,b>)"
shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"
using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)"]
proof (induct "<u,v>" arbitrary: u v rule: wf_induct)
@@ -449,8 +449,8 @@
lemma posDivAlg_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;
- ~ (a $< b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b) |] ==> P(a,b)"
+ and ih: "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
+ \<not> (a $< b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b)\<rbrakk> \<Longrightarrow> P(a,b)"
shows "P(u,v)"
apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)")
apply simp
@@ -461,28 +461,28 @@
done
(*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!!*)
+ then this rewrite can work for all constants\<And>*)
lemma intify_eq_0_iff_zle: "intify(m) = #0 \<longleftrightarrow> (m $\<le> #0 & #0 $\<le> m)"
by (simp add: int_eq_iff_zle)
subsection\<open>Some convenient biconditionals for products of signs\<close>
-lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j"
+lemma zmult_pos: "\<lbrakk>#0 $< i; #0 $< j\<rbrakk> \<Longrightarrow> #0 $< i $* j"
by (drule zmult_zless_mono1, auto)
-lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j"
+lemma zmult_neg: "\<lbrakk>i $< #0; j $< #0\<rbrakk> \<Longrightarrow> #0 $< i $* j"
by (drule zmult_zless_mono1_neg, auto)
-lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0"
+lemma zmult_pos_neg: "\<lbrakk>#0 $< i; j $< #0\<rbrakk> \<Longrightarrow> i $* j $< #0"
by (drule zmult_zless_mono1_neg, auto)
(** Inequality reasoning **)
lemma int_0_less_lemma:
- "[| x \<in> int; y \<in> int |]
- ==> (#0 $< x $* y) \<longleftrightarrow> (#0 $< x & #0 $< y | x $< #0 & y $< #0)"
+ "\<lbrakk>x \<in> int; y \<in> int\<rbrakk>
+ \<Longrightarrow> (#0 $< x $* y) \<longleftrightarrow> (#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)
@@ -502,8 +502,8 @@
done
lemma int_0_le_lemma:
- "[| x \<in> int; y \<in> int |]
- ==> (#0 $\<le> x $* y) \<longleftrightarrow> (#0 $\<le> x & #0 $\<le> y | x $\<le> #0 & y $\<le> #0)"
+ "\<lbrakk>x \<in> int; y \<in> int\<rbrakk>
+ \<Longrightarrow> (#0 $\<le> x $* y) \<longleftrightarrow> (#0 $\<le> x & #0 $\<le> y | x $\<le> #0 & y $\<le> #0)"
by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff)
lemma int_0_le_mult_iff:
@@ -526,7 +526,7 @@
(*Typechecking for posDivAlg*)
lemma posDivAlg_type [rule_format]:
- "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> posDivAlg(<a,b>) \<in> int * int"
apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
apply assumption+
apply (case_tac "#0 $< ba")
@@ -541,8 +541,8 @@
(*Correctness of posDivAlg: it computes quotients correctly*)
lemma posDivAlg_correct [rule_format]:
- "[| a \<in> int; b \<in> int |]
- ==> #0 $\<le> a \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, posDivAlg(<a,b>))"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk>
+ \<Longrightarrow> #0 $\<le> 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)
@@ -565,8 +565,8 @@
subsection\<open>Correctness of negDivAlg, the division algorithm for a<0 and b>0\<close>
lemma negDivAlg_termination:
- "[| #0 $< b; a $+ b $< #0 |]
- ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)"
+ "\<lbrakk>#0 $< b; a $+ b $< #0\<rbrakk>
+ \<Longrightarrow> 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]
zless_zminus)
@@ -575,7 +575,7 @@
lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure]
lemma negDivAlg_eqn:
- "[| #0 $< b; a \<in> int; b \<in> int |] ==>
+ "\<lbrakk>#0 $< b; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow>
negDivAlg(<a,b>) =
(if #0 $\<le> a$+b then <#-1,a$+b>
else adjust(b, negDivAlg (<a, #2$*b>)))"
@@ -586,9 +586,9 @@
lemma negDivAlg_induct_lemma [rule_format]:
assumes prem:
- "!!a b. [| a \<in> int; b \<in> int;
- ~ (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>) |]
- ==> P(<a,b>)"
+ "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
+ \<not> (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>)\<rbrakk>
+ \<Longrightarrow> P(<a,b>)"
shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"
using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)"]
proof (induct "<u,v>" arbitrary: u v rule: wf_induct)
@@ -605,9 +605,9 @@
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 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b) |]
- ==> P(a,b)"
+ and ih: "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
+ \<not> (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b)\<rbrakk>
+ \<Longrightarrow> P(a,b)"
shows "P(u,v)"
apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)")
apply simp
@@ -620,7 +620,7 @@
(*Typechecking for negDivAlg*)
lemma negDivAlg_type:
- "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> negDivAlg(<a,b>) \<in> int * int"
apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
apply assumption+
apply (case_tac "#0 $< ba")
@@ -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 \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, negDivAlg(<a,b>))"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk>
+ \<Longrightarrow> 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)
@@ -661,7 +661,7 @@
subsection\<open>Existence shown by proving the division algorithm to be correct\<close>
(*the case a=0*)
-lemma quorem_0: "[|b \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)"
+lemma quorem_0: "\<lbrakk>b \<noteq> #0; b \<in> int\<rbrakk> \<Longrightarrow> quorem (<#0,b>, <#0,#0>)"
by (force simp add: quorem_def neq_iff_zless)
lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>"
@@ -676,7 +676,7 @@
(*Needed below. Actually it's an equivalence.*)
-lemma linear_arith_lemma: "~ (#0 $\<le> #-1 $+ b) ==> (b $\<le> #0)"
+lemma linear_arith_lemma: "\<not> (#0 $\<le> #-1 $+ b) \<Longrightarrow> (b $\<le> #0)"
apply (simp add: not_zle_iff_zless)
apply (drule zminus_zless_zminus [THEN iffD2])
apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus)
@@ -692,14 +692,14 @@
apply auto
done
-lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int"
+lemma negateSnd_type: "qr \<in> int * int \<Longrightarrow> negateSnd (qr) \<in> int * int"
apply (unfold negateSnd_def)
apply auto
done
lemma quorem_neg:
- "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|]
- ==> quorem (<a,b>, negateSnd(qr))"
+ "\<lbrakk>quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int\<rbrakk>
+ \<Longrightarrow> quorem (<a,b>, negateSnd(qr))"
apply clarify
apply (auto elim: zless_asym simp add: quorem_def zless_zminus)
txt\<open>linear arithmetic from here on\<close>
@@ -711,7 +711,7 @@
done
lemma divAlg_correct:
- "[|b \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))"
+ "\<lbrakk>b \<noteq> #0; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> 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)
@@ -720,7 +720,7 @@
apply (auto simp add: zle_def)
done
-lemma divAlg_type: "[|a \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> int * int"
+lemma divAlg_type: "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> divAlg(<a,b>) \<in> int * int"
apply (auto simp add: divAlg_def)
apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type)
done
@@ -765,7 +765,7 @@
(** Basic laws about division and remainder **)
lemma raw_zmod_zdiv_equality:
- "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> 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 (cut_tac a = "a" and b = "b" in divAlg_correct)
@@ -778,7 +778,7 @@
apply auto
done
-lemma pos_mod: "#0 $< b ==> #0 $\<le> a zmod b & a zmod b $< b"
+lemma pos_mod: "#0 $< b \<Longrightarrow> #0 $\<le> a zmod b & a zmod b $< b"
apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
apply (blast dest: zle_zless_trans)+
@@ -787,7 +787,7 @@
lemmas pos_mod_sign = pos_mod [THEN conjunct1]
and pos_mod_bound = pos_mod [THEN conjunct2]
-lemma neg_mod: "b $< #0 ==> a zmod b $\<le> #0 & b $< a zmod b"
+lemma neg_mod: "b $< #0 \<Longrightarrow> a zmod b $\<le> #0 & b $< a zmod b"
apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct)
apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def)
apply (blast dest: zle_zless_trans)
@@ -801,8 +801,8 @@
(** proving general properties of zdiv and zmod **)
lemma quorem_div_mod:
- "[|b \<noteq> #0; a \<in> int; b \<in> int |]
- ==> quorem (<a,b>, <a zdiv b, a zmod b>)"
+ "\<lbrakk>b \<noteq> #0; a \<in> int; b \<in> int\<rbrakk>
+ \<Longrightarrow> 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
neg_mod_sign neg_mod_bound)
@@ -810,58 +810,58 @@
(*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 |]
- ==> a zdiv b = q"
+ "\<lbrakk>quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int\<rbrakk>
+ \<Longrightarrow> 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 |]
- ==> a zmod b = r"
+ "\<lbrakk>quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int\<rbrakk>
+ \<Longrightarrow> a zmod b = r"
by (blast intro: quorem_div_mod [THEN unique_remainder])
lemma zdiv_pos_pos_trivial_raw:
- "[| a \<in> int; b \<in> int; #0 $\<le> a; a $< b |] ==> a zdiv b = #0"
+ "\<lbrakk>a \<in> int; b \<in> int; #0 $\<le> a; a $< b\<rbrakk> \<Longrightarrow> a zdiv b = #0"
apply (rule quorem_div)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zle_zless_trans)+
done
-lemma zdiv_pos_pos_trivial: "[| #0 $\<le> a; a $< b |] ==> a zdiv b = #0"
+lemma zdiv_pos_pos_trivial: "\<lbrakk>#0 $\<le> a; a $< b\<rbrakk> \<Longrightarrow> a zdiv b = #0"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_pos_pos_trivial_raw)
apply auto
done
lemma zdiv_neg_neg_trivial_raw:
- "[| a \<in> int; b \<in> int; a $\<le> #0; b $< a |] ==> a zdiv b = #0"
+ "\<lbrakk>a \<in> int; b \<in> int; a $\<le> #0; b $< a\<rbrakk> \<Longrightarrow> a zdiv b = #0"
apply (rule_tac r = "a" in quorem_div)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zle_zless_trans zless_trans)+
done
-lemma zdiv_neg_neg_trivial: "[| a $\<le> #0; b $< a |] ==> a zdiv b = #0"
+lemma zdiv_neg_neg_trivial: "\<lbrakk>a $\<le> #0; b $< a\<rbrakk> \<Longrightarrow> a zdiv b = #0"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_neg_neg_trivial_raw)
apply auto
done
-lemma zadd_le_0_lemma: "[| a$+b $\<le> #0; #0 $< a; #0 $< b |] ==> False"
+lemma zadd_le_0_lemma: "\<lbrakk>a$+b $\<le> #0; #0 $< a; #0 $< b\<rbrakk> \<Longrightarrow> False"
apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono)
apply (auto simp add: zle_def)
apply (blast dest: zless_trans)
done
lemma zdiv_pos_neg_trivial_raw:
- "[| a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0 |] ==> a zdiv b = #-1"
+ "\<lbrakk>a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0\<rbrakk> \<Longrightarrow> a zdiv b = #-1"
apply (rule_tac r = "a $+ b" in quorem_div)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
done
-lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $\<le> #0 |] ==> a zdiv b = #-1"
+lemma zdiv_pos_neg_trivial: "\<lbrakk>#0 $< a; a$+b $\<le> #0\<rbrakk> \<Longrightarrow> a zdiv b = #-1"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zdiv_pos_neg_trivial_raw)
apply auto
@@ -871,42 +871,42 @@
lemma zmod_pos_pos_trivial_raw:
- "[| a \<in> int; b \<in> int; #0 $\<le> a; a $< b |] ==> a zmod b = a"
+ "\<lbrakk>a \<in> int; b \<in> int; #0 $\<le> a; a $< b\<rbrakk> \<Longrightarrow> a zmod b = a"
apply (rule_tac q = "#0" in quorem_mod)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zle_zless_trans)+
done
-lemma zmod_pos_pos_trivial: "[| #0 $\<le> a; a $< b |] ==> a zmod b = intify(a)"
+lemma zmod_pos_pos_trivial: "\<lbrakk>#0 $\<le> a; a $< b\<rbrakk> \<Longrightarrow> a zmod b = intify(a)"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_pos_pos_trivial_raw)
apply auto
done
lemma zmod_neg_neg_trivial_raw:
- "[| a \<in> int; b \<in> int; a $\<le> #0; b $< a |] ==> a zmod b = a"
+ "\<lbrakk>a \<in> int; b \<in> int; a $\<le> #0; b $< a\<rbrakk> \<Longrightarrow> a zmod b = a"
apply (rule_tac q = "#0" in quorem_mod)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zle_zless_trans zless_trans)+
done
-lemma zmod_neg_neg_trivial: "[| a $\<le> #0; b $< a |] ==> a zmod b = intify(a)"
+lemma zmod_neg_neg_trivial: "\<lbrakk>a $\<le> #0; b $< a\<rbrakk> \<Longrightarrow> a zmod b = intify(a)"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_neg_neg_trivial_raw)
apply auto
done
lemma zmod_pos_neg_trivial_raw:
- "[| a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0 |] ==> a zmod b = a$+b"
+ "\<lbrakk>a \<in> int; b \<in> int; #0 $< a; a$+b $\<le> #0\<rbrakk> \<Longrightarrow> a zmod b = a$+b"
apply (rule_tac q = "#-1" in quorem_mod)
apply (auto simp add: quorem_def)
(*linear arithmetic*)
apply (blast dest: zadd_le_0_lemma zle_zless_trans)+
done
-lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $\<le> #0 |] ==> a zmod b = a$+b"
+lemma zmod_pos_neg_trivial: "\<lbrakk>#0 $< a; a$+b $\<le> #0\<rbrakk> \<Longrightarrow> a zmod b = a$+b"
apply (cut_tac a = "intify (a)" and b = "intify (b)"
in zmod_pos_neg_trivial_raw)
apply auto
@@ -918,7 +918,7 @@
(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*)
lemma zdiv_zminus_zminus_raw:
- "[|a \<in> int; b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> ($-a) zdiv ($-b) = a zdiv b"
apply (case_tac "b = #0")
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div])
@@ -932,7 +932,7 @@
(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*)
lemma zmod_zminus_zminus_raw:
- "[|a \<in> int; b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> ($-a) zmod ($-b) = $- (a zmod b)"
apply (case_tac "b = #0")
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod])
@@ -947,7 +947,7 @@
subsection\<open>division of a number by itself\<close>
-lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $\<le> q"
+lemma self_quotient_aux1: "\<lbrakk>#0 $< a; a = r $+ a$*q; r $< a\<rbrakk> \<Longrightarrow> #1 $\<le> q"
apply (subgoal_tac "#0 $< a$*q")
apply (cut_tac w = "#0" and z = "q" in add1_zle_iff)
apply (simp add: int_0_less_mult_iff)
@@ -958,7 +958,7 @@
apply (simp add: zcompare_rls)
done
-lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $\<le> r |] ==> q $\<le> #1"
+lemma self_quotient_aux2: "\<lbrakk>#0 $< a; a = r $+ a$*q; #0 $\<le> r\<rbrakk> \<Longrightarrow> q $\<le> #1"
apply (subgoal_tac "#0 $\<le> a$* (#1$-q)")
apply (simp add: int_0_le_mult_iff zcompare_rls)
apply (blast dest: zle_zless_trans)
@@ -968,7 +968,7 @@
done
lemma self_quotient:
- "[| quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #0|] ==> q = #1"
+ "\<lbrakk>quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #0\<rbrakk> \<Longrightarrow> q = #1"
apply (simp add: split_ifs quorem_def neq_iff_zless)
apply (rule zle_anti_sym)
apply safe
@@ -984,22 +984,22 @@
done
lemma self_remainder:
- "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0"
+ "\<lbrakk>quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0\<rbrakk> \<Longrightarrow> r = #0"
apply (frule self_quotient)
apply (auto simp add: quorem_def)
done
-lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1"
+lemma zdiv_self_raw: "\<lbrakk>a \<noteq> #0; a \<in> int\<rbrakk> \<Longrightarrow> a zdiv a = #1"
apply (blast intro: quorem_div_mod [THEN self_quotient])
done
-lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1"
+lemma zdiv_self [simp]: "intify(a) \<noteq> #0 \<Longrightarrow> a zdiv a = #1"
apply (drule zdiv_self_raw)
apply auto
done
(*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"
+lemma zmod_self_raw: "a \<in> int \<Longrightarrow> a zmod a = #0"
apply (case_tac "a = #0")
apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD)
apply (blast intro: quorem_div_mod [THEN self_remainder])
@@ -1016,29 +1016,29 @@
lemma zdiv_zero [simp]: "#0 zdiv b = #0"
by (simp add: zdiv_def divAlg_def)
-lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
+lemma zdiv_eq_minus1: "#0 $< b \<Longrightarrow> #-1 zdiv b = #-1"
by (simp (no_asm_simp) add: zdiv_def divAlg_def)
lemma zmod_zero [simp]: "#0 zmod b = #0"
by (simp add: zmod_def divAlg_def)
-lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1"
+lemma zdiv_minus1: "#0 $< b \<Longrightarrow> #-1 zdiv b = #-1"
by (simp add: zdiv_def divAlg_def)
-lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1"
+lemma zmod_minus1: "#0 $< b \<Longrightarrow> #-1 zmod b = b $- #1"
by (simp add: zmod_def divAlg_def)
(** a positive, b positive **)
-lemma zdiv_pos_pos: "[| #0 $< a; #0 $\<le> b |]
- ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))"
+lemma zdiv_pos_pos: "\<lbrakk>#0 $< a; #0 $\<le> b\<rbrakk>
+ \<Longrightarrow> 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 $\<le> b |]
- ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))"
+ "\<lbrakk>#0 $< a; #0 $\<le> b\<rbrakk>
+ \<Longrightarrow> 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)
done
@@ -1046,15 +1046,15 @@
(** a negative, b positive **)
lemma zdiv_neg_pos:
- "[| a $< #0; #0 $< b |]
- ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))"
+ "\<lbrakk>a $< #0; #0 $< b\<rbrakk>
+ \<Longrightarrow> 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 zmod b = snd (negDivAlg(<intify(a), intify(b)>))"
+ "\<lbrakk>a $< #0; #0 $< b\<rbrakk>
+ \<Longrightarrow> 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)
done
@@ -1062,8 +1062,8 @@
(** a positive, b negative **)
lemma zdiv_pos_neg:
- "[| #0 $< a; b $< #0 |]
- ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))"
+ "\<lbrakk>#0 $< a; b $< #0\<rbrakk>
+ \<Longrightarrow> 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
apply (blast dest: zle_zless_trans)+
@@ -1072,8 +1072,8 @@
done
lemma zmod_pos_neg:
- "[| #0 $< a; b $< #0 |]
- ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))"
+ "\<lbrakk>#0 $< a; b $< #0\<rbrakk>
+ \<Longrightarrow> 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
apply (blast dest: zle_zless_trans)+
@@ -1084,16 +1084,16 @@
(** a negative, b negative **)
lemma zdiv_neg_neg:
- "[| a $< #0; b $\<le> #0 |]
- ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))"
+ "\<lbrakk>a $< #0; b $\<le> #0\<rbrakk>
+ \<Longrightarrow> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zdiv_def divAlg_def)
apply auto
apply (blast dest!: zle_zless_trans)+
done
lemma zmod_neg_neg:
- "[| a $< #0; b $\<le> #0 |]
- ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))"
+ "\<lbrakk>a $< #0; b $\<le> #0\<rbrakk>
+ \<Longrightarrow> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))"
apply (simp (no_asm_simp) add: zmod_def divAlg_def)
apply auto
apply (blast dest!: zle_zless_trans)+
@@ -1138,7 +1138,7 @@
apply auto
done
-lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a"
+lemma zdiv_minus1_right_raw: "a \<in> int \<Longrightarrow> a zdiv #-1 = $-a"
apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality)
apply auto
apply (rule equation_zminus [THEN iffD2])
@@ -1154,7 +1154,7 @@
subsection\<open>Monotonicity in the first argument (divisor)\<close>
-lemma zdiv_mono1: "[| a $\<le> a'; #0 $< b |] ==> a zdiv b $\<le> a' zdiv b"
+lemma zdiv_mono1: "\<lbrakk>a $\<le> a'; #0 $< b\<rbrakk> \<Longrightarrow> a zdiv b $\<le> a' zdiv b"
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
apply (rule unique_quotient_lemma)
@@ -1163,7 +1163,7 @@
apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound)
done
-lemma zdiv_mono1_neg: "[| a $\<le> a'; b $< #0 |] ==> a' zdiv b $\<le> a zdiv b"
+lemma zdiv_mono1_neg: "\<lbrakk>a $\<le> a'; b $< #0\<rbrakk> \<Longrightarrow> a' zdiv b $\<le> a zdiv b"
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality)
apply (rule unique_quotient_lemma_neg)
@@ -1176,7 +1176,7 @@
subsection\<open>Monotonicity in the second argument (dividend)\<close>
lemma q_pos_lemma:
- "[| #0 $\<le> b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $\<le> q'"
+ "\<lbrakk>#0 $\<le> b'$*q' $+ r'; r' $< b'; #0 $< b'\<rbrakk> \<Longrightarrow> #0 $\<le> q'"
apply (subgoal_tac "#0 $< b'$* (q' $+ #1)")
apply (simp add: int_0_less_mult_iff)
apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1])
@@ -1186,9 +1186,9 @@
done
lemma zdiv_mono2_lemma:
- "[| b$*q $+ r = b'$*q' $+ r'; #0 $\<le> b'$*q' $+ r';
- r' $< b'; #0 $\<le> r; #0 $< b'; b' $\<le> b |]
- ==> q $\<le> q'"
+ "\<lbrakk>b$*q $+ r = b'$*q' $+ r'; #0 $\<le> b'$*q' $+ r';
+ r' $< b'; #0 $\<le> r; #0 $< b'; b' $\<le> b\<rbrakk>
+ \<Longrightarrow> q $\<le> q'"
apply (frule q_pos_lemma, assumption+)
apply (subgoal_tac "b$*q $< b$* (q' $+ #1)")
apply (simp add: zmult_zless_cancel1)
@@ -1207,8 +1207,8 @@
lemma zdiv_mono2_raw:
- "[| #0 $\<le> a; #0 $< b'; b' $\<le> b; a \<in> int |]
- ==> a zdiv b $\<le> a zdiv b'"
+ "\<lbrakk>#0 $\<le> a; #0 $< b'; b' $\<le> b; a \<in> int\<rbrakk>
+ \<Longrightarrow> a zdiv b $\<le> a zdiv b'"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
@@ -1220,14 +1220,14 @@
done
lemma zdiv_mono2:
- "[| #0 $\<le> a; #0 $< b'; b' $\<le> b |]
- ==> a zdiv b $\<le> a zdiv b'"
+ "\<lbrakk>#0 $\<le> a; #0 $< b'; b' $\<le> b\<rbrakk>
+ \<Longrightarrow> a zdiv b $\<le> a zdiv b'"
apply (cut_tac a = "intify (a)" in zdiv_mono2_raw)
apply auto
done
lemma q_neg_lemma:
- "[| b'$*q' $+ r' $< #0; #0 $\<le> r'; #0 $< b' |] ==> q' $< #0"
+ "\<lbrakk>b'$*q' $+ r' $< #0; #0 $\<le> r'; #0 $< b'\<rbrakk> \<Longrightarrow> q' $< #0"
apply (subgoal_tac "b'$*q' $< #0")
prefer 2 apply (force intro: zle_zless_trans)
apply (simp add: zmult_less_0_iff)
@@ -1237,9 +1237,9 @@
lemma zdiv_mono2_neg_lemma:
- "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0;
- r $< b; #0 $\<le> r'; #0 $< b'; b' $\<le> b |]
- ==> q' $\<le> q"
+ "\<lbrakk>b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0;
+ r $< b; #0 $\<le> r'; #0 $< b'; b' $\<le> b\<rbrakk>
+ \<Longrightarrow> q' $\<le> q"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
apply (frule q_neg_lemma, assumption+)
@@ -1266,8 +1266,8 @@
done
lemma zdiv_mono2_neg_raw:
- "[| a $< #0; #0 $< b'; b' $\<le> b; a \<in> int |]
- ==> a zdiv b' $\<le> a zdiv b"
+ "\<lbrakk>a $< #0; #0 $< b'; b' $\<le> b; a \<in> int\<rbrakk>
+ \<Longrightarrow> a zdiv b' $\<le> a zdiv b"
apply (subgoal_tac "#0 $< b")
prefer 2 apply (blast dest: zless_zle_trans)
apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
@@ -1278,8 +1278,8 @@
apply (simp_all add: pos_mod_sign pos_mod_bound)
done
-lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $\<le> b |]
- ==> a zdiv b' $\<le> a zdiv b"
+lemma zdiv_mono2_neg: "\<lbrakk>a $< #0; #0 $< b'; b' $\<le> b\<rbrakk>
+ \<Longrightarrow> a zdiv b' $\<le> a zdiv b"
apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw)
apply auto
done
@@ -1291,16 +1291,16 @@
(** 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 (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)"
+ "\<lbrakk>quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0\<rbrakk>
+ \<Longrightarrow> 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)
done
lemma zdiv_zmult1_eq_raw:
- "[|b \<in> int; c \<in> int|]
- ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c"
+ "\<lbrakk>b \<in> int; c \<in> int\<rbrakk>
+ \<Longrightarrow> (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 (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div])
@@ -1313,7 +1313,7 @@
done
lemma zmod_zmult1_eq_raw:
- "[|b \<in> int; c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c"
+ "\<lbrakk>b \<in> int; c \<in> int\<rbrakk> \<Longrightarrow> (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 (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod])
@@ -1337,10 +1337,10 @@
apply (rule zmod_zmult1_eq)
done
-lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)"
+lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 \<Longrightarrow> (a$*b) zdiv b = intify(a)"
by (simp add: zdiv_zmult1_eq)
-lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)"
+lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 \<Longrightarrow> (b$*a) zdiv b = intify(a)"
by (simp add: zmult_commute)
lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0"
@@ -1354,9 +1354,9 @@
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$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)"
+ "\<lbrakk>quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>);
+ c \<in> int; c \<noteq> #0\<rbrakk>
+ \<Longrightarrow> 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)
apply (auto intro: raw_zmod_zdiv_equality)
@@ -1364,7 +1364,7 @@
(*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|] ==>
+ "\<lbrakk>a \<in> int; b \<in> int; c \<in> int\<rbrakk> \<Longrightarrow>
(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)
@@ -1380,8 +1380,8 @@
done
lemma zmod_zadd1_eq_raw:
- "[|a \<in> int; b \<in> int; c \<in> int|]
- ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c"
+ "\<lbrakk>a \<in> int; b \<in> int; c \<in> int\<rbrakk>
+ \<Longrightarrow> (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,
@@ -1395,7 +1395,7 @@
done
lemma zmod_div_trivial_raw:
- "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> (a zmod b) zdiv b = #0"
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
@@ -1408,7 +1408,7 @@
done
lemma zmod_mod_trivial_raw:
- "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b"
+ "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> (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
@@ -1436,11 +1436,11 @@
lemma zdiv_zadd_self1 [simp]:
- "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1"
+ "intify(a) \<noteq> #0 \<Longrightarrow> (a$+b) zdiv a = b zdiv a $+ #1"
by (simp (no_asm_simp) add: zdiv_zadd1_eq)
lemma zdiv_zadd_self2 [simp]:
- "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1"
+ "intify(a) \<noteq> #0 \<Longrightarrow> (b$+a) zdiv a = b zdiv a $+ #1"
by (simp (no_asm_simp) add: zdiv_zadd1_eq)
lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a"
@@ -1458,14 +1458,14 @@
subsection\<open>proving a zdiv (b*c) = (a zdiv b) zdiv c\<close>
-(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but
- 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems
+(*The condition c>0 seems necessary. Consider that 7 zdiv \<not>6 = \<not>2 but
+ 7 zdiv 2 zdiv \<not>3 = 3 zdiv \<not>3 = \<not>1. The subcase (a zdiv b) zmod c = 0 seems
to cause particular problems.*)
(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **)
lemma zdiv_zmult2_aux1:
- "[| #0 $< c; b $< r; r $\<le> #0 |] ==> b$*c $< b$*(q zmod c) $+ r"
+ "\<lbrakk>#0 $< c; b $< r; r $\<le> #0\<rbrakk> \<Longrightarrow> b$*c $< b$*(q zmod c) $+ r"
apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1")
apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
apply (rule zle_zless_trans)
@@ -1476,7 +1476,7 @@
done
lemma zdiv_zmult2_aux2:
- "[| #0 $< c; b $< r; r $\<le> #0 |] ==> b $* (q zmod c) $+ r $\<le> #0"
+ "\<lbrakk>#0 $< c; b $< r; r $\<le> #0\<rbrakk> \<Longrightarrow> b $* (q zmod c) $+ r $\<le> #0"
apply (subgoal_tac "b $* (q zmod c) $\<le> #0")
prefer 2
apply (simp add: zmult_le_0_iff pos_mod_sign)
@@ -1488,7 +1488,7 @@
done
lemma zdiv_zmult2_aux3:
- "[| #0 $< c; #0 $\<le> r; r $< b |] ==> #0 $\<le> b $* (q zmod c) $+ r"
+ "\<lbrakk>#0 $< c; #0 $\<le> r; r $< b\<rbrakk> \<Longrightarrow> #0 $\<le> b $* (q zmod c) $+ r"
apply (subgoal_tac "#0 $\<le> b $* (q zmod c)")
prefer 2
apply (simp add: int_0_le_mult_iff pos_mod_sign)
@@ -1500,7 +1500,7 @@
done
lemma zdiv_zmult2_aux4:
- "[| #0 $< c; #0 $\<le> r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c"
+ "\<lbrakk>#0 $< c; #0 $\<le> r; r $< b\<rbrakk> \<Longrightarrow> b $* (q zmod c) $+ r $< b $* c"
apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)")
apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls)
apply (rule zless_zle_trans)
@@ -1511,8 +1511,8 @@
done
lemma zdiv_zmult2_lemma:
- "[| 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>)"
+ "\<lbrakk>quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c\<rbrakk>
+ \<Longrightarrow> 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
zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2
@@ -1521,7 +1521,7 @@
done
lemma zdiv_zmult2_eq_raw:
- "[|#0 $< c; a \<in> int; b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
+ "\<lbrakk>#0 $< c; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> 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 (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div])
@@ -1529,14 +1529,14 @@
apply (blast dest: zle_zless_trans)
done
-lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c"
+lemma zdiv_zmult2_eq: "#0 $< c \<Longrightarrow> a zdiv (b$*c) = (a zdiv b) zdiv c"
apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw)
apply auto
done
lemma zmod_zmult2_eq_raw:
- "[|#0 $< c; a \<in> int; b \<in> int|]
- ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
+ "\<lbrakk>#0 $< c; a \<in> int; b \<in> int\<rbrakk>
+ \<Longrightarrow> 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 (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod])
@@ -1545,7 +1545,7 @@
done
lemma zmod_zmult2_eq:
- "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
+ "#0 $< c \<Longrightarrow> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b"
apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw)
apply auto
done
@@ -1553,32 +1553,32 @@
subsection\<open>Cancellation of common factors in "zdiv"\<close>
lemma zdiv_zmult_zmult1_aux1:
- "[| #0 $< b; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
+ "\<lbrakk>#0 $< b; intify(c) \<noteq> #0\<rbrakk> \<Longrightarrow> (c$*a) zdiv (c$*b) = a zdiv b"
apply (subst zdiv_zmult2_eq)
apply auto
done
lemma zdiv_zmult_zmult1_aux2:
- "[| b $< #0; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b"
+ "\<lbrakk>b $< #0; intify(c) \<noteq> #0\<rbrakk> \<Longrightarrow> (c$*a) zdiv (c$*b) = a zdiv b"
apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)")
apply (rule_tac [2] zdiv_zmult_zmult1_aux1)
apply auto
done
lemma zdiv_zmult_zmult1_raw:
- "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b"
+ "\<lbrakk>intify(c) \<noteq> #0; b \<in> int\<rbrakk> \<Longrightarrow> (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 (auto simp add: neq_iff_zless [of b]
zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2)
done
-lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b"
+lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 \<Longrightarrow> (c$*a) zdiv (c$*b) = a zdiv b"
apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw)
apply auto
done
-lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b"
+lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 \<Longrightarrow> (a$*c) zdiv (b$*c) = a zdiv b"
apply (drule zdiv_zmult_zmult1)
apply (auto simp add: zmult_commute)
done
@@ -1587,22 +1587,22 @@
subsection\<open>Distribution of factors over "zmod"\<close>
lemma zmod_zmult_zmult1_aux1:
- "[| #0 $< b; intify(c) \<noteq> #0 |]
- ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
+ "\<lbrakk>#0 $< b; intify(c) \<noteq> #0\<rbrakk>
+ \<Longrightarrow> (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 |]
- ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
+ "\<lbrakk>b $< #0; intify(c) \<noteq> #0\<rbrakk>
+ \<Longrightarrow> (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)
apply auto
done
lemma zmod_zmult_zmult1_raw:
- "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)"
+ "\<lbrakk>b \<in> int; c \<in> int\<rbrakk> \<Longrightarrow> (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 (case_tac "c = #0")
@@ -1624,7 +1624,7 @@
(** Quotients of signs **)
-lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0"
+lemma zdiv_neg_pos_less0: "\<lbrakk>a $< #0; #0 $< b\<rbrakk> \<Longrightarrow> a zdiv b $< #0"
apply (subgoal_tac "a zdiv b $\<le> #-1")
apply (erule zle_zless_trans)
apply (simp (no_asm))
@@ -1635,12 +1635,12 @@
apply (auto simp add: zdiv_minus1)
done
-lemma zdiv_nonneg_neg_le0: "[| #0 $\<le> a; b $< #0 |] ==> a zdiv b $\<le> #0"
+lemma zdiv_nonneg_neg_le0: "\<lbrakk>#0 $\<le> a; b $< #0\<rbrakk> \<Longrightarrow> a zdiv b $\<le> #0"
apply (drule zdiv_mono1_neg)
apply auto
done
-lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $\<le> a zdiv b) \<longleftrightarrow> (#0 $\<le> a)"
+lemma pos_imp_zdiv_nonneg_iff: "#0 $< b \<Longrightarrow> (#0 $\<le> a zdiv b) \<longleftrightarrow> (#0 $\<le> a)"
apply auto
apply (drule_tac [2] zdiv_mono1)
apply (auto simp add: neq_iff_zless)
@@ -1648,7 +1648,7 @@
apply (blast intro: zdiv_neg_pos_less0)
done
-lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $\<le> a zdiv b) \<longleftrightarrow> (a $\<le> #0)"
+lemma neg_imp_zdiv_nonneg_iff: "b $< #0 \<Longrightarrow> (#0 $\<le> a zdiv b) \<longleftrightarrow> (a $\<le> #0)"
apply (subst zdiv_zminus_zminus [symmetric])
apply (rule iff_trans)
apply (rule pos_imp_zdiv_nonneg_iff)
@@ -1656,13 +1656,13 @@
done
(*But not (a zdiv b $\<le> 0 iff a$\<le>0); consider a=1, b=2 when a zdiv b = 0.*)
-lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) \<longleftrightarrow> (a $< #0)"
+lemma pos_imp_zdiv_neg_iff: "#0 $< b \<Longrightarrow> (a zdiv b $< #0) \<longleftrightarrow> (a $< #0)"
apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
apply (erule pos_imp_zdiv_nonneg_iff)
done
(*Again the law fails for $\<le>: consider a = -1, b = -2 when a zdiv b = 0*)
-lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) \<longleftrightarrow> (#0 $< a)"
+lemma neg_imp_zdiv_neg_iff: "b $< #0 \<Longrightarrow> (a zdiv b $< #0) \<longleftrightarrow> (#0 $< a)"
apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym])
apply (erule neg_imp_zdiv_nonneg_iff)
done
@@ -1674,7 +1674,7 @@
(** computing "zdiv" by shifting **)
- lemma pos_zdiv_mult_2: "#0 $\<le> a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a"
+ lemma pos_zdiv_mult_2: "#0 $\<le> a \<Longrightarrow> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a"
apply (case_tac "a = #0")
apply (subgoal_tac "#1 $\<le> a")
apply (arith_tac 2)
@@ -1694,7 +1694,7 @@
done
- lemma neg_zdiv_mult_2: "a $\<le> #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) \<longleftrightarrow> (b$+#1) zdiv a"
+ lemma neg_zdiv_mult_2: "a $\<le> #0 \<Longrightarrow> (#1 $+ #2$*b) zdiv (#2$*a) \<longleftrightarrow> (b$+#1) zdiv a"
apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) \<longleftrightarrow> ($-b-#1) zdiv ($-a)")
apply (rule_tac [2] pos_zdiv_mult_2)
apply (auto simp add: zmult_zminus_right)
@@ -1706,12 +1706,12 @@
(*Not clear why this must be proved separately; probably integ_of causes
simplification problems*)
- lemma lemma: "~ #0 $\<le> x ==> x $\<le> #0"
+ lemma lemma: "\<not> #0 $\<le> x \<Longrightarrow> x $\<le> #0"
apply auto
done
lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) =
- (if ~b | #0 $\<le> integ_of w
+ (if \<not>b | #0 $\<le> integ_of w
then integ_of v zdiv (integ_of w)
else (integ_of v $+ #1) zdiv (integ_of w))"
apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT)
@@ -1723,7 +1723,7 @@
(** computing "zmod" by shifting (proofs resemble those for "zdiv") **)
- lemma pos_zmod_mult_2: "#0 $\<le> a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)"
+ lemma pos_zmod_mult_2: "#0 $\<le> a \<Longrightarrow> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)"
apply (case_tac "a = #0")
apply (subgoal_tac "#1 $\<le> a")
apply (arith_tac 2)
@@ -1743,7 +1743,7 @@
done
- lemma neg_zmod_mult_2: "a $\<le> #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1"
+ lemma neg_zmod_mult_2: "a $\<le> #0 \<Longrightarrow> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1"
apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))")
apply (rule_tac [2] pos_zmod_mult_2)
apply (auto simp add: zmult_zminus_right)
--- a/src/ZF/List.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/List.thy Tue Sep 27 16:51:35 2022 +0100
@@ -93,19 +93,19 @@
definition
(* Function `take' returns the first n elements of a list *)
take :: "[i,i]=>i" where
- "take(n, as) == list_rec(\<lambda>n\<in>nat. [],
+ "take(n, as) \<equiv> 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
\<comment> \<open>returns the (n+1)th element of a list, or 0 if the
list is too short.\<close>
- "nth(n, as) == list_rec(\<lambda>n\<in>nat. 0,
+ "nth(n, as) \<equiv> 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(\<lambda>n\<in>nat. Nil,
+ "list_update(xs, i, v) \<equiv> 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
@@ -123,11 +123,11 @@
definition
min :: "[i,i] =>i" where
- "min(x, y) == (if x \<le> y then x else y)"
+ "min(x, y) \<equiv> (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) \<equiv> (if x \<le> y then y else x)"
(*** Aspects of the datatype definition ***)
@@ -143,7 +143,7 @@
lemma Cons_iff: "Cons(a,l)=Cons(a',l') \<longleftrightarrow> a=a' & l=l'"
by auto
-lemma Nil_Cons_iff: "~ Nil=Cons(a,l)"
+lemma Nil_Cons_iff: "\<not> Nil=Cons(a,l)"
by auto
lemma list_unfold: "list(A) = {0} + (A * list(A))"
@@ -153,7 +153,7 @@
(** Lemmas to justify using "list" in other recursive type definitions **)
-lemma list_mono: "A<=B ==> list(A) \<subseteq> list(B)"
+lemma list_mono: "A<=B \<Longrightarrow> list(A) \<subseteq> list(B)"
apply (unfold list.defs )
apply (rule lfp_mono)
apply (simp_all add: list.bnd_mono)
@@ -171,14 +171,14 @@
(*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 \<in> list(A); A \<subseteq> univ(B) |] ==> l \<in> univ(B)"
+lemma list_into_univ: "\<lbrakk>l \<in> list(A); A \<subseteq> univ(B)\<rbrakk> \<Longrightarrow> l \<in> univ(B)"
by (blast intro: list_subset_univ [THEN subsetD])
lemma list_case_type:
- "[| l \<in> list(A);
+ "\<lbrakk>l \<in> list(A);
c \<in> C(Nil);
- !!x y. [| x \<in> A; y \<in> list(A) |] ==> h(x,y): C(Cons(x,y))
- |] ==> list_case(c,h,l) \<in> C(l)"
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> list(A)\<rbrakk> \<Longrightarrow> h(x,y): C(Cons(x,y))
+\<rbrakk> \<Longrightarrow> list_case(c,h,l) \<in> C(l)"
by (erule list.induct, auto)
lemma list_0_triv: "list(0) = {Nil}"
@@ -189,26 +189,26 @@
(*** List functions ***)
-lemma tl_type: "l \<in> list(A) ==> tl(l) \<in> list(A)"
+lemma tl_type: "l \<in> list(A) \<Longrightarrow> tl(l) \<in> list(A)"
apply (induct_tac "l")
apply (simp_all (no_asm_simp) add: list.intros)
done
(** drop **)
-lemma drop_Nil [simp]: "i \<in> nat ==> drop(i, Nil) = Nil"
+lemma drop_Nil [simp]: "i \<in> nat \<Longrightarrow> drop(i, Nil) = Nil"
apply (induct_tac "i")
apply (simp_all (no_asm_simp))
done
-lemma drop_succ_Cons [simp]: "i \<in> nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"
+lemma drop_succ_Cons [simp]: "i \<in> nat \<Longrightarrow> drop(succ(i), Cons(a,l)) = drop(i,l)"
apply (rule sym)
apply (induct_tac "i")
apply (simp (no_asm))
apply (simp (no_asm_simp))
done
-lemma drop_type [simp,TC]: "[| i \<in> nat; l \<in> list(A) |] ==> drop(i,l) \<in> list(A)"
+lemma drop_type [simp,TC]: "\<lbrakk>i \<in> nat; l \<in> list(A)\<rbrakk> \<Longrightarrow> drop(i,l) \<in> list(A)"
apply (induct_tac "i")
apply (simp_all (no_asm_simp) add: tl_type)
done
@@ -219,60 +219,60 @@
(** Type checking -- proved by induction, as usual **)
lemma list_rec_type [TC]:
- "[| l \<in> list(A);
+ "\<lbrakk>l \<in> list(A);
c \<in> C(Nil);
- !!x y r. [| x \<in> A; y \<in> list(A); r \<in> C(y) |] ==> h(x,y,r): C(Cons(x,y))
- |] ==> list_rec(c,h,l) \<in> C(l)"
+ \<And>x y r. \<lbrakk>x \<in> A; y \<in> list(A); r \<in> C(y)\<rbrakk> \<Longrightarrow> h(x,y,r): C(Cons(x,y))
+\<rbrakk> \<Longrightarrow> list_rec(c,h,l) \<in> C(l)"
by (induct_tac "l", auto)
(** map **)
lemma map_type [TC]:
- "[| l \<in> list(A); !!x. x \<in> A ==> h(x): B |] ==> map(h,l) \<in> list(B)"
+ "\<lbrakk>l \<in> list(A); \<And>x. x \<in> A \<Longrightarrow> h(x): B\<rbrakk> \<Longrightarrow> 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 \<in> list(A) ==> map(h,l) \<in> list({h(u). u \<in> A})"
+lemma map_type2 [TC]: "l \<in> list(A) \<Longrightarrow> map(h,l) \<in> list({h(u). u \<in> A})"
apply (erule map_type)
apply (erule RepFunI)
done
(** length **)
-lemma length_type [TC]: "l \<in> list(A) ==> length(l) \<in> nat"
+lemma length_type [TC]: "l \<in> list(A) \<Longrightarrow> 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"
+ "\<lbrakk>x < length(xs); xs \<in> list(A)\<rbrakk> \<Longrightarrow> x \<in> nat"
by (frule lt_nat_in_nat, typecheck)
(** app **)
-lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys \<in> list(A)"
+lemma app_type [TC]: "\<lbrakk>xs: list(A); ys: list(A)\<rbrakk> \<Longrightarrow> xs@ys \<in> list(A)"
by (simp add: app_list_def)
(** rev **)
-lemma rev_type [TC]: "xs: list(A) ==> rev(xs) \<in> list(A)"
+lemma rev_type [TC]: "xs: list(A) \<Longrightarrow> rev(xs) \<in> list(A)"
by (simp add: rev_list_def)
(** flat **)
-lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) \<in> list(A)"
+lemma flat_type [TC]: "ls: list(list(A)) \<Longrightarrow> flat(ls) \<in> list(A)"
by (simp add: flat_list_def)
(** set_of_list **)
-lemma set_of_list_type [TC]: "l \<in> list(A) ==> set_of_list(l) \<in> Pow(A)"
+lemma set_of_list_type [TC]: "l \<in> list(A) \<Longrightarrow> 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) \<union> set_of_list(ys)"
+ "xs: list(A) \<Longrightarrow> 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,34 +280,34 @@
(** list_add **)
-lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \<in> nat"
+lemma list_add_type [TC]: "xs: list(nat) \<Longrightarrow> list_add(xs) \<in> nat"
by (simp add: list_add_list_def)
(*** theorems about map ***)
-lemma map_ident [simp]: "l \<in> list(A) ==> map(%u. u, l) = l"
+lemma map_ident [simp]: "l \<in> list(A) \<Longrightarrow> map(%u. u, l) = l"
apply (induct_tac "l")
apply (simp_all (no_asm_simp))
done
-lemma map_compose: "l \<in> list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"
+lemma map_compose: "l \<in> list(A) \<Longrightarrow> map(h, map(j,l)) = map(%u. h(j(u)), l)"
apply (induct_tac "l")
apply (simp_all (no_asm_simp))
done
-lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
+lemma map_app_distrib: "xs: list(A) \<Longrightarrow> map(h, xs@ys) = map(h,xs) @ map(h,ys)"
apply (induct_tac "xs")
apply (simp_all (no_asm_simp))
done
-lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"
+lemma map_flat: "ls: list(list(A)) \<Longrightarrow> map(h, flat(ls)) = flat(map(map(h),ls))"
apply (induct_tac "ls")
apply (simp_all (no_asm_simp) add: map_app_distrib)
done
lemma list_rec_map:
- "l \<in> list(A) ==>
+ "l \<in> list(A) \<Longrightarrow>
list_rec(c, d, map(h,l)) =
list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"
apply (induct_tac "l")
@@ -316,31 +316,31 @@
(** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **)
-(* @{term"c \<in> list(Collect(B,P)) ==> c \<in> list"} *)
+(* @{term"c \<in> list(Collect(B,P)) \<Longrightarrow> c \<in> list"} *)
lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD]
-lemma map_list_Collect: "l \<in> list({x \<in> A. h(x)=j(x)}) ==> map(h,l) = map(j,l)"
+lemma map_list_Collect: "l \<in> list({x \<in> A. h(x)=j(x)}) \<Longrightarrow> map(h,l) = map(j,l)"
apply (induct_tac "l")
apply (simp_all (no_asm_simp))
done
(*** theorems about length ***)
-lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)"
+lemma length_map [simp]: "xs: list(A) \<Longrightarrow> length(map(h,xs)) = length(xs)"
by (induct_tac "xs", simp_all)
lemma length_app [simp]:
- "[| xs: list(A); ys: list(A) |]
- ==> length(xs@ys) = length(xs) #+ length(ys)"
+ "\<lbrakk>xs: list(A); ys: list(A)\<rbrakk>
+ \<Longrightarrow> length(xs@ys) = length(xs) #+ length(ys)"
by (induct_tac "xs", simp_all)
-lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)"
+lemma length_rev [simp]: "xs: list(A) \<Longrightarrow> length(rev(xs)) = length(xs)"
apply (induct_tac "xs")
apply (simp_all (no_asm_simp) add: length_app)
done
lemma length_flat:
- "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"
+ "ls: list(list(A)) \<Longrightarrow> length(flat(ls)) = list_add(map(length,ls))"
apply (induct_tac "ls")
apply (simp_all (no_asm_simp) add: length_app)
done
@@ -349,12 +349,12 @@
(*Lemma for the inductive step of drop_length*)
lemma drop_length_Cons [rule_format]:
- "xs: list(A) ==>
+ "xs: list(A) \<Longrightarrow>
\<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 \<in> list(A) ==> \<forall>i \<in> length(l). (\<exists>z zs. drop(i,l) = Cons(z,zs))"
+ "l \<in> list(A) \<Longrightarrow> \<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)
@@ -365,20 +365,20 @@
(*** theorems about app ***)
-lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs"
+lemma app_right_Nil [simp]: "xs: list(A) \<Longrightarrow> xs@Nil=xs"
by (erule list.induct, simp_all)
-lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"
+lemma app_assoc: "xs: list(A) \<Longrightarrow> (xs@ys)@zs = xs@(ys@zs)"
by (induct_tac "xs", simp_all)
-lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"
+lemma flat_app_distrib: "ls: list(list(A)) \<Longrightarrow> flat(ls@ms) = flat(ls)@flat(ms)"
apply (induct_tac "ls")
apply (simp_all (no_asm_simp) add: app_assoc)
done
(*** theorems about rev ***)
-lemma rev_map_distrib: "l \<in> list(A) ==> rev(map(h,l)) = map(h,rev(l))"
+lemma rev_map_distrib: "l \<in> list(A) \<Longrightarrow> rev(map(h,l)) = map(h,rev(l))"
apply (induct_tac "l")
apply (simp_all (no_asm_simp) add: map_app_distrib)
done
@@ -388,17 +388,17 @@
rev_type does not instantiate ?A. Only the premises do.
*)
lemma rev_app_distrib:
- "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"
+ "\<lbrakk>xs: list(A); ys: list(A)\<rbrakk> \<Longrightarrow> rev(xs@ys) = rev(ys)@rev(xs)"
apply (erule list.induct)
apply (simp_all add: app_assoc)
done
-lemma rev_rev_ident [simp]: "l \<in> list(A) ==> rev(rev(l))=l"
+lemma rev_rev_ident [simp]: "l \<in> list(A) \<Longrightarrow> rev(rev(l))=l"
apply (induct_tac "l")
apply (simp_all (no_asm_simp) add: rev_app_distrib)
done
-lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"
+lemma rev_flat: "ls: list(list(A)) \<Longrightarrow> rev(flat(ls)) = flat(map(rev,rev(ls)))"
apply (induct_tac "ls")
apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib)
done
@@ -407,18 +407,18 @@
(*** theorems about list_add ***)
lemma list_add_app:
- "[| xs: list(nat); ys: list(nat) |]
- ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
+ "\<lbrakk>xs: list(nat); ys: list(nat)\<rbrakk>
+ \<Longrightarrow> list_add(xs@ys) = list_add(ys) #+ list_add(xs)"
apply (induct_tac "xs", simp_all)
done
-lemma list_add_rev: "l \<in> list(nat) ==> list_add(rev(l)) = list_add(l)"
+lemma list_add_rev: "l \<in> list(nat) \<Longrightarrow> list_add(rev(l)) = list_add(l)"
apply (induct_tac "l")
apply (simp_all (no_asm_simp) add: list_add_app)
done
lemma list_add_flat:
- "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"
+ "ls: list(list(nat)) \<Longrightarrow> list_add(flat(ls)) = list_add(map(list_add,ls))"
apply (induct_tac "ls")
apply (simp_all (no_asm_simp) add: list_add_app)
done
@@ -426,30 +426,30 @@
(** New induction rules **)
lemma list_append_induct [case_names Nil snoc, consumes 1]:
- "[| l \<in> list(A);
+ "\<lbrakk>l \<in> list(A);
P(Nil);
- !!x y. [| x \<in> A; y \<in> list(A); P(y) |] ==> P(y @ [x])
- |] ==> P(l)"
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> list(A); P(y)\<rbrakk> \<Longrightarrow> P(y @ [x])
+\<rbrakk> \<Longrightarrow> P(l)"
apply (subgoal_tac "P(rev(rev(l)))", simp)
apply (erule rev_type [THEN list.induct], simp_all)
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) \<longrightarrow> P(l')|]
- ==> P(l)"
- shows "n \<in> nat ==> \<forall>l \<in> list(A). length(l) < n \<longrightarrow> P(l)"
+ "\<And>l. \<lbrakk>l \<in> list(A);
+ \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')\<rbrakk>
+ \<Longrightarrow> P(l)"
+ shows "n \<in> nat \<Longrightarrow> \<forall>l \<in> list(A). length(l) < n \<longrightarrow> P(l)"
apply (induct_tac n, simp)
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) \<longrightarrow> P(l')|]
- ==> P(l)
- |] ==> P(l)"
+ "\<lbrakk>l \<in> list(A);
+ \<And>l. \<lbrakk>l \<in> list(A);
+ \<forall>l' \<in> list(A). length(l') < length(l) \<longrightarrow> P(l')\<rbrakk>
+ \<Longrightarrow> P(l)
+\<rbrakk> \<Longrightarrow> P(l)"
apply (rule list_complete_induct_lemma [of A])
prefer 4 apply (rule le_refl, simp)
apply blast
@@ -462,31 +462,31 @@
(** min FIXME: replace by Int! **)
(* Min theorems are also true for i, j ordinals *)
-lemma min_sym: "[| i \<in> nat; j \<in> nat |] ==> min(i,j)=min(j,i)"
+lemma min_sym: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> min(i,j)=min(j,i)"
apply (unfold min_def)
apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym)
done
-lemma min_type [simp,TC]: "[| i \<in> nat; j \<in> nat |] ==> min(i,j):nat"
+lemma min_type [simp,TC]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> min(i,j):nat"
by (unfold min_def, auto)
-lemma min_0 [simp]: "i \<in> nat ==> min(0,i) = 0"
+lemma min_0 [simp]: "i \<in> nat \<Longrightarrow> min(0,i) = 0"
apply (unfold min_def)
apply (auto dest: not_lt_imp_le)
done
-lemma min_02 [simp]: "i \<in> nat ==> min(i, 0) = 0"
+lemma min_02 [simp]: "i \<in> nat \<Longrightarrow> min(i, 0) = 0"
apply (unfold min_def)
apply (auto dest: not_lt_imp_le)
done
-lemma lt_min_iff: "[| i \<in> nat; j \<in> nat; k \<in> nat |] ==> i<min(j,k) \<longleftrightarrow> i<j & i<k"
+lemma lt_min_iff: "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> i<min(j,k) \<longleftrightarrow> i<j & i<k"
apply (unfold min_def)
apply (auto dest!: not_lt_imp_le intro: lt_trans2 lt_trans)
done
lemma min_succ_succ [simp]:
- "[| i \<in> nat; j \<in> nat |] ==> min(succ(i), succ(j))= succ(min(i, j))"
+ "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> min(succ(i), succ(j))= succ(min(i, j))"
apply (unfold min_def, auto)
done
@@ -495,65 +495,65 @@
(** filter **)
lemma filter_append [simp]:
- "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
+ "xs:list(A) \<Longrightarrow> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)"
by (induct_tac "xs", auto)
-lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)"
+lemma filter_type [simp,TC]: "xs:list(A) \<Longrightarrow> 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) \<Longrightarrow> 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)) \<subseteq> set_of_list(xs)"
+lemma filter_is_subset: "xs:list(A) \<Longrightarrow> 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"
+lemma filter_False [simp]: "xs:list(A) \<Longrightarrow> filter(%p. False, xs) = Nil"
by (induct_tac "xs", auto)
-lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs"
+lemma filter_True [simp]: "xs:list(A) \<Longrightarrow> filter(%p. True, xs) = xs"
by (induct_tac "xs", auto)
(** length **)
-lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 \<longleftrightarrow> xs=Nil"
+lemma length_is_0_iff [simp]: "xs:list(A) \<Longrightarrow> length(xs)=0 \<longleftrightarrow> xs=Nil"
by (erule list.induct, auto)
-lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) \<longleftrightarrow> xs=Nil"
+lemma length_is_0_iff2 [simp]: "xs:list(A) \<Longrightarrow> 0 = length(xs) \<longleftrightarrow> xs=Nil"
by (erule list.induct, auto)
-lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1"
+lemma length_tl [simp]: "xs:list(A) \<Longrightarrow> length(tl(xs)) = length(xs) #- 1"
by (erule list.induct, auto)
-lemma length_greater_0_iff: "xs:list(A) ==> 0<length(xs) \<longleftrightarrow> xs \<noteq> Nil"
+lemma length_greater_0_iff: "xs:list(A) \<Longrightarrow> 0<length(xs) \<longleftrightarrow> xs \<noteq> Nil"
by (erule list.induct, auto)
-lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) \<longleftrightarrow> (\<exists>y ys. xs=Cons(y, ys) & length(ys)=n)"
+lemma length_succ_iff: "xs:list(A) \<Longrightarrow> length(xs)=succ(n) \<longleftrightarrow> (\<exists>y ys. xs=Cons(y, ys) & length(ys)=n)"
by (erule list.induct, auto)
(** more theorems about append **)
lemma append_is_Nil_iff [simp]:
- "xs:list(A) ==> (xs@ys = Nil) \<longleftrightarrow> (xs=Nil & ys = Nil)"
+ "xs:list(A) \<Longrightarrow> (xs@ys = Nil) \<longleftrightarrow> (xs=Nil & ys = Nil)"
by (erule list.induct, auto)
lemma append_is_Nil_iff2 [simp]:
- "xs:list(A) ==> (Nil = xs@ys) \<longleftrightarrow> (xs=Nil & ys = Nil)"
+ "xs:list(A) \<Longrightarrow> (Nil = xs@ys) \<longleftrightarrow> (xs=Nil & ys = Nil)"
by (erule list.induct, auto)
lemma append_left_is_self_iff [simp]:
- "xs:list(A) ==> (xs@ys = xs) \<longleftrightarrow> (ys = Nil)"
+ "xs:list(A) \<Longrightarrow> (xs@ys = xs) \<longleftrightarrow> (ys = Nil)"
by (erule list.induct, auto)
lemma append_left_is_self_iff2 [simp]:
- "xs:list(A) ==> (xs = xs@ys) \<longleftrightarrow> (ys = Nil)"
+ "xs:list(A) \<Longrightarrow> (xs = xs@ys) \<longleftrightarrow> (ys = Nil)"
by (erule list.induct, auto)
(*TOO SLOW as a default simprule!*)
lemma append_left_is_Nil_iff [rule_format]:
- "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
+ "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
length(ys)=length(zs) \<longrightarrow> (xs@ys=zs \<longleftrightarrow> (xs=Nil & ys=zs))"
apply (erule list.induct)
apply (auto simp add: length_app)
@@ -561,14 +561,14 @@
(*TOO SLOW as a default simprule!*)
lemma append_left_is_Nil_iff2 [rule_format]:
- "[| xs:list(A); ys:list(A); zs:list(A) |] ==>
+ "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow>
length(ys)=length(zs) \<longrightarrow> (zs=ys@xs \<longleftrightarrow> (xs=Nil & ys=zs))"
apply (erule list.induct)
apply (auto simp add: length_app)
done
lemma append_eq_append_iff [rule_format]:
- "xs:list(A) ==> \<forall>ys \<in> list(A).
+ "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A).
length(xs)=length(ys) \<longrightarrow> (xs@us = ys@vs) \<longleftrightarrow> (xs=ys & us=vs)"
apply (erule list.induct)
apply (simp (no_asm_simp))
@@ -578,7 +578,7 @@
declare append_eq_append_iff [simp]
lemma append_eq_append [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<forall>ys \<in> list(A). \<forall>us \<in> list(A). \<forall>vs \<in> list(A).
length(us) = length(vs) \<longrightarrow> (xs@us = ys@vs) \<longrightarrow> (xs=ys & us=vs)"
apply (induct_tac "xs")
@@ -589,24 +589,24 @@
done
lemma append_eq_append_iff2 [simp]:
- "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |]
- ==> xs@us = ys@vs \<longleftrightarrow> (xs=ys & us=vs)"
+ "\<lbrakk>xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs)\<rbrakk>
+ \<Longrightarrow> xs@us = ys@vs \<longleftrightarrow> (xs=ys & us=vs)"
apply (rule iffI)
apply (rule append_eq_append, auto)
done
lemma append_self_iff [simp]:
- "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs \<longleftrightarrow> ys=zs"
+ "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow> xs@ys=xs@zs \<longleftrightarrow> ys=zs"
by simp
lemma append_self_iff2 [simp]:
- "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs \<longleftrightarrow> ys=zs"
+ "\<lbrakk>xs:list(A); ys:list(A); zs:list(A)\<rbrakk> \<Longrightarrow> ys@xs=zs@xs \<longleftrightarrow> ys=zs"
by simp
(* Can also be proved from append_eq_append_iff2,
but the proof requires two more hypotheses: x \<in> A and y \<in> A *)
lemma append1_eq_iff [rule_format]:
- "xs:list(A) ==> \<forall>ys \<in> list(A). xs@[x] = ys@[y] \<longleftrightarrow> (xs = ys & x=y)"
+ "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). xs@[x] = ys@[y] \<longleftrightarrow> (xs = ys & x=y)"
apply (erule list.induct)
apply clarify
apply (erule list.cases)
@@ -618,41 +618,41 @@
declare append1_eq_iff [simp]
lemma append_right_is_self_iff [simp]:
- "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) \<longleftrightarrow> (xs=Nil)"
+ "\<lbrakk>xs:list(A); ys:list(A)\<rbrakk> \<Longrightarrow> (xs@ys = ys) \<longleftrightarrow> (xs=Nil)"
by (simp (no_asm_simp) add: append_left_is_Nil_iff)
lemma append_right_is_self_iff2 [simp]:
- "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) \<longleftrightarrow> (xs=Nil)"
+ "\<lbrakk>xs:list(A); ys:list(A)\<rbrakk> \<Longrightarrow> (ys = xs@ys) \<longleftrightarrow> (xs=Nil)"
apply (rule iffI)
apply (drule sym, auto)
done
lemma hd_append [rule_format]:
- "xs:list(A) ==> xs \<noteq> Nil \<longrightarrow> hd(xs @ ys) = hd(xs)"
+ "xs:list(A) \<Longrightarrow> xs \<noteq> Nil \<longrightarrow> hd(xs @ ys) = hd(xs)"
by (induct_tac "xs", auto)
declare hd_append [simp]
lemma tl_append [rule_format]:
- "xs:list(A) ==> xs\<noteq>Nil \<longrightarrow> tl(xs @ ys) = tl(xs)@ys"
+ "xs:list(A) \<Longrightarrow> xs\<noteq>Nil \<longrightarrow> tl(xs @ ys) = tl(xs)@ys"
by (induct_tac "xs", auto)
declare tl_append [simp]
(** rev **)
-lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil \<longleftrightarrow> xs = Nil)"
+lemma rev_is_Nil_iff [simp]: "xs:list(A) \<Longrightarrow> (rev(xs) = Nil \<longleftrightarrow> xs = Nil)"
by (erule list.induct, auto)
-lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) \<longleftrightarrow> xs = Nil)"
+lemma Nil_is_rev_iff [simp]: "xs:list(A) \<Longrightarrow> (Nil = rev(xs) \<longleftrightarrow> xs = Nil)"
by (erule list.induct, auto)
lemma rev_is_rev_iff [rule_format]:
- "xs:list(A) ==> \<forall>ys \<in> list(A). rev(xs)=rev(ys) \<longleftrightarrow> xs=ys"
+ "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(A). rev(xs)=rev(ys) \<longleftrightarrow> xs=ys"
apply (erule list.induct, force, clarify)
apply (erule_tac a = ys in list.cases, auto)
done
declare rev_is_rev_iff [simp]
lemma rev_list_elim [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
(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)
@@ -660,63 +660,63 @@
(** more theorems about drop **)
lemma length_drop [rule_format]:
- "n \<in> nat ==> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
+ "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(drop(n, xs)) = length(xs) #- n"
apply (erule nat_induct)
apply (auto elim: list.cases)
done
declare length_drop [simp]
lemma drop_all [rule_format]:
- "n \<in> nat ==> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> drop(n, xs)=Nil"
+ "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> drop(n, xs)=Nil"
apply (erule nat_induct)
apply (auto elim: list.cases)
done
declare drop_all [simp]
lemma drop_append [rule_format]:
- "n \<in> nat ==>
+ "n \<in> nat \<Longrightarrow>
\<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)
done
lemma drop_drop:
- "m \<in> nat ==> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
+ "m \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). \<forall>n \<in> nat. drop(n, drop(m, xs))=drop(n #+ m, xs)"
apply (induct_tac "m")
apply (auto elim: list.cases)
done
(** take **)
-lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) = Nil"
+lemma take_0 [simp]: "xs:list(A) \<Longrightarrow> take(0, xs) = Nil"
apply (unfold take_def)
apply (erule list.induct, auto)
done
lemma take_succ_Cons [simp]:
- "n \<in> nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
+ "n \<in> nat \<Longrightarrow> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))"
by (simp add: take_def)
(* Needed for proving take_all *)
-lemma take_Nil [simp]: "n \<in> nat ==> take(n, Nil) = Nil"
+lemma take_Nil [simp]: "n \<in> nat \<Longrightarrow> take(n, Nil) = Nil"
by (unfold take_def, auto)
lemma take_all [rule_format]:
- "n \<in> nat ==> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> take(n, xs) = xs"
+ "n \<in> nat \<Longrightarrow> \<forall>xs \<in> list(A). length(xs) \<le> n \<longrightarrow> take(n, xs) = xs"
apply (erule nat_induct)
apply (auto elim: list.cases)
done
declare take_all [simp]
lemma take_type [rule_format]:
- "xs:list(A) ==> \<forall>n \<in> nat. take(n, xs):list(A)"
+ "xs:list(A) \<Longrightarrow> \<forall>n \<in> nat. take(n, xs):list(A)"
apply (erule list.induct, simp, clarify)
apply (erule natE, auto)
done
declare take_type [simp,TC]
lemma take_append [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<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)
@@ -725,7 +725,7 @@
declare take_append [simp]
lemma take_take [rule_format]:
- "m \<in> nat ==>
+ "m \<in> nat \<Longrightarrow>
\<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)
@@ -739,14 +739,14 @@
lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a"
by (simp add: nth_def)
-lemma nth_Cons [simp]: "n \<in> nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)"
+lemma nth_Cons [simp]: "n \<in> nat \<Longrightarrow> nth(succ(n), Cons(a,l)) = nth(n,l)"
by (simp add: nth_def)
lemma nth_empty [simp]: "nth(n, Nil) = 0"
by (simp add: nth_def)
lemma nth_type [rule_format]:
- "xs:list(A) ==> \<forall>n. n < length(xs) \<longrightarrow> nth(n,xs) \<in> A"
+ "xs:list(A) \<Longrightarrow> \<forall>n. n < length(xs) \<longrightarrow> nth(n,xs) \<in> A"
apply (erule list.induct, simp, clarify)
apply (subgoal_tac "n \<in> nat")
apply (erule natE, auto dest!: le_in_nat)
@@ -754,13 +754,13 @@
declare nth_type [simp,TC]
lemma nth_eq_0 [rule_format]:
- "xs:list(A) ==> \<forall>n \<in> nat. length(xs) \<le> n \<longrightarrow> nth(n,xs) = 0"
+ "xs:list(A) \<Longrightarrow> \<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) \<Longrightarrow>
\<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)
@@ -769,7 +769,7 @@
lemma set_of_list_conv_nth:
"xs:list(A)
- ==> set_of_list(xs) = {x \<in> A. \<exists>i\<in>nat. i<length(xs) & x = nth(i,xs)}"
+ \<Longrightarrow> set_of_list(xs) = {x \<in> 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)
@@ -779,7 +779,7 @@
(* Other theorems about lists *)
lemma nth_take_lemma [rule_format]:
- "k \<in> nat ==>
+ "k \<in> nat \<Longrightarrow>
\<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")
@@ -797,9 +797,9 @@
done
lemma nth_equalityI [rule_format]:
- "[| xs:list(A); ys:list(A); length(xs) = length(ys);
- \<forall>i \<in> nat. i < length(xs) \<longrightarrow> nth(i,xs) = nth(i,ys) |]
- ==> xs = ys"
+ "\<lbrakk>xs:list(A); ys:list(A); length(xs) = length(ys);
+ \<forall>i \<in> nat. i < length(xs) \<longrightarrow> nth(i,xs) = nth(i,ys)\<rbrakk>
+ \<Longrightarrow> 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 (simp_all add: take_all)
@@ -808,8 +808,8 @@
(*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 = ys"
+ "\<lbrakk>xs:list(A); ys:list(A); (\<forall>i \<in> nat. take(i, xs) = take(i,ys))\<rbrakk>
+ \<Longrightarrow> xs = 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)
@@ -821,14 +821,14 @@
done
lemma nth_drop [rule_format]:
- "n \<in> nat ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)"
+ "n \<in> nat \<Longrightarrow> \<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)
done
lemma take_succ [rule_format]:
"xs\<in>list(A)
- ==> \<forall>i. i < length(xs) \<longrightarrow> take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]"
+ \<Longrightarrow> \<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 (erule natE)
@@ -836,14 +836,14 @@
done
lemma take_add [rule_format]:
- "[|xs\<in>list(A); j\<in>nat|]
- ==> \<forall>i\<in>nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))"
+ "\<lbrakk>xs\<in>list(A); j\<in>nat\<rbrakk>
+ \<Longrightarrow> \<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)
done
lemma length_take:
- "l\<in>list(A) ==> \<forall>n\<in>nat. length(take(n,l)) = min(n, length(l))"
+ "l\<in>list(A) \<Longrightarrow> \<forall>n\<in>nat. length(take(n,l)) = min(n, length(l))"
apply (induct_tac "l", safe, simp_all)
apply (erule natE, simp_all)
done
@@ -865,29 +865,29 @@
definition
zip :: "[i, i]=>i" where
- "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys"
+ "zip(xs, ys) \<equiv> zip_aux(set_of_list(ys),xs)`ys"
(* zip equations *)
-lemma list_on_set_of_list: "xs \<in> list(A) ==> xs \<in> list(set_of_list(xs))"
+lemma list_on_set_of_list: "xs \<in> list(A) \<Longrightarrow> xs \<in> list(set_of_list(xs))"
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"
+lemma zip_Nil [simp]: "ys:list(A) \<Longrightarrow> zip(Nil, ys)=Nil"
apply (simp add: zip_def list_on_set_of_list [of _ A])
apply (erule list.cases, simp_all)
done
-lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil"
+lemma zip_Nil2 [simp]: "xs:list(A) \<Longrightarrow> zip(xs, Nil)=Nil"
apply (simp add: zip_def list_on_set_of_list [of _ A])
apply (erule list.cases, simp_all)
done
lemma zip_aux_unique [rule_format]:
- "[|B<=C; xs \<in> list(A)|]
- ==> \<forall>ys \<in> list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys"
+ "\<lbrakk>B<=C; xs \<in> list(A)\<rbrakk>
+ \<Longrightarrow> \<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)
@@ -896,7 +896,7 @@
done
lemma zip_Cons_Cons [simp]:
- "[| xs:list(A); ys:list(B); x \<in> A; y \<in> B |] ==>
+ "\<lbrakk>xs:list(A); ys:list(B); x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow>
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)
@@ -905,7 +905,7 @@
done
lemma zip_type [rule_format]:
- "xs:list(A) ==> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
+ "xs:list(A) \<Longrightarrow> \<forall>ys \<in> list(B). zip(xs, ys):list(A*B)"
apply (induct_tac "xs")
apply (simp (no_asm))
apply clarify
@@ -915,7 +915,7 @@
(* zip length *)
lemma length_zip [rule_format]:
- "xs:list(A) ==> \<forall>ys \<in> list(B). length(zip(xs,ys)) =
+ "xs:list(A) \<Longrightarrow> \<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)
@@ -924,7 +924,7 @@
declare length_zip [simp]
lemma zip_append1 [rule_format]:
- "[| ys:list(A); zs:list(B) |] ==>
+ "\<lbrakk>ys:list(A); zs:list(B)\<rbrakk> \<Longrightarrow>
\<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)
@@ -932,21 +932,21 @@
done
lemma zip_append2 [rule_format]:
- "[| xs:list(A); zs:list(B) |] ==> \<forall>ys \<in> list(B). zip(xs, ys@zs) =
+ "\<lbrakk>xs:list(A); zs:list(B)\<rbrakk> \<Longrightarrow> \<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 (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) |]
- ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)"
+ "\<lbrakk>length(xs) = length(us); length(ys) = length(vs);
+ xs:list(A); us:list(B); ys:list(A); vs:list(B)\<rbrakk>
+ \<Longrightarrow> 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]:
- "ys:list(B) ==> \<forall>xs \<in> list(A).
+ "ys:list(B) \<Longrightarrow> \<forall>xs \<in> list(A).
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)
@@ -955,7 +955,7 @@
declare zip_rev [simp]
lemma nth_zip [rule_format]:
- "ys:list(B) ==> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
+ "ys:list(B) \<Longrightarrow> \<forall>i \<in> nat. \<forall>xs \<in> list(A).
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)
@@ -965,28 +965,28 @@
declare nth_zip [simp]
lemma set_of_list_zip [rule_format]:
- "[| xs:list(A); ys:list(B); i \<in> nat |]
- ==> set_of_list(zip(xs, ys)) =
+ "\<lbrakk>xs:list(A); ys:list(B); i \<in> nat\<rbrakk>
+ \<Longrightarrow> set_of_list(zip(xs, 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)
(** list_update **)
-lemma list_update_Nil [simp]: "i \<in> nat ==>list_update(Nil, i, v) = Nil"
+lemma list_update_Nil [simp]: "i \<in> nat \<Longrightarrow>list_update(Nil, i, v) = Nil"
by (unfold list_update_def, auto)
lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)"
by (unfold list_update_def, auto)
lemma list_update_Cons_succ [simp]:
- "n \<in> nat ==>
+ "n \<in> nat \<Longrightarrow>
list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))"
apply (unfold list_update_def, auto)
done
lemma list_update_type [rule_format]:
- "[| xs:list(A); v \<in> A |] ==> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
+ "\<lbrakk>xs:list(A); v \<in> A\<rbrakk> \<Longrightarrow> \<forall>n \<in> nat. list_update(xs, n, v):list(A)"
apply (induct_tac "xs")
apply (simp (no_asm))
apply clarify
@@ -995,7 +995,7 @@
declare list_update_type [simp,TC]
lemma length_list_update [rule_format]:
- "xs:list(A) ==> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
+ "xs:list(A) \<Longrightarrow> \<forall>i \<in> nat. length(list_update(xs, i, v))=length(xs)"
apply (induct_tac "xs")
apply (simp (no_asm))
apply clarify
@@ -1004,7 +1004,7 @@
declare length_list_update [simp]
lemma nth_list_update [rule_format]:
- "[| xs:list(A) |] ==> \<forall>i \<in> nat. \<forall>j \<in> nat. i < length(xs) \<longrightarrow>
+ "\<lbrakk>xs:list(A)\<rbrakk> \<Longrightarrow> \<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
@@ -1016,12 +1016,12 @@
done
lemma nth_list_update_eq [simp]:
- "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x"
+ "\<lbrakk>i < length(xs); xs:list(A)\<rbrakk> \<Longrightarrow> nth(i, list_update(xs, i,x)) = x"
by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update)
lemma nth_list_update_neq [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<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))
@@ -1033,7 +1033,7 @@
declare nth_list_update_neq [simp]
lemma list_update_overwrite [rule_format]:
- "xs:list(A) ==> \<forall>i \<in> nat. i < length(xs)
+ "xs:list(A) \<Longrightarrow> \<forall>i \<in> nat. i < length(xs)
\<longrightarrow> list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)"
apply (induct_tac "xs")
apply (simp (no_asm))
@@ -1043,7 +1043,7 @@
declare list_update_overwrite [simp]
lemma list_update_same_conv [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<forall>i \<in> nat. i < length(xs) \<longrightarrow>
(list_update(xs, i, x) = xs) \<longleftrightarrow> (nth(i, xs) = x)"
apply (induct_tac "xs")
@@ -1053,7 +1053,7 @@
done
lemma update_zip [rule_format]:
- "ys:list(B) ==>
+ "ys:list(B) \<Longrightarrow>
\<forall>i \<in> nat. \<forall>xy \<in> A*B. \<forall>xs \<in> list(A).
length(xs) = length(ys) \<longrightarrow>
list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)),
@@ -1065,7 +1065,7 @@
done
lemma set_update_subset_cons [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<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
@@ -1074,8 +1074,8 @@
done
lemma set_of_list_update_subsetI:
- "[| set_of_list(xs) \<subseteq> A; xs:list(A); x \<in> A; i \<in> nat|]
- ==> set_of_list(list_update(xs, i,x)) \<subseteq> A"
+ "\<lbrakk>set_of_list(xs) \<subseteq> A; xs:list(A); x \<in> A; i \<in> nat\<rbrakk>
+ \<Longrightarrow> set_of_list(list_update(xs, i,x)) \<subseteq> A"
apply (rule subset_trans)
apply (rule set_update_subset_cons, auto)
done
@@ -1083,13 +1083,13 @@
(** upt **)
lemma upt_rec:
- "j \<in> nat ==> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
+ "j \<in> nat \<Longrightarrow> upt(i,j) = (if i<j then Cons(i, upt(succ(i), j)) else Nil)"
apply (induct_tac "j", auto)
apply (drule not_lt_imp_le)
apply (auto simp: lt_Ord intro: le_anti_sym)
done
-lemma upt_conv_Nil [simp]: "[| j \<le> i; j \<in> nat |] ==> upt(i,j) = Nil"
+lemma upt_conv_Nil [simp]: "\<lbrakk>j \<le> i; j \<in> nat\<rbrakk> \<Longrightarrow> upt(i,j) = Nil"
apply (subst upt_rec, auto)
apply (auto simp add: le_iff)
apply (drule lt_asym [THEN notE], auto)
@@ -1097,34 +1097,34 @@
(*Only needed if upt_Suc is deleted from the simpset*)
lemma upt_succ_append:
- "[| i \<le> j; j \<in> nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]"
+ "\<lbrakk>i \<le> j; j \<in> nat\<rbrakk> \<Longrightarrow> upt(i,succ(j)) = upt(i, j)@[j]"
by simp
lemma upt_conv_Cons:
- "[| i<j; j \<in> nat |] ==> upt(i,j) = Cons(i,upt(succ(i),j))"
+ "\<lbrakk>i<j; j \<in> nat\<rbrakk> \<Longrightarrow> upt(i,j) = Cons(i,upt(succ(i),j))"
apply (rule trans)
apply (rule upt_rec, auto)
done
-lemma upt_type [simp,TC]: "j \<in> nat ==> upt(i,j):list(nat)"
+lemma upt_type [simp,TC]: "j \<in> nat \<Longrightarrow> upt(i,j):list(nat)"
by (induct_tac "j", auto)
(*LOOPS as a simprule, since j<=j*)
lemma upt_add_eq_append:
- "[| i \<le> j; j \<in> nat; k \<in> nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
+ "\<lbrakk>i \<le> j; j \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)"
apply (induct_tac "k")
apply (auto simp add: app_assoc app_type)
apply (rule_tac j = j in le_trans, auto)
done
-lemma length_upt [simp]: "[| i \<in> nat; j \<in> nat |] ==>length(upt(i,j)) = j #- i"
+lemma length_upt [simp]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow>length(upt(i,j)) = j #- i"
apply (induct_tac "j")
apply (rule_tac [2] sym)
apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff)
done
lemma nth_upt [simp]:
- "[| i \<in> nat; j \<in> nat; k \<in> nat; i #+ k < j |] ==> nth(k, upt(i,j)) = i #+ k"
+ "\<lbrakk>i \<in> nat; j \<in> nat; k \<in> nat; i #+ k < j\<rbrakk> \<Longrightarrow> nth(k, upt(i,j)) = i #+ k"
apply (rotate_tac -1, erule rev_mp)
apply (induct_tac "j", simp)
apply (auto dest!: not_lt_imp_le
@@ -1132,7 +1132,7 @@
done
lemma take_upt [rule_format]:
- "[| m \<in> nat; n \<in> nat |] ==>
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow>
\<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)
@@ -1147,13 +1147,13 @@
declare take_upt [simp]
lemma map_succ_upt:
- "[| m \<in> nat; n \<in> nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> map(succ, upt(m,n))= upt(succ(m), succ(n))"
apply (induct_tac "n")
apply (auto simp add: map_app_distrib)
done
lemma nth_map [rule_format]:
- "xs:list(A) ==>
+ "xs:list(A) \<Longrightarrow>
\<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)
@@ -1162,7 +1162,7 @@
declare nth_map [simp]
lemma nth_map_upt [rule_format]:
- "[| m \<in> nat; n \<in> nat |] ==>
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow>
\<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)
@@ -1180,17 +1180,17 @@
definition
sublist :: "[i, i] => i" where
- "sublist(xs, A)==
+ "sublist(xs, A)\<equiv>
map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
-lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil"
+lemma sublist_0 [simp]: "xs:list(A) \<Longrightarrow>sublist(xs, 0) =Nil"
by (unfold sublist_def, auto)
lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil"
by (unfold sublist_def, auto)
lemma sublist_shift_lemma:
- "[| xs:list(B); i \<in> nat |] ==>
+ "\<lbrakk>xs:list(B); i \<in> nat\<rbrakk> \<Longrightarrow>
map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
map(fst, filter(%p. snd(p):nat & snd(p) #+ i \<in> A, zip(xs,upt(0,length(xs)))))"
apply (erule list_append_induct)
@@ -1199,18 +1199,18 @@
done
lemma sublist_type [simp,TC]:
- "xs:list(B) ==> sublist(xs, A):list(B)"
+ "xs:list(B) \<Longrightarrow> sublist(xs, A):list(B)"
apply (unfold sublist_def)
apply (induct_tac "xs")
apply (auto simp add: filter_append map_app_distrib)
done
lemma upt_add_eq_append2:
- "[| i \<in> nat; j \<in> nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
+ "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)"
by (simp add: upt_add_eq_append [of 0] nat_0_le)
lemma sublist_append:
- "[| xs:list(B); ys:list(B) |] ==>
+ "\<lbrakk>xs:list(B); ys:list(B)\<rbrakk> \<Longrightarrow>
sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \<in> nat. j #+ length(xs): A})"
apply (unfold sublist_def)
apply (erule_tac l = ys in list_append_induct, simp)
@@ -1221,7 +1221,7 @@
lemma sublist_Cons:
- "[| xs:list(B); x \<in> B |] ==>
+ "\<lbrakk>xs:list(B); x \<in> B\<rbrakk> \<Longrightarrow>
sublist(Cons(x, xs), A) =
(if 0 \<in> A then [x] else []) @ sublist(xs, {j \<in> nat. succ(j) \<in> A})"
apply (erule_tac l = xs in list_append_induct)
@@ -1234,7 +1234,7 @@
by (simp add: sublist_Cons)
lemma sublist_upt_eq_take [rule_format]:
- "xs:list(A) ==> \<forall>n\<in>nat. sublist(xs,n) = take(n,xs)"
+ "xs:list(A) \<Longrightarrow> \<forall>n\<in>nat. sublist(xs,n) = take(n,xs)"
apply (erule list.induct, simp)
apply (clarify )
apply (erule natE)
@@ -1243,7 +1243,7 @@
declare sublist_upt_eq_take [simp]
lemma sublist_Int_eq:
- "xs \<in> list(B) ==> sublist(xs, A \<inter> nat) = sublist(xs, A)"
+ "xs \<in> list(B) \<Longrightarrow> sublist(xs, A \<inter> nat) = sublist(xs, A)"
apply (erule list.induct)
apply (simp_all add: sublist_Cons)
done
@@ -1256,15 +1256,15 @@
"repeat(a,succ(n)) = Cons(a,repeat(a,n))"
-lemma length_repeat: "n \<in> nat ==> length(repeat(a,n)) = n"
+lemma length_repeat: "n \<in> nat \<Longrightarrow> length(repeat(a,n)) = n"
by (induct_tac n, auto)
-lemma repeat_succ_app: "n \<in> nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]"
+lemma repeat_succ_app: "n \<in> nat \<Longrightarrow> repeat(a,succ(n)) = repeat(a,n) @ [a]"
apply (induct_tac n)
apply (simp_all del: app_Cons add: app_Cons [symmetric])
done
-lemma repeat_type [TC]: "[|a \<in> A; n \<in> nat|] ==> repeat(a,n) \<in> list(A)"
+lemma repeat_type [TC]: "\<lbrakk>a \<in> A; n \<in> nat\<rbrakk> \<Longrightarrow> repeat(a,n) \<in> list(A)"
by (induct_tac n, auto)
end
--- a/src/ZF/Nat.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Nat.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,43 +9,43 @@
definition
nat :: i where
- "nat == lfp(Inf, %X. {0} \<union> {succ(i). i \<in> X})"
+ "nat \<equiv> lfp(Inf, %X. {0} \<union> {succ(i). i \<in> X})"
definition
quasinat :: "i => o" where
- "quasinat(n) == n=0 | (\<exists>m. n = succ(m))"
+ "quasinat(n) \<equiv> n=0 | (\<exists>m. n = succ(m))"
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 | (\<exists>x. k=succ(x) & y=b(x))"
+ "nat_case(a,b,k) \<equiv> 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) \<equiv>
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 \<equiv> {<x,y>:nat*nat. x \<le> y}"
definition
Lt :: i where
- "Lt == {<x, y>:nat*nat. x < y}"
+ "Lt \<equiv> {<x, y>:nat*nat. x < y}"
definition
Ge :: i where
- "Ge == {<x,y>:nat*nat. y \<le> x}"
+ "Ge \<equiv> {<x,y>:nat*nat. y \<le> x}"
definition
Gt :: i where
- "Gt == {<x,y>:nat*nat. y < x}"
+ "Gt \<equiv> {<x,y>:nat*nat. y < x}"
definition
greater_than :: "i=>i" where
- "greater_than(n) == {i \<in> nat. n < i}"
+ "greater_than(n) \<equiv> {i \<in> nat. n < i}"
text\<open>No need for a less-than operator: a natural number is its list of
predecessors!\<close>
@@ -66,7 +66,7 @@
apply (rule singletonI [THEN UnI1])
done
-lemma nat_succI [intro!,TC]: "n \<in> nat ==> succ(n) \<in> nat"
+lemma nat_succI [intro!,TC]: "n \<in> nat \<Longrightarrow> succ(n) \<in> nat"
apply (subst nat_unfold)
apply (erule RepFunI [THEN UnI2])
done
@@ -87,7 +87,7 @@
(*Mathematical induction*)
lemma nat_induct [case_names 0 succ, induct set: nat]:
- "[| n \<in> nat; P(0); !!x. [| x \<in> nat; P(x) |] ==> P(succ(x)) |] ==> P(n)"
+ "\<lbrakk>n \<in> nat; P(0); \<And>x. \<lbrakk>x \<in> nat; P(x)\<rbrakk> \<Longrightarrow> P(succ(x))\<rbrakk> \<Longrightarrow> P(n)"
by (erule def_induct [OF nat_def nat_bnd_mono], blast)
lemma natE:
@@ -96,13 +96,13 @@
using assms
by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto
-lemma nat_into_Ord [simp]: "n \<in> nat ==> Ord(n)"
+lemma nat_into_Ord [simp]: "n \<in> nat \<Longrightarrow> Ord(n)"
by (erule nat_induct, auto)
-(* @{term"i \<in> nat ==> 0 \<le> i"}; same thing as @{term"0<succ(i)"} *)
+(* @{term"i \<in> nat \<Longrightarrow> 0 \<le> i"}; same thing as @{term"0<succ(i)"} *)
lemmas nat_0_le = nat_into_Ord [THEN Ord_0_le]
-(* @{term"i \<in> nat ==> i \<le> i"}; same thing as @{term"i<succ(i)"} *)
+(* @{term"i \<in> nat \<Longrightarrow> 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)"
@@ -119,16 +119,16 @@
apply (erule ltD)
done
-lemma naturals_not_limit: "a \<in> nat ==> ~ Limit(a)"
+lemma naturals_not_limit: "a \<in> nat \<Longrightarrow> \<not> Limit(a)"
by (induct a rule: nat_induct, auto)
-lemma succ_natD: "succ(i): nat ==> i \<in> nat"
+lemma succ_natD: "succ(i): nat \<Longrightarrow> i \<in> nat"
by (rule Ord_trans [OF succI1], auto)
lemma nat_succ_iff [iff]: "succ(n): nat \<longleftrightarrow> n \<in> nat"
by (blast dest!: succ_natD)
-lemma nat_le_Limit: "Limit(i) ==> nat \<le> i"
+lemma nat_le_Limit: "Limit(i) \<Longrightarrow> nat \<le> i"
apply (rule subset_imp_le)
apply (simp_all add: Limit_is_Ord)
apply (rule subsetI)
@@ -137,15 +137,15 @@
apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord)
done
-(* [| succ(i): k; k \<in> nat |] ==> i \<in> k *)
+(* \<lbrakk>succ(i): k; k \<in> nat\<rbrakk> \<Longrightarrow> i \<in> k *)
lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord]
-lemma lt_nat_in_nat: "[| m<n; n \<in> nat |] ==> m \<in> nat"
+lemma lt_nat_in_nat: "\<lbrakk>m<n; n \<in> nat\<rbrakk> \<Longrightarrow> m \<in> nat"
apply (erule ltE)
apply (erule Ord_trans, assumption, simp)
done
-lemma le_in_nat: "[| m \<le> n; n \<in> nat |] ==> m \<in> nat"
+lemma le_in_nat: "\<lbrakk>m \<le> n; n \<in> nat\<rbrakk> \<Longrightarrow> m \<in> nat"
by (blast dest!: lt_nat_in_nat)
@@ -163,7 +163,7 @@
lemma nat_induct_from:
assumes "m \<le> n" "m \<in> nat" "n \<in> nat"
and "P(m)"
- and "!!x. [| x \<in> nat; m \<le> x; P(x) |] ==> P(succ(x))"
+ and "\<And>x. \<lbrakk>x \<in> nat; m \<le> x; P(x)\<rbrakk> \<Longrightarrow> P(succ(x))"
shows "P(n)"
proof -
from assms(3) have "m \<le> n \<longrightarrow> P(m) \<longrightarrow> P(n)"
@@ -173,11 +173,11 @@
(*Induction suitable for subtraction and less-than*)
lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]:
- "[| m \<in> nat; n \<in> nat;
- !!x. x \<in> nat ==> P(x,0);
- !!y. y \<in> nat ==> P(0,succ(y));
- !!x y. [| x \<in> nat; y \<in> nat; P(x,y) |] ==> P(succ(x),succ(y)) |]
- ==> P(m,n)"
+ "\<lbrakk>m \<in> nat; n \<in> nat;
+ \<And>x. x \<in> nat \<Longrightarrow> P(x,0);
+ \<And>y. y \<in> nat \<Longrightarrow> P(0,succ(y));
+ \<And>x y. \<lbrakk>x \<in> nat; y \<in> nat; P(x,y)\<rbrakk> \<Longrightarrow> P(succ(x),succ(y))\<rbrakk>
+ \<Longrightarrow> P(m,n)"
apply (erule_tac x = m in rev_bspec)
apply (erule nat_induct, simp)
apply (rule ballI)
@@ -189,7 +189,7 @@
(** Induction principle analogous to trancl_induct **)
lemma succ_lt_induct_lemma [rule_format]:
- "m \<in> nat ==> P(m,succ(m)) \<longrightarrow> (\<forall>x\<in>nat. P(m,x) \<longrightarrow> P(m,succ(x))) \<longrightarrow>
+ "m \<in> nat \<Longrightarrow> 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])
@@ -198,10 +198,10 @@
done
lemma succ_lt_induct:
- "[| m<n; n \<in> nat;
+ "\<lbrakk>m<n; n \<in> nat;
P(m,succ(m));
- !!x. [| x \<in> nat; P(m,x) |] ==> P(m,succ(x)) |]
- ==> P(m,n)"
+ \<And>x. \<lbrakk>x \<in> nat; P(m,x)\<rbrakk> \<Longrightarrow> P(m,succ(x))\<rbrakk>
+ \<Longrightarrow> P(m,n)"
by (blast intro: succ_lt_induct_lemma lt_nat_in_nat)
subsection\<open>quasinat: to allow a case-split rule for \<^term>\<open>nat_case\<close>\<close>
@@ -213,20 +213,20 @@
lemma [iff]: "quasinat(succ(x))"
by (simp add: quasinat_def)
-lemma nat_imp_quasinat: "n \<in> nat ==> quasinat(n)"
+lemma nat_imp_quasinat: "n \<in> nat \<Longrightarrow> quasinat(n)"
by (erule natE, simp_all)
-lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0"
+lemma non_nat_case: "\<not> quasinat(x) \<Longrightarrow> 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)"
+lemma nat_cases_disj: "k=0 | (\<exists>y. k = succ(y)) | \<not> quasinat(k)"
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"
+ "\<lbrakk>k=0 \<Longrightarrow> P; \<And>y. k = succ(y) \<Longrightarrow> P; \<not> quasinat(k) \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (insert nat_cases_disj [of k], blast)
(** nat_case **)
@@ -238,13 +238,13 @@
by (simp add: nat_case_def)
lemma nat_case_type [TC]:
- "[| n \<in> nat; a \<in> C(0); !!m. m \<in> nat ==> b(m): C(succ(m)) |]
- ==> nat_case(a,b,n) \<in> C(n)"
+ "\<lbrakk>n \<in> nat; a \<in> C(0); \<And>m. m \<in> nat \<Longrightarrow> b(m): C(succ(m))\<rbrakk>
+ \<Longrightarrow> nat_case(a,b,n) \<in> C(n)"
by (erule nat_induct, auto)
lemma split_nat_case:
"P(nat_case(a,b,k)) \<longleftrightarrow>
- ((k=0 \<longrightarrow> P(a)) & (\<forall>x. k=succ(x) \<longrightarrow> P(b(x))) & (~ quasinat(k) \<longrightarrow> P(0)))"
+ ((k=0 \<longrightarrow> P(a)) & (\<forall>x. k=succ(x) \<longrightarrow> P(b(x))) & (\<not> quasinat(k) \<longrightarrow> P(0)))"
apply (rule nat_cases [of k])
apply (auto simp add: non_nat_case)
done
@@ -261,7 +261,7 @@
apply (rule nat_case_0)
done
-lemma nat_rec_succ: "m \<in> nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"
+lemma nat_rec_succ: "m \<in> nat \<Longrightarrow> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"
apply (rule nat_rec_def [THEN def_wfrec, THEN trans])
apply (rule wf_Memrel)
apply (simp add: vimage_singleton_iff)
@@ -269,12 +269,12 @@
(** The union of two natural numbers is a natural number -- their maximum **)
-lemma Un_nat_type [TC]: "[| i \<in> nat; j \<in> nat |] ==> i \<union> j \<in> nat"
+lemma Un_nat_type [TC]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> i \<union> j \<in> nat"
apply (rule Un_least_lt [THEN ltD])
apply (simp_all add: lt_def)
done
-lemma Int_nat_type [TC]: "[| i \<in> nat; j \<in> nat |] ==> i \<inter> j \<in> nat"
+lemma Int_nat_type [TC]: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> i \<inter> j \<in> nat"
apply (rule Int_greatest_lt [THEN ltD])
apply (simp_all add: lt_def)
done
@@ -284,7 +284,7 @@
by blast
text\<open>A natural number is the set of its predecessors\<close>
-lemma nat_eq_Collect_lt: "i \<in> nat ==> {j\<in>nat. j<i} = i"
+lemma nat_eq_Collect_lt: "i \<in> nat \<Longrightarrow> {j\<in>nat. j<i} = i"
apply (rule equalityI)
apply (blast dest: ltD)
apply (auto simp add: Ord_mem_iff_lt)
--- a/src/ZF/OrdQuant.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/OrdQuant.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,16 +11,16 @@
definition
(* Ordinal Quantifiers *)
oall :: "[i, i => o] => o" where
- "oall(A, P) == \<forall>x. x<A \<longrightarrow> P(x)"
+ "oall(A, P) \<equiv> \<forall>x. x<A \<longrightarrow> P(x)"
definition
oex :: "[i, i => o] => o" where
- "oex(A, P) == \<exists>x. x<A & P(x)"
+ "oex(A, P) \<equiv> \<exists>x. x<A & P(x)"
definition
(* Ordinal Union *)
OUnion :: "[i, i => i] => i" where
- "OUnion(i,B) == {z: \<Union>x\<in>i. B(x). Ord(i)}"
+ "OUnion(i,B) \<equiv> {z: \<Union>x\<in>i. B(x). Ord(i)}"
syntax
"_oall" :: "[idt, i, o] => o" (\<open>(3\<forall>_<_./ _)\<close> 10)
@@ -37,11 +37,11 @@
(*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!*)
+ x < 0 \<Longrightarrow> P(x) \<equiv> True, which causes dire effects!*)
lemma [simp]: "(\<forall>x<0. P(x))"
by (simp add: oall_def)
-lemma [simp]: "~(\<exists>x<0. P(x))"
+lemma [simp]: "\<not>(\<exists>x<0. P(x))"
by (simp add: oex_def)
lemma [simp]: "(\<forall>x<succ(i). P(x)) <-> (Ord(i) \<longrightarrow> P(i) & (\<forall>x<i. P(x)))"
@@ -57,66 +57,66 @@
subsubsection \<open>Union over ordinals\<close>
lemma Ord_OUN [intro,simp]:
- "[| !!x. x<A ==> Ord(B(x)) |] ==> Ord(\<Union>x<A. B(x))"
+ "\<lbrakk>\<And>x. x<A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x<A. B(x))"
by (simp add: OUnion_def ltI Ord_UN)
lemma OUN_upper_lt:
- "[| a<A; i < b(a); Ord(\<Union>x<A. b(x)) |] ==> i < (\<Union>x<A. b(x))"
+ "\<lbrakk>a<A; i < b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x<A. b(x))"
by (unfold OUnion_def lt_def, blast )
lemma OUN_upper_le:
- "[| a<A; i\<le>b(a); Ord(\<Union>x<A. b(x)) |] ==> i \<le> (\<Union>x<A. b(x))"
+ "\<lbrakk>a<A; i\<le>b(a); Ord(\<Union>x<A. b(x))\<rbrakk> \<Longrightarrow> i \<le> (\<Union>x<A. b(x))"
apply (unfold OUnion_def, auto)
apply (rule UN_upper_le )
apply (auto simp add: lt_def)
done
-lemma Limit_OUN_eq: "Limit(i) ==> (\<Union>x<i. x) = i"
+lemma Limit_OUN_eq: "Limit(i) \<Longrightarrow> (\<Union>x<i. x) = i"
by (simp add: OUnion_def Limit_Union_eq Limit_is_Ord)
(* 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"
+ "(\<And>x. x<A \<Longrightarrow> B(x) \<subseteq> C) \<Longrightarrow> (\<Union>x<A. B(x)) \<subseteq> C"
by (simp add: OUnion_def UN_least ltI)
lemma OUN_least_le:
- "[| Ord(i); !!x. x<A ==> b(x) \<le> i |] ==> (\<Union>x<A. b(x)) \<le> i"
+ "\<lbrakk>Ord(i); \<And>x. x<A \<Longrightarrow> b(x) \<le> i\<rbrakk> \<Longrightarrow> (\<Union>x<A. b(x)) \<le> i"
by (simp add: OUnion_def UN_least_le ltI Ord_0_le)
lemma le_implies_OUN_le_OUN:
- "[| !!x. x<A ==> c(x) \<le> d(x) |] ==> (\<Union>x<A. c(x)) \<le> (\<Union>x<A. d(x))"
+ "\<lbrakk>\<And>x. x<A \<Longrightarrow> c(x) \<le> d(x)\<rbrakk> \<Longrightarrow> (\<Union>x<A. c(x)) \<le> (\<Union>x<A. d(x))"
by (blast intro: OUN_least_le OUN_upper_le le_Ord2 Ord_OUN)
lemma OUN_UN_eq:
- "(!!x. x \<in> A ==> Ord(B(x)))
- ==> (\<Union>z < (\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z < B(x). C(z))"
+ "(\<And>x. x \<in> A \<Longrightarrow> Ord(B(x)))
+ \<Longrightarrow> (\<Union>z < (\<Union>x\<in>A. B(x)). C(z)) = (\<Union>x\<in>A. \<Union>z < B(x). C(z))"
by (simp add: OUnion_def)
lemma OUN_Union_eq:
- "(!!x. x \<in> X ==> Ord(x))
- ==> (\<Union>z < \<Union>(X). C(z)) = (\<Union>x\<in>X. \<Union>z < x. C(z))"
+ "(\<And>x. x \<in> X \<Longrightarrow> Ord(x))
+ \<Longrightarrow> (\<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 this quantifier...*)
lemma atomize_oall [symmetric, rulify]:
- "(!!x. x<A ==> P(x)) == Trueprop (\<forall>x<A. P(x))"
+ "(\<And>x. x<A \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x<A. P(x))"
by (simp add: oall_def atomize_all atomize_imp)
subsubsection \<open>universal quantifier for ordinals\<close>
lemma oallI [intro!]:
- "[| !!x. x<A ==> P(x) |] ==> \<forall>x<A. P(x)"
+ "\<lbrakk>\<And>x. x<A \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> \<forall>x<A. P(x)"
by (simp add: oall_def)
-lemma ospec: "[| \<forall>x<A. P(x); x<A |] ==> P(x)"
+lemma ospec: "\<lbrakk>\<forall>x<A. P(x); x<A\<rbrakk> \<Longrightarrow> P(x)"
by (simp add: oall_def)
lemma oallE:
- "[| \<forall>x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q"
+ "\<lbrakk>\<forall>x<A. P(x); P(x) \<Longrightarrow> Q; \<not>x<A \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: oall_def, blast)
lemma rev_oallE [elim]:
- "[| \<forall>x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q"
+ "\<lbrakk>\<forall>x<A. P(x); \<not>x<A \<Longrightarrow> Q; P(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: oall_def, blast)
@@ -126,43 +126,43 @@
(*Congruence rule for rewriting*)
lemma oall_cong [cong]:
- "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |]
- ==> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))"
+ "\<lbrakk>a=a'; \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk>
+ \<Longrightarrow> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))"
by (simp add: oall_def)
subsubsection \<open>existential quantifier for ordinals\<close>
lemma oexI [intro]:
- "[| P(x); x<A |] ==> \<exists>x<A. P(x)"
+ "\<lbrakk>P(x); x<A\<rbrakk> \<Longrightarrow> \<exists>x<A. P(x)"
apply (simp add: oex_def, blast)
done
(*Not of the general form for such rules... *)
lemma oexCI:
- "[| \<forall>x<A. ~P(x) ==> P(a); a<A |] ==> \<exists>x<A. P(x)"
+ "\<lbrakk>\<forall>x<A. \<not>P(x) \<Longrightarrow> P(a); a<A\<rbrakk> \<Longrightarrow> \<exists>x<A. P(x)"
apply (simp add: oex_def, blast)
done
lemma oexE [elim!]:
- "[| \<exists>x<A. P(x); !!x. [| x<A; P(x) |] ==> Q |] ==> Q"
+ "\<lbrakk>\<exists>x<A. P(x); \<And>x. \<lbrakk>x<A; P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
apply (simp add: oex_def, blast)
done
lemma oex_cong [cong]:
- "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |]
- ==> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))"
+ "\<lbrakk>a=a'; \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk>
+ \<Longrightarrow> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))"
apply (simp add: oex_def cong add: conj_cong)
done
subsubsection \<open>Rules for Ordinal-Indexed Unions\<close>
-lemma OUN_I [intro]: "[| a<i; b \<in> B(a) |] ==> b: (\<Union>z<i. B(z))"
+lemma OUN_I [intro]: "\<lbrakk>a<i; b \<in> B(a)\<rbrakk> \<Longrightarrow> b: (\<Union>z<i. B(z))"
by (unfold OUnion_def lt_def, blast)
lemma OUN_E [elim!]:
- "[| b \<in> (\<Union>z<i. B(z)); !!a.[| b \<in> B(a); a<i |] ==> R |] ==> R"
+ "\<lbrakk>b \<in> (\<Union>z<i. B(z)); \<And>a.\<lbrakk>b \<in> B(a); a<i\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
apply (unfold OUnion_def lt_def, blast)
done
@@ -170,11 +170,11 @@
by (unfold OUnion_def oex_def lt_def, blast)
lemma OUN_cong [cong]:
- "[| i=j; !!x. x<j ==> C(x)=D(x) |] ==> (\<Union>x<i. C(x)) = (\<Union>x<j. D(x))"
+ "\<lbrakk>i=j; \<And>x. x<j \<Longrightarrow> C(x)=D(x)\<rbrakk> \<Longrightarrow> (\<Union>x<i. C(x)) = (\<Union>x<j. D(x))"
by (simp add: OUnion_def lt_def OUN_iff)
lemma lt_induct:
- "[| i<k; !!x.[| x<k; \<forall>y<x. P(y) |] ==> P(x) |] ==> P(i)"
+ "\<lbrakk>i<k; \<And>x.\<lbrakk>x<k; \<forall>y<x. P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> P(i)"
apply (simp add: lt_def oall_def)
apply (erule conjE)
apply (erule Ord_induct, assumption, blast)
@@ -185,11 +185,11 @@
definition
"rall" :: "[i=>o, i=>o] => o" where
- "rall(M, P) == \<forall>x. M(x) \<longrightarrow> P(x)"
+ "rall(M, P) \<equiv> \<forall>x. M(x) \<longrightarrow> P(x)"
definition
"rex" :: "[i=>o, i=>o] => o" where
- "rex(M, P) == \<exists>x. M(x) & P(x)"
+ "rex(M, P) \<equiv> \<exists>x. M(x) & P(x)"
syntax
"_rall" :: "[pttrn, i=>o, o] => o" (\<open>(3\<forall>_[_]./ _)\<close> 10)
@@ -201,18 +201,18 @@
subsubsection\<open>Relativized universal quantifier\<close>
-lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> \<forall>x[M]. P(x)"
+lemma rallI [intro!]: "\<lbrakk>\<And>x. M(x) \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> \<forall>x[M]. P(x)"
by (simp add: rall_def)
-lemma rspec: "[| \<forall>x[M]. P(x); M(x) |] ==> P(x)"
+lemma rspec: "\<lbrakk>\<forall>x[M]. P(x); M(x)\<rbrakk> \<Longrightarrow> P(x)"
by (simp add: rall_def)
(*Instantiates x first: better for automatic theorem proving?*)
lemma rev_rallE [elim]:
- "[| \<forall>x[M]. P(x); ~ M(x) ==> Q; P(x) ==> Q |] ==> Q"
+ "\<lbrakk>\<forall>x[M]. P(x); \<not> M(x) \<Longrightarrow> Q; P(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: rall_def, blast)
-lemma rallE: "[| \<forall>x[M]. P(x); P(x) ==> Q; ~ M(x) ==> Q |] ==> Q"
+lemma rallE: "\<lbrakk>\<forall>x[M]. P(x); P(x) \<Longrightarrow> Q; \<not> M(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by blast
(*Trival rewrite rule; (\<forall>x[M].P)<->P holds only if A is nonempty!*)
@@ -221,32 +221,32 @@
(*Congruence rule for rewriting*)
lemma rall_cong [cong]:
- "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\<forall>x[M]. P(x)) <-> (\<forall>x[M]. P'(x))"
+ "(\<And>x. M(x) \<Longrightarrow> P(x) <-> P'(x)) \<Longrightarrow> (\<forall>x[M]. P(x)) <-> (\<forall>x[M]. P'(x))"
by (simp add: rall_def)
subsubsection\<open>Relativized existential quantifier\<close>
-lemma rexI [intro]: "[| P(x); M(x) |] ==> \<exists>x[M]. P(x)"
+lemma rexI [intro]: "\<lbrakk>P(x); M(x)\<rbrakk> \<Longrightarrow> \<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) |] ==> \<exists>x[M]. P(x)"
+lemma rev_rexI: "\<lbrakk>M(x); P(x)\<rbrakk> \<Longrightarrow> \<exists>x[M]. P(x)"
by blast
(*Not of the general form for such rules... *)
-lemma rexCI: "[| \<forall>x[M]. ~P(x) ==> P(a); M(a) |] ==> \<exists>x[M]. P(x)"
+lemma rexCI: "\<lbrakk>\<forall>x[M]. \<not>P(x) \<Longrightarrow> P(a); M(a)\<rbrakk> \<Longrightarrow> \<exists>x[M]. P(x)"
by blast
-lemma rexE [elim!]: "[| \<exists>x[M]. P(x); !!x. [| M(x); P(x) |] ==> Q |] ==> Q"
+lemma rexE [elim!]: "\<lbrakk>\<exists>x[M]. P(x); \<And>x. \<lbrakk>M(x); P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: rex_def, blast)
-(*We do not even have (\<exists>x[M]. True) <-> True unless A is nonempty!!*)
+(*We do not even have (\<exists>x[M]. True) <-> True unless A is nonempty\<And>*)
lemma rex_triv [simp]: "(\<exists>x[M]. P) \<longleftrightarrow> ((\<exists>x. M(x)) \<and> P)"
by (simp add: rex_def)
lemma rex_cong [cong]:
- "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\<exists>x[M]. P(x)) <-> (\<exists>x[M]. P'(x))"
+ "(\<And>x. M(x) \<Longrightarrow> P(x) <-> P'(x)) \<Longrightarrow> (\<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))"
@@ -255,7 +255,7 @@
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 (\<forall>x[M]. P(x))"
+lemma atomize_rall: "(\<And>x. M(x) \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x[M]. P(x))"
by (simp add: rall_def atomize_all atomize_imp)
declare atomize_rall [symmetric, rulify]
@@ -264,7 +264,7 @@
"(\<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))"
+ "(\<not>(\<forall>x[M]. P(x))) <-> (\<exists>x[M]. \<not>P(x))"
by blast+
lemma rall_simps2:
@@ -283,7 +283,7 @@
"(\<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))"
+ "(\<not>(\<exists>x[M]. P(x))) <-> (\<forall>x[M]. \<not>P(x))"
by blast+
lemma rex_simps2:
@@ -324,7 +324,7 @@
definition
setclass :: "[i,i] => o" (\<open>##_\<close> [40] 40) where
- "setclass(A) == %x. x \<in> A"
+ "setclass(A) \<equiv> %x. x \<in> A"
lemma setclass_iff [simp]: "setclass(A,x) <-> x \<in> A"
by (simp add: setclass_def)
--- a/src/ZF/Order.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Order.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,15 +17,15 @@
definition
part_ord :: "[i,i]=>o" (*Strict partial ordering*) where
- "part_ord(A,r) == irrefl(A,r) & trans[A](r)"
+ "part_ord(A,r) \<equiv> irrefl(A,r) & trans[A](r)"
definition
linear :: "[i,i]=>o" (*Strict total ordering*) where
- "linear(A,r) == (\<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r | x=y | <y,x>:r)"
+ "linear(A,r) \<equiv> (\<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
- "tot_ord(A,r) == part_ord(A,r) & linear(A,r)"
+ "tot_ord(A,r) \<equiv> part_ord(A,r) & linear(A,r)"
definition
"preorder_on(A, r) \<equiv> refl(A, r) \<and> trans[A](r)"
@@ -41,62 +41,62 @@
definition
well_ord :: "[i,i]=>o" (*Well-ordering*) where
- "well_ord(A,r) == tot_ord(A,r) & wf[A](r)"
+ "well_ord(A,r) \<equiv> tot_ord(A,r) & wf[A](r)"
definition
mono_map :: "[i,i,i,i]=>i" (*Order-preserving maps*) where
- "mono_map(A,r,B,s) ==
+ "mono_map(A,r,B,s) \<equiv>
{f \<in> 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" (\<open>(\<langle>_, _\<rangle> \<cong>/ \<langle>_, _\<rangle>)\<close> 51) (*Order isomorphisms*) where
- "\<langle>A,r\<rangle> \<cong> \<langle>B,s\<rangle> ==
+ "\<langle>A,r\<rangle> \<cong> \<langle>B,s\<rangle> \<equiv>
{f \<in> bij(A,B). \<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r \<longleftrightarrow> <f`x,f`y>:s}"
definition
pred :: "[i,i,i]=>i" (*Set of predecessors*) where
- "pred(A,x,r) == {y \<in> A. <y,x>:r}"
+ "pred(A,x,r) \<equiv> {y \<in> A. <y,x>:r}"
definition
ord_iso_map :: "[i,i,i,i]=>i" (*Construction for linearity theorem*) where
- "ord_iso_map(A,r,B,s) ==
+ "ord_iso_map(A,r,B,s) \<equiv>
\<Union>x\<in>A. \<Union>y\<in>B. \<Union>f \<in> ord_iso(pred(A,x,r), r, pred(B,y,s), s). {<x,y>}"
definition
first :: "[i, i, i] => o" where
- "first(u, X, R) == u \<in> X & (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
+ "first(u, X, R) \<equiv> u \<in> X & (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
subsection\<open>Immediate Consequences of the Definitions\<close>
lemma part_ord_Imp_asym:
- "part_ord(A,r) ==> asym(r \<inter> A*A)"
+ "part_ord(A,r) \<Longrightarrow> asym(r \<inter> A*A)"
by (unfold part_ord_def irrefl_def trans_on_def asym_def, blast)
lemma linearE:
- "[| linear(A,r); x \<in> A; y \<in> A;
- <x,y>:r ==> P; x=y ==> P; <y,x>:r ==> P |]
- ==> P"
+ "\<lbrakk>linear(A,r); x \<in> A; y \<in> A;
+ <x,y>:r \<Longrightarrow> P; x=y \<Longrightarrow> P; <y,x>:r \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by (simp add: linear_def, blast)
(** General properties of well_ord **)
lemma well_ordI:
- "[| wf[A](r); linear(A,r) |] ==> well_ord(A,r)"
+ "\<lbrakk>wf[A](r); linear(A,r)\<rbrakk> \<Longrightarrow> well_ord(A,r)"
apply (simp add: irrefl_def part_ord_def tot_ord_def
trans_on_def well_ord_def wf_on_not_refl)
apply (fast elim: linearE wf_on_asym wf_on_chain3)
done
lemma well_ord_is_wf:
- "well_ord(A,r) ==> wf[A](r)"
+ "well_ord(A,r) \<Longrightarrow> wf[A](r)"
by (unfold well_ord_def, safe)
lemma well_ord_is_trans_on:
- "well_ord(A,r) ==> trans[A](r)"
+ "well_ord(A,r) \<Longrightarrow> trans[A](r)"
by (unfold well_ord_def tot_ord_def part_ord_def, safe)
-lemma well_ord_is_linear: "well_ord(A,r) ==> linear(A,r)"
+lemma well_ord_is_linear: "well_ord(A,r) \<Longrightarrow> linear(A,r)"
by (unfold well_ord_def tot_ord_def, blast)
@@ -107,7 +107,7 @@
lemmas predI = conjI [THEN pred_iff [THEN iffD2]]
-lemma predE: "[| y \<in> pred(A,x,r); [| y \<in> A; <y,x>:r |] ==> P |] ==> P"
+lemma predE: "\<lbrakk>y \<in> pred(A,x,r); \<lbrakk>y \<in> A; <y,x>:r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp add: pred_def)
lemma pred_subset_under: "pred(A,x,r) \<subseteq> r -`` {x}"
@@ -121,8 +121,8 @@
by (simp add: pred_def, blast)
lemma trans_pred_pred_eq:
- "[| trans[A](r); <y,x>:r; x \<in> A; y \<in> A |]
- ==> pred(pred(A,x,r), y, r) = pred(A,y,r)"
+ "\<lbrakk>trans[A](r); <y,x>:r; x \<in> A; y \<in> A\<rbrakk>
+ \<Longrightarrow> pred(pred(A,x,r), y, r) = pred(A,y,r)"
by (unfold trans_on_def pred_def, blast)
@@ -133,21 +133,21 @@
(*Note: a relation s such that s<=r need not be a partial ordering*)
lemma part_ord_subset:
- "[| part_ord(A,r); B<=A |] ==> part_ord(B,r)"
+ "\<lbrakk>part_ord(A,r); B<=A\<rbrakk> \<Longrightarrow> part_ord(B,r)"
by (unfold part_ord_def irrefl_def trans_on_def, blast)
lemma linear_subset:
- "[| linear(A,r); B<=A |] ==> linear(B,r)"
+ "\<lbrakk>linear(A,r); B<=A\<rbrakk> \<Longrightarrow> linear(B,r)"
by (unfold linear_def, blast)
lemma tot_ord_subset:
- "[| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)"
+ "\<lbrakk>tot_ord(A,r); B<=A\<rbrakk> \<Longrightarrow> tot_ord(B,r)"
apply (unfold tot_ord_def)
apply (fast elim!: part_ord_subset linear_subset)
done
lemma well_ord_subset:
- "[| well_ord(A,r); B<=A |] ==> well_ord(B,r)"
+ "\<lbrakk>well_ord(A,r); B<=A\<rbrakk> \<Longrightarrow> well_ord(B,r)"
apply (unfold well_ord_def)
apply (fast elim!: tot_ord_subset wf_on_subset_A)
done
@@ -239,40 +239,40 @@
(** Order-preserving (monotone) maps **)
-lemma mono_map_is_fun: "f \<in> mono_map(A,r,B,s) ==> f \<in> A->B"
+lemma mono_map_is_fun: "f \<in> mono_map(A,r,B,s) \<Longrightarrow> f \<in> A->B"
by (simp add: mono_map_def)
lemma mono_map_is_inj:
- "[| linear(A,r); wf[B](s); f \<in> mono_map(A,r,B,s) |] ==> f \<in> inj(A,B)"
+ "\<lbrakk>linear(A,r); wf[B](s); f \<in> mono_map(A,r,B,s)\<rbrakk> \<Longrightarrow> f \<in> inj(A,B)"
apply (unfold mono_map_def inj_def, clarify)
apply (erule_tac x=w and y=x in linearE, assumption+)
apply (force intro: apply_type dest: wf_on_not_refl)+
done
lemma ord_isoI:
- "[| f \<in> bij(A, B);
- !!x y. [| x \<in> A; y \<in> A |] ==> <x, y> \<in> r \<longleftrightarrow> <f`x, f`y> \<in> s |]
- ==> f \<in> ord_iso(A,r,B,s)"
+ "\<lbrakk>f \<in> bij(A, B);
+ \<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <x, y> \<in> r \<longleftrightarrow> <f`x, f`y> \<in> s\<rbrakk>
+ \<Longrightarrow> f \<in> ord_iso(A,r,B,s)"
by (simp add: ord_iso_def)
lemma ord_iso_is_mono_map:
- "f \<in> ord_iso(A,r,B,s) ==> f \<in> mono_map(A,r,B,s)"
+ "f \<in> ord_iso(A,r,B,s) \<Longrightarrow> f \<in> mono_map(A,r,B,s)"
apply (simp add: ord_iso_def mono_map_def)
apply (blast dest!: bij_is_fun)
done
lemma ord_iso_is_bij:
- "f \<in> ord_iso(A,r,B,s) ==> f \<in> bij(A,B)"
+ "f \<in> ord_iso(A,r,B,s) \<Longrightarrow> f \<in> bij(A,B)"
by (simp add: ord_iso_def)
(*Needed? But ord_iso_converse is!*)
lemma ord_iso_apply:
- "[| f \<in> ord_iso(A,r,B,s); <x,y>: r; x \<in> A; y \<in> A |] ==> <f`x, f`y> \<in> s"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); <x,y>: r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <f`x, f`y> \<in> s"
by (simp add: ord_iso_def)
lemma ord_iso_converse:
- "[| f \<in> ord_iso(A,r,B,s); <x,y>: s; x \<in> B; y \<in> B |]
- ==> <converse(f) ` x, converse(f) ` y> \<in> r"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); <x,y>: s; x \<in> B; y \<in> B\<rbrakk>
+ \<Longrightarrow> <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])+
@@ -287,7 +287,7 @@
by (rule id_bij [THEN ord_isoI], simp)
(*Symmetry of similarity*)
-lemma ord_iso_sym: "f \<in> ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)"
+lemma ord_iso_sym: "f \<in> ord_iso(A,r,B,s) \<Longrightarrow> converse(f): ord_iso(B,s,A,r)"
apply (simp add: ord_iso_def)
apply (auto simp add: right_inverse_bij bij_converse_bij
bij_is_fun [THEN apply_funtype])
@@ -295,16 +295,16 @@
(*Transitivity of similarity*)
lemma mono_map_trans:
- "[| g \<in> mono_map(A,r,B,s); f \<in> mono_map(B,s,C,t) |]
- ==> (f O g): mono_map(A,r,C,t)"
+ "\<lbrakk>g \<in> mono_map(A,r,B,s); f \<in> mono_map(B,s,C,t)\<rbrakk>
+ \<Longrightarrow> (f O g): mono_map(A,r,C,t)"
apply (unfold mono_map_def)
apply (auto simp add: comp_fun)
done
(*Transitivity of similarity: the order-isomorphism relation*)
lemma ord_iso_trans:
- "[| g \<in> ord_iso(A,r,B,s); f \<in> ord_iso(B,s,C,t) |]
- ==> (f O g): ord_iso(A,r,C,t)"
+ "\<lbrakk>g \<in> ord_iso(A,r,B,s); f \<in> ord_iso(B,s,C,t)\<rbrakk>
+ \<Longrightarrow> (f O g): ord_iso(A,r,C,t)"
apply (unfold ord_iso_def, clarify)
apply (frule bij_is_fun [of f])
apply (frule bij_is_fun [of g])
@@ -314,8 +314,8 @@
(** Two monotone maps can make an order-isomorphism **)
lemma mono_ord_isoI:
- "[| f \<in> mono_map(A,r,B,s); g \<in> mono_map(B,s,A,r);
- f O g = id(B); g O f = id(A) |] ==> f \<in> ord_iso(A,r,B,s)"
+ "\<lbrakk>f \<in> mono_map(A,r,B,s); g \<in> mono_map(B,s,A,r);
+ f O g = id(B); g O f = id(A)\<rbrakk> \<Longrightarrow> f \<in> 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) > \<in> r")
@@ -324,9 +324,9 @@
done
lemma well_ord_mono_ord_isoI:
- "[| well_ord(A,r); well_ord(B,s);
- f \<in> mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) |]
- ==> f \<in> ord_iso(A,r,B,s)"
+ "\<lbrakk>well_ord(A,r); well_ord(B,s);
+ f \<in> mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r)\<rbrakk>
+ \<Longrightarrow> f \<in> ord_iso(A,r,B,s)"
apply (intro mono_ord_isoI, auto)
apply (frule mono_map_is_fun [THEN fun_is_rel])
apply (erule converse_converse [THEN subst], rule left_comp_inverse)
@@ -338,13 +338,13 @@
(** Order-isomorphisms preserve the ordering's properties **)
lemma part_ord_ord_iso:
- "[| part_ord(B,s); f \<in> ord_iso(A,r,B,s) |] ==> part_ord(A,r)"
+ "\<lbrakk>part_ord(B,s); f \<in> ord_iso(A,r,B,s)\<rbrakk> \<Longrightarrow> part_ord(A,r)"
apply (simp add: part_ord_def irrefl_def trans_on_def ord_iso_def)
apply (fast intro: bij_is_fun [THEN apply_type])
done
lemma linear_ord_iso:
- "[| linear(B,s); f \<in> ord_iso(A,r,B,s) |] ==> linear(A,r)"
+ "\<lbrakk>linear(B,s); f \<in> ord_iso(A,r,B,s)\<rbrakk> \<Longrightarrow> linear(A,r)"
apply (simp add: linear_def ord_iso_def, safe)
apply (drule_tac x1 = "f`x" and x = "f`y" in bspec [THEN bspec])
apply (safe elim!: bij_is_fun [THEN apply_type])
@@ -353,7 +353,7 @@
done
lemma wf_on_ord_iso:
- "[| wf[B](s); f \<in> ord_iso(A,r,B,s) |] ==> wf[A](r)"
+ "\<lbrakk>wf[B](s); f \<in> ord_iso(A,r,B,s)\<rbrakk> \<Longrightarrow> wf[A](r)"
apply (simp add: wf_on_def wf_def ord_iso_def, safe)
apply (drule_tac x = "{f`z. z \<in> Z \<inter> A}" in spec)
apply (safe intro!: equalityI)
@@ -361,7 +361,7 @@
done
lemma well_ord_ord_iso:
- "[| well_ord(B,s); f \<in> ord_iso(A,r,B,s) |] ==> well_ord(A,r)"
+ "\<lbrakk>well_ord(B,s); f \<in> ord_iso(A,r,B,s)\<rbrakk> \<Longrightarrow> well_ord(A,r)"
apply (unfold well_ord_def tot_ord_def)
apply (fast elim!: part_ord_ord_iso linear_ord_iso wf_on_ord_iso)
done
@@ -372,8 +372,8 @@
(*Inductive argument for Kunen's Lemma 6.1, etc.
Simple proof from Halmos, page 72*)
lemma well_ord_iso_subset_lemma:
- "[| well_ord(A,r); f \<in> ord_iso(A,r, A',r); A'<= A; y \<in> A |]
- ==> ~ <f`y, y>: r"
+ "\<lbrakk>well_ord(A,r); f \<in> ord_iso(A,r, A',r); A'<= A; y \<in> A\<rbrakk>
+ \<Longrightarrow> \<not> <f`y, y>: r"
apply (simp add: well_ord_def ord_iso_def)
apply (elim conjE CollectE)
apply (rule_tac a=y in wf_on_induct, assumption+)
@@ -383,7 +383,7 @@
(*Kunen's Lemma 6.1 \<in> there's no order-isomorphism to an initial segment
of a well-ordering*)
lemma well_ord_iso_predE:
- "[| well_ord(A,r); f \<in> ord_iso(A, r, pred(A,x,r), r); x \<in> A |] ==> P"
+ "\<lbrakk>well_ord(A,r); f \<in> ord_iso(A, r, pred(A,x,r), r); x \<in> A\<rbrakk> \<Longrightarrow> 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 *)
@@ -394,8 +394,8 @@
(*Simple consequence of Lemma 6.1*)
lemma well_ord_iso_pred_eq:
- "[| well_ord(A,r); f \<in> ord_iso(pred(A,a,r), r, pred(A,c,r), r);
- a \<in> A; c \<in> A |] ==> a=c"
+ "\<lbrakk>well_ord(A,r); f \<in> ord_iso(pred(A,a,r), r, pred(A,c,r), r);
+ a \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> a=c"
apply (frule well_ord_is_trans_on)
apply (frule well_ord_is_linear)
apply (erule_tac x=a and y=c in linearE, assumption+)
@@ -408,7 +408,7 @@
(*Does not assume r is a wellordering!*)
lemma ord_iso_image_pred:
- "[|f \<in> ord_iso(A,r,B,s); a \<in> A|] ==> f `` pred(A,a,r) = pred(B, f`a, s)"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); a \<in> A\<rbrakk> \<Longrightarrow> 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])
@@ -420,8 +420,8 @@
done
lemma ord_iso_restrict_image:
- "[| f \<in> ord_iso(A,r,B,s); C<=A |]
- ==> restrict(f,C) \<in> ord_iso(C, r, f``C, s)"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); C<=A\<rbrakk>
+ \<Longrightarrow> 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
@@ -429,18 +429,18 @@
(*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 \<in> ord_iso(A,r,B,s); a \<in> A |]
- ==> restrict(f, pred(A,a,r)) \<in> ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); a \<in> A\<rbrakk>
+ \<Longrightarrow> 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;
+ "\<lbrakk>well_ord(A,r); well_ord(B,s); <a,c>: r;
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 \<in> A; c \<in> A; b \<in> B; d \<in> B |] ==> <b,d>: s"
+ a \<in> A; c \<in> A; b \<in> B; d \<in> B\<rbrakk> \<Longrightarrow> <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")
apply (simp (no_asm_simp))
@@ -452,9 +452,9 @@
(*See Halmos, page 72*)
lemma well_ord_iso_unique_lemma:
- "[| well_ord(A,r);
- f \<in> ord_iso(A,r, B,s); g \<in> ord_iso(A,r, B,s); y \<in> A |]
- ==> ~ <g`y, f`y> \<in> s"
+ "\<lbrakk>well_ord(A,r);
+ f \<in> ord_iso(A,r, B,s); g \<in> ord_iso(A,r, B,s); y \<in> A\<rbrakk>
+ \<Longrightarrow> \<not> <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
@@ -470,8 +470,8 @@
(*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*)
-lemma well_ord_iso_unique: "[| well_ord(A,r);
- f \<in> ord_iso(A,r, B,s); g \<in> ord_iso(A,r, B,s) |] ==> f = g"
+lemma well_ord_iso_unique: "\<lbrakk>well_ord(A,r);
+ f \<in> ord_iso(A,r, B,s); g \<in> ord_iso(A,r, B,s)\<rbrakk> \<Longrightarrow> f = g"
apply (rule fun_extension)
apply (erule ord_iso_is_bij [THEN bij_is_fun])+
apply (subgoal_tac "f`x \<in> B & g`x \<in> B & linear(B,s)")
@@ -499,19 +499,19 @@
done
lemma function_ord_iso_map:
- "well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))"
+ "well_ord(B,s) \<Longrightarrow> function(ord_iso_map(A,r,B,s))"
apply (unfold ord_iso_map_def function_def)
apply (blast intro: well_ord_iso_pred_eq ord_iso_sym ord_iso_trans)
done
-lemma ord_iso_map_fun: "well_ord(B,s) ==> ord_iso_map(A,r,B,s)
+lemma ord_iso_map_fun: "well_ord(B,s) \<Longrightarrow> 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)
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> ord_iso_map(A,r,B,s)
\<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)
@@ -524,7 +524,7 @@
done
lemma ord_iso_map_ord_iso:
- "[| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s)
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk> \<Longrightarrow> ord_iso_map(A,r,B,s)
\<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)
@@ -539,9 +539,9 @@
(*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 \<in> A; a \<notin> domain(ord_iso_map(A,r,B,s)) |]
- ==> domain(ord_iso_map(A,r,B,s)) \<subseteq> pred(A, a, r)"
+ "\<lbrakk>well_ord(A,r); well_ord(B,s);
+ a \<in> A; a \<notin> domain(ord_iso_map(A,r,B,s))\<rbrakk>
+ \<Longrightarrow> 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 *)
@@ -563,8 +563,8 @@
(*For the 4-way case analysis in the main result*)
lemma domain_ord_iso_map_cases:
- "[| well_ord(A,r); well_ord(B,s) |]
- ==> domain(ord_iso_map(A,r,B,s)) = A |
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> domain(ord_iso_map(A,r,B,s)) = A |
(\<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)
@@ -582,8 +582,8 @@
(*As above, by duality*)
lemma range_ord_iso_map_cases:
- "[| well_ord(A,r); well_ord(B,s) |]
- ==> range(ord_iso_map(A,r,B,s)) = B |
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> range(ord_iso_map(A,r,B,s)) = B |
(\<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)
@@ -591,8 +591,8 @@
text\<open>Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets\<close>
theorem well_ord_trichotomy:
- "[| well_ord(A,r); well_ord(B,s) |]
- ==> ord_iso_map(A,r,B,s) \<in> ord_iso(A, r, B, s) |
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> 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)
@@ -614,21 +614,21 @@
(** Properties of converse(r) **)
-lemma irrefl_converse: "irrefl(A,r) ==> irrefl(A,converse(r))"
+lemma irrefl_converse: "irrefl(A,r) \<Longrightarrow> irrefl(A,converse(r))"
by (unfold irrefl_def, blast)
-lemma trans_on_converse: "trans[A](r) ==> trans[A](converse(r))"
+lemma trans_on_converse: "trans[A](r) \<Longrightarrow> trans[A](converse(r))"
by (unfold trans_on_def, blast)
-lemma part_ord_converse: "part_ord(A,r) ==> part_ord(A,converse(r))"
+lemma part_ord_converse: "part_ord(A,r) \<Longrightarrow> part_ord(A,converse(r))"
apply (unfold part_ord_def)
apply (blast intro!: irrefl_converse trans_on_converse)
done
-lemma linear_converse: "linear(A,r) ==> linear(A,converse(r))"
+lemma linear_converse: "linear(A,r) \<Longrightarrow> linear(A,converse(r))"
by (unfold linear_def, blast)
-lemma tot_ord_converse: "tot_ord(A,r) ==> tot_ord(A,converse(r))"
+lemma tot_ord_converse: "tot_ord(A,r) \<Longrightarrow> tot_ord(A,converse(r))"
apply (unfold tot_ord_def)
apply (blast intro!: part_ord_converse linear_converse)
done
@@ -637,11 +637,11 @@
(** By Krzysztof Grabczewski.
Lemmas involving the first element of a well ordered set **)
-lemma first_is_elem: "first(b,B,r) ==> b \<in> B"
+lemma first_is_elem: "first(b,B,r) \<Longrightarrow> b \<in> B"
by (unfold first_def, blast)
lemma well_ord_imp_ex1_first:
- "[| well_ord(A,r); B<=A; B\<noteq>0 |] ==> (\<exists>!b. first(b,B,r))"
+ "\<lbrakk>well_ord(A,r); B<=A; B\<noteq>0\<rbrakk> \<Longrightarrow> (\<exists>!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)
@@ -650,7 +650,7 @@
done
lemma the_first_in:
- "[| well_ord(A,r); B<=A; B\<noteq>0 |] ==> (THE b. first(b,B,r)) \<in> B"
+ "\<lbrakk>well_ord(A,r); B<=A; B\<noteq>0\<rbrakk> \<Longrightarrow> (THE b. first(b,B,r)) \<in> B"
apply (drule well_ord_imp_ex1_first, assumption+)
apply (rule first_is_elem)
apply (erule theI)
@@ -660,7 +660,7 @@
subsection \<open>Lemmas for the Reflexive Orders\<close>
lemma subset_vimage_vimage_iff:
- "[| Preorder(r); A \<subseteq> field(r); B \<subseteq> field(r) |] ==>
+ "\<lbrakk>Preorder(r); A \<subseteq> field(r); B \<subseteq> field(r)\<rbrakk> \<Longrightarrow>
r -`` A \<subseteq> r -`` B \<longleftrightarrow> (\<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
@@ -673,12 +673,12 @@
done
lemma subset_vimage1_vimage1_iff:
- "[| Preorder(r); a \<in> field(r); b \<in> field(r) |] ==>
+ "\<lbrakk>Preorder(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
r -`` {a} \<subseteq> r -`` {b} \<longleftrightarrow> <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 \<in> field(r); b \<in> field(r) |] ==>
+ "\<lbrakk>refl(field(r), r); antisym(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
r `` {a} = r `` {b} \<longleftrightarrow> a = b"
apply rule
apply (frule equality_iffD)
@@ -689,13 +689,13 @@
done
lemma Partial_order_eq_Image1_Image1_iff:
- "[| Partial_order(r); a \<in> field(r); b \<in> field(r) |] ==>
+ "\<lbrakk>Partial_order(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
r `` {a} = r `` {b} \<longleftrightarrow> 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 \<in> field(r); b \<in> field(r) |] ==>
+ "\<lbrakk>refl(field(r), r); antisym(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
r -`` {a} = r -`` {b} \<longleftrightarrow> a = b"
apply rule
apply (frule equality_iffD)
@@ -706,7 +706,7 @@
done
lemma Partial_order_eq_vimage1_vimage1_iff:
- "[| Partial_order(r); a \<in> field(r); b \<in> field(r) |] ==>
+ "\<lbrakk>Partial_order(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
r -`` {a} = r -`` {b} \<longleftrightarrow> a = b"
by (simp add: partial_order_on_def preorder_on_def
Refl_antisym_eq_vimage1_vimage1_iff)
--- a/src/ZF/OrderArith.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/OrderArith.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,7 +10,7 @@
definition
(*disjoint sum of two relations; underlies ordinal addition*)
radd :: "[i,i,i,i]=>i" where
- "radd(A,r,B,s) ==
+ "radd(A,r,B,s) \<equiv>
{z: (A+B) * (A+B).
(\<exists>x y. z = <Inl(x), Inr(y)>) |
(\<exists>x' x. z = <Inl(x'), Inl(x)> & <x',x>:r) |
@@ -19,7 +19,7 @@
definition
(*lexicographic product of two relations; underlies ordinal multiplication*)
rmult :: "[i,i,i,i]=>i" where
- "rmult(A,r,B,s) ==
+ "rmult(A,r,B,s) \<equiv>
{z: (A*B) * (A*B).
\<exists>x' y' x y. z = <<x',y'>, <x,y>> &
(<x',x>: r | (x'=x & <y',y>: s))}"
@@ -27,11 +27,11 @@
definition
(*inverse image of a relation*)
rvimage :: "[i,i,i]=>i" where
- "rvimage(A,f,r) == {z \<in> A*A. \<exists>x y. z = <x,y> & <f`x,f`y>: r}"
+ "rvimage(A,f,r) \<equiv> {z \<in> A*A. \<exists>x y. z = <x,y> & <f`x,f`y>: r}"
definition
measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i" where
- "measure(A,f) == {<x,y>: A*A. f(x) < f(y)}"
+ "measure(A,f) \<equiv> {<x,y>: A*A. f(x) < f(y)}"
subsection\<open>Addition of Relations -- Disjoint Sum\<close>
@@ -59,11 +59,11 @@
subsubsection\<open>Elimination Rule\<close>
lemma raddE:
- "[| <p',p> \<in> radd(A,r,B,s);
- !!x y. [| p'=Inl(x); x \<in> A; p=Inr(y); y \<in> B |] ==> Q;
- !!x' x. [| p'=Inl(x'); p=Inl(x); <x',x>: r; x':A; x \<in> A |] ==> Q;
- !!y' y. [| p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y \<in> B |] ==> Q
- |] ==> Q"
+ "\<lbrakk><p',p> \<in> radd(A,r,B,s);
+ \<And>x y. \<lbrakk>p'=Inl(x); x \<in> A; p=Inr(y); y \<in> B\<rbrakk> \<Longrightarrow> Q;
+ \<And>x' x. \<lbrakk>p'=Inl(x'); p=Inl(x); <x',x>: r; x':A; x \<in> A\<rbrakk> \<Longrightarrow> Q;
+ \<And>y' y. \<lbrakk>p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y \<in> B\<rbrakk> \<Longrightarrow> Q
+\<rbrakk> \<Longrightarrow> Q"
by (unfold radd_def, blast)
subsubsection\<open>Type checking\<close>
@@ -78,13 +78,13 @@
subsubsection\<open>Linearity\<close>
lemma linear_radd:
- "[| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))"
+ "\<lbrakk>linear(A,r); linear(B,s)\<rbrakk> \<Longrightarrow> linear(A+B,radd(A,r,B,s))"
by (unfold linear_def, blast)
subsubsection\<open>Well-foundedness\<close>
-lemma wf_on_radd: "[| wf[A](r); wf[B](s) |] ==> wf[A+B](radd(A,r,B,s))"
+lemma wf_on_radd: "\<lbrakk>wf[A](r); wf[B](s)\<rbrakk> \<Longrightarrow> wf[A+B](radd(A,r,B,s))"
apply (rule wf_onI2)
apply (subgoal_tac "\<forall>x\<in>A. Inl (x) \<in> Ba")
\<comment> \<open>Proving the lemma, which is needed twice!\<close>
@@ -99,14 +99,14 @@
apply (erule_tac r = s and a = ya in wf_on_induct, assumption, blast)
done
-lemma wf_radd: "[| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))"
+lemma wf_radd: "\<lbrakk>wf(r); wf(s)\<rbrakk> \<Longrightarrow> wf(radd(field(r),r,field(s),s))"
apply (simp add: wf_iff_wf_on_field)
apply (rule wf_on_subset_A [OF _ field_radd])
apply (blast intro: wf_on_radd)
done
lemma well_ord_radd:
- "[| well_ord(A,r); well_ord(B,s) |] ==> well_ord(A+B, radd(A,r,B,s))"
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk> \<Longrightarrow> well_ord(A+B, radd(A,r,B,s))"
apply (rule well_ordI)
apply (simp add: well_ord_def wf_on_radd)
apply (simp add: well_ord_def tot_ord_def linear_radd)
@@ -115,8 +115,8 @@
subsubsection\<open>An \<^term>\<open>ord_iso\<close> congruence law\<close>
lemma sum_bij:
- "[| f \<in> bij(A,C); g \<in> bij(B,D) |]
- ==> (\<lambda>z\<in>A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \<in> bij(A+B, C+D)"
+ "\<lbrakk>f \<in> bij(A,C); g \<in> bij(B,D)\<rbrakk>
+ \<Longrightarrow> (\<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)
@@ -124,7 +124,7 @@
done
lemma sum_ord_iso_cong:
- "[| f \<in> ord_iso(A,r,A',r'); g \<in> ord_iso(B,s,B',s') |] ==>
+ "\<lbrakk>f \<in> ord_iso(A,r,A',r'); g \<in> ord_iso(B,s,B',s')\<rbrakk> \<Longrightarrow>
(\<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)
@@ -135,7 +135,7 @@
(*Could we prove an ord_iso result? Perhaps
ord_iso(A+B, radd(A,r,B,s), A \<union> B, r \<union> s) *)
-lemma sum_disjoint_bij: "A \<inter> B = 0 ==>
+lemma sum_disjoint_bij: "A \<inter> B = 0 \<Longrightarrow>
(\<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 \<in> A then Inl (z) else Inr (z) " in lam_bijective)
apply auto
@@ -170,10 +170,10 @@
by (unfold rmult_def, blast)
lemma rmultE:
- "[| <<a',b'>, <a,b>> \<in> rmult(A,r,B,s);
- [| <a',a>: r; a':A; a \<in> A; b':B; b \<in> B |] ==> Q;
- [| <b',b>: s; a \<in> A; a'=a; b':B; b \<in> B |] ==> Q
- |] ==> Q"
+ "\<lbrakk><<a',b'>, <a,b>> \<in> rmult(A,r,B,s);
+ \<lbrakk><a',a>: r; a':A; a \<in> A; b':B; b \<in> B\<rbrakk> \<Longrightarrow> Q;
+ \<lbrakk><b',b>: s; a \<in> A; a'=a; b':B; b \<in> B\<rbrakk> \<Longrightarrow> Q
+\<rbrakk> \<Longrightarrow> Q"
by blast
subsubsection\<open>Type checking\<close>
@@ -186,12 +186,12 @@
subsubsection\<open>Linearity\<close>
lemma linear_rmult:
- "[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))"
+ "\<lbrakk>linear(A,r); linear(B,s)\<rbrakk> \<Longrightarrow> linear(A*B,rmult(A,r,B,s))"
by (simp add: linear_def, blast)
subsubsection\<open>Well-foundedness\<close>
-lemma wf_on_rmult: "[| wf[A](r); wf[B](s) |] ==> wf[A*B](rmult(A,r,B,s))"
+lemma wf_on_rmult: "\<lbrakk>wf[A](r); wf[B](s)\<rbrakk> \<Longrightarrow> wf[A*B](rmult(A,r,B,s))"
apply (rule wf_onI2)
apply (erule SigmaE)
apply (erule ssubst)
@@ -203,14 +203,14 @@
done
-lemma wf_rmult: "[| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))"
+lemma wf_rmult: "\<lbrakk>wf(r); wf(s)\<rbrakk> \<Longrightarrow> wf(rmult(field(r),r,field(s),s))"
apply (simp add: wf_iff_wf_on_field)
apply (rule wf_on_subset_A [OF _ field_rmult])
apply (blast intro: wf_on_rmult)
done
lemma well_ord_rmult:
- "[| well_ord(A,r); well_ord(B,s) |] ==> well_ord(A*B, rmult(A,r,B,s))"
+ "\<lbrakk>well_ord(A,r); well_ord(B,s)\<rbrakk> \<Longrightarrow> well_ord(A*B, rmult(A,r,B,s))"
apply (rule well_ordI)
apply (simp add: well_ord_def wf_on_rmult)
apply (simp add: well_ord_def tot_ord_def linear_rmult)
@@ -220,8 +220,8 @@
subsubsection\<open>An \<^term>\<open>ord_iso\<close> congruence law\<close>
lemma prod_bij:
- "[| f \<in> bij(A,C); g \<in> bij(B,D) |]
- ==> (lam <x,y>:A*B. <f`x, g`y>) \<in> bij(A*B, C*D)"
+ "\<lbrakk>f \<in> bij(A,C); g \<in> bij(B,D)\<rbrakk>
+ \<Longrightarrow> (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)
@@ -229,8 +229,8 @@
done
lemma prod_ord_iso_cong:
- "[| f \<in> ord_iso(A,r,A',r'); g \<in> ord_iso(B,s,B',s') |]
- ==> (lam <x,y>:A*B. <f`x, g`y>)
+ "\<lbrakk>f \<in> ord_iso(A,r,A',r'); g \<in> ord_iso(B,s,B',s')\<rbrakk>
+ \<Longrightarrow> (lam <x,y>:A*B. <f`x, g`y>)
\<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)
@@ -243,7 +243,7 @@
(*Used??*)
lemma singleton_prod_ord_iso:
- "well_ord({x},xr) ==>
+ "well_ord({x},xr) \<Longrightarrow>
(\<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))
@@ -253,7 +253,7 @@
(*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\<notin>C ==>
+ "a\<notin>C \<Longrightarrow>
(\<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)
@@ -267,7 +267,7 @@
done
lemma prod_sum_singleton_ord_iso:
- "[| a \<in> A; well_ord(A,r) |] ==>
+ "\<lbrakk>a \<in> A; well_ord(A,r)\<rbrakk> \<Longrightarrow>
(\<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),
@@ -324,19 +324,19 @@
subsubsection\<open>Partial Ordering Properties\<close>
lemma irrefl_rvimage:
- "[| f \<in> inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); irrefl(B,r)\<rbrakk> \<Longrightarrow> irrefl(A, rvimage(A,f,r))"
apply (unfold irrefl_def rvimage_def)
apply (blast intro: inj_is_fun [THEN apply_type])
done
lemma trans_on_rvimage:
- "[| f \<in> inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); trans[B](r)\<rbrakk> \<Longrightarrow> trans[A](rvimage(A,f,r))"
apply (unfold trans_on_def rvimage_def)
apply (blast intro: inj_is_fun [THEN apply_type])
done
lemma part_ord_rvimage:
- "[| f \<in> inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); part_ord(B,r)\<rbrakk> \<Longrightarrow> part_ord(A, rvimage(A,f,r))"
apply (unfold part_ord_def)
apply (blast intro!: irrefl_rvimage trans_on_rvimage)
done
@@ -344,13 +344,13 @@
subsubsection\<open>Linearity\<close>
lemma linear_rvimage:
- "[| f \<in> inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); linear(B,r)\<rbrakk> \<Longrightarrow> linear(A,rvimage(A,f,r))"
apply (simp add: inj_def linear_def rvimage_iff)
apply (blast intro: apply_funtype)
done
lemma tot_ord_rvimage:
- "[| f \<in> inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); tot_ord(B,r)\<rbrakk> \<Longrightarrow> tot_ord(A, rvimage(A,f,r))"
apply (unfold tot_ord_def)
apply (blast intro!: part_ord_rvimage linear_rvimage)
done
@@ -358,7 +358,7 @@
subsubsection\<open>Well-foundedness\<close>
-lemma wf_rvimage [intro!]: "wf(r) ==> wf(rvimage(A,f,r))"
+lemma wf_rvimage [intro!]: "wf(r) \<Longrightarrow> wf(rvimage(A,f,r))"
apply (simp (no_asm_use) add: rvimage_def wf_eq_minimal)
apply clarify
apply (subgoal_tac "\<exists>w. w \<in> {w: {f`x. x \<in> Q}. \<exists>x. x \<in> Q & (f`x = w) }")
@@ -370,8 +370,8 @@
done
text\<open>But note that the combination of \<open>wf_imp_wf_on\<close> and
- \<open>wf_rvimage\<close> gives \<^prop>\<open>wf(r) ==> wf[C](rvimage(A,f,r))\<close>\<close>
-lemma wf_on_rvimage: "[| f \<in> A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))"
+ \<open>wf_rvimage\<close> gives \<^prop>\<open>wf(r) \<Longrightarrow> wf[C](rvimage(A,f,r))\<close>\<close>
+lemma wf_on_rvimage: "\<lbrakk>f \<in> A->B; wf[B](r)\<rbrakk> \<Longrightarrow> wf[A](rvimage(A,f,r))"
apply (rule wf_onI2)
apply (subgoal_tac "\<forall>z\<in>A. f`z=f`y \<longrightarrow> z \<in> Ba")
apply blast
@@ -382,7 +382,7 @@
(*Note that we need only wf[A](...) and linear(A,...) to get the result!*)
lemma well_ord_rvimage:
- "[| f \<in> inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))"
+ "\<lbrakk>f \<in> inj(A,B); well_ord(B,r)\<rbrakk> \<Longrightarrow> well_ord(A, rvimage(A,f,r))"
apply (rule well_ordI)
apply (unfold well_ord_def tot_ord_def)
apply (blast intro!: wf_on_rvimage inj_is_fun)
@@ -390,13 +390,13 @@
done
lemma ord_iso_rvimage:
- "f \<in> bij(A,B) ==> f \<in> ord_iso(A, rvimage(A,f,s), B, s)"
+ "f \<in> bij(A,B) \<Longrightarrow> f \<in> ord_iso(A, rvimage(A,f,s), B, s)"
apply (unfold ord_iso_def)
apply (simp add: rvimage_iff)
done
lemma ord_iso_rvimage_eq:
- "f \<in> ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r \<inter> A*A"
+ "f \<in> ord_iso(A,r, B,s) \<Longrightarrow> rvimage(A,f,s) = r \<inter> A*A"
by (unfold ord_iso_def rvimage_def, blast)
@@ -409,28 +409,28 @@
definition
wfrank :: "[i,i]=>i" where
- "wfrank(r,a) == wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))"
+ "wfrank(r,a) \<equiv> wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))"
definition
wftype :: "i=>i" where
- "wftype(r) == \<Union>y \<in> range(r). succ(wfrank(r,y))"
+ "wftype(r) \<equiv> \<Union>y \<in> range(r). succ(wfrank(r,y))"
-lemma wfrank: "wf(r) ==> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))"
+lemma wfrank: "wf(r) \<Longrightarrow> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))"
by (subst wfrank_def [THEN def_wfrec], simp_all)
-lemma Ord_wfrank: "wf(r) ==> Ord(wfrank(r,a))"
+lemma Ord_wfrank: "wf(r) \<Longrightarrow> Ord(wfrank(r,a))"
apply (rule_tac a=a in wf_induct, assumption)
apply (subst wfrank, assumption)
apply (rule Ord_succ [THEN Ord_UN], blast)
done
-lemma wfrank_lt: "[|wf(r); <a,b> \<in> r|] ==> wfrank(r,a) < wfrank(r,b)"
+lemma wfrank_lt: "\<lbrakk>wf(r); <a,b> \<in> r\<rbrakk> \<Longrightarrow> wfrank(r,a) < wfrank(r,b)"
apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption)
apply (rule UN_I [THEN ltI])
apply (simp add: Ord_wfrank vimage_iff)+
done
-lemma Ord_wftype: "wf(r) ==> Ord(wftype(r))"
+lemma Ord_wftype: "wf(r) \<Longrightarrow> Ord(wftype(r))"
by (simp add: wftype_def Ord_wfrank)
lemma wftypeI: "\<lbrakk>wf(r); x \<in> field(r)\<rbrakk> \<Longrightarrow> wfrank(r,x) \<in> wftype(r)"
@@ -440,7 +440,7 @@
lemma wf_imp_subset_rvimage:
- "[|wf(r); r \<subseteq> A*A|] ==> \<exists>i f. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i))"
+ "\<lbrakk>wf(r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> \<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,19 +450,19 @@
done
theorem wf_iff_subset_rvimage:
- "relation(r) ==> wf(r) \<longleftrightarrow> (\<exists>i f A. Ord(i) & r \<subseteq> rvimage(A, f, Memrel(i)))"
+ "relation(r) \<Longrightarrow> wf(r) \<longleftrightarrow> (\<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\<open>Other Results\<close>
-lemma wf_times: "A \<inter> B = 0 ==> wf(A*B)"
+lemma wf_times: "A \<inter> B = 0 \<Longrightarrow> wf(A*B)"
by (simp add: wf_def, blast)
text\<open>Could also be used to prove \<open>wf_radd\<close>\<close>
lemma wf_Un:
- "[| range(r) \<inter> domain(s) = 0; wf(r); wf(s) |] ==> wf(r \<union> s)"
+ "\<lbrakk>range(r) \<inter> domain(s) = 0; wf(r); wf(s)\<rbrakk> \<Longrightarrow> wf(r \<union> s)"
apply (simp add: wf_def, clarify)
apply (rule equalityI)
prefer 2 apply blast
@@ -500,8 +500,8 @@
by (simp (no_asm) add: measure_def)
lemma linear_measure:
- assumes Ordf: "!!x. x \<in> A ==> Ord(f(x))"
- and inj: "!!x y. [|x \<in> A; y \<in> A; f(x) = f(y) |] ==> x=y"
+ assumes Ordf: "\<And>x. x \<in> A \<Longrightarrow> Ord(f(x))"
+ and inj: "\<And>x y. \<lbrakk>x \<in> A; y \<in> A; f(x) = f(y)\<rbrakk> \<Longrightarrow> x=y"
shows "linear(A, measure(A,f))"
apply (auto simp add: linear_def)
apply (rule_tac i="f(x)" and j="f(y)" in Ord_linear_lt)
@@ -513,8 +513,8 @@
by (rule wf_imp_wf_on [OF wf_measure])
lemma well_ord_measure:
- assumes Ordf: "!!x. x \<in> A ==> Ord(f(x))"
- and inj: "!!x y. [|x \<in> A; y \<in> A; f(x) = f(y) |] ==> x=y"
+ assumes Ordf: "\<And>x. x \<in> A \<Longrightarrow> Ord(f(x))"
+ and inj: "\<And>x y. \<lbrakk>x \<in> A; y \<in> A; f(x) = f(y)\<rbrakk> \<Longrightarrow> x=y"
shows "well_ord(A, measure(A,f))"
apply (rule well_ordI)
apply (rule wf_on_measure)
@@ -528,9 +528,9 @@
lemma wf_on_Union:
assumes wfA: "wf[A](r)"
- and wfB: "!!a. a\<in>A ==> wf[B(a)](s)"
- and ok: "!!a u v. [|<u,v> \<in> s; v \<in> B(a); a \<in> A|]
- ==> (\<exists>a'\<in>A. <a',a> \<in> r & u \<in> B(a')) | u \<in> B(a)"
+ and wfB: "\<And>a. a\<in>A \<Longrightarrow> wf[B(a)](s)"
+ and ok: "\<And>a u v. \<lbrakk><u,v> \<in> s; v \<in> B(a); a \<in> A\<rbrakk>
+ \<Longrightarrow> (\<exists>a'\<in>A. <a',a> \<in> r & u \<in> B(a')) | u \<in> B(a)"
shows "wf[\<Union>a\<in>A. B(a)](s)"
apply (rule wf_onI2)
apply (erule UN_E)
--- a/src/ZF/OrderType.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/OrderType.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,45 +13,45 @@
definition
ordermap :: "[i,i]=>i" where
- "ordermap(A,r) == \<lambda>x\<in>A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
+ "ordermap(A,r) \<equiv> \<lambda>x\<in>A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
definition
ordertype :: "[i,i]=>i" where
- "ordertype(A,r) == ordermap(A,r)``A"
+ "ordertype(A,r) \<equiv> ordermap(A,r)``A"
definition
(*alternative definition of ordinal numbers*)
Ord_alt :: "i => o" where
- "Ord_alt(X) == well_ord(X, Memrel(X)) & (\<forall>u\<in>X. u=pred(X, u, Memrel(X)))"
+ "Ord_alt(X) \<equiv> well_ord(X, Memrel(X)) & (\<forall>u\<in>X. u=pred(X, u, Memrel(X)))"
definition
(*coercion to ordinal: if not, just 0*)
ordify :: "i=>i" where
- "ordify(x) == if Ord(x) then x else 0"
+ "ordify(x) \<equiv> if Ord(x) then x else 0"
definition
(*ordinal multiplication*)
omult :: "[i,i]=>i" (infixl \<open>**\<close> 70) where
- "i ** j == ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))"
+ "i ** j \<equiv> ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))"
definition
(*ordinal addition*)
raw_oadd :: "[i,i]=>i" where
- "raw_oadd(i,j) == ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))"
+ "raw_oadd(i,j) \<equiv> ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))"
definition
oadd :: "[i,i]=>i" (infixl \<open>++\<close> 65) where
- "i ++ j == raw_oadd(ordify(i),ordify(j))"
+ "i ++ j \<equiv> raw_oadd(ordify(i),ordify(j))"
definition
(*ordinal subtraction*)
odiff :: "[i,i]=>i" (infixl \<open>--\<close> 65) where
- "i -- j == ordertype(i-j, Memrel(i))"
+ "i -- j \<equiv> ordertype(i-j, Memrel(i))"
subsection\<open>Proofs needing the combination of Ordinal.thy and Order.thy\<close>
-lemma le_well_ord_Memrel: "j \<le> i ==> well_ord(j, Memrel(i))"
+lemma le_well_ord_Memrel: "j \<le> i \<Longrightarrow> 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
@@ -60,23 +60,23 @@
apply (blast intro: Ord_in_Ord lt_Ord)+
done
-(*"Ord(i) ==> well_ord(i, Memrel(i))"*)
+(*"Ord(i) \<Longrightarrow> well_ord(i, Memrel(i))"*)
lemmas well_ord_Memrel = le_refl [THEN le_well_ord_Memrel]
(*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:
- "j<i ==> pred(i, j, Memrel(i)) = j"
+ "j<i \<Longrightarrow> pred(i, j, Memrel(i)) = j"
apply (simp add: pred_def lt_def)
apply (blast intro: Ord_trans)
done
lemma pred_Memrel:
- "x \<in> A ==> pred(A, x, Memrel(A)) = A \<inter> x"
+ "x \<in> A \<Longrightarrow> pred(A, x, Memrel(A)) = A \<inter> x"
by (unfold pred_def Memrel_def, blast)
lemma Ord_iso_implies_eq_lemma:
- "[| j<i; f \<in> ord_iso(i,Memrel(i),j,Memrel(j)) |] ==> R"
+ "\<lbrakk>j<i; f \<in> ord_iso(i,Memrel(i),j,Memrel(j))\<rbrakk> \<Longrightarrow> R"
apply (frule lt_pred_Memrel)
apply (erule ltE)
apply (rule well_ord_Memrel [THEN well_ord_iso_predE, of i f j], auto)
@@ -88,8 +88,8 @@
(*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*)
lemma Ord_iso_implies_eq:
- "[| Ord(i); Ord(j); f \<in> ord_iso(i,Memrel(i),j,Memrel(j)) |]
- ==> i=j"
+ "\<lbrakk>Ord(i); Ord(j); f \<in> ord_iso(i,Memrel(i),j,Memrel(j))\<rbrakk>
+ \<Longrightarrow> 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)+
done
@@ -108,8 +108,8 @@
(*Useful for cardinality reasoning; see CardinalArith.ML*)
lemma ordermap_eq_image:
- "[| wf[A](r); x \<in> A |]
- ==> ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)"
+ "\<lbrakk>wf[A](r); x \<in> A\<rbrakk>
+ \<Longrightarrow> ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)"
apply (unfold ordermap_def pred_def)
apply (simp (no_asm_simp))
apply (erule wfrec_on [THEN trans], assumption)
@@ -118,8 +118,8 @@
(*Useful for rewriting PROVIDED pred is not unfolded until later!*)
lemma ordermap_pred_unfold:
- "[| wf[A](r); x \<in> A |]
- ==> ordermap(A,r) ` x = {ordermap(A,r)`y . y \<in> pred(A,x,r)}"
+ "\<lbrakk>wf[A](r); x \<in> A\<rbrakk>
+ \<Longrightarrow> 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!*)
@@ -127,8 +127,8 @@
(*The theorem above is
-[| wf[A](r); x \<in> A |]
-==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \<in> A . <y,x> \<in> r}}
+\<lbrakk>wf[A](r); x \<in> A\<rbrakk>
+\<Longrightarrow> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \<in> 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
@@ -144,7 +144,7 @@
subsubsection\<open>Showing that ordermap, ordertype yield ordinals\<close>
lemma Ord_ordermap:
- "[| well_ord(A,r); x \<in> A |] ==> Ord(ordermap(A,r) ` x)"
+ "\<lbrakk>well_ord(A,r); x \<in> A\<rbrakk> \<Longrightarrow> Ord(ordermap(A,r) ` x)"
apply (unfold well_ord_def tot_ord_def part_ord_def, safe)
apply (rule_tac a=x in wf_on_induct, assumption+)
apply (simp (no_asm_simp) add: ordermap_pred_unfold)
@@ -155,7 +155,7 @@
done
lemma Ord_ordertype:
- "well_ord(A,r) ==> Ord(ordertype(A,r))"
+ "well_ord(A,r) \<Longrightarrow> Ord(ordertype(A,r))"
apply (unfold ordertype_def)
apply (subst image_fun [OF ordermap_type subset_refl])
apply (rule OrdI [OF _ Ord_is_Transset])
@@ -169,15 +169,15 @@
subsubsection\<open>ordermap preserves the orderings in both directions\<close>
lemma ordermap_mono:
- "[| <w,x>: r; wf[A](r); w \<in> A; x \<in> A |]
- ==> ordermap(A,r)`w \<in> ordermap(A,r)`x"
+ "\<lbrakk><w,x>: r; wf[A](r); w \<in> A; x \<in> A\<rbrakk>
+ \<Longrightarrow> 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 \<in> ordermap(A,r)`x; well_ord(A,r); w \<in> A; x \<in> A |]
- ==> <w,x>: r"
+ "\<lbrakk>ordermap(A,r)`w \<in> ordermap(A,r)`x; well_ord(A,r); w \<in> A; x \<in> A\<rbrakk>
+ \<Longrightarrow> <w,x>: r"
apply (unfold well_ord_def tot_ord_def, safe)
apply (erule_tac x=w and y=x in linearE, assumption+)
apply (blast elim!: mem_not_refl [THEN notE])
@@ -189,7 +189,7 @@
by (rule surj_image) (rule ordermap_type)
lemma ordermap_bij:
- "well_ord(A,r) ==> ordermap(A,r) \<in> bij(A, ordertype(A,r))"
+ "well_ord(A,r) \<Longrightarrow> 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
@@ -200,7 +200,7 @@
lemma ordertype_ord_iso:
"well_ord(A,r)
- ==> ordermap(A,r) \<in> ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))"
+ \<Longrightarrow> 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
intro!: ordermap_type [THEN apply_type] ordermap_mono ordermap_bij)
@@ -208,16 +208,16 @@
done
lemma ordertype_eq:
- "[| f \<in> ord_iso(A,r,B,s); well_ord(B,s) |]
- ==> ordertype(A,r) = ordertype(B,s)"
+ "\<lbrakk>f \<in> ord_iso(A,r,B,s); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> ordertype(A,r) = ordertype(B,s)"
apply (frule well_ord_ord_iso, assumption)
apply (rule Ord_iso_implies_eq, (erule Ord_ordertype)+)
apply (blast intro: ord_iso_trans ord_iso_sym ordertype_ord_iso)
done
lemma ordertype_eq_imp_ord_iso:
- "[| ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s) |]
- ==> \<exists>f. f \<in> ord_iso(A,r,B,s)"
+ "\<lbrakk>ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> \<exists>f. f \<in> ord_iso(A,r,B,s)"
apply (rule exI)
apply (rule ordertype_ord_iso [THEN ord_iso_trans], assumption)
apply (erule ssubst)
@@ -227,7 +227,7 @@
subsubsection\<open>Basic equalities for ordertype\<close>
(*Ordertype of Memrel*)
-lemma le_ordertype_Memrel: "j \<le> i ==> ordertype(j,Memrel(i)) = j"
+lemma le_ordertype_Memrel: "j \<le> i \<Longrightarrow> 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)
@@ -238,7 +238,7 @@
apply (fast elim: ltE Ord_in_Ord Ord_trans)
done
-(*"Ord(i) ==> ordertype(i, Memrel(i)) = i"*)
+(*"Ord(i) \<Longrightarrow> ordertype(i, Memrel(i)) = i"*)
lemmas ordertype_Memrel = le_refl [THEN le_ordertype_Memrel]
lemma ordertype_0 [simp]: "ordertype(0,r) = 0"
@@ -248,7 +248,7 @@
apply (rule Ord_0 [THEN ordertype_Memrel])
done
-(*Ordertype of rvimage: [| f \<in> bij(A,B); well_ord(B,s) |] ==>
+(*Ordertype of rvimage: \<lbrakk>f \<in> bij(A,B); well_ord(B,s)\<rbrakk> \<Longrightarrow>
ordertype(A, rvimage(A,f,s)) = ordertype(B,s) *)
lemmas bij_ordertype_vimage = ord_iso_rvimage [THEN ordertype_eq]
@@ -256,8 +256,8 @@
(*Ordermap returns the same result if applied to an initial segment*)
lemma ordermap_pred_eq_ordermap:
- "[| well_ord(A,r); y \<in> A; z \<in> pred(A,y,r) |]
- ==> ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"
+ "\<lbrakk>well_ord(A,r); y \<in> A; z \<in> pred(A,y,r)\<rbrakk>
+ \<Longrightarrow> ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"
apply (frule wf_on_subset_A [OF well_ord_is_wf pred_subset])
apply (rule_tac a=z in wf_on_induct, assumption+)
apply (safe elim!: predE)
@@ -278,15 +278,15 @@
text\<open>Theorems by Krzysztof Grabczewski; proofs simplified by lcp\<close>
-lemma ordertype_pred_subset: "[| well_ord(A,r); x \<in> A |] ==>
+lemma ordertype_pred_subset: "\<lbrakk>well_ord(A,r); x \<in> A\<rbrakk> \<Longrightarrow>
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
lemma ordertype_pred_lt:
- "[| well_ord(A,r); x \<in> A |]
- ==> ordertype(pred(A,x,r),r) < ordertype(A,r)"
+ "\<lbrakk>well_ord(A,r); x \<in> A\<rbrakk>
+ \<Longrightarrow> ordertype(pred(A,x,r),r) < ordertype(A,r)"
apply (rule ordertype_pred_subset [THEN subset_imp_le, THEN leE])
apply (simp_all add: Ord_ordertype well_ord_subset [OF _ pred_subset])
apply (erule sym [THEN ordertype_eq_imp_ord_iso, THEN exE])
@@ -298,7 +298,7 @@
well_ord(pred(A,x,r), r) *)
lemma ordertype_pred_unfold:
"well_ord(A,r)
- ==> ordertype(A,r) = {ordertype(pred(A,x,r),r). x \<in> A}"
+ \<Longrightarrow> ordertype(A,r) = {ordertype(pred(A,x,r),r). x \<in> A}"
apply (rule equalityI)
apply (safe intro!: ordertype_pred_lt [THEN ltD])
apply (auto simp add: ordertype_def well_ord_is_wf [THEN ordermap_eq_image]
@@ -310,7 +310,7 @@
subsection\<open>Alternative definition of ordinal\<close>
(*proof by Krzysztof Grabczewski*)
-lemma Ord_is_Ord_alt: "Ord(i) ==> Ord_alt(i)"
+lemma Ord_is_Ord_alt: "Ord(i) \<Longrightarrow> Ord_alt(i)"
apply (unfold Ord_alt_def)
apply (rule conjI)
apply (erule well_ord_Memrel)
@@ -319,7 +319,7 @@
(*proof by lcp*)
lemma Ord_alt_is_Ord:
- "Ord_alt(i) ==> Ord(i)"
+ "Ord_alt(i) \<Longrightarrow> Ord(i)"
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)
@@ -339,7 +339,7 @@
done
lemma ordertype_sum_0_eq:
- "well_ord(A,r) ==> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)"
+ "well_ord(A,r) \<Longrightarrow> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)"
apply (rule bij_sum_0 [THEN ord_isoI, THEN ordertype_eq])
prefer 2 apply assumption
apply force
@@ -351,7 +351,7 @@
done
lemma ordertype_0_sum_eq:
- "well_ord(A,r) ==> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)"
+ "well_ord(A,r) \<Longrightarrow> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)"
apply (rule bij_0_sum [THEN ord_isoI, THEN ordertype_eq])
prefer 2 apply assumption
apply force
@@ -361,7 +361,7 @@
(*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *)
lemma pred_Inl_bij:
- "a \<in> A ==> (\<lambda>x\<in>pred(A,a,r). Inl(x))
+ "a \<in> A \<Longrightarrow> (\<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)
@@ -369,8 +369,8 @@
done
lemma ordertype_pred_Inl_eq:
- "[| a \<in> A; well_ord(A,r) |]
- ==> ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) =
+ "\<lbrakk>a \<in> A; well_ord(A,r)\<rbrakk>
+ \<Longrightarrow> 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])
@@ -378,7 +378,7 @@
done
lemma pred_Inr_bij:
- "b \<in> B ==>
+ "b \<in> B \<Longrightarrow>
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)
@@ -386,8 +386,8 @@
done
lemma ordertype_pred_Inr_eq:
- "[| b \<in> 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)) =
+ "\<lbrakk>b \<in> B; well_ord(A,r); well_ord(B,s)\<rbrakk>
+ \<Longrightarrow> 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)
@@ -407,7 +407,7 @@
subsubsection\<open>Basic laws for ordinal addition\<close>
-lemma Ord_raw_oadd: "[|Ord(i); Ord(j)|] ==> Ord(raw_oadd(i,j))"
+lemma Ord_raw_oadd: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> Ord(raw_oadd(i,j))"
by (simp add: raw_oadd_def ordify_def Ord_ordertype well_ord_radd
well_ord_Memrel)
@@ -417,19 +417,19 @@
text\<open>Ordinal addition with zero\<close>
-lemma raw_oadd_0: "Ord(i) ==> raw_oadd(i,0) = i"
+lemma raw_oadd_0: "Ord(i) \<Longrightarrow> raw_oadd(i,0) = i"
by (simp add: raw_oadd_def ordify_def ordertype_sum_0_eq
ordertype_Memrel well_ord_Memrel)
-lemma oadd_0 [simp]: "Ord(i) ==> i++0 = i"
+lemma oadd_0 [simp]: "Ord(i) \<Longrightarrow> i++0 = i"
apply (simp (no_asm_simp) add: oadd_def raw_oadd_0 ordify_def)
done
-lemma raw_oadd_0_left: "Ord(i) ==> raw_oadd(0,i) = i"
+lemma raw_oadd_0_left: "Ord(i) \<Longrightarrow> raw_oadd(0,i) = i"
by (simp add: raw_oadd_def ordify_def ordertype_0_sum_eq ordertype_Memrel
well_ord_Memrel)
-lemma oadd_0_left [simp]: "Ord(i) ==> 0++i = i"
+lemma oadd_0_left [simp]: "Ord(i) \<Longrightarrow> 0++i = i"
by (simp add: oadd_def raw_oadd_0_left ordify_def)
@@ -438,14 +438,14 @@
else (if Ord(j) then j else 0))"
by (simp add: oadd_def ordify_def raw_oadd_0_left raw_oadd_0)
-lemma raw_oadd_eq_oadd: "[|Ord(i); Ord(j)|] ==> raw_oadd(i,j) = i++j"
+lemma raw_oadd_eq_oadd: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> raw_oadd(i,j) = i++j"
by (simp add: oadd_def ordify_def)
(*** Further properties of ordinal addition. Statements by Grabczewski,
proofs by lcp. ***)
(*Surely also provable by transfinite induction on j?*)
-lemma lt_oadd1: "k<i ==> k < i++j"
+lemma lt_oadd1: "k<i \<Longrightarrow> k < i++j"
apply (simp add: oadd_def ordify_def lt_Ord2 raw_oadd_0, clarify)
apply (simp add: raw_oadd_def)
apply (rule ltE, assumption)
@@ -455,31 +455,31 @@
apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel)
done
-(*Thus also we obtain the rule @{term"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 \<Longrightarrow> i \<le> k"} *)
+lemma oadd_le_self: "Ord(i) \<Longrightarrow> i \<le> i++j"
apply (rule all_lt_imp_le)
apply (auto simp add: Ord_oadd lt_oadd1)
done
text\<open>Various other results\<close>
-lemma id_ord_iso_Memrel: "A<=B ==> id(A) \<in> ord_iso(A, Memrel(A), A, Memrel(B))"
+lemma id_ord_iso_Memrel: "A<=B \<Longrightarrow> 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
done
lemma subset_ord_iso_Memrel:
- "[| f \<in> ord_iso(A,Memrel(B),C,r); A<=B |] ==> f \<in> ord_iso(A,Memrel(A),C,r)"
+ "\<lbrakk>f \<in> ord_iso(A,Memrel(B),C,r); A<=B\<rbrakk> \<Longrightarrow> f \<in> ord_iso(A,Memrel(A),C,r)"
apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN fun_is_rel])
apply (frule ord_iso_trans [OF id_ord_iso_Memrel], assumption)
apply (simp add: right_comp_id)
done
lemma restrict_ord_iso:
- "[| f \<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)"
+ "\<lbrakk>f \<in> ord_iso(i, Memrel(i), Order.pred(A,a,r), r); a \<in> A; j < i;
+ trans[A](r)\<rbrakk>
+ \<Longrightarrow> 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)
@@ -488,15 +488,15 @@
done
lemma restrict_ord_iso2:
- "[| 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))
+ "\<lbrakk>f \<in> ord_iso(Order.pred(A,a,r), r, i, Memrel(i)); a \<in> A;
+ j < i; trans[A](r)\<rbrakk>
+ \<Longrightarrow> 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))) =
+ "\<lbrakk>well_ord(A,r); k<j\<rbrakk>
+ \<Longrightarrow> 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])
@@ -504,7 +504,7 @@
apply (simp_all add: well_ord_radd well_ord_Memrel)
done
-lemma oadd_lt_mono2: "k<j ==> i++k < i++j"
+lemma oadd_lt_mono2: "k<j \<Longrightarrow> i++k < i++j"
apply (simp add: oadd_def ordify_def raw_oadd_0_left lt_Ord lt_Ord2, clarify)
apply (simp add: raw_oadd_def)
apply (rule ltE, assumption)
@@ -516,7 +516,7 @@
leI [THEN le_ordertype_Memrel] ordertype_sum_Memrel)
done
-lemma oadd_lt_cancel2: "[| i++j < i++k; Ord(j) |] ==> j<k"
+lemma oadd_lt_cancel2: "\<lbrakk>i++j < i++k; Ord(j)\<rbrakk> \<Longrightarrow> j<k"
apply (simp (asm_lr) add: oadd_eq_if_raw_oadd split: split_if_asm)
prefer 2
apply (frule_tac i = i and j = j in oadd_le_self)
@@ -526,17 +526,17 @@
apply (blast dest: oadd_lt_mono2 elim: lt_irrefl lt_asym)+
done
-lemma oadd_lt_iff2: "Ord(j) ==> i++j < i++k \<longleftrightarrow> j<k"
+lemma oadd_lt_iff2: "Ord(j) \<Longrightarrow> i++j < i++k \<longleftrightarrow> j<k"
by (blast intro!: oadd_lt_mono2 dest!: oadd_lt_cancel2)
-lemma oadd_inject: "[| i++j = i++k; Ord(j); Ord(k) |] ==> j=k"
+lemma oadd_inject: "\<lbrakk>i++j = i++k; Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> j=k"
apply (simp add: oadd_eq_if_raw_oadd split: split_if_asm)
apply (simp add: raw_oadd_eq_oadd)
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 | (\<exists>l\<in>j. k = i++l )"
+lemma lt_oadd_disj: "k < i++j \<Longrightarrow> k<i | (\<exists>l\<in>j. k = i++l )"
apply (simp add: Ord_in_Ord' [of _ j] oadd_eq_if_raw_oadd
split: split_if_asm)
prefer 2
@@ -564,7 +564,7 @@
apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel)+
done
-lemma oadd_unfold: "[| Ord(i); Ord(j) |] ==> i++j = i \<union> (\<Union>k\<in>j. {i++k})"
+lemma oadd_unfold: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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)
@@ -572,12 +572,12 @@
apply (force intro: lt_oadd1 oadd_lt_mono2 simp add: Ord_mem_iff_lt)
done
-lemma oadd_1: "Ord(i) ==> i++1 = succ(i)"
+lemma oadd_1: "Ord(i) \<Longrightarrow> i++1 = succ(i)"
apply (simp (no_asm_simp) add: oadd_unfold Ord_1 oadd_0)
apply blast
done
-lemma oadd_succ [simp]: "Ord(j) ==> i++succ(j) = succ(i++j)"
+lemma oadd_succ [simp]: "Ord(j) \<Longrightarrow> i++succ(j) = succ(i++j)"
apply (simp add: oadd_eq_if_raw_oadd, clarify)
apply (simp add: raw_oadd_eq_oadd)
apply (simp add: oadd_1 [of j, symmetric] oadd_1 [of "i++j", symmetric]
@@ -588,28 +588,28 @@
text\<open>Ordinal addition with limit ordinals\<close>
lemma oadd_UN:
- "[| !!x. x \<in> A ==> Ord(j(x)); a \<in> A |]
- ==> i ++ (\<Union>x\<in>A. j(x)) = (\<Union>x\<in>A. i++j(x))"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> Ord(j(x)); a \<in> A\<rbrakk>
+ \<Longrightarrow> 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]
elim!: ltE dest!: ltI [THEN lt_oadd_disj])
-lemma oadd_Limit: "Limit(j) ==> i++j = (\<Union>k\<in>j. i++k)"
+lemma oadd_Limit: "Limit(j) \<Longrightarrow> 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]
Union_eq_UN [symmetric] Limit_Union_eq)
done
-lemma oadd_eq_0_iff: "[| Ord(i); Ord(j) |] ==> (i ++ j) = 0 \<longleftrightarrow> i=0 & j=0"
+lemma oadd_eq_0_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> (i ++ j) = 0 \<longleftrightarrow> i=0 & j=0"
apply (erule trans_induct3 [of j])
apply (simp_all add: oadd_Limit)
apply (simp add: Union_empty_iff Limit_def lt_def, blast)
done
-lemma oadd_eq_lt_iff: "[| Ord(i); Ord(j) |] ==> 0 < (i ++ j) \<longleftrightarrow> 0<i | 0<j"
+lemma oadd_eq_lt_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 0 < (i ++ j) \<longleftrightarrow> 0<i | 0<j"
by (simp add: Ord_0_lt_iff [symmetric] oadd_eq_0_iff)
-lemma oadd_LimitI: "[| Ord(i); Limit(j) |] ==> Limit(i ++ j)"
+lemma oadd_LimitI: "\<lbrakk>Ord(i); Limit(j)\<rbrakk> \<Longrightarrow> Limit(i ++ j)"
apply (simp add: oadd_Limit)
apply (frule Limit_has_1 [THEN ltD])
apply (rule increasing_LimitI)
@@ -624,7 +624,7 @@
text\<open>Order/monotonicity properties of ordinal addition\<close>
-lemma oadd_le_self2: "Ord(i) ==> i \<le> j++i"
+lemma oadd_le_self2: "Ord(i) \<Longrightarrow> i \<le> j++i"
proof (induct i rule: trans_induct3)
case 0 thus ?case by (simp add: Ord_0_le)
next
@@ -639,7 +639,7 @@
thus ?case using limit.hyps by (simp add: oadd_Limit)
qed
-lemma oadd_le_mono1: "k \<le> j ==> k++i \<le> j++i"
+lemma oadd_le_mono1: "k \<le> j \<Longrightarrow> k++i \<le> j++i"
apply (frule lt_Ord)
apply (frule le_Ord2)
apply (simp add: oadd_eq_if_raw_oadd, clarify)
@@ -651,16 +651,16 @@
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: "\<lbrakk>i' \<le> i; j'<j\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>i' \<le> i; j' \<le> j\<rbrakk> \<Longrightarrow> 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 \<longleftrightarrow> j \<le> k"
+lemma oadd_le_iff2: "\<lbrakk>Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> i++j \<le> i++k \<longleftrightarrow> 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"
+lemma oadd_lt_self: "\<lbrakk>Ord(i); 0<j\<rbrakk> \<Longrightarrow> i < i++j"
apply (rule lt_trans2)
apply (erule le_refl)
apply (simp only: lt_Ord2 oadd_1 [of i, symmetric])
@@ -668,12 +668,12 @@
done
text\<open>Every ordinal is exceeded by some limit ordinal.\<close>
-lemma Ord_imp_greater_Limit: "Ord(i) ==> \<exists>k. i<k & Limit(k)"
+lemma Ord_imp_greater_Limit: "Ord(i) \<Longrightarrow> \<exists>k. i<k & Limit(k)"
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)"
+lemma Ord2_imp_greater_Limit: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> \<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)
done
@@ -685,7 +685,7 @@
It's probably simpler to define the difference recursively!\<close>
lemma bij_sum_Diff:
- "A<=B ==> (\<lambda>y\<in>B. if(y \<in> A, Inl(y), Inr(y))) \<in> bij(B, A+(B-A))"
+ "A<=B \<Longrightarrow> (\<lambda>y\<in>B. if(y \<in> 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)
@@ -694,7 +694,7 @@
done
lemma ordertype_sum_Diff:
- "i \<le> j ==>
+ "i \<le> j \<Longrightarrow>
ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) =
ordertype(j, Memrel(j))"
apply (safe dest!: le_subset_iff [THEN iffD1])
@@ -708,7 +708,7 @@
done
lemma Ord_odiff [simp,TC]:
- "[| Ord(i); Ord(j) |] ==> Ord(i--j)"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> Ord(i--j)"
apply (unfold odiff_def)
apply (blast intro: Ord_ordertype Diff_subset well_ord_subset well_ord_Memrel)
done
@@ -716,7 +716,7 @@
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)))"
+ \<Longrightarrow> 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])
apply (rule sum_ord_iso_cong [THEN ordertype_eq])
@@ -725,19 +725,19 @@
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 \<Longrightarrow> 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])
(*By oadd_inject, the difference between i and j is unique. Note that we get
- i++j = k ==> j = k--i. *)
-lemma odiff_oadd_inverse: "[| Ord(i); Ord(j) |] ==> (i++j) -- i = j"
+ i++j = k \<Longrightarrow> j = k--i. *)
+lemma odiff_oadd_inverse: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> (i++j) -- i = j"
apply (rule oadd_inject)
apply (blast intro: oadd_odiff_inverse oadd_le_self)
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: "\<lbrakk>i<j; k \<le> i\<rbrakk> \<Longrightarrow> i--k < j--k"
apply (rule_tac i = k in oadd_lt_cancel2)
apply (simp add: oadd_odiff_inverse)
apply (subst oadd_odiff_inverse)
@@ -749,7 +749,7 @@
subsection\<open>Ordinal Multiplication\<close>
lemma Ord_omult [simp,TC]:
- "[| Ord(i); Ord(j) |] ==> Ord(i**j)"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> Ord(i**j)"
apply (unfold omult_def)
apply (blast intro: Ord_ordertype well_ord_rmult well_ord_Memrel)
done
@@ -757,13 +757,13 @@
subsubsection\<open>A useful unfolding law\<close>
lemma pred_Pair_eq:
- "[| a \<in> A; b \<in> B |] ==> pred(A*B, <a,b>, rmult(A,r,B,s)) =
+ "\<lbrakk>a \<in> A; b \<in> B\<rbrakk> \<Longrightarrow> 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 \<in> A; b \<in> B; well_ord(A,r); well_ord(B,s) |] ==>
+ "\<lbrakk>a \<in> A; b \<in> B; well_ord(A,r); well_ord(B,s)\<rbrakk> \<Longrightarrow>
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))"
@@ -776,8 +776,8 @@
done
lemma ordertype_pred_Pair_lemma:
- "[| i'<i; j'<j |]
- ==> ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))),
+ "\<lbrakk>i'<i; j'<j\<rbrakk>
+ \<Longrightarrow> 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)
@@ -795,8 +795,8 @@
done
lemma lt_omult:
- "[| Ord(i); Ord(j); k<j**i |]
- ==> \<exists>j' i'. k = j**i' ++ j' & j'<j & i'<i"
+ "\<lbrakk>Ord(i); Ord(j); k<j**i\<rbrakk>
+ \<Longrightarrow> \<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)
@@ -806,7 +806,7 @@
done
lemma omult_oadd_lt:
- "[| j'<j; i'<i |] ==> j**i' ++ j' < j**i"
+ "\<lbrakk>j'<j; i'<i\<rbrakk> \<Longrightarrow> j**i' ++ j' < j**i"
apply (unfold omult_def)
apply (rule ltI)
prefer 2
@@ -820,7 +820,7 @@
done
lemma omult_unfold:
- "[| Ord(i); Ord(j) |] ==> j**i = (\<Union>j'\<in>j. \<Union>i'\<in>i. {j**i' ++ j'})"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> j**i = (\<Union>j'\<in>j. \<Union>i'\<in>i. {j**i' ++ j'})"
apply (rule subsetI [THEN equalityI])
apply (rule lt_omult [THEN exE])
apply (erule_tac [3] ltI)
@@ -845,7 +845,7 @@
text\<open>Ordinal multiplication by 1\<close>
-lemma omult_1 [simp]: "Ord(i) ==> i**1 = i"
+lemma omult_1 [simp]: "Ord(i) \<Longrightarrow> i**1 = i"
apply (unfold omult_def)
apply (rule_tac s1="Memrel(i)"
in ord_isoI [THEN ordertype_eq, THEN trans])
@@ -853,7 +853,7 @@
apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel)
done
-lemma omult_1_left [simp]: "Ord(i) ==> 1**i = i"
+lemma omult_1_left [simp]: "Ord(i) \<Longrightarrow> 1**i = i"
apply (unfold omult_def)
apply (rule_tac s1="Memrel(i)"
in ord_isoI [THEN ordertype_eq, THEN trans])
@@ -864,7 +864,7 @@
text\<open>Distributive law for ordinal multiplication and addition\<close>
lemma oadd_omult_distrib:
- "[| Ord(i); Ord(j); Ord(k) |] ==> i**(j++k) = (i**j)++(i**k)"
+ "\<lbrakk>Ord(i); Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> i**(j++k) = (i**j)++(i**k)"
apply (simp add: oadd_eq_if_raw_oadd)
apply (simp add: omult_def raw_oadd_def)
apply (rule ordertype_eq [THEN trans])
@@ -879,13 +879,13 @@
Ord_ordertype)
done
-lemma omult_succ: "[| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i"
+lemma omult_succ: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> i**succ(j) = (i**j)++i"
by (simp del: oadd_succ add: oadd_1 [of j, symmetric] oadd_omult_distrib)
text\<open>Associative law\<close>
lemma omult_assoc:
- "[| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)"
+ "\<lbrakk>Ord(i); Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> (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
@@ -902,24 +902,24 @@
text\<open>Ordinal multiplication with limit ordinals\<close>
lemma omult_UN:
- "[| Ord(i); !!x. x \<in> A ==> Ord(j(x)) |]
- ==> i ** (\<Union>x\<in>A. j(x)) = (\<Union>x\<in>A. i**j(x))"
+ "\<lbrakk>Ord(i); \<And>x. x \<in> A \<Longrightarrow> Ord(j(x))\<rbrakk>
+ \<Longrightarrow> 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)"
+lemma omult_Limit: "\<lbrakk>Ord(i); Limit(j)\<rbrakk> \<Longrightarrow> i**j = (\<Union>k\<in>j. i**k)"
by (simp add: Limit_is_Ord [THEN Ord_in_Ord] omult_UN [symmetric]
Union_eq_UN [symmetric] Limit_Union_eq)
subsubsection\<open>Ordering/monotonicity properties of ordinal multiplication\<close>
-(*As a special case we have "[| 0<i; 0<j |] ==> 0 < i**j" *)
-lemma lt_omult1: "[| k<i; 0<j |] ==> k < i**j"
+(*As a special case we have "\<lbrakk>0<i; 0<j\<rbrakk> \<Longrightarrow> 0 < i**j" *)
+lemma lt_omult1: "\<lbrakk>k<i; 0<j\<rbrakk> \<Longrightarrow> k < i**j"
apply (safe elim!: ltE intro!: ltI Ord_omult)
apply (force simp add: omult_unfold)
done
-lemma omult_le_self: "[| Ord(i); 0<j |] ==> i \<le> i**j"
+lemma omult_le_self: "\<lbrakk>Ord(i); 0<j\<rbrakk> \<Longrightarrow> i \<le> i**j"
by (blast intro: all_lt_imp_le Ord_omult lt_omult1 lt_Ord2)
lemma omult_le_mono1:
@@ -940,24 +940,24 @@
qed
qed
-lemma omult_lt_mono2: "[| k<j; 0<i |] ==> i**k < i**j"
+lemma omult_lt_mono2: "\<lbrakk>k<j; 0<i\<rbrakk> \<Longrightarrow> i**k < i**j"
apply (rule ltI)
apply (simp (no_asm_simp) add: omult_unfold lt_Ord2)
apply (safe elim!: ltE intro!: Ord_omult)
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: "\<lbrakk>k \<le> j; Ord(i)\<rbrakk> \<Longrightarrow> 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)
done
-lemma omult_le_mono: "[| i' \<le> i; j' \<le> j |] ==> i'**j' \<le> i**j"
+lemma omult_le_mono: "\<lbrakk>i' \<le> i; j' \<le> j\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>i' \<le> i; j'<j; 0<i\<rbrakk> \<Longrightarrow> i'**j' < i**j"
by (blast intro: lt_trans1 omult_le_mono1 omult_lt_mono2 Ord_succD elim: ltE)
lemma omult_le_self2:
@@ -988,7 +988,7 @@
text\<open>Further properties of ordinal multiplication\<close>
-lemma omult_inject: "[| i**j = i**k; 0<i; Ord(j); Ord(k) |] ==> j=k"
+lemma omult_inject: "\<lbrakk>i**j = i**k; 0<i; Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> j=k"
apply (rule Ord_linear_lt)
prefer 4 apply assumption
apply auto
--- a/src/ZF/Ordinal.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Ordinal.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,27 +9,27 @@
definition
Memrel :: "i=>i" where
- "Memrel(A) == {z\<in>A*A . \<exists>x y. z=<x,y> & x\<in>y }"
+ "Memrel(A) \<equiv> {z\<in>A*A . \<exists>x y. z=<x,y> & x\<in>y }"
definition
Transset :: "i=>o" where
- "Transset(i) == \<forall>x\<in>i. x<=i"
+ "Transset(i) \<equiv> \<forall>x\<in>i. x<=i"
definition
Ord :: "i=>o" where
- "Ord(i) == Transset(i) & (\<forall>x\<in>i. Transset(x))"
+ "Ord(i) \<equiv> Transset(i) & (\<forall>x\<in>i. Transset(x))"
definition
lt :: "[i,i] => o" (infixl \<open><\<close> 50) (*less-than on ordinals*) where
- "i<j == i\<in>j & Ord(j)"
+ "i<j \<equiv> i\<in>j & Ord(j)"
definition
Limit :: "i=>o" where
- "Limit(i) == Ord(i) & 0<i & (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
+ "Limit(i) \<equiv> Ord(i) & 0<i & (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
abbreviation
le (infixl \<open>\<le>\<close> 50) where
- "x \<le> y == x < succ(y)"
+ "x \<le> y \<equiv> x < succ(y)"
subsection\<open>Rules for Transset\<close>
@@ -50,21 +50,21 @@
subsubsection\<open>Consequences of Downwards Closure\<close>
lemma Transset_doubleton_D:
- "[| Transset(C); {a,b}: C |] ==> a\<in>C & b\<in>C"
+ "\<lbrakk>Transset(C); {a,b}: C\<rbrakk> \<Longrightarrow> a\<in>C & b\<in>C"
by (unfold Transset_def, blast)
lemma Transset_Pair_D:
- "[| Transset(C); <a,b>\<in>C |] ==> a\<in>C & b\<in>C"
+ "\<lbrakk>Transset(C); <a,b>\<in>C\<rbrakk> \<Longrightarrow> a\<in>C & b\<in>C"
apply (simp add: Pair_def)
apply (blast dest: Transset_doubleton_D)
done
lemma Transset_includes_domain:
- "[| Transset(C); A*B \<subseteq> C; b \<in> B |] ==> A \<subseteq> C"
+ "\<lbrakk>Transset(C); A*B \<subseteq> C; b \<in> B\<rbrakk> \<Longrightarrow> A \<subseteq> C"
by (blast dest: Transset_Pair_D)
lemma Transset_includes_range:
- "[| Transset(C); A*B \<subseteq> C; a \<in> A |] ==> B \<subseteq> C"
+ "\<lbrakk>Transset(C); A*B \<subseteq> C; a \<in> A\<rbrakk> \<Longrightarrow> B \<subseteq> C"
by (blast dest: Transset_Pair_D)
subsubsection\<open>Closure Properties\<close>
@@ -73,73 +73,73 @@
by (unfold Transset_def, blast)
lemma Transset_Un:
- "[| Transset(i); Transset(j) |] ==> Transset(i \<union> j)"
+ "\<lbrakk>Transset(i); Transset(j)\<rbrakk> \<Longrightarrow> Transset(i \<union> j)"
by (unfold Transset_def, blast)
lemma Transset_Int:
- "[| Transset(i); Transset(j) |] ==> Transset(i \<inter> j)"
+ "\<lbrakk>Transset(i); Transset(j)\<rbrakk> \<Longrightarrow> Transset(i \<inter> j)"
by (unfold Transset_def, blast)
-lemma Transset_succ: "Transset(i) ==> Transset(succ(i))"
+lemma Transset_succ: "Transset(i) \<Longrightarrow> Transset(succ(i))"
by (unfold Transset_def, blast)
-lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))"
+lemma Transset_Pow: "Transset(i) \<Longrightarrow> Transset(Pow(i))"
by (unfold Transset_def, blast)
-lemma Transset_Union: "Transset(A) ==> Transset(\<Union>(A))"
+lemma Transset_Union: "Transset(A) \<Longrightarrow> Transset(\<Union>(A))"
by (unfold Transset_def, blast)
lemma Transset_Union_family:
- "[| !!i. i\<in>A ==> Transset(i) |] ==> Transset(\<Union>(A))"
+ "\<lbrakk>\<And>i. i\<in>A \<Longrightarrow> Transset(i)\<rbrakk> \<Longrightarrow> Transset(\<Union>(A))"
by (unfold Transset_def, blast)
lemma Transset_Inter_family:
- "[| !!i. i\<in>A ==> Transset(i) |] ==> Transset(\<Inter>(A))"
+ "\<lbrakk>\<And>i. i\<in>A \<Longrightarrow> Transset(i)\<rbrakk> \<Longrightarrow> 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))"
+ "(\<And>x. x \<in> A \<Longrightarrow> Transset(B(x))) \<Longrightarrow> Transset (\<Union>x\<in>A. B(x))"
by (rule Transset_Union_family, auto)
lemma Transset_INT:
- "(!!x. x \<in> A ==> Transset(B(x))) ==> Transset (\<Inter>x\<in>A. B(x))"
+ "(\<And>x. x \<in> A \<Longrightarrow> Transset(B(x))) \<Longrightarrow> Transset (\<Inter>x\<in>A. B(x))"
by (rule Transset_Inter_family, auto)
subsection\<open>Lemmas for Ordinals\<close>
lemma OrdI:
- "[| Transset(i); !!x. x\<in>i ==> Transset(x) |] ==> Ord(i)"
+ "\<lbrakk>Transset(i); \<And>x. x\<in>i \<Longrightarrow> Transset(x)\<rbrakk> \<Longrightarrow> Ord(i)"
by (simp add: Ord_def)
-lemma Ord_is_Transset: "Ord(i) ==> Transset(i)"
+lemma Ord_is_Transset: "Ord(i) \<Longrightarrow> Transset(i)"
by (simp add: Ord_def)
lemma Ord_contains_Transset:
- "[| Ord(i); j\<in>i |] ==> Transset(j) "
+ "\<lbrakk>Ord(i); j\<in>i\<rbrakk> \<Longrightarrow> Transset(j) "
by (unfold Ord_def, blast)
-lemma Ord_in_Ord: "[| Ord(i); j\<in>i |] ==> Ord(j)"
+lemma Ord_in_Ord: "\<lbrakk>Ord(i); j\<in>i\<rbrakk> \<Longrightarrow> Ord(j)"
by (unfold Ord_def Transset_def, blast)
(*suitable for rewriting PROVIDED i has been fixed*)
-lemma Ord_in_Ord': "[| j\<in>i; Ord(i) |] ==> Ord(j)"
+lemma Ord_in_Ord': "\<lbrakk>j\<in>i; Ord(i)\<rbrakk> \<Longrightarrow> Ord(j)"
by (blast intro: Ord_in_Ord)
-(* Ord(succ(j)) ==> Ord(j) *)
+(* Ord(succ(j)) \<Longrightarrow> Ord(j) *)
lemmas Ord_succD = Ord_in_Ord [OF _ succI1]
-lemma Ord_subset_Ord: "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)"
+lemma Ord_subset_Ord: "\<lbrakk>Ord(i); Transset(j); j<=i\<rbrakk> \<Longrightarrow> Ord(j)"
by (simp add: Ord_def Transset_def, blast)
-lemma OrdmemD: "[| j\<in>i; Ord(i) |] ==> j<=i"
+lemma OrdmemD: "\<lbrakk>j\<in>i; Ord(i)\<rbrakk> \<Longrightarrow> j<=i"
by (unfold Ord_def Transset_def, blast)
-lemma Ord_trans: "[| i\<in>j; j\<in>k; Ord(k) |] ==> i\<in>k"
+lemma Ord_trans: "\<lbrakk>i\<in>j; j\<in>k; Ord(k)\<rbrakk> \<Longrightarrow> i\<in>k"
by (blast dest: OrdmemD)
-lemma Ord_succ_subsetI: "[| i\<in>j; Ord(j) |] ==> succ(i) \<subseteq> j"
+lemma Ord_succ_subsetI: "\<lbrakk>i\<in>j; Ord(j)\<rbrakk> \<Longrightarrow> succ(i) \<subseteq> j"
by (blast dest: OrdmemD)
@@ -148,7 +148,7 @@
lemma Ord_0 [iff,TC]: "Ord(0)"
by (blast intro: OrdI Transset_0)
-lemma Ord_succ [TC]: "Ord(i) ==> Ord(succ(i))"
+lemma Ord_succ [TC]: "Ord(i) \<Longrightarrow> Ord(succ(i))"
by (blast intro: OrdI Transset_succ Ord_is_Transset Ord_contains_Transset)
lemmas Ord_1 = Ord_0 [THEN Ord_succ]
@@ -156,18 +156,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 \<union> j)"
+lemma Ord_Un [intro,simp,TC]: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> Ord(i \<union> j)"
apply (unfold Ord_def)
apply (blast intro!: Transset_Un)
done
-lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \<inter> j)"
+lemma Ord_Int [TC]: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> Ord(i \<inter> j)"
apply (unfold Ord_def)
apply (blast intro!: Transset_Int)
done
text\<open>There is no set of all ordinals, for then it would contain itself\<close>
-lemma ON_class: "~ (\<forall>i. i\<in>X <-> Ord(i))"
+lemma ON_class: "\<not> (\<forall>i. i\<in>X <-> Ord(i))"
proof (rule notI)
assume X: "\<forall>i. i \<in> X \<longleftrightarrow> Ord(i)"
have "\<forall>x y. x\<in>X \<longrightarrow> y\<in>x \<longrightarrow> y\<in>X"
@@ -183,63 +183,63 @@
subsection\<open>< is 'less Than' for Ordinals\<close>
-lemma ltI: "[| i\<in>j; Ord(j) |] ==> i<j"
+lemma ltI: "\<lbrakk>i\<in>j; Ord(j)\<rbrakk> \<Longrightarrow> i<j"
by (unfold lt_def, blast)
lemma ltE:
- "[| i<j; [| i\<in>j; Ord(i); Ord(j) |] ==> P |] ==> P"
+ "\<lbrakk>i<j; \<lbrakk>i\<in>j; Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
apply (unfold lt_def)
apply (blast intro: Ord_in_Ord)
done
-lemma ltD: "i<j ==> i\<in>j"
+lemma ltD: "i<j \<Longrightarrow> i\<in>j"
by (erule ltE, assumption)
-lemma not_lt0 [simp]: "~ i<0"
+lemma not_lt0 [simp]: "\<not> i<0"
by (unfold lt_def, blast)
-lemma lt_Ord: "j<i ==> Ord(j)"
+lemma lt_Ord: "j<i \<Longrightarrow> Ord(j)"
by (erule ltE, assumption)
-lemma lt_Ord2: "j<i ==> Ord(i)"
+lemma lt_Ord2: "j<i \<Longrightarrow> Ord(i)"
by (erule ltE, assumption)
-(* @{term"ja \<le> j ==> Ord(j)"} *)
+(* @{term"ja \<le> j \<Longrightarrow> Ord(j)"} *)
lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD]
-(* i<0 ==> R *)
+(* i<0 \<Longrightarrow> R *)
lemmas lt0E = not_lt0 [THEN notE, elim!]
-lemma lt_trans [trans]: "[| i<j; j<k |] ==> i<k"
+lemma lt_trans [trans]: "\<lbrakk>i<j; j<k\<rbrakk> \<Longrightarrow> i<k"
by (blast intro!: ltI elim!: ltE intro: Ord_trans)
-lemma lt_not_sym: "i<j ==> ~ (j<i)"
+lemma lt_not_sym: "i<j \<Longrightarrow> \<not> (j<i)"
apply (unfold lt_def)
apply (blast elim: mem_asym)
done
-(* [| i<j; ~P ==> j<i |] ==> P *)
+(* \<lbrakk>i<j; \<not>P \<Longrightarrow> j<i\<rbrakk> \<Longrightarrow> P *)
lemmas lt_asym = lt_not_sym [THEN swap]
-lemma lt_irrefl [elim!]: "i<i ==> P"
+lemma lt_irrefl [elim!]: "i<i \<Longrightarrow> P"
by (blast intro: lt_asym)
-lemma lt_not_refl: "~ i<i"
+lemma lt_not_refl: "\<not> i<i"
apply (rule notI)
apply (erule lt_irrefl)
done
-text\<open>Recall that \<^term>\<open>i \<le> j\<close> abbreviates \<^term>\<open>i<succ(j)\<close> !!\<close>
+text\<open>Recall that \<^term>\<open>i \<le> j\<close> abbreviates \<^term>\<open>i<succ(j)\<close> \<And>\<close>
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"
+(*Equivalently, i<j \<Longrightarrow> i < succ(j)*)
+lemma leI: "i<j \<Longrightarrow> i \<le> j"
by (simp add: le_iff)
-lemma le_eqI: "[| i=j; Ord(j) |] ==> i \<le> j"
+lemma le_eqI: "\<lbrakk>i=j; Ord(j)\<rbrakk> \<Longrightarrow> i \<le> j"
by (simp add: le_iff)
lemmas le_refl = refl [THEN le_eqI]
@@ -247,14 +247,14 @@
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: "(\<not> (i=j & Ord(j)) \<Longrightarrow> i<j) \<Longrightarrow> i \<le> j"
by (simp add: le_iff, blast)
lemma leE:
- "[| i \<le> j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P"
+ "\<lbrakk>i \<le> j; i<j \<Longrightarrow> P; \<lbrakk>i=j; Ord(j)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp add: le_iff, blast)
-lemma le_anti_sym: "[| i \<le> j; j \<le> i |] ==> i=j"
+lemma le_anti_sym: "\<lbrakk>i \<le> j; j \<le> i\<rbrakk> \<Longrightarrow> i=j"
apply (simp add: le_iff)
apply (blast elim: lt_asym)
done
@@ -270,19 +270,19 @@
lemma Memrel_iff [simp]: "<a,b> \<in> Memrel(A) <-> a\<in>b & a\<in>A & b\<in>A"
by (unfold Memrel_def, blast)
-lemma MemrelI [intro!]: "[| a \<in> b; a \<in> A; b \<in> A |] ==> <a,b> \<in> Memrel(A)"
+lemma MemrelI [intro!]: "\<lbrakk>a \<in> b; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> <a,b> \<in> Memrel(A)"
by auto
lemma MemrelE [elim!]:
- "[| <a,b> \<in> Memrel(A);
- [| a \<in> A; b \<in> A; a\<in>b |] ==> P |]
- ==> P"
+ "\<lbrakk><a,b> \<in> Memrel(A);
+ \<lbrakk>a \<in> A; b \<in> A; a\<in>b\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by auto
lemma Memrel_type: "Memrel(A) \<subseteq> A*A"
by (unfold Memrel_def, blast)
-lemma Memrel_mono: "A<=B ==> Memrel(A) \<subseteq> Memrel(B)"
+lemma Memrel_mono: "A<=B \<Longrightarrow> Memrel(A) \<subseteq> Memrel(B)"
by (unfold Memrel_def, blast)
lemma Memrel_0 [simp]: "Memrel(0) = 0"
@@ -303,17 +303,17 @@
text\<open>The premise \<^term>\<open>Ord(i)\<close> does not suffice.\<close>
lemma trans_Memrel:
- "Ord(i) ==> trans(Memrel(i))"
+ "Ord(i) \<Longrightarrow> trans(Memrel(i))"
by (unfold Ord_def Transset_def trans_def, blast)
text\<open>However, the following premise is strong enough.\<close>
lemma Transset_trans_Memrel:
- "\<forall>j\<in>i. Transset(j) ==> trans(Memrel(i))"
+ "\<forall>j\<in>i. Transset(j) \<Longrightarrow> 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> \<in> Memrel(A) <-> a\<in>b & b\<in>A"
+ "Transset(A) \<Longrightarrow> <a,b> \<in> Memrel(A) <-> a\<in>b & b\<in>A"
by (unfold Transset_def, blast)
@@ -321,9 +321,9 @@
(*Epsilon induction over a transitive set*)
lemma Transset_induct:
- "[| i \<in> k; Transset(k);
- !!x.[| x \<in> k; \<forall>y\<in>x. P(y) |] ==> P(x) |]
- ==> P(i)"
+ "\<lbrakk>i \<in> k; Transset(k);
+ \<And>x.\<lbrakk>x \<in> k; \<forall>y\<in>x. P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> P(i)"
apply (simp add: Transset_def)
apply (erule wf_Memrel [THEN wf_induct2], blast+)
done
@@ -382,34 +382,34 @@
apply (blast intro: leI le_eqI o) +
done
-lemma le_imp_not_lt: "j \<le> i ==> ~ i<j"
+lemma le_imp_not_lt: "j \<le> i \<Longrightarrow> \<not> 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: "\<lbrakk>\<not> i<j; Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> j \<le> i"
by (rule_tac i = i and j = j in Ord_linear2, auto)
subsubsection \<open>Some Rewrite Rules for \<open><\<close>, \<open>\<le>\<close>\<close>
-lemma Ord_mem_iff_lt: "Ord(j) ==> i\<in>j <-> i<j"
+lemma Ord_mem_iff_lt: "Ord(j) \<Longrightarrow> i\<in>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: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> \<not> 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: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> \<not> 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) \<Longrightarrow> 0 \<le> i"
by (erule not_lt_iff_le [THEN iffD1], auto)
-lemma Ord_0_lt: "[| Ord(i); i\<noteq>0 |] ==> 0<i"
+lemma Ord_0_lt: "\<lbrakk>Ord(i); i\<noteq>0\<rbrakk> \<Longrightarrow> 0<i"
apply (erule not_le_iff_lt [THEN iffD1])
apply (rule Ord_0, blast)
done
-lemma Ord_0_lt_iff: "Ord(i) ==> i\<noteq>0 <-> 0<i"
+lemma Ord_0_lt_iff: "Ord(i) \<Longrightarrow> i\<noteq>0 <-> 0<i"
by (blast intro: Ord_0_lt)
@@ -420,12 +420,12 @@
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: "\<lbrakk>j<=i; Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> i<=j"
by (blast dest: OrdmemD elim: ltE leE)
lemma le_subset_iff: "j \<le> i <-> j<=i & Ord(i) & Ord(j)"
@@ -437,27 +437,27 @@
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: "\<lbrakk>Ord(i); Ord(j); \<And>x. x<j \<Longrightarrow> x<i\<rbrakk> \<Longrightarrow> j \<le> i"
by (blast intro: not_lt_imp_le dest: lt_irrefl)
subsubsection\<open>Transitivity Laws\<close>
-lemma lt_trans1: "[| i \<le> j; j<k |] ==> i<k"
+lemma lt_trans1: "\<lbrakk>i \<le> j; j<k\<rbrakk> \<Longrightarrow> i<k"
by (blast elim!: leE intro: lt_trans)
-lemma lt_trans2: "[| i<j; j \<le> k |] ==> i<k"
+lemma lt_trans2: "\<lbrakk>i<j; j \<le> k\<rbrakk> \<Longrightarrow> i<k"
by (blast elim!: leE intro: lt_trans)
-lemma le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> k"
+lemma le_trans: "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> k"
by (blast intro: lt_trans1)
-lemma succ_leI: "i<j ==> succ(i) \<le> j"
+lemma succ_leI: "i<j \<Longrightarrow> 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"
+(*Identical to succ(i) < succ(j) \<Longrightarrow> i<j *)
+lemma succ_leE: "succ(i) \<le> j \<Longrightarrow> i<j"
apply (rule not_le_iff_lt [THEN iffD1])
apply (blast elim: ltE leE lt_asym)+
done
@@ -465,15 +465,15 @@
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) \<Longrightarrow> i \<le> j"
by (blast dest!: succ_leE)
-lemma lt_subset_trans: "[| i \<subseteq> j; j<k; Ord(i) |] ==> i<k"
+lemma lt_subset_trans: "\<lbrakk>i \<subseteq> j; j<k; Ord(i)\<rbrakk> \<Longrightarrow> 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"
+lemma lt_imp_0_lt: "j<i \<Longrightarrow> 0<i"
by (blast intro: lt_trans1 Ord_0_le [OF lt_Ord])
lemma succ_lt_iff: "succ(i) < j <-> i<j & succ(i) \<noteq> j"
@@ -487,86 +487,86 @@
apply (blast dest: lt_asym)
done
-lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \<in> succ(j) <-> i\<in>j"
+lemma Ord_succ_mem_iff: "Ord(j) \<Longrightarrow> succ(i) \<in> succ(j) <-> i\<in>j"
apply (insert succ_le_iff [of i j])
apply (simp add: lt_def)
done
subsubsection\<open>Union and Intersection\<close>
-lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i \<le> i \<union> j"
+lemma Un_upper1_le: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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 \<union> j"
+lemma Un_upper2_le: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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 \<union> j < k"
+lemma Un_least_lt: "\<lbrakk>i<k; j<k\<rbrakk> \<Longrightarrow> 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)
done
-lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i \<union> j < k <-> i<k & j<k"
+lemma Un_least_lt_iff: "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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)
done
lemma Un_least_mem_iff:
- "[| Ord(i); Ord(j); Ord(k) |] ==> i \<union> j \<in> k <-> i\<in>k & j\<in>k"
+ "\<lbrakk>Ord(i); Ord(j); Ord(k)\<rbrakk> \<Longrightarrow> i \<union> j \<in> k <-> i\<in>k & j\<in>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 \<inter> j < k"
+lemma Int_greatest_lt: "\<lbrakk>i<k; j<k\<rbrakk> \<Longrightarrow> 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)
done
lemma Ord_Un_if:
- "[| Ord(i); Ord(j) |] ==> i \<union> j = (if j<i then i else j)"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> 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])
lemma succ_Un_distrib:
- "[| Ord(i); Ord(j) |] ==> succ(i \<union> j) = succ(i) \<union> succ(j)"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> succ(i \<union> j) = succ(i) \<union> succ(j)"
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"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> k < i \<union> j <-> k < i | k < j"
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"
+ "\<lbrakk>Ord(i); Ord(j)\<rbrakk> \<Longrightarrow> k \<le> i \<union> j <-> k \<le> i | k \<le> j"
by (simp add: succ_Un_distrib lt_Un_iff [symmetric])
-lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i \<union> j"
+lemma Un_upper1_lt: "\<lbrakk>k < i; Ord(j)\<rbrakk> \<Longrightarrow> k < i \<union> j"
by (simp add: lt_Un_iff lt_Ord2)
-lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i \<union> j"
+lemma Un_upper2_lt: "\<lbrakk>k < j; Ord(i)\<rbrakk> \<Longrightarrow> 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"
+lemma Ord_Union_succ_eq: "Ord(i) \<Longrightarrow> \<Union>(succ(i)) = i"
by (blast intro: Ord_trans)
subsection\<open>Results about Limits\<close>
-lemma Ord_Union [intro,simp,TC]: "[| !!i. i\<in>A ==> Ord(i) |] ==> Ord(\<Union>(A))"
+lemma Ord_Union [intro,simp,TC]: "\<lbrakk>\<And>i. i\<in>A \<Longrightarrow> Ord(i)\<rbrakk> \<Longrightarrow> Ord(\<Union>(A))"
apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI])
apply (blast intro: Ord_contains_Transset)+
done
lemma Ord_UN [intro,simp,TC]:
- "[| !!x. x\<in>A ==> Ord(B(x)) |] ==> Ord(\<Union>x\<in>A. B(x))"
+ "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Union>x\<in>A. B(x))"
by (rule Ord_Union, blast)
lemma Ord_Inter [intro,simp,TC]:
- "[| !!i. i\<in>A ==> Ord(i) |] ==> Ord(\<Inter>(A))"
+ "\<lbrakk>\<And>i. i\<in>A \<Longrightarrow> Ord(i)\<rbrakk> \<Longrightarrow> Ord(\<Inter>(A))"
apply (rule Transset_Inter_family [THEN OrdI])
apply (blast intro: Ord_is_Transset)
apply (simp add: Inter_def)
@@ -574,95 +574,95 @@
done
lemma Ord_INT [intro,simp,TC]:
- "[| !!x. x\<in>A ==> Ord(B(x)) |] ==> Ord(\<Inter>x\<in>A. B(x))"
+ "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> Ord(B(x))\<rbrakk> \<Longrightarrow> Ord(\<Inter>x\<in>A. B(x))"
by (rule Ord_Inter, blast)
(* No < version of this theorem: consider that @{term"(\<Union>i\<in>nat.i)=nat"}! *)
lemma UN_least_le:
- "[| Ord(i); !!x. x\<in>A ==> b(x) \<le> i |] ==> (\<Union>x\<in>A. b(x)) \<le> i"
+ "\<lbrakk>Ord(i); \<And>x. x\<in>A \<Longrightarrow> b(x) \<le> i\<rbrakk> \<Longrightarrow> (\<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
lemma UN_succ_least_lt:
- "[| j<i; !!x. x\<in>A ==> b(x)<j |] ==> (\<Union>x\<in>A. succ(b(x))) < i"
+ "\<lbrakk>j<i; \<And>x. x\<in>A \<Longrightarrow> b(x)<j\<rbrakk> \<Longrightarrow> (\<Union>x\<in>A. succ(b(x))) < i"
apply (rule ltE, assumption)
apply (rule UN_least_le [THEN lt_trans2])
apply (blast intro: succ_leI)+
done
lemma UN_upper_lt:
- "[| a\<in>A; i < b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i < (\<Union>x\<in>A. b(x))"
+ "\<lbrakk>a\<in>A; i < b(a); Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> i < (\<Union>x\<in>A. b(x))"
by (unfold lt_def, blast)
lemma UN_upper_le:
- "[| a \<in> A; i \<le> b(a); Ord(\<Union>x\<in>A. b(x)) |] ==> i \<le> (\<Union>x\<in>A. b(x))"
+ "\<lbrakk>a \<in> A; i \<le> b(a); Ord(\<Union>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> 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)+
done
-lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) ==> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
+lemma lt_Union_iff: "\<forall>i\<in>A. Ord(i) \<Longrightarrow> (j < \<Union>(A)) <-> (\<exists>i\<in>A. j<i)"
by (auto simp: lt_def Ord_Union)
lemma Union_upper_le:
- "[| j \<in> J; i\<le>j; Ord(\<Union>(J)) |] ==> i \<le> \<Union>J"
+ "\<lbrakk>j \<in> J; i\<le>j; Ord(\<Union>(J))\<rbrakk> \<Longrightarrow> i \<le> \<Union>J"
apply (subst Union_eq_UN)
apply (rule UN_upper_le, auto)
done
lemma le_implies_UN_le_UN:
- "[| !!x. x\<in>A ==> c(x) \<le> d(x) |] ==> (\<Union>x\<in>A. c(x)) \<le> (\<Union>x\<in>A. d(x))"
+ "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> c(x) \<le> d(x)\<rbrakk> \<Longrightarrow> (\<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)+
done
-lemma Ord_equality: "Ord(i) ==> (\<Union>y\<in>i. succ(y)) = i"
+lemma Ord_equality: "Ord(i) \<Longrightarrow> (\<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) \<subseteq> i"
+lemma Ord_Union_subset: "Ord(i) \<Longrightarrow> \<Union>(i) \<subseteq> i"
by (blast intro: Ord_trans)
subsection\<open>Limit Ordinals -- General Properties\<close>
-lemma Limit_Union_eq: "Limit(i) ==> \<Union>(i) = i"
+lemma Limit_Union_eq: "Limit(i) \<Longrightarrow> \<Union>(i) = i"
apply (unfold Limit_def)
apply (fast intro!: ltI elim!: ltE elim: Ord_trans)
done
-lemma Limit_is_Ord: "Limit(i) ==> Ord(i)"
+lemma Limit_is_Ord: "Limit(i) \<Longrightarrow> Ord(i)"
apply (unfold Limit_def)
apply (erule conjunct1)
done
-lemma Limit_has_0: "Limit(i) ==> 0 < i"
+lemma Limit_has_0: "Limit(i) \<Longrightarrow> 0 < i"
apply (unfold Limit_def)
apply (erule conjunct2 [THEN conjunct1])
done
-lemma Limit_nonzero: "Limit(i) ==> i \<noteq> 0"
+lemma Limit_nonzero: "Limit(i) \<Longrightarrow> i \<noteq> 0"
by (drule Limit_has_0, blast)
-lemma Limit_has_succ: "[| Limit(i); j<i |] ==> succ(j) < i"
+lemma Limit_has_succ: "\<lbrakk>Limit(i); j<i\<rbrakk> \<Longrightarrow> succ(j) < i"
by (unfold Limit_def, blast)
-lemma Limit_succ_lt_iff [simp]: "Limit(i) ==> succ(j) < i <-> (j<i)"
+lemma Limit_succ_lt_iff [simp]: "Limit(i) \<Longrightarrow> succ(j) < i <-> (j<i)"
apply (safe intro!: Limit_has_succ)
apply (frule lt_Ord)
apply (blast intro: lt_trans)
done
-lemma zero_not_Limit [iff]: "~ Limit(0)"
+lemma zero_not_Limit [iff]: "\<not> Limit(0)"
by (simp add: Limit_def)
-lemma Limit_has_1: "Limit(i) ==> 1 < i"
+lemma Limit_has_1: "Limit(i) \<Longrightarrow> 1 < i"
by (blast intro: Limit_has_0 Limit_has_succ)
-lemma increasing_LimitI: "[| 0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y |] ==> Limit(l)"
+lemma increasing_LimitI: "\<lbrakk>0<l; \<forall>x\<in>l. \<exists>y\<in>l. x<y\<rbrakk> \<Longrightarrow> Limit(l)"
apply (unfold Limit_def, simp add: lt_Ord2, clarify)
apply (drule_tac i=y in ltD)
apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2)
@@ -676,28 +676,28 @@
{ fix y
assume yi: "y<i"
hence Osy: "Ord(succ(y))" by (simp add: lt_Ord Ord_succ)
- have "~ i \<le> y" using yi by (blast dest: le_imp_not_lt)
+ have "\<not> i \<le> y" using yi by (blast dest: le_imp_not_lt)
hence "succ(y) < i" using nsucc [of y]
by (blast intro: Ord_linear_lt [OF Osy Oi]) }
thus ?thesis using i Oi by (auto simp add: Limit_def)
qed
-lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P"
+lemma succ_LimitE [elim!]: "Limit(succ(i)) \<Longrightarrow> P"
apply (rule lt_irrefl)
apply (rule Limit_has_succ, assumption)
apply (erule Limit_is_Ord [THEN Ord_succD, THEN le_refl])
done
-lemma not_succ_Limit [simp]: "~ Limit(succ(i))"
+lemma not_succ_Limit [simp]: "\<not> Limit(succ(i))"
by blast
-lemma Limit_le_succD: "[| Limit(i); i \<le> succ(j) |] ==> i \<le> j"
+lemma Limit_le_succD: "\<lbrakk>Limit(i); i \<le> succ(j)\<rbrakk> \<Longrightarrow> i \<le> j"
by (blast elim!: leE)
subsubsection\<open>Traditional 3-Way Case Analysis on Ordinals\<close>
-lemma Ord_cases_disj: "Ord(i) ==> i=0 | (\<exists>j. Ord(j) & i=succ(j)) | Limit(i)"
+lemma Ord_cases_disj: "Ord(i) \<Longrightarrow> i=0 | (\<exists>j. Ord(j) & i=succ(j)) | Limit(i)"
by (blast intro!: non_succ_LimitI Ord_0_lt)
lemma Ord_cases:
@@ -706,11 +706,11 @@
by (insert Ord_cases_disj [OF i], auto)
lemma trans_induct3_raw:
- "[| Ord(i);
+ "\<lbrakk>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)"
+ \<And>x. \<lbrakk>Ord(x); P(x)\<rbrakk> \<Longrightarrow> P(succ(x));
+ \<And>x. \<lbrakk>Limit(x); \<forall>y\<in>x. P(y)\<rbrakk> \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> P(i)"
apply (erule trans_induct)
apply (erule Ord_cases, blast+)
done
@@ -722,7 +722,7 @@
text\<open>A set of ordinals is either empty, contains its own union, or its
union is a limit ordinal.\<close>
-lemma Union_le: "[| !!x. x\<in>I ==> x\<le>j; Ord(j) |] ==> \<Union>(I) \<le> j"
+lemma Union_le: "\<lbrakk>\<And>x. x\<in>I \<Longrightarrow> x\<le>j; Ord(j)\<rbrakk> \<Longrightarrow> \<Union>(I) \<le> j"
by (auto simp add: le_subset_iff Union_least)
lemma Ord_set_cases:
@@ -754,10 +754,10 @@
qed
text\<open>If the union of a set of ordinals is a successor, then it is an element of that set.\<close>
-lemma Ord_Union_eq_succD: "[|\<forall>x\<in>X. Ord(x); \<Union>X = succ(j)|] ==> succ(j) \<in> X"
+lemma Ord_Union_eq_succD: "\<lbrakk>\<forall>x\<in>X. Ord(x); \<Union>X = succ(j)\<rbrakk> \<Longrightarrow> succ(j) \<in> X"
by (drule Ord_set_cases, auto)
-lemma Limit_Union [rule_format]: "[| I \<noteq> 0; (\<And>i. i\<in>I \<Longrightarrow> Limit(i)) |] ==> Limit(\<Union>I)"
+lemma Limit_Union [rule_format]: "\<lbrakk>I \<noteq> 0; (\<And>i. i\<in>I \<Longrightarrow> Limit(i))\<rbrakk> \<Longrightarrow> Limit(\<Union>I)"
apply (simp add: Limit_def lt_def)
apply (blast intro!: equalityI)
done
--- a/src/ZF/Perm.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Perm.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,43 +15,43 @@
definition
(*composition of relations and functions; NOT Suppes's relative product*)
comp :: "[i,i]=>i" (infixr \<open>O\<close> 60) where
- "r O s == {xz \<in> domain(s)*range(r) .
+ "r O s \<equiv> {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) == (\<lambda>x\<in>A. x)"
+ "id(A) \<equiv> (\<lambda>x\<in>A. x)"
definition
(*one-to-one functions from A to B*)
inj :: "[i,i]=>i" where
- "inj(A,B) == { f \<in> A->B. \<forall>w\<in>A. \<forall>x\<in>A. f`w=f`x \<longrightarrow> w=x}"
+ "inj(A,B) \<equiv> { f \<in> 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 \<in> A->B . \<forall>y\<in>B. \<exists>x\<in>A. f`x=y}"
+ "surj(A,B) \<equiv> { f \<in> 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) \<inter> surj(A,B)"
+ "bij(A,B) \<equiv> inj(A,B) \<inter> surj(A,B)"
subsection\<open>Surjective Function Space\<close>
-lemma surj_is_fun: "f \<in> surj(A,B) ==> f \<in> A->B"
+lemma surj_is_fun: "f \<in> surj(A,B) \<Longrightarrow> f \<in> A->B"
apply (unfold surj_def)
apply (erule CollectD1)
done
-lemma fun_is_surj: "f \<in> Pi(A,B) ==> f \<in> surj(A,range(f))"
+lemma fun_is_surj: "f \<in> Pi(A,B) \<Longrightarrow> f \<in> surj(A,range(f))"
apply (unfold surj_def)
apply (blast intro: apply_equality range_of_fun domain_type)
done
-lemma surj_range: "f \<in> surj(A,B) ==> range(f)=B"
+lemma surj_range: "f \<in> surj(A,B) \<Longrightarrow> range(f)=B"
apply (unfold surj_def)
apply (best intro: apply_Pair elim: range_type)
done
@@ -59,15 +59,15 @@
text\<open>A function with a right inverse is a surjection\<close>
lemma f_imp_surjective:
- "[| f \<in> A->B; !!y. y \<in> B ==> d(y): A; !!y. y \<in> B ==> f`d(y) = y |]
- ==> f \<in> surj(A,B)"
+ "\<lbrakk>f \<in> A->B; \<And>y. y \<in> B \<Longrightarrow> d(y): A; \<And>y. y \<in> B \<Longrightarrow> f`d(y) = y\<rbrakk>
+ \<Longrightarrow> f \<in> surj(A,B)"
by (simp add: surj_def, blast)
lemma lam_surjective:
- "[| !!x. x \<in> A ==> c(x): B;
- !!y. y \<in> B ==> d(y): A;
- !!y. y \<in> B ==> c(d(y)) = y
- |] ==> (\<lambda>x\<in>A. c(x)) \<in> surj(A,B)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> c(x): B;
+ \<And>y. y \<in> B \<Longrightarrow> d(y): A;
+ \<And>y. y \<in> B \<Longrightarrow> c(d(y)) = y
+\<rbrakk> \<Longrightarrow> (\<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
@@ -82,63 +82,63 @@
subsection\<open>Injective Function Space\<close>
-lemma inj_is_fun: "f \<in> inj(A,B) ==> f \<in> A->B"
+lemma inj_is_fun: "f \<in> inj(A,B) \<Longrightarrow> f \<in> A->B"
apply (unfold inj_def)
apply (erule CollectD1)
done
text\<open>Good for dealing with sets of pairs, but a bit ugly in use [used in AC]\<close>
lemma inj_equality:
- "[| <a,b>:f; <c,b>:f; f \<in> inj(A,B) |] ==> a=c"
+ "\<lbrakk><a,b>:f; <c,b>:f; f \<in> inj(A,B)\<rbrakk> \<Longrightarrow> a=c"
apply (unfold inj_def)
apply (blast dest: Pair_mem_PiD)
done
-lemma inj_apply_equality: "[| f \<in> inj(A,B); f`a=f`b; a \<in> A; b \<in> A |] ==> a=b"
+lemma inj_apply_equality: "\<lbrakk>f \<in> inj(A,B); f`a=f`b; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a=b"
by (unfold inj_def, blast)
text\<open>A function with a left inverse is an injection\<close>
-lemma f_imp_injective: "[| f \<in> A->B; \<forall>x\<in>A. d(f`x)=x |] ==> f \<in> inj(A,B)"
+lemma f_imp_injective: "\<lbrakk>f \<in> A->B; \<forall>x\<in>A. d(f`x)=x\<rbrakk> \<Longrightarrow> f \<in> inj(A,B)"
apply (simp (no_asm_simp) add: inj_def)
apply (blast intro: subst_context [THEN box_equals])
done
lemma lam_injective:
- "[| !!x. x \<in> A ==> c(x): B;
- !!x. x \<in> A ==> d(c(x)) = x |]
- ==> (\<lambda>x\<in>A. c(x)) \<in> inj(A,B)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> c(x): B;
+ \<And>x. x \<in> A \<Longrightarrow> d(c(x)) = x\<rbrakk>
+ \<Longrightarrow> (\<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
subsection\<open>Bijections\<close>
-lemma bij_is_inj: "f \<in> bij(A,B) ==> f \<in> inj(A,B)"
+lemma bij_is_inj: "f \<in> bij(A,B) \<Longrightarrow> f \<in> inj(A,B)"
apply (unfold bij_def)
apply (erule IntD1)
done
-lemma bij_is_surj: "f \<in> bij(A,B) ==> f \<in> surj(A,B)"
+lemma bij_is_surj: "f \<in> bij(A,B) \<Longrightarrow> f \<in> surj(A,B)"
apply (unfold bij_def)
apply (erule IntD2)
done
-lemma bij_is_fun: "f \<in> bij(A,B) ==> f \<in> A->B"
+lemma bij_is_fun: "f \<in> bij(A,B) \<Longrightarrow> f \<in> A->B"
by (rule bij_is_inj [THEN inj_is_fun])
lemma lam_bijective:
- "[| !!x. x \<in> A ==> c(x): B;
- !!y. y \<in> B ==> d(y): A;
- !!x. x \<in> A ==> d(c(x)) = x;
- !!y. y \<in> B ==> c(d(y)) = y
- |] ==> (\<lambda>x\<in>A. c(x)) \<in> bij(A,B)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> c(x): B;
+ \<And>y. y \<in> B \<Longrightarrow> d(y): A;
+ \<And>x. x \<in> A \<Longrightarrow> d(c(x)) = x;
+ \<And>y. y \<in> B \<Longrightarrow> c(d(y)) = y
+\<rbrakk> \<Longrightarrow> (\<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: "(\<forall>y\<in>x. \<exists>!y'. f(y') = f(y))
- ==> (\<lambda>z\<in>{f(y). y \<in> x}. THE y. f(y) = z) \<in> bij({f(y). y \<in> x}, x)"
+ \<Longrightarrow> (\<lambda>z\<in>{f(y). y \<in> x}. THE y. f(y) = z) \<in> bij({f(y). y \<in> x}, x)"
apply (rule_tac d = f in lam_bijective)
apply (auto simp add: the_equality2)
done
@@ -146,12 +146,12 @@
subsection\<open>Identity Function\<close>
-lemma idI [intro!]: "a \<in> A ==> <a,a> \<in> id(A)"
+lemma idI [intro!]: "a \<in> A \<Longrightarrow> <a,a> \<in> id(A)"
apply (unfold id_def)
apply (erule lamI)
done
-lemma idE [elim!]: "[| p \<in> id(A); !!x.[| x \<in> A; p=<x,x> |] ==> P |] ==> P"
+lemma idE [elim!]: "\<lbrakk>p \<in> id(A); \<And>x.\<lbrakk>x \<in> A; p=<x,x>\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp add: id_def lam_def, blast)
lemma id_type: "id(A) \<in> A->A"
@@ -159,17 +159,17 @@
apply (rule lam_type, assumption)
done
-lemma id_conv [simp]: "x \<in> A ==> id(A)`x = x"
+lemma id_conv [simp]: "x \<in> A \<Longrightarrow> id(A)`x = x"
apply (unfold id_def)
apply (simp (no_asm_simp))
done
-lemma id_mono: "A<=B ==> id(A) \<subseteq> id(B)"
+lemma id_mono: "A<=B \<Longrightarrow> id(A) \<subseteq> id(B)"
apply (unfold id_def)
apply (erule lam_mono)
done
-lemma id_subset_inj: "A<=B ==> id(A): inj(A,B)"
+lemma id_subset_inj: "A<=B \<Longrightarrow> id(A): inj(A,B)"
apply (simp add: inj_def id_def)
apply (blast intro: lam_type)
done
@@ -198,7 +198,7 @@
subsection\<open>Converse of a Function\<close>
-lemma inj_converse_fun: "f \<in> inj(A,B) ==> converse(f) \<in> range(f)->A"
+lemma inj_converse_fun: "f \<in> inj(A,B) \<Longrightarrow> converse(f) \<in> range(f)->A"
apply (unfold inj_def)
apply (simp (no_asm_simp) add: Pi_iff function_def)
apply (erule CollectE)
@@ -210,34 +210,34 @@
text\<open>The premises are equivalent to saying that f is injective...\<close>
lemma left_inverse_lemma:
- "[| f \<in> A->B; converse(f): C->A; a \<in> A |] ==> converse(f)`(f`a) = a"
+ "\<lbrakk>f \<in> A->B; converse(f): C->A; a \<in> A\<rbrakk> \<Longrightarrow> converse(f)`(f`a) = a"
by (blast intro: apply_Pair apply_equality converseI)
-lemma left_inverse [simp]: "[| f \<in> inj(A,B); a \<in> A |] ==> converse(f)`(f`a) = a"
+lemma left_inverse [simp]: "\<lbrakk>f \<in> inj(A,B); a \<in> A\<rbrakk> \<Longrightarrow> converse(f)`(f`a) = a"
by (blast intro: left_inverse_lemma inj_converse_fun inj_is_fun)
lemma left_inverse_eq:
- "[|f \<in> inj(A,B); f ` x = y; x \<in> A|] ==> converse(f) ` y = x"
+ "\<lbrakk>f \<in> inj(A,B); f ` x = y; x \<in> A\<rbrakk> \<Longrightarrow> converse(f) ` y = x"
by auto
lemmas left_inverse_bij = bij_is_inj [THEN left_inverse]
lemma right_inverse_lemma:
- "[| f \<in> A->B; converse(f): C->A; b \<in> C |] ==> f`(converse(f)`b) = b"
+ "\<lbrakk>f \<in> A->B; converse(f): C->A; b \<in> C\<rbrakk> \<Longrightarrow> f`(converse(f)`b) = b"
by (rule apply_Pair [THEN converseD [THEN apply_equality]], auto)
(*Should the premises be f \<in> surj(A,B), b \<in> B for symmetry with left_inverse?
No: they would not imply that converse(f) was a function! *)
lemma right_inverse [simp]:
- "[| f \<in> inj(A,B); b \<in> range(f) |] ==> f`(converse(f)`b) = b"
+ "\<lbrakk>f \<in> inj(A,B); b \<in> range(f)\<rbrakk> \<Longrightarrow> f`(converse(f)`b) = b"
by (blast intro: right_inverse_lemma inj_converse_fun inj_is_fun)
-lemma right_inverse_bij: "[| f \<in> bij(A,B); b \<in> B |] ==> f`(converse(f)`b) = b"
+lemma right_inverse_bij: "\<lbrakk>f \<in> bij(A,B); b \<in> B\<rbrakk> \<Longrightarrow> f`(converse(f)`b) = b"
by (force simp add: bij_def surj_range)
subsection\<open>Converses of Injections, Surjections, Bijections\<close>
-lemma inj_converse_inj: "f \<in> inj(A,B) ==> converse(f): inj(range(f), A)"
+lemma inj_converse_inj: "f \<in> inj(A,B) \<Longrightarrow> converse(f): inj(range(f), A)"
apply (rule f_imp_injective)
apply (erule inj_converse_fun, clarify)
apply (rule right_inverse)
@@ -245,12 +245,12 @@
apply blast
done
-lemma inj_converse_surj: "f \<in> inj(A,B) ==> converse(f): surj(range(f), A)"
+lemma inj_converse_surj: "f \<in> inj(A,B) \<Longrightarrow> converse(f): surj(range(f), A)"
by (blast intro: f_imp_surjective inj_converse_fun left_inverse inj_is_fun
range_of_fun [THEN apply_type])
text\<open>Adding this as an intro! rule seems to cause looping\<close>
-lemma bij_converse_bij [TC]: "f \<in> bij(A,B) ==> converse(f): bij(B,A)"
+lemma bij_converse_bij [TC]: "f \<in> bij(A,B) \<Longrightarrow> converse(f): bij(B,A)"
apply (unfold bij_def)
apply (fast elim: surj_range [THEN subst] inj_converse_inj inj_converse_surj)
done
@@ -261,19 +261,19 @@
text\<open>The inductive definition package could derive these theorems for \<^term>\<open>r O s\<close>\<close>
-lemma compI [intro]: "[| <a,b>:s; <b,c>:r |] ==> <a,c> \<in> r O s"
+lemma compI [intro]: "\<lbrakk><a,b>:s; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c> \<in> r O s"
by (unfold comp_def, blast)
lemma compE [elim!]:
- "[| xz \<in> r O s;
- !!x y z. [| xz=<x,z>; <x,y>:s; <y,z>:r |] ==> P |]
- ==> P"
+ "\<lbrakk>xz \<in> r O s;
+ \<And>x y z. \<lbrakk>xz=<x,z>; <x,y>:s; <y,z>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by (unfold comp_def, blast)
lemma compEpair:
- "[| <a,c> \<in> r O s;
- !!y. [| <a,y>:s; <y,c>:r |] ==> P |]
- ==> P"
+ "\<lbrakk><a,c> \<in> r O s;
+ \<And>y. \<lbrakk><a,y>:s; <y,c>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by (erule compE, simp)
lemma converse_comp: "converse(R O S) = converse(S) O converse(R)"
@@ -286,32 +286,32 @@
lemma range_comp: "range(r O s) \<subseteq> range(r)"
by blast
-lemma range_comp_eq: "domain(r) \<subseteq> range(s) ==> range(r O s) = range(r)"
+lemma range_comp_eq: "domain(r) \<subseteq> range(s) \<Longrightarrow> range(r O s) = range(r)"
by (rule range_comp [THEN equalityI], blast)
lemma domain_comp: "domain(r O s) \<subseteq> domain(s)"
by blast
-lemma domain_comp_eq: "range(s) \<subseteq> domain(r) ==> domain(r O s) = domain(s)"
+lemma domain_comp_eq: "range(s) \<subseteq> domain(r) \<Longrightarrow> 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 \<in> inj(A,B) ==> f \<in> inj(A,range(f))"
+lemma inj_inj_range: "f \<in> inj(A,B) \<Longrightarrow> f \<in> inj(A,range(f))"
by (auto simp add: inj_def Pi_iff function_def)
-lemma inj_bij_range: "f \<in> inj(A,B) ==> f \<in> bij(A,range(f))"
+lemma inj_bij_range: "f \<in> inj(A,B) \<Longrightarrow> f \<in> bij(A,range(f))"
by (auto simp add: bij_def intro: inj_inj_range inj_is_fun fun_is_surj)
subsection\<open>Other Results\<close>
-lemma comp_mono: "[| r'<=r; s'<=s |] ==> (r' O s') \<subseteq> (r O s)"
+lemma comp_mono: "\<lbrakk>r'<=r; s'<=s\<rbrakk> \<Longrightarrow> (r' O s') \<subseteq> (r O s)"
by blast
text\<open>composition preserves relations\<close>
-lemma comp_rel: "[| s<=A*B; r<=B*C |] ==> (r O s) \<subseteq> A*C"
+lemma comp_rel: "\<lbrakk>s<=A*B; r<=B*C\<rbrakk> \<Longrightarrow> (r O s) \<subseteq> A*C"
by blast
text\<open>associative law for composition\<close>
@@ -320,31 +320,31 @@
(*left identity of composition; provable inclusions are
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"
+ and \<lbrakk>r<=A*B; B<=C\<rbrakk> \<Longrightarrow> r \<subseteq> id(C) O r *)
+lemma left_comp_id: "r<=A*B \<Longrightarrow> id(B) O r = r"
by blast
(*right identity of composition; provable inclusions are
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"
+ and \<lbrakk>r<=A*B; A<=C\<rbrakk> \<Longrightarrow> r \<subseteq> r O id(C) *)
+lemma right_comp_id: "r<=A*B \<Longrightarrow> r O id(A) = r"
by blast
subsection\<open>Composition Preserves Functions, Injections, and Surjections\<close>
-lemma comp_function: "[| function(g); function(f) |] ==> function(f O g)"
+lemma comp_function: "\<lbrakk>function(g); function(f)\<rbrakk> \<Longrightarrow> function(f O g)"
by (unfold function_def, blast)
text\<open>Don't think the premises can be weakened much\<close>
-lemma comp_fun: "[| g \<in> A->B; f \<in> B->C |] ==> (f O g) \<in> A->C"
+lemma comp_fun: "\<lbrakk>g \<in> A->B; f \<in> B->C\<rbrakk> \<Longrightarrow> (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
(*Thanks to the new definition of "apply", the premise f \<in> B->C is gone!*)
lemma comp_fun_apply [simp]:
- "[| g \<in> A->B; a \<in> A |] ==> (f O g)`a = f`(g`a)"
+ "\<lbrakk>g \<in> A->B; a \<in> A\<rbrakk> \<Longrightarrow> (f O g)`a = f`(g`a)"
apply (frule apply_Pair, assumption)
apply (simp add: apply_def image_comp)
apply (blast dest: apply_equality)
@@ -352,8 +352,8 @@
text\<open>Simplifies compositions of lambda-abstractions\<close>
lemma comp_lam:
- "[| !!x. x \<in> A ==> b(x): B |]
- ==> (\<lambda>y\<in>B. c(y)) O (\<lambda>x\<in>A. b(x)) = (\<lambda>x\<in>A. c(b(x)))"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> b(x): B\<rbrakk>
+ \<Longrightarrow> (\<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)
@@ -363,7 +363,7 @@
done
lemma comp_inj:
- "[| g \<in> inj(A,B); f \<in> inj(B,C) |] ==> (f O g) \<in> inj(A,C)"
+ "\<lbrakk>g \<in> inj(A,B); f \<in> inj(B,C)\<rbrakk> \<Longrightarrow> (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 \<in> surj(A,B); f \<in> surj(B,C) |] ==> (f O g) \<in> surj(A,C)"
+ "\<lbrakk>g \<in> surj(A,B); f \<in> surj(B,C)\<rbrakk> \<Longrightarrow> (f O g) \<in> surj(A,C)"
apply (unfold surj_def)
apply (blast intro!: comp_fun comp_fun_apply)
done
lemma comp_bij:
- "[| g \<in> bij(A,B); f \<in> bij(B,C) |] ==> (f O g) \<in> bij(A,C)"
+ "\<lbrakk>g \<in> bij(A,B); f \<in> bij(B,C)\<rbrakk> \<Longrightarrow> (f O g) \<in> bij(A,C)"
apply (unfold bij_def)
apply (blast intro: comp_inj comp_surj)
done
@@ -390,11 +390,11 @@
Artificial Intelligence, 10:1--27, 1978.\<close>
lemma comp_mem_injD1:
- "[| (f O g): inj(A,C); g \<in> A->B; f \<in> B->C |] ==> g \<in> inj(A,B)"
+ "\<lbrakk>(f O g): inj(A,C); g \<in> A->B; f \<in> B->C\<rbrakk> \<Longrightarrow> g \<in> inj(A,B)"
by (unfold inj_def, force)
lemma comp_mem_injD2:
- "[| (f O g): inj(A,C); g \<in> surj(A,B); f \<in> B->C |] ==> f \<in> inj(B,C)"
+ "\<lbrakk>(f O g): inj(A,C); g \<in> surj(A,B); f \<in> B->C\<rbrakk> \<Longrightarrow> f \<in> inj(B,C)"
apply (unfold inj_def surj_def, safe)
apply (rule_tac x1 = x in bspec [THEN bexE])
apply (erule_tac [3] x1 = w in bspec [THEN bexE], assumption+, safe)
@@ -404,14 +404,14 @@
done
lemma comp_mem_surjD1:
- "[| (f O g): surj(A,C); g \<in> A->B; f \<in> B->C |] ==> f \<in> surj(B,C)"
+ "\<lbrakk>(f O g): surj(A,C); g \<in> A->B; f \<in> B->C\<rbrakk> \<Longrightarrow> f \<in> surj(B,C)"
apply (unfold surj_def)
apply (blast intro!: comp_fun_apply [symmetric] apply_funtype)
done
lemma comp_mem_surjD2:
- "[| (f O g): surj(A,C); g \<in> A->B; f \<in> inj(B,C) |] ==> g \<in> surj(A,B)"
+ "\<lbrakk>(f O g): surj(A,C); g \<in> A->B; f \<in> inj(B,C)\<rbrakk> \<Longrightarrow> g \<in> surj(A,B)"
apply (unfold inj_def surj_def, safe)
apply (drule_tac x = "f`y" in bspec, auto)
apply (blast intro: apply_funtype)
@@ -420,17 +420,17 @@
subsubsection\<open>Inverses of Composition\<close>
text\<open>left inverse of composition; one inclusion is
- \<^term>\<open>f \<in> A->B ==> id(A) \<subseteq> converse(f) O f\<close>\<close>
-lemma left_comp_inverse: "f \<in> inj(A,B) ==> converse(f) O f = id(A)"
+ \<^term>\<open>f \<in> A->B \<Longrightarrow> id(A) \<subseteq> converse(f) O f\<close>\<close>
+lemma left_comp_inverse: "f \<in> inj(A,B) \<Longrightarrow> converse(f) O f = id(A)"
apply (unfold inj_def, clarify)
apply (rule equalityI)
apply (auto simp add: apply_iff, blast)
done
text\<open>right inverse of composition; one inclusion is
- \<^term>\<open>f \<in> A->B ==> f O converse(f) \<subseteq> id(B)\<close>\<close>
+ \<^term>\<open>f \<in> A->B \<Longrightarrow> f O converse(f) \<subseteq> id(B)\<close>\<close>
lemma right_comp_inverse:
- "f \<in> surj(A,B) ==> f O converse(f) = id(B)"
+ "f \<in> surj(A,B) \<Longrightarrow> f O converse(f) = id(B)"
apply (simp add: surj_def, clarify)
apply (rule equalityI)
apply (best elim: domain_type range_type dest: apply_equality2)
@@ -441,7 +441,7 @@
subsubsection\<open>Proving that a Function is a Bijection\<close>
lemma comp_eq_id_iff:
- "[| f \<in> A->B; g \<in> B->A |] ==> f O g = id(B) \<longleftrightarrow> (\<forall>y\<in>B. f`(g`y)=y)"
+ "\<lbrakk>f \<in> A->B; g \<in> B->A\<rbrakk> \<Longrightarrow> f O g = id(B) \<longleftrightarrow> (\<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 \<in> A->B; g \<in> B->A; f O g = id(B); g O f = id(A) |] ==> f \<in> bij(A,B)"
+ "\<lbrakk>f \<in> A->B; g \<in> B->A; f O g = id(B); g O f = id(A)\<rbrakk> \<Longrightarrow> 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 \<in> A->A; f O f = id(A) |] ==> f \<in> bij(A,A)"
+lemma nilpotent_imp_bijective: "\<lbrakk>f \<in> A->A; f O f = id(A)\<rbrakk> \<Longrightarrow> f \<in> bij(A,A)"
by (blast intro: fg_imp_bijective)
lemma invertible_imp_bijective:
- "[| converse(f): B->A; f \<in> A->B |] ==> f \<in> bij(A,B)"
+ "\<lbrakk>converse(f): B->A; f \<in> A->B\<rbrakk> \<Longrightarrow> 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\<open>Theorem by KG, proof by LCP\<close>
lemma inj_disjoint_Un:
- "[| f \<in> inj(A,B); g \<in> inj(C,D); B \<inter> D = 0 |]
- ==> (\<lambda>a\<in>A \<union> C. if a \<in> A then f`a else g`a) \<in> inj(A \<union> C, B \<union> D)"
+ "\<lbrakk>f \<in> inj(A,B); g \<in> inj(C,D); B \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> (\<lambda>a\<in>A \<union> C. if a \<in> A then f`a else g`a) \<in> inj(A \<union> C, B \<union> D)"
apply (rule_tac d = "%z. if z \<in> 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 \<in> surj(A,B); g \<in> surj(C,D); A \<inter> C = 0 |]
- ==> (f \<union> g) \<in> surj(A \<union> C, B \<union> D)"
+ "\<lbrakk>f \<in> surj(A,B); g \<in> surj(C,D); A \<inter> C = 0\<rbrakk>
+ \<Longrightarrow> (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\<open>A simple, high-level proof; the version for injections follows from it,
using \<^term>\<open>f \<in> inj(A,B) \<longleftrightarrow> f \<in> bij(A,range(f))\<close>\<close>
lemma bij_disjoint_Un:
- "[| f \<in> bij(A,B); g \<in> bij(C,D); A \<inter> C = 0; B \<inter> D = 0 |]
- ==> (f \<union> g) \<in> bij(A \<union> C, B \<union> D)"
+ "\<lbrakk>f \<in> bij(A,B); g \<in> bij(C,D); A \<inter> C = 0; B \<inter> D = 0\<rbrakk>
+ \<Longrightarrow> (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)
@@ -500,30 +500,30 @@
subsubsection\<open>Restrictions as Surjections and Bijections\<close>
lemma surj_image:
- "f \<in> Pi(A,B) ==> f \<in> surj(A, f``A)"
+ "f \<in> Pi(A,B) \<Longrightarrow> f \<in> surj(A, f``A)"
apply (simp add: surj_def)
apply (blast intro: apply_equality apply_Pair Pi_type)
done
-lemma surj_image_eq: "f \<in> surj(A, B) ==> f``A = B"
+lemma surj_image_eq: "f \<in> surj(A, B) \<Longrightarrow> f``A = B"
by (auto simp add: surj_def image_fun) (blast dest: apply_type)
lemma restrict_image [simp]: "restrict(f,A) `` B = f `` (A \<inter> B)"
by (auto simp add: restrict_def)
lemma restrict_inj:
- "[| f \<in> inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"
+ "\<lbrakk>f \<in> inj(A,B); C<=A\<rbrakk> \<Longrightarrow> restrict(f,C): inj(C,B)"
apply (unfold inj_def)
apply (safe elim!: restrict_type2, auto)
done
-lemma restrict_surj: "[| f \<in> Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)"
+lemma restrict_surj: "\<lbrakk>f \<in> Pi(A,B); C<=A\<rbrakk> \<Longrightarrow> restrict(f,C): surj(C, f``C)"
apply (insert restrict_type2 [THEN surj_image])
apply (simp add: restrict_image)
done
lemma restrict_bij:
- "[| f \<in> inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"
+ "\<lbrakk>f \<in> inj(A,B); C<=A\<rbrakk> \<Longrightarrow> restrict(f,C): bij(C, f``C)"
apply (simp add: inj_def bij_def)
apply (blast intro: restrict_surj surj_is_fun)
done
@@ -531,13 +531,13 @@
subsubsection\<open>Lemmas for Ramsey's Theorem\<close>
-lemma inj_weaken_type: "[| f \<in> inj(A,B); B<=D |] ==> f \<in> inj(A,D)"
+lemma inj_weaken_type: "\<lbrakk>f \<in> inj(A,B); B<=D\<rbrakk> \<Longrightarrow> f \<in> inj(A,D)"
apply (unfold inj_def)
apply (blast intro: fun_weaken_type)
done
lemma inj_succ_restrict:
- "[| f \<in> inj(succ(m), A) |] ==> restrict(f,m) \<in> inj(m, A-{f`m})"
+ "\<lbrakk>f \<in> inj(succ(m), A)\<rbrakk> \<Longrightarrow> 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)
@@ -545,8 +545,8 @@
lemma inj_extend:
- "[| f \<in> inj(A,B); a\<notin>A; b\<notin>B |]
- ==> cons(<a,b>,f) \<in> inj(cons(a,A), cons(b,B))"
+ "\<lbrakk>f \<in> inj(A,B); a\<notin>A; b\<notin>B\<rbrakk>
+ \<Longrightarrow> 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 Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/QPair.thy Tue Sep 27 16:51:35 2022 +0100
@@ -22,27 +22,27 @@
definition
QPair :: "[i, i] => i" (\<open><(_;/ _)>\<close>) where
- "<a;b> == a+b"
+ "<a;b> \<equiv> a+b"
definition
qfst :: "i => i" where
- "qfst(p) == THE a. \<exists>b. p=<a;b>"
+ "qfst(p) \<equiv> THE a. \<exists>b. p=<a;b>"
definition
qsnd :: "i => i" where
- "qsnd(p) == THE b. \<exists>a. p=<a;b>"
+ "qsnd(p) \<equiv> THE b. \<exists>a. p=<a;b>"
definition
qsplit :: "[[i, i] => 'a, i] => 'a::{}" (*for pattern-matching*) where
- "qsplit(c,p) == c(qfst(p), qsnd(p))"
+ "qsplit(c,p) \<equiv> c(qfst(p), qsnd(p))"
definition
qconverse :: "i => i" where
- "qconverse(r) == {z. w \<in> r, \<exists>x y. w=<x;y> & z=<y;x>}"
+ "qconverse(r) \<equiv> {z. w \<in> r, \<exists>x y. w=<x;y> & z=<y;x>}"
definition
QSigma :: "[i, i => i] => i" where
- "QSigma(A,B) == \<Union>x\<in>A. \<Union>y\<in>B(x). {<x;y>}"
+ "QSigma(A,B) \<equiv> \<Union>x\<in>A. \<Union>y\<in>B(x). {<x;y>}"
syntax
"_QSUM" :: "[idt, i, i] => i" (\<open>(3QSUM _ \<in> _./ _)\<close> 10)
@@ -51,23 +51,23 @@
abbreviation
qprod (infixr \<open><*>\<close> 80) where
- "A <*> B == QSigma(A, %_. B)"
+ "A <*> B \<equiv> QSigma(A, %_. B)"
definition
qsum :: "[i,i]=>i" (infixr \<open><+>\<close> 65) where
- "A <+> B == ({0} <*> A) \<union> ({1} <*> B)"
+ "A <+> B \<equiv> ({0} <*> A) \<union> ({1} <*> B)"
definition
QInl :: "i=>i" where
- "QInl(a) == <0;a>"
+ "QInl(a) \<equiv> <0;a>"
definition
QInr :: "i=>i" where
- "QInr(b) == <1;b>"
+ "QInr(b) \<equiv> <1;b>"
definition
qcase :: "[i=>i, i=>i, i]=>i" where
- "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))"
+ "qcase(c,d) \<equiv> qsplit(%y z. cond(y, d(z), c(z)))"
subsection\<open>Quine ordered pairing\<close>
@@ -84,40 +84,40 @@
lemmas QPair_inject = QPair_iff [THEN iffD1, THEN conjE, elim!]
-lemma QPair_inject1: "<a;b> = <c;d> ==> a=c"
+lemma QPair_inject1: "<a;b> = <c;d> \<Longrightarrow> a=c"
by blast
-lemma QPair_inject2: "<a;b> = <c;d> ==> b=d"
+lemma QPair_inject2: "<a;b> = <c;d> \<Longrightarrow> b=d"
by blast
subsubsection\<open>QSigma: Disjoint union of a family of sets
Generalizes Cartesian product\<close>
-lemma QSigmaI [intro!]: "[| a \<in> A; b \<in> B(a) |] ==> <a;b> \<in> QSigma(A,B)"
+lemma QSigmaI [intro!]: "\<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> <a;b> \<in> QSigma(A,B)"
by (simp add: QSigma_def)
(** Elimination rules for <a;b>:A*B -- introducing no eigenvariables **)
lemma QSigmaE [elim!]:
- "[| c \<in> QSigma(A,B);
- !!x y.[| x \<in> A; y \<in> B(x); c=<x;y> |] ==> P
- |] ==> P"
+ "\<lbrakk>c \<in> QSigma(A,B);
+ \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x); c=<x;y>\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: QSigma_def, blast)
lemma QSigmaE2 [elim!]:
- "[| <a;b>: QSigma(A,B); [| a \<in> A; b \<in> B(a) |] ==> P |] ==> P"
+ "\<lbrakk><a;b>: QSigma(A,B); \<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp add: QSigma_def)
-lemma QSigmaD1: "<a;b> \<in> QSigma(A,B) ==> a \<in> A"
+lemma QSigmaD1: "<a;b> \<in> QSigma(A,B) \<Longrightarrow> a \<in> A"
by blast
-lemma QSigmaD2: "<a;b> \<in> QSigma(A,B) ==> b \<in> B(a)"
+lemma QSigmaD2: "<a;b> \<in> QSigma(A,B) \<Longrightarrow> b \<in> B(a)"
by blast
lemma QSigma_cong:
- "[| A=A'; !!x. x \<in> A' ==> B(x)=B'(x) |] ==>
+ "\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow>
QSigma(A,B) = QSigma(A',B')"
by (simp add: QSigma_def)
@@ -136,69 +136,69 @@
lemma qsnd_conv [simp]: "qsnd(<a;b>) = b"
by (simp add: qsnd_def)
-lemma qfst_type [TC]: "p \<in> QSigma(A,B) ==> qfst(p) \<in> A"
+lemma qfst_type [TC]: "p \<in> QSigma(A,B) \<Longrightarrow> qfst(p) \<in> A"
by auto
-lemma qsnd_type [TC]: "p \<in> QSigma(A,B) ==> qsnd(p) \<in> B(qfst(p))"
+lemma qsnd_type [TC]: "p \<in> QSigma(A,B) \<Longrightarrow> qsnd(p) \<in> B(qfst(p))"
by auto
-lemma QPair_qfst_qsnd_eq: "a \<in> QSigma(A,B) ==> <qfst(a); qsnd(a)> = a"
+lemma QPair_qfst_qsnd_eq: "a \<in> QSigma(A,B) \<Longrightarrow> <qfst(a); qsnd(a)> = a"
by auto
subsubsection\<open>Eliminator: qsplit\<close>
(*A META-equality, so that it applies to higher types as well...*)
-lemma qsplit [simp]: "qsplit(%x y. c(x,y), <a;b>) == c(a,b)"
+lemma qsplit [simp]: "qsplit(%x y. c(x,y), <a;b>) \<equiv> c(a,b)"
by (simp add: qsplit_def)
lemma qsplit_type [elim!]:
- "[| p \<in> QSigma(A,B);
- !!x y.[| x \<in> A; y \<in> B(x) |] ==> c(x,y):C(<x;y>)
- |] ==> qsplit(%x y. c(x,y), p) \<in> C(p)"
+ "\<lbrakk>p \<in> QSigma(A,B);
+ \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(<x;y>)
+\<rbrakk> \<Longrightarrow> qsplit(%x y. c(x,y), p) \<in> C(p)"
by auto
lemma expand_qsplit:
- "u \<in> A<*>B ==> R(qsplit(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x;y> \<longrightarrow> R(c(x,y)))"
+ "u \<in> A<*>B \<Longrightarrow> R(qsplit(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x;y> \<longrightarrow> R(c(x,y)))"
apply (simp add: qsplit_def, auto)
done
subsubsection\<open>qsplit for predicates: result type o\<close>
-lemma qsplitI: "R(a,b) ==> qsplit(R, <a;b>)"
+lemma qsplitI: "R(a,b) \<Longrightarrow> qsplit(R, <a;b>)"
by (simp add: qsplit_def)
lemma qsplitE:
- "[| qsplit(R,z); z \<in> QSigma(A,B);
- !!x y. [| z = <x;y>; R(x,y) |] ==> P
- |] ==> P"
+ "\<lbrakk>qsplit(R,z); z \<in> QSigma(A,B);
+ \<And>x y. \<lbrakk>z = <x;y>; R(x,y)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: qsplit_def, auto)
-lemma qsplitD: "qsplit(R,<a;b>) ==> R(a,b)"
+lemma qsplitD: "qsplit(R,<a;b>) \<Longrightarrow> R(a,b)"
by (simp add: qsplit_def)
subsubsection\<open>qconverse\<close>
-lemma qconverseI [intro!]: "<a;b>:r ==> <b;a>:qconverse(r)"
+lemma qconverseI [intro!]: "<a;b>:r \<Longrightarrow> <b;a>:qconverse(r)"
by (simp add: qconverse_def, blast)
-lemma qconverseD [elim!]: "<a;b> \<in> qconverse(r) ==> <b;a> \<in> r"
+lemma qconverseD [elim!]: "<a;b> \<in> qconverse(r) \<Longrightarrow> <b;a> \<in> r"
by (simp add: qconverse_def, blast)
lemma qconverseE [elim!]:
- "[| yx \<in> qconverse(r);
- !!x y. [| yx=<y;x>; <x;y>:r |] ==> P
- |] ==> P"
+ "\<lbrakk>yx \<in> qconverse(r);
+ \<And>x y. \<lbrakk>yx=<y;x>; <x;y>:r\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: qconverse_def, blast)
-lemma qconverse_qconverse: "r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r"
+lemma qconverse_qconverse: "r<=QSigma(A,B) \<Longrightarrow> qconverse(qconverse(r)) = r"
by blast
-lemma qconverse_type: "r \<subseteq> A <*> B ==> qconverse(r) \<subseteq> B <*> A"
+lemma qconverse_type: "r \<subseteq> A <*> B \<Longrightarrow> qconverse(r) \<subseteq> B <*> A"
by blast
lemma qconverse_prod: "qconverse(A <*> B) = B <*> A"
@@ -214,19 +214,19 @@
(** Introduction rules for the injections **)
-lemma QInlI [intro!]: "a \<in> A ==> QInl(a) \<in> A <+> B"
+lemma QInlI [intro!]: "a \<in> A \<Longrightarrow> QInl(a) \<in> A <+> B"
by (simp add: qsum_defs, blast)
-lemma QInrI [intro!]: "b \<in> B ==> QInr(b) \<in> A <+> B"
+lemma QInrI [intro!]: "b \<in> B \<Longrightarrow> QInr(b) \<in> A <+> B"
by (simp add: qsum_defs, blast)
(** Elimination rules **)
lemma qsumE [elim!]:
- "[| u \<in> A <+> B;
- !!x. [| x \<in> A; u=QInl(x) |] ==> P;
- !!y. [| y \<in> B; u=QInr(y) |] ==> P
- |] ==> P"
+ "\<lbrakk>u \<in> A <+> B;
+ \<And>x. \<lbrakk>x \<in> A; u=QInl(x)\<rbrakk> \<Longrightarrow> P;
+ \<And>y. \<lbrakk>y \<in> B; u=QInr(y)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: qsum_defs, blast)
@@ -254,10 +254,10 @@
lemmas QInl_neq_QInr = QInl_QInr_iff [THEN iffD1, THEN FalseE, elim!]
lemmas QInr_neq_QInl = QInr_QInl_iff [THEN iffD1, THEN FalseE, elim!]
-lemma QInlD: "QInl(a): A<+>B ==> a \<in> A"
+lemma QInlD: "QInl(a): A<+>B \<Longrightarrow> a \<in> A"
by blast
-lemma QInrD: "QInr(b): A<+>B ==> b \<in> B"
+lemma QInrD: "QInr(b): A<+>B \<Longrightarrow> b \<in> B"
by blast
(** <+> is itself injective... who cares?? **)
@@ -284,10 +284,10 @@
by (simp add: qsum_defs )
lemma qcase_type:
- "[| u \<in> A <+> B;
- !!x. x \<in> A ==> c(x): C(QInl(x));
- !!y. y \<in> B ==> d(y): C(QInr(y))
- |] ==> qcase(c,d,u) \<in> C(u)"
+ "\<lbrakk>u \<in> A <+> B;
+ \<And>x. x \<in> A \<Longrightarrow> c(x): C(QInl(x));
+ \<And>y. y \<in> B \<Longrightarrow> d(y): C(QInr(y))
+\<rbrakk> \<Longrightarrow> 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 \<in> Part(B,h)}"
by blast
-lemma Part_qsum_equality: "C \<subseteq> A <+> B ==> Part(C,QInl) \<union> Part(C,QInr) = C"
+lemma Part_qsum_equality: "C \<subseteq> A <+> B \<Longrightarrow> Part(C,QInl) \<union> Part(C,QInr) = C"
by blast
subsubsection\<open>Monotonicity\<close>
-lemma QPair_mono: "[| a<=c; b<=d |] ==> <a;b> \<subseteq> <c;d>"
+lemma QPair_mono: "\<lbrakk>a<=c; b<=d\<rbrakk> \<Longrightarrow> <a;b> \<subseteq> <c;d>"
by (simp add: QPair_def sum_mono)
lemma QSigma_mono [rule_format]:
- "[| A<=C; \<forall>x\<in>A. B(x) \<subseteq> D(x) |] ==> QSigma(A,B) \<subseteq> QSigma(C,D)"
+ "\<lbrakk>A<=C; \<forall>x\<in>A. B(x) \<subseteq> D(x)\<rbrakk> \<Longrightarrow> QSigma(A,B) \<subseteq> QSigma(C,D)"
by blast
-lemma QInl_mono: "a<=b ==> QInl(a) \<subseteq> QInl(b)"
+lemma QInl_mono: "a<=b \<Longrightarrow> QInl(a) \<subseteq> QInl(b)"
by (simp add: QInl_def subset_refl [THEN QPair_mono])
-lemma QInr_mono: "a<=b ==> QInr(a) \<subseteq> QInr(b)"
+lemma QInr_mono: "a<=b \<Longrightarrow> QInr(a) \<subseteq> QInr(b)"
by (simp add: QInr_def subset_refl [THEN QPair_mono])
-lemma qsum_mono: "[| A<=C; B<=D |] ==> A <+> B \<subseteq> C <+> D"
+lemma qsum_mono: "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> A <+> B \<subseteq> C <+> D"
by blast
end
--- a/src/ZF/QUniv.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/QUniv.thy Tue Sep 27 16:51:35 2022 +0100
@@ -21,32 +21,32 @@
definition
quniv :: "i => i" where
- "quniv(A) == Pow(univ(eclose(A)))"
+ "quniv(A) \<equiv> Pow(univ(eclose(A)))"
subsection\<open>Properties involving Transset and Sum\<close>
lemma Transset_includes_summands:
- "[| Transset(C); A+B \<subseteq> C |] ==> A \<subseteq> C & B \<subseteq> C"
+ "\<lbrakk>Transset(C); A+B \<subseteq> C\<rbrakk> \<Longrightarrow> 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) \<inter> C \<subseteq> (A \<inter> C) + (B \<inter> C)"
+ "Transset(C) \<Longrightarrow> (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\<open>Introduction and Elimination Rules\<close>
-lemma qunivI: "X \<subseteq> univ(eclose(A)) ==> X \<in> quniv(A)"
+lemma qunivI: "X \<subseteq> univ(eclose(A)) \<Longrightarrow> X \<in> quniv(A)"
by (simp add: quniv_def)
-lemma qunivD: "X \<in> quniv(A) ==> X \<subseteq> univ(eclose(A))"
+lemma qunivD: "X \<in> quniv(A) \<Longrightarrow> X \<subseteq> univ(eclose(A))"
by (simp add: quniv_def)
-lemma quniv_mono: "A<=B ==> quniv(A) \<subseteq> quniv(B)"
+lemma quniv_mono: "A<=B \<Longrightarrow> quniv(A) \<subseteq> quniv(B)"
apply (unfold quniv_def)
apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono])
done
@@ -86,12 +86,12 @@
(*Quine ordered pairs*)
lemma QPair_subset_univ:
- "[| a \<subseteq> univ(A); b \<subseteq> univ(A) |] ==> <a;b> \<subseteq> univ(A)"
+ "\<lbrakk>a \<subseteq> univ(A); b \<subseteq> univ(A)\<rbrakk> \<Longrightarrow> <a;b> \<subseteq> univ(A)"
by (simp add: QPair_def sum_subset_univ)
subsection\<open>Quine Disjoint Sum\<close>
-lemma QInl_subset_univ: "a \<subseteq> univ(A) ==> QInl(a) \<subseteq> univ(A)"
+lemma QInl_subset_univ: "a \<subseteq> univ(A) \<Longrightarrow> QInl(a) \<subseteq> univ(A)"
apply (unfold QInl_def)
apply (erule empty_subsetI [THEN QPair_subset_univ])
done
@@ -102,7 +102,7 @@
lemmas naturals_subset_univ =
subset_trans [OF naturals_subset_nat nat_subset_univ]
-lemma QInr_subset_univ: "a \<subseteq> univ(A) ==> QInr(a) \<subseteq> univ(A)"
+lemma QInr_subset_univ: "a \<subseteq> univ(A) \<Longrightarrow> QInr(a) \<subseteq> univ(A)"
apply (unfold QInr_def)
apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ])
done
@@ -111,7 +111,7 @@
(*Quine ordered pairs*)
lemma QPair_in_quniv:
- "[| a: quniv(A); b: quniv(A) |] ==> <a;b> \<in> quniv(A)"
+ "\<lbrakk>a: quniv(A); b: quniv(A)\<rbrakk> \<Longrightarrow> <a;b> \<in> quniv(A)"
by (simp add: quniv_def QPair_def sum_subset_univ)
lemma QSigma_quniv: "quniv(A) <*> quniv(A) \<subseteq> quniv(A)"
@@ -121,7 +121,7 @@
(*The opposite inclusion*)
lemma quniv_QPair_D:
- "<a;b> \<in> quniv(A) ==> a: quniv(A) & b: quniv(A)"
+ "<a;b> \<in> quniv(A) \<Longrightarrow> a: quniv(A) & b: quniv(A)"
apply (unfold quniv_def QPair_def)
apply (rule Transset_includes_summands [THEN conjE])
apply (rule Transset_eclose [THEN Transset_univ])
@@ -136,10 +136,10 @@
subsection\<open>Quine Disjoint Sum\<close>
-lemma QInl_in_quniv: "a: quniv(A) ==> QInl(a) \<in> quniv(A)"
+lemma QInl_in_quniv: "a: quniv(A) \<Longrightarrow> 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) \<in> quniv(A)"
+lemma QInr_in_quniv: "b: quniv(A) \<Longrightarrow> QInr(b) \<in> quniv(A)"
by (simp add: QInr_def one_in_quniv QPair_in_quniv)
lemma qsum_quniv: "quniv(C) <+> quniv(C) \<subseteq> quniv(C)"
@@ -152,7 +152,7 @@
lemmas nat_subset_quniv = subset_trans [OF nat_subset_univ univ_subset_quniv]
-(* n:nat ==> n:quniv(A) *)
+(* n:nat \<Longrightarrow> n:quniv(A) *)
lemmas nat_into_quniv = nat_subset_quniv [THEN subsetD]
lemmas bool_subset_quniv = subset_trans [OF bool_subset_univ univ_subset_quniv]
@@ -163,7 +163,7 @@
(*Intersecting <a;b> with Vfrom...*)
lemma QPair_Int_Vfrom_succ_subset:
- "Transset(X) ==>
+ "Transset(X) \<Longrightarrow>
<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]
@@ -176,18 +176,18 @@
(*Rule for level i -- preserving the level, not decreasing it*)
lemma QPair_Int_Vfrom_subset:
- "Transset(X) ==>
+ "Transset(X) \<Longrightarrow>
<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
-(*@{term"[| a \<inter> Vset(i) \<subseteq> c; b \<inter> Vset(i) \<subseteq> d |] ==> <a;b> \<inter> Vset(i) \<subseteq> <c;d>"}*)
+(*@{term"\<lbrakk>a \<inter> Vset(i) \<subseteq> c; b \<inter> Vset(i) \<subseteq> d\<rbrakk> \<Longrightarrow> <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> \<inter> Vset(i) \<subseteq> (\<Union>j\<in>i. <a \<inter> Vset(j); b \<inter> Vset(j)>)"
+ "Ord(i) \<Longrightarrow> <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)
--- a/src/ZF/Resid/Confluence.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Resid/Confluence.thy Tue Sep 27 16:51:35 2022 +0100
@@ -7,13 +7,13 @@
definition
confluence :: "i=>o" where
- "confluence(R) ==
+ "confluence(R) \<equiv>
\<forall>x y. <x,y> \<in> R \<longrightarrow> (\<forall>z.<x,z> \<in> R \<longrightarrow> (\<exists>u.<y,u> \<in> R & <z,u> \<in> R))"
definition
strip :: "o" where
- "strip == \<forall>x y. (x ===> y) \<longrightarrow>
- (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) & (z===>u)))"
+ "strip \<equiv> \<forall>x y. (x =\<Longrightarrow> y) \<longrightarrow>
+ (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) & (z=\<Longrightarrow>u)))"
(* ------------------------------------------------------------------------- *)
@@ -21,7 +21,7 @@
(* ------------------------------------------------------------------------- *)
lemma strip_lemma_r:
- "[|confluence(Spar_red1)|]==> strip"
+ "\<lbrakk>confluence(Spar_red1)\<rbrakk>\<Longrightarrow> strip"
apply (unfold confluence_def strip_def)
apply (rule impI [THEN allI, THEN allI])
apply (erule Spar_red.induct, fast)
@@ -30,7 +30,7 @@
lemma strip_lemma_l:
- "strip==> confluence(Spar_red)"
+ "strip\<Longrightarrow> confluence(Spar_red)"
apply (unfold confluence_def strip_def)
apply (rule impI [THEN allI, THEN allI])
apply (erule Spar_red.induct, blast)
@@ -54,7 +54,7 @@
lemmas confluence_parallel_reduction =
parallel_moves [THEN strip_lemma_r, THEN strip_lemma_l]
-lemma lemma1: "[|confluence(Spar_red)|]==> confluence(Sred)"
+lemma lemma1: "\<lbrakk>confluence(Spar_red)\<rbrakk>\<Longrightarrow> confluence(Sred)"
by (unfold confluence_def, blast intro: par_red_red red_par_red)
lemmas confluence_beta_reduction =
@@ -69,17 +69,17 @@
abbreviation
Sconv1_rel (infixl \<open><-1->\<close> 50) where
- "a<-1->b == <a,b> \<in> Sconv1"
+ "a<-1->b \<equiv> <a,b> \<in> Sconv1"
abbreviation
Sconv_rel (infixl \<open><-\<longrightarrow>\<close> 50) where
- "a<-\<longrightarrow>b == <a,b> \<in> Sconv"
+ "a<-\<longrightarrow>b \<equiv> <a,b> \<in> Sconv"
inductive
domains "Sconv1" \<subseteq> "lambda*lambda"
intros
- red1: "m -1-> n ==> m<-1->n"
- expl: "n -1-> m ==> m<-1->n"
+ red1: "m -1-> n \<Longrightarrow> m<-1->n"
+ expl: "n -1-> m \<Longrightarrow> m<-1->n"
type_intros red1D1 red1D2 lambda.intros bool_typechecks
declare Sconv1.intros [intro]
@@ -87,14 +87,14 @@
inductive
domains "Sconv" \<subseteq> "lambda*lambda"
intros
- one_step: "m<-1->n ==> m<-\<longrightarrow>n"
- refl: "m \<in> lambda ==> m<-\<longrightarrow>m"
- trans: "[|m<-\<longrightarrow>n; n<-\<longrightarrow>p|] ==> m<-\<longrightarrow>p"
+ one_step: "m<-1->n \<Longrightarrow> m<-\<longrightarrow>n"
+ refl: "m \<in> lambda \<Longrightarrow> m<-\<longrightarrow>m"
+ trans: "\<lbrakk>m<-\<longrightarrow>n; n<-\<longrightarrow>p\<rbrakk> \<Longrightarrow> m<-\<longrightarrow>p"
type_intros Sconv1.dom_subset [THEN subsetD] lambda.intros bool_typechecks
declare Sconv.intros [intro]
-lemma conv_sym: "m<-\<longrightarrow>n ==> n<-\<longrightarrow>m"
+lemma conv_sym: "m<-\<longrightarrow>n \<Longrightarrow> n<-\<longrightarrow>m"
apply (erule Sconv.induct)
apply (erule Sconv1.induct, blast+)
done
@@ -103,7 +103,7 @@
(* Church_Rosser Theorem *)
(* ------------------------------------------------------------------------- *)
-lemma Church_Rosser: "m<-\<longrightarrow>n ==> \<exists>p.(m -\<longrightarrow>p) & (n -\<longrightarrow> p)"
+lemma Church_Rosser: "m<-\<longrightarrow>n \<Longrightarrow> \<exists>p.(m -\<longrightarrow>p) & (n -\<longrightarrow> p)"
apply (erule Sconv.induct)
apply (erule Sconv1.induct)
apply (blast intro: red1D1 redD2)
--- a/src/ZF/Resid/Redex.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Resid/Redex.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,14 +19,14 @@
abbreviation
Ssub_rel (infixl \<open>\<Longleftarrow>\<close> 70) where
- "a \<Longleftarrow> b == <a,b> \<in> Ssub"
+ "a \<Longleftarrow> b \<equiv> <a,b> \<in> Ssub"
abbreviation
Scomp_rel (infixl \<open>\<sim>\<close> 70) where
- "a \<sim> b == <a,b> \<in> Scomp"
+ "a \<sim> b \<equiv> <a,b> \<in> Scomp"
abbreviation
- "regular(a) == a \<in> Sreg"
+ "regular(a) \<equiv> a \<in> Sreg"
consts union_aux :: "i=>i"
primrec (*explicit lambda is required because both arguments of "\<squnion>" vary*)
@@ -44,35 +44,35 @@
definition
union (infixl \<open>\<squnion>\<close> 70) where
- "u \<squnion> v == union_aux(u)`v"
+ "u \<squnion> v \<equiv> union_aux(u)`v"
inductive
domains "Ssub" \<subseteq> "redexes*redexes"
intros
- Sub_Var: "n \<in> nat ==> Var(n) \<Longleftarrow> Var(n)"
- Sub_Fun: "[|u \<Longleftarrow> v|]==> Fun(u) \<Longleftarrow> Fun(v)"
- Sub_App1: "[|u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2; b \<in> bool|]==>
+ Sub_Var: "n \<in> nat \<Longrightarrow> Var(n) \<Longleftarrow> Var(n)"
+ Sub_Fun: "\<lbrakk>u \<Longleftarrow> v\<rbrakk>\<Longrightarrow> Fun(u) \<Longleftarrow> Fun(v)"
+ Sub_App1: "\<lbrakk>u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2; b \<in> bool\<rbrakk>\<Longrightarrow>
App(0,u1,u2) \<Longleftarrow> App(b,v1,v2)"
- Sub_App2: "[|u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2|]==> App(1,u1,u2) \<Longleftarrow> App(1,v1,v2)"
+ Sub_App2: "\<lbrakk>u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2\<rbrakk>\<Longrightarrow> App(1,u1,u2) \<Longleftarrow> App(1,v1,v2)"
type_intros redexes.intros bool_typechecks
inductive
domains "Scomp" \<subseteq> "redexes*redexes"
intros
- Comp_Var: "n \<in> nat ==> Var(n) \<sim> Var(n)"
- Comp_Fun: "[|u \<sim> v|]==> Fun(u) \<sim> Fun(v)"
- Comp_App: "[|u1 \<sim> v1; u2 \<sim> v2; b1 \<in> bool; b2 \<in> bool|]
- ==> App(b1,u1,u2) \<sim> App(b2,v1,v2)"
+ Comp_Var: "n \<in> nat \<Longrightarrow> Var(n) \<sim> Var(n)"
+ Comp_Fun: "\<lbrakk>u \<sim> v\<rbrakk>\<Longrightarrow> Fun(u) \<sim> Fun(v)"
+ Comp_App: "\<lbrakk>u1 \<sim> v1; u2 \<sim> v2; b1 \<in> bool; b2 \<in> bool\<rbrakk>
+ \<Longrightarrow> App(b1,u1,u2) \<sim> App(b2,v1,v2)"
type_intros redexes.intros bool_typechecks
inductive
domains "Sreg" \<subseteq> redexes
intros
- Reg_Var: "n \<in> nat ==> regular(Var(n))"
- Reg_Fun: "[|regular(u)|]==> regular(Fun(u))"
- Reg_App1: "[|regular(Fun(u)); regular(v) |]==>regular(App(1,Fun(u),v))"
- Reg_App2: "[|regular(u); regular(v) |]==>regular(App(0,u,v))"
+ Reg_Var: "n \<in> nat \<Longrightarrow> regular(Var(n))"
+ Reg_Fun: "\<lbrakk>regular(u)\<rbrakk>\<Longrightarrow> regular(Fun(u))"
+ Reg_App1: "\<lbrakk>regular(Fun(u)); regular(v)\<rbrakk>\<Longrightarrow>regular(App(1,Fun(u),v))"
+ Reg_App2: "\<lbrakk>regular(u); regular(v)\<rbrakk>\<Longrightarrow>regular(App(0,u,v))"
type_intros redexes.intros bool_typechecks
@@ -90,15 +90,15 @@
(* Equality rules for union *)
(* ------------------------------------------------------------------------- *)
-lemma union_Var [simp]: "n \<in> nat ==> Var(n) \<squnion> Var(n)=Var(n)"
+lemma union_Var [simp]: "n \<in> nat \<Longrightarrow> Var(n) \<squnion> Var(n)=Var(n)"
by (simp add: union_def)
-lemma union_Fun [simp]: "v \<in> redexes ==> Fun(u) \<squnion> Fun(v) = Fun(u \<squnion> v)"
+lemma union_Fun [simp]: "v \<in> redexes \<Longrightarrow> Fun(u) \<squnion> Fun(v) = Fun(u \<squnion> v)"
by (simp add: union_def)
lemma union_App [simp]:
- "[|b2 \<in> bool; u2 \<in> redexes; v2 \<in> redexes|]
- ==> App(b1,u1,v1) \<squnion> App(b2,u2,v2)=App(b1 or b2,u1 \<squnion> u2,v1 \<squnion> v2)"
+ "\<lbrakk>b2 \<in> bool; u2 \<in> redexes; v2 \<in> redexes\<rbrakk>
+ \<Longrightarrow> App(b1,u1,v1) \<squnion> App(b2,u2,v2)=App(b1 or b2,u1 \<squnion> u2,v1 \<squnion> v2)"
by (simp add: union_def)
@@ -135,16 +135,16 @@
(* comp proofs *)
(* ------------------------------------------------------------------------- *)
-lemma comp_refl [simp]: "u \<in> redexes ==> u \<sim> u"
+lemma comp_refl [simp]: "u \<in> redexes \<Longrightarrow> u \<sim> u"
by (erule redexes.induct, blast+)
-lemma comp_sym: "u \<sim> v ==> v \<sim> u"
+lemma comp_sym: "u \<sim> v \<Longrightarrow> v \<sim> u"
by (erule Scomp.induct, blast+)
lemma comp_sym_iff: "u \<sim> v \<longleftrightarrow> v \<sim> u"
by (blast intro: comp_sym)
-lemma comp_trans [rule_format]: "u \<sim> v ==> \<forall>w. v \<sim> w\<longrightarrow>u \<sim> w"
+lemma comp_trans [rule_format]: "u \<sim> v \<Longrightarrow> \<forall>w. v \<sim> w\<longrightarrow>u \<sim> w"
by (erule Scomp.induct, blast+)
(* ------------------------------------------------------------------------- *)
--- a/src/ZF/Resid/Reduction.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Resid/Reduction.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,14 +13,14 @@
abbreviation
Apl :: "[i,i]=>i" where
- "Apl(n,m) == App(0,n,m)"
+ "Apl(n,m) \<equiv> App(0,n,m)"
inductive
domains "lambda" \<subseteq> redexes
intros
- Lambda_Var: " n \<in> nat ==> Var(n) \<in> lambda"
- Lambda_Fun: " u \<in> lambda ==> Fun(u) \<in> lambda"
- Lambda_App: "[|u \<in> lambda; v \<in> lambda|] ==> Apl(u,v) \<in> lambda"
+ Lambda_Var: " n \<in> nat \<Longrightarrow> Var(n) \<in> lambda"
+ Lambda_Fun: " u \<in> lambda \<Longrightarrow> Fun(u) \<in> lambda"
+ Lambda_App: "\<lbrakk>u \<in> lambda; v \<in> lambda\<rbrakk> \<Longrightarrow> Apl(u,v) \<in> lambda"
type_intros redexes.intros bool_typechecks
declare lambda.intros [intro]
@@ -40,10 +40,10 @@
(* ------------------------------------------------------------------------- *)
lemma unmark_type [intro, simp]:
- "u \<in> redexes ==> unmark(u) \<in> lambda"
+ "u \<in> redexes \<Longrightarrow> unmark(u) \<in> lambda"
by (erule redexes.induct, simp_all)
-lemma lambda_unmark: "u \<in> lambda ==> unmark(u) = u"
+lemma lambda_unmark: "u \<in> lambda \<Longrightarrow> unmark(u) = u"
by (erule lambda.induct, simp_all)
@@ -52,11 +52,11 @@
(* ------------------------------------------------------------------------- *)
lemma liftL_type [rule_format]:
- "v \<in> lambda ==> \<forall>k \<in> nat. lift_rec(v,k) \<in> lambda"
+ "v \<in> lambda \<Longrightarrow> \<forall>k \<in> nat. lift_rec(v,k) \<in> lambda"
by (erule lambda.induct, simp_all add: lift_rec_Var)
lemma substL_type [rule_format, simp]:
- "v \<in> lambda ==> \<forall>n \<in> nat. \<forall>u \<in> lambda. subst_rec(u,v,n) \<in> lambda"
+ "v \<in> lambda \<Longrightarrow> \<forall>n \<in> nat. \<forall>u \<in> lambda. subst_rec(u,v,n) \<in> lambda"
by (erule lambda.induct, simp_all add: liftL_type subst_Var)
@@ -75,28 +75,28 @@
abbreviation
Sred1_rel (infixl \<open>-1->\<close> 50) where
- "a -1-> b == <a,b> \<in> Sred1"
+ "a -1-> b \<equiv> <a,b> \<in> Sred1"
abbreviation
Sred_rel (infixl \<open>-\<longrightarrow>\<close> 50) where
- "a -\<longrightarrow> b == <a,b> \<in> Sred"
+ "a -\<longrightarrow> b \<equiv> <a,b> \<in> Sred"
abbreviation
Spar_red1_rel (infixl \<open>=1=>\<close> 50) where
- "a =1=> b == <a,b> \<in> Spar_red1"
+ "a =1=> b \<equiv> <a,b> \<in> Spar_red1"
abbreviation
- Spar_red_rel (infixl \<open>===>\<close> 50) where
- "a ===> b == <a,b> \<in> Spar_red"
+ Spar_red_rel (infixl \<open>=\<Longrightarrow>\<close> 50) where
+ "a =\<Longrightarrow> b \<equiv> <a,b> \<in> Spar_red"
inductive
domains "Sred1" \<subseteq> "lambda*lambda"
intros
- beta: "[|m \<in> lambda; n \<in> lambda|] ==> Apl(Fun(m),n) -1-> n/m"
- rfun: "[|m -1-> n|] ==> Fun(m) -1-> Fun(n)"
- apl_l: "[|m2 \<in> lambda; m1 -1-> n1|] ==> Apl(m1,m2) -1-> Apl(n1,m2)"
- apl_r: "[|m1 \<in> lambda; m2 -1-> n2|] ==> Apl(m1,m2) -1-> Apl(m1,n2)"
+ beta: "\<lbrakk>m \<in> lambda; n \<in> lambda\<rbrakk> \<Longrightarrow> Apl(Fun(m),n) -1-> n/m"
+ rfun: "\<lbrakk>m -1-> n\<rbrakk> \<Longrightarrow> Fun(m) -1-> Fun(n)"
+ apl_l: "\<lbrakk>m2 \<in> lambda; m1 -1-> n1\<rbrakk> \<Longrightarrow> Apl(m1,m2) -1-> Apl(n1,m2)"
+ apl_r: "\<lbrakk>m1 \<in> lambda; m2 -1-> n2\<rbrakk> \<Longrightarrow> Apl(m1,m2) -1-> Apl(m1,n2)"
type_intros red_typechecks
declare Sred1.intros [intro, simp]
@@ -104,9 +104,9 @@
inductive
domains "Sred" \<subseteq> "lambda*lambda"
intros
- one_step: "m-1->n ==> m-\<longrightarrow>n"
- refl: "m \<in> lambda==>m -\<longrightarrow>m"
- trans: "[|m-\<longrightarrow>n; n-\<longrightarrow>p|] ==>m-\<longrightarrow>p"
+ one_step: "m-1->n \<Longrightarrow> m-\<longrightarrow>n"
+ refl: "m \<in> lambda\<Longrightarrow>m -\<longrightarrow>m"
+ trans: "\<lbrakk>m-\<longrightarrow>n; n-\<longrightarrow>p\<rbrakk> \<Longrightarrow>m-\<longrightarrow>p"
type_intros Sred1.dom_subset [THEN subsetD] red_typechecks
declare Sred.one_step [intro, simp]
@@ -115,10 +115,10 @@
inductive
domains "Spar_red1" \<subseteq> "lambda*lambda"
intros
- beta: "[|m =1=> m'; n =1=> n'|] ==> Apl(Fun(m),n) =1=> n'/m'"
- rvar: "n \<in> nat ==> Var(n) =1=> Var(n)"
- rfun: "m =1=> m' ==> Fun(m) =1=> Fun(m')"
- rapl: "[|m =1=> m'; n =1=> n'|] ==> Apl(m,n) =1=> Apl(m',n')"
+ beta: "\<lbrakk>m =1=> m'; n =1=> n'\<rbrakk> \<Longrightarrow> Apl(Fun(m),n) =1=> n'/m'"
+ rvar: "n \<in> nat \<Longrightarrow> Var(n) =1=> Var(n)"
+ rfun: "m =1=> m' \<Longrightarrow> Fun(m) =1=> Fun(m')"
+ rapl: "\<lbrakk>m =1=> m'; n =1=> n'\<rbrakk> \<Longrightarrow> Apl(m,n) =1=> Apl(m',n')"
type_intros red_typechecks
declare Spar_red1.intros [intro, simp]
@@ -126,8 +126,8 @@
inductive
domains "Spar_red" \<subseteq> "lambda*lambda"
intros
- one_step: "m =1=> n ==> m ===> n"
- trans: "[|m===>n; n===>p|] ==> m===>p"
+ one_step: "m =1=> n \<Longrightarrow> m =\<Longrightarrow> n"
+ trans: "\<lbrakk>m=\<Longrightarrow>n; n=\<Longrightarrow>p\<rbrakk> \<Longrightarrow> m=\<Longrightarrow>p"
type_intros Spar_red1.dom_subset [THEN subsetD] red_typechecks
declare Spar_red.one_step [intro, simp]
@@ -158,27 +158,27 @@
(* Lemmas for reduction *)
(* ------------------------------------------------------------------------- *)
-lemma red_Fun: "m-\<longrightarrow>n ==> Fun(m) -\<longrightarrow> Fun(n)"
+lemma red_Fun: "m-\<longrightarrow>n \<Longrightarrow> Fun(m) -\<longrightarrow> Fun(n)"
apply (erule Sred.induct)
apply (rule_tac [3] Sred.trans, simp_all)
done
-lemma red_Apll: "[|n \<in> lambda; m -\<longrightarrow> m'|] ==> Apl(m,n)-\<longrightarrow>Apl(m',n)"
+lemma red_Apll: "\<lbrakk>n \<in> lambda; m -\<longrightarrow> m'\<rbrakk> \<Longrightarrow> Apl(m,n)-\<longrightarrow>Apl(m',n)"
apply (erule Sred.induct)
apply (rule_tac [3] Sred.trans, simp_all)
done
-lemma red_Aplr: "[|n \<in> lambda; m -\<longrightarrow> m'|] ==> Apl(n,m)-\<longrightarrow>Apl(n,m')"
+lemma red_Aplr: "\<lbrakk>n \<in> lambda; m -\<longrightarrow> m'\<rbrakk> \<Longrightarrow> Apl(n,m)-\<longrightarrow>Apl(n,m')"
apply (erule Sred.induct)
apply (rule_tac [3] Sred.trans, simp_all)
done
-lemma red_Apl: "[|m -\<longrightarrow> m'; n-\<longrightarrow>n'|] ==> Apl(m,n)-\<longrightarrow>Apl(m',n')"
+lemma red_Apl: "\<lbrakk>m -\<longrightarrow> m'; n-\<longrightarrow>n'\<rbrakk> \<Longrightarrow> Apl(m,n)-\<longrightarrow>Apl(m',n')"
apply (rule_tac n = "Apl (m',n) " in Sred.trans)
apply (simp_all add: red_Apll red_Aplr)
done
-lemma red_beta: "[|m \<in> lambda; m':lambda; n \<in> lambda; n':lambda; m -\<longrightarrow> m'; n-\<longrightarrow>n'|] ==>
+lemma red_beta: "\<lbrakk>m \<in> lambda; m':lambda; n \<in> lambda; n':lambda; m -\<longrightarrow> m'; n-\<longrightarrow>n'\<rbrakk> \<Longrightarrow>
Apl(Fun(m),n)-\<longrightarrow> n'/m'"
apply (rule_tac n = "Apl (Fun (m'),n') " in Sred.trans)
apply (simp_all add: red_Apl red_Fun)
@@ -190,19 +190,19 @@
(* ------------------------------------------------------------------------- *)
-lemma refl_par_red1: "m \<in> lambda==> m =1=> m"
+lemma refl_par_red1: "m \<in> lambda\<Longrightarrow> m =1=> m"
by (erule lambda.induct, simp_all)
-lemma red1_par_red1: "m-1->n ==> m=1=>n"
+lemma red1_par_red1: "m-1->n \<Longrightarrow> m=1=>n"
by (erule Sred1.induct, simp_all add: refl_par_red1)
-lemma red_par_red: "m-\<longrightarrow>n ==> m===>n"
+lemma red_par_red: "m-\<longrightarrow>n \<Longrightarrow> m=\<Longrightarrow>n"
apply (erule Sred.induct)
apply (rule_tac [3] Spar_red.trans)
apply (simp_all add: refl_par_red1 red1_par_red1)
done
-lemma par_red_red: "m===>n ==> m-\<longrightarrow>n"
+lemma par_red_red: "m=\<Longrightarrow>n \<Longrightarrow> m-\<longrightarrow>n"
apply (erule Spar_red.induct)
apply (erule Spar_red1.induct)
apply (rule_tac [5] Sred.trans)
@@ -214,7 +214,7 @@
(* Simulation *)
(* ------------------------------------------------------------------------- *)
-lemma simulation: "m=1=>n ==> \<exists>v. m|>v = n & m \<sim> v & regular(v)"
+lemma simulation: "m=1=>n \<Longrightarrow> \<exists>v. m|>v = n & m \<sim> v & regular(v)"
by (erule Spar_red1.induct, force+)
@@ -223,11 +223,11 @@
(* ------------------------------------------------------------------------- *)
lemma unmmark_lift_rec:
- "u \<in> redexes ==> \<forall>k \<in> nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)"
+ "u \<in> redexes \<Longrightarrow> \<forall>k \<in> nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)"
by (erule redexes.induct, simp_all add: lift_rec_Var)
lemma unmmark_subst_rec:
- "v \<in> redexes ==> \<forall>k \<in> nat. \<forall>u \<in> redexes.
+ "v \<in> redexes \<Longrightarrow> \<forall>k \<in> nat. \<forall>u \<in> redexes.
unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)"
by (erule redexes.induct, simp_all add: unmmark_lift_rec subst_Var)
@@ -237,12 +237,12 @@
(* ------------------------------------------------------------------------- *)
lemma completeness_l [rule_format]:
- "u \<sim> v ==> regular(v) \<longrightarrow> unmark(u) =1=> unmark(u|>v)"
+ "u \<sim> v \<Longrightarrow> regular(v) \<longrightarrow> unmark(u) =1=> unmark(u|>v)"
apply (erule Scomp.induct)
apply (auto simp add: unmmark_subst_rec)
done
-lemma completeness: "[|u \<in> lambda; u \<sim> v; regular(v)|] ==> u =1=> unmark(u|>v)"
+lemma completeness: "\<lbrakk>u \<in> lambda; u \<sim> v; regular(v)\<rbrakk> \<Longrightarrow> u =1=> unmark(u|>v)"
by (drule completeness_l, simp_all add: lambda_unmark)
end
--- a/src/ZF/Resid/Residuals.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Resid/Residuals.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,25 +9,25 @@
Sres :: "i"
abbreviation
- "residuals(u,v,w) == <u,v,w> \<in> Sres"
+ "residuals(u,v,w) \<equiv> <u,v,w> \<in> Sres"
inductive
domains "Sres" \<subseteq> "redexes*redexes*redexes"
intros
- Res_Var: "n \<in> nat ==> residuals(Var(n),Var(n),Var(n))"
- Res_Fun: "[|residuals(u,v,w)|]==>
+ Res_Var: "n \<in> nat \<Longrightarrow> residuals(Var(n),Var(n),Var(n))"
+ Res_Fun: "\<lbrakk>residuals(u,v,w)\<rbrakk>\<Longrightarrow>
residuals(Fun(u),Fun(v),Fun(w))"
- Res_App: "[|residuals(u1,v1,w1);
- residuals(u2,v2,w2); b \<in> bool|]==>
+ Res_App: "\<lbrakk>residuals(u1,v1,w1);
+ residuals(u2,v2,w2); b \<in> bool\<rbrakk>\<Longrightarrow>
residuals(App(b,u1,u2),App(0,v1,v2),App(b,w1,w2))"
- Res_redex: "[|residuals(u1,v1,w1);
- residuals(u2,v2,w2); b \<in> bool|]==>
+ Res_redex: "\<lbrakk>residuals(u1,v1,w1);
+ residuals(u2,v2,w2); b \<in> bool\<rbrakk>\<Longrightarrow>
residuals(App(b,Fun(u1),u2),App(1,Fun(v1),v2),w2/w1)"
type_intros subst_type nat_typechecks redexes.intros bool_typechecks
definition
res_func :: "[i,i]=>i" (infixl \<open>|>\<close> 70) where
- "u |> v == THE w. residuals(u,v,w)"
+ "u |> v \<equiv> THE w. residuals(u,v,w)"
subsection\<open>Setting up rule lists\<close>
@@ -64,15 +64,15 @@
subsection\<open>residuals is a partial function\<close>
lemma residuals_function [rule_format]:
- "residuals(u,v,w) ==> \<forall>w1. residuals(u,v,w1) \<longrightarrow> w1 = w"
+ "residuals(u,v,w) \<Longrightarrow> \<forall>w1. residuals(u,v,w1) \<longrightarrow> w1 = w"
by (erule Sres.induct, force+)
lemma residuals_intro [rule_format]:
- "u \<sim> v ==> regular(v) \<longrightarrow> (\<exists>w. residuals(u,v,w))"
+ "u \<sim> v \<Longrightarrow> regular(v) \<longrightarrow> (\<exists>w. residuals(u,v,w))"
by (erule Scomp.induct, force+)
lemma comp_resfuncD:
- "[| u \<sim> v; regular(v) |] ==> residuals(u, v, THE w. residuals(u, v, w))"
+ "\<lbrakk>u \<sim> v; regular(v)\<rbrakk> \<Longrightarrow> residuals(u, v, THE w. residuals(u, v, w))"
apply (frule residuals_intro, assumption, clarify)
apply (subst the_equality)
apply (blast intro: residuals_function)+
@@ -80,32 +80,32 @@
subsection\<open>Residual function\<close>
-lemma res_Var [simp]: "n \<in> nat ==> Var(n) |> Var(n) = Var(n)"
+lemma res_Var [simp]: "n \<in> nat \<Longrightarrow> Var(n) |> Var(n) = Var(n)"
by (unfold res_func_def, blast)
lemma res_Fun [simp]:
- "[|s \<sim> t; regular(t)|]==> Fun(s) |> Fun(t) = Fun(s |> t)"
+ "\<lbrakk>s \<sim> t; regular(t)\<rbrakk>\<Longrightarrow> Fun(s) |> Fun(t) = Fun(s |> t)"
apply (unfold res_func_def)
apply (blast intro: comp_resfuncD residuals_function)
done
lemma res_App [simp]:
- "[|s \<sim> u; regular(u); t \<sim> v; regular(v); b \<in> bool|]
- ==> App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)"
+ "\<lbrakk>s \<sim> u; regular(u); t \<sim> v; regular(v); b \<in> bool\<rbrakk>
+ \<Longrightarrow> App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)"
apply (unfold res_func_def)
apply (blast dest!: comp_resfuncD intro: residuals_function)
done
lemma res_redex [simp]:
- "[|s \<sim> u; regular(u); t \<sim> v; regular(v); b \<in> bool|]
- ==> App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)"
+ "\<lbrakk>s \<sim> u; regular(u); t \<sim> v; regular(v); b \<in> bool\<rbrakk>
+ \<Longrightarrow> App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)"
apply (unfold res_func_def)
apply (blast elim!: redexes.free_elims dest!: comp_resfuncD
intro: residuals_function)
done
lemma resfunc_type [simp]:
- "[|s \<sim> t; regular(t)|]==> regular(t) \<longrightarrow> s |> t \<in> redexes"
+ "\<lbrakk>s \<sim> t; regular(t)\<rbrakk>\<Longrightarrow> regular(t) \<longrightarrow> s |> t \<in> redexes"
by (erule Scomp.induct, auto)
subsection\<open>Commutation theorem\<close>
@@ -117,14 +117,14 @@
"u \<Longleftarrow> v \<Longrightarrow> regular(v) \<longrightarrow> regular(u)"
by (erule Ssub.induct, auto)
-lemma residuals_lift_rec: "[|u \<sim> v; k \<in> nat|]==> regular(v)\<longrightarrow> (\<forall>n \<in> nat.
+lemma residuals_lift_rec: "\<lbrakk>u \<sim> v; k \<in> nat\<rbrakk>\<Longrightarrow> regular(v)\<longrightarrow> (\<forall>n \<in> nat.
lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))"
apply (erule Scomp.induct, safe)
apply (simp_all add: lift_rec_Var subst_Var lift_subst)
done
lemma residuals_subst_rec:
- "u1 \<sim> u2 ==> \<forall>v1 v2. v1 \<sim> v2 \<longrightarrow> regular(v2) \<longrightarrow> regular(u2) \<longrightarrow>
+ "u1 \<sim> u2 \<Longrightarrow> \<forall>v1 v2. v1 \<sim> v2 \<longrightarrow> regular(v2) \<longrightarrow> regular(u2) \<longrightarrow>
(\<forall>n \<in> nat. subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) =
subst_rec(v1 |> v2, u1 |> u2,n))"
apply (erule Scomp.induct, safe)
@@ -135,29 +135,29 @@
lemma commutation [simp]:
- "[|u1 \<sim> u2; v1 \<sim> v2; regular(u2); regular(v2)|]
- ==> (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)"
+ "\<lbrakk>u1 \<sim> u2; v1 \<sim> v2; regular(u2); regular(v2)\<rbrakk>
+ \<Longrightarrow> (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)"
by (simp add: residuals_subst_rec)
subsection\<open>Residuals are comp and regular\<close>
lemma residuals_preserve_comp [rule_format, simp]:
- "u \<sim> v ==> \<forall>w. u \<sim> w \<longrightarrow> v \<sim> w \<longrightarrow> regular(w) \<longrightarrow> (u|>w) \<sim> (v|>w)"
+ "u \<sim> v \<Longrightarrow> \<forall>w. u \<sim> w \<longrightarrow> v \<sim> w \<longrightarrow> regular(w) \<longrightarrow> (u|>w) \<sim> (v|>w)"
by (erule Scomp.induct, force+)
lemma residuals_preserve_reg [rule_format, simp]:
- "u \<sim> v ==> regular(u) \<longrightarrow> regular(v) \<longrightarrow> regular(u|>v)"
+ "u \<sim> v \<Longrightarrow> regular(u) \<longrightarrow> regular(v) \<longrightarrow> regular(u|>v)"
apply (erule Scomp.induct, auto)
done
subsection\<open>Preservation lemma\<close>
-lemma union_preserve_comp: "u \<sim> v ==> v \<sim> (u \<squnion> v)"
+lemma union_preserve_comp: "u \<sim> v \<Longrightarrow> v \<sim> (u \<squnion> v)"
by (erule Scomp.induct, simp_all)
lemma preservation [rule_format]:
- "u \<sim> v ==> regular(v) \<longrightarrow> u|>v = (u \<squnion> v)|>v"
+ "u \<sim> v \<Longrightarrow> regular(v) \<longrightarrow> u|>v = (u \<squnion> v)|>v"
apply (erule Scomp.induct, safe)
apply (drule_tac [3] psi = "Fun (u) |> v = w" for u v w in asm_rl)
apply (auto simp add: union_preserve_comp comp_sym_iff)
@@ -175,7 +175,7 @@
w |> u = (w|>v) |> (u|>v))"
by (erule Ssub.induct, force+)
-lemma prism: "[|v \<Longleftarrow> u; regular(u); w \<sim> v|] ==> w |> u = (w|>v) |> (u|>v)"
+lemma prism: "\<lbrakk>v \<Longleftarrow> u; regular(u); w \<sim> v\<rbrakk> \<Longrightarrow> w |> u = (w|>v) |> (u|>v)"
apply (rule prism_l)
apply (rule_tac [4] comp_trans, auto)
done
@@ -183,7 +183,7 @@
subsection\<open>Levy's Cube Lemma\<close>
-lemma cube: "[|u \<sim> v; regular(v); regular(u); w \<sim> u|]==>
+lemma cube: "\<lbrakk>u \<sim> v; regular(v); regular(u); w \<sim> u\<rbrakk>\<Longrightarrow>
(w|>u) |> (v|>u) = (w|>v) |> (u|>v)"
apply (subst preservation [of u], assumption, assumption)
apply (subst preservation [of v], erule comp_sym, assumption)
@@ -198,7 +198,7 @@
subsection\<open>paving theorem\<close>
-lemma paving: "[|w \<sim> u; w \<sim> v; regular(u); regular(v)|]==>
+lemma paving: "\<lbrakk>w \<sim> u; w \<sim> v; regular(u); regular(v)\<rbrakk>\<Longrightarrow>
\<exists>uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u) \<sim> vu \<and>
regular(vu) & (w|>v) \<sim> uv \<and> regular(uv)"
apply (subgoal_tac "u \<sim> v")
--- a/src/ZF/Resid/Substitution.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Resid/Substitution.thy Tue Sep 27 16:51:35 2022 +0100
@@ -18,11 +18,11 @@
definition
lift_rec :: "[i,i]=> i" where
- "lift_rec(r,k) == lift_aux(r)`k"
+ "lift_rec(r,k) \<equiv> lift_aux(r)`k"
abbreviation
lift :: "i=>i" where
- "lift(r) == lift_rec(r,0)"
+ "lift(r) \<equiv> lift_rec(r,0)"
consts
@@ -39,11 +39,11 @@
definition
subst_rec :: "[i,i,i]=> i" (**NOTE THE ARGUMENT ORDER BELOW**) where
- "subst_rec(u,r,k) == subst_aux(r)`u`k"
+ "subst_rec(u,r,k) \<equiv> subst_aux(r)`u`k"
abbreviation
subst :: "[i,i]=>i" (infixl \<open>'/\<close> 70) where
- "u/v == subst_rec(u,v,0)"
+ "u/v \<equiv> subst_rec(u,v,0)"
@@ -51,18 +51,18 @@
(* Arithmetic extensions *)
(* ------------------------------------------------------------------------- *)
-lemma gt_not_eq: "p < n ==> n\<noteq>p"
+lemma gt_not_eq: "p < n \<Longrightarrow> n\<noteq>p"
by blast
-lemma succ_pred [rule_format, simp]: "j \<in> nat ==> i < j \<longrightarrow> succ(j #- 1) = j"
+lemma succ_pred [rule_format, simp]: "j \<in> nat \<Longrightarrow> i < j \<longrightarrow> succ(j #- 1) = j"
by (induct_tac "j", auto)
-lemma lt_pred: "[|succ(x)<n; n \<in> nat|] ==> x < n #- 1 "
+lemma lt_pred: "\<lbrakk>succ(x)<n; n \<in> nat\<rbrakk> \<Longrightarrow> x < n #- 1 "
apply (rule succ_leE)
apply (simp add: succ_pred)
done
-lemma gt_pred: "[|n < succ(x); p<n; n \<in> nat|] ==> n #- 1 < x "
+lemma gt_pred: "\<lbrakk>n < succ(x); p<n; n \<in> nat\<rbrakk> \<Longrightarrow> n #- 1 < x "
apply (rule succ_leE)
apply (simp add: succ_pred)
done
@@ -75,22 +75,22 @@
(* lift_rec equality rules *)
(* ------------------------------------------------------------------------- *)
lemma lift_rec_Var:
- "n \<in> nat ==> lift_rec(Var(i),n) = (if i<n then Var(i) else Var(succ(i)))"
+ "n \<in> nat \<Longrightarrow> lift_rec(Var(i),n) = (if i<n then Var(i) else Var(succ(i)))"
by (simp add: lift_rec_def)
lemma lift_rec_le [simp]:
- "[|i \<in> nat; k\<le>i|] ==> lift_rec(Var(i),k) = Var(succ(i))"
+ "\<lbrakk>i \<in> nat; k\<le>i\<rbrakk> \<Longrightarrow> lift_rec(Var(i),k) = Var(succ(i))"
by (simp add: lift_rec_def le_in_nat)
-lemma lift_rec_gt [simp]: "[| k \<in> nat; i<k |] ==> lift_rec(Var(i),k) = Var(i)"
+lemma lift_rec_gt [simp]: "\<lbrakk>k \<in> nat; i<k\<rbrakk> \<Longrightarrow> lift_rec(Var(i),k) = Var(i)"
by (simp add: lift_rec_def)
lemma lift_rec_Fun [simp]:
- "k \<in> nat ==> lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))"
+ "k \<in> nat \<Longrightarrow> lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))"
by (simp add: lift_rec_def)
lemma lift_rec_App [simp]:
- "k \<in> nat ==> lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))"
+ "k \<in> nat \<Longrightarrow> lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))"
by (simp add: lift_rec_def)
@@ -99,43 +99,43 @@
(* ------------------------------------------------------------------------- *)
lemma subst_Var:
- "[|k \<in> nat; u \<in> redexes|]
- ==> subst_rec(u,Var(i),k) =
+ "\<lbrakk>k \<in> nat; u \<in> redexes\<rbrakk>
+ \<Longrightarrow> subst_rec(u,Var(i),k) =
(if k<i then Var(i #- 1) else if k=i then u else Var(i))"
by (simp add: subst_rec_def gt_not_eq leI)
lemma subst_eq [simp]:
- "[|n \<in> nat; u \<in> redexes|] ==> subst_rec(u,Var(n),n) = u"
+ "\<lbrakk>n \<in> nat; u \<in> redexes\<rbrakk> \<Longrightarrow> subst_rec(u,Var(n),n) = u"
by (simp add: subst_rec_def)
lemma subst_gt [simp]:
- "[|u \<in> redexes; p \<in> nat; p<n|] ==> subst_rec(u,Var(n),p) = Var(n #- 1)"
+ "\<lbrakk>u \<in> redexes; p \<in> nat; p<n\<rbrakk> \<Longrightarrow> subst_rec(u,Var(n),p) = Var(n #- 1)"
by (simp add: subst_rec_def)
lemma subst_lt [simp]:
- "[|u \<in> redexes; p \<in> nat; n<p|] ==> subst_rec(u,Var(n),p) = Var(n)"
+ "\<lbrakk>u \<in> redexes; p \<in> nat; n<p\<rbrakk> \<Longrightarrow> subst_rec(u,Var(n),p) = Var(n)"
by (simp add: subst_rec_def gt_not_eq leI lt_nat_in_nat)
lemma subst_Fun [simp]:
- "[|p \<in> nat; u \<in> redexes|]
- ==> subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) "
+ "\<lbrakk>p \<in> nat; u \<in> redexes\<rbrakk>
+ \<Longrightarrow> subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) "
by (simp add: subst_rec_def)
lemma subst_App [simp]:
- "[|p \<in> nat; u \<in> redexes|]
- ==> subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))"
+ "\<lbrakk>p \<in> nat; u \<in> redexes\<rbrakk>
+ \<Longrightarrow> subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))"
by (simp add: subst_rec_def)
lemma lift_rec_type [rule_format, simp]:
- "u \<in> redexes ==> \<forall>k \<in> nat. lift_rec(u,k) \<in> redexes"
+ "u \<in> redexes \<Longrightarrow> \<forall>k \<in> nat. lift_rec(u,k) \<in> redexes"
apply (erule redexes.induct)
apply (simp_all add: lift_rec_Var lift_rec_Fun lift_rec_App)
done
lemma subst_type [rule_format, simp]:
- "v \<in> redexes ==> \<forall>n \<in> nat. \<forall>u \<in> redexes. subst_rec(u,v,n) \<in> redexes"
+ "v \<in> redexes \<Longrightarrow> \<forall>n \<in> nat. \<forall>u \<in> redexes. subst_rec(u,v,n) \<in> redexes"
apply (erule redexes.induct)
apply (simp_all add: subst_Var lift_rec_type)
done
@@ -148,7 +148,7 @@
(*The i\<in>nat is redundant*)
lemma lift_lift_rec [rule_format]:
"u \<in> redexes
- ==> \<forall>n \<in> nat. \<forall>i \<in> nat. i\<le>n \<longrightarrow>
+ \<Longrightarrow> \<forall>n \<in> nat. \<forall>i \<in> nat. i\<le>n \<longrightarrow>
(lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))"
apply (erule redexes.induct, auto)
apply (case_tac "n < i")
@@ -157,15 +157,15 @@
done
lemma lift_lift:
- "[|u \<in> redexes; n \<in> nat|]
- ==> lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))"
+ "\<lbrakk>u \<in> redexes; n \<in> nat\<rbrakk>
+ \<Longrightarrow> lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))"
by (simp add: lift_lift_rec)
-lemma lt_not_m1_lt: "\<lbrakk>m < n; n \<in> nat; m \<in> nat\<rbrakk>\<Longrightarrow> ~ n #- 1 < m"
+lemma lt_not_m1_lt: "\<lbrakk>m < n; n \<in> nat; m \<in> nat\<rbrakk>\<Longrightarrow> \<not> n #- 1 < m"
by (erule natE, auto)
lemma lift_rec_subst_rec [rule_format]:
- "v \<in> redexes ==>
+ "v \<in> redexes \<Longrightarrow>
\<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. n\<le>m\<longrightarrow>
lift_rec(subst_rec(u,v,n),m) =
subst_rec(lift_rec(u,m),lift_rec(v,succ(m)),n)"
@@ -181,13 +181,13 @@
lemma lift_subst:
- "[|v \<in> redexes; u \<in> redexes; n \<in> nat|]
- ==> lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))"
+ "\<lbrakk>v \<in> redexes; u \<in> redexes; n \<in> nat\<rbrakk>
+ \<Longrightarrow> lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))"
by (simp add: lift_rec_subst_rec)
lemma lift_rec_subst_rec_lt [rule_format]:
- "v \<in> redexes ==>
+ "v \<in> redexes \<Longrightarrow>
\<forall>n \<in> nat. \<forall>m \<in> nat. \<forall>u \<in> redexes. m\<le>n\<longrightarrow>
lift_rec(subst_rec(u,v,n),m) =
subst_rec(lift_rec(u,m),lift_rec(v,m),succ(n))"
@@ -204,14 +204,14 @@
lemma subst_rec_lift_rec [rule_format]:
- "u \<in> redexes ==>
+ "u \<in> redexes \<Longrightarrow>
\<forall>n \<in> nat. \<forall>v \<in> redexes. subst_rec(v,lift_rec(u,n),n) = u"
apply (erule redexes.induct, auto)
apply (case_tac "n < na", auto)
done
lemma subst_rec_subst_rec [rule_format]:
- "v \<in> redexes ==>
+ "v \<in> redexes \<Longrightarrow>
\<forall>m \<in> nat. \<forall>n \<in> nat. \<forall>u \<in> redexes. \<forall>w \<in> redexes. m\<le>n \<longrightarrow>
subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m) =
subst_rec(w,subst_rec(u,v,m),n)"
@@ -237,8 +237,8 @@
lemma substitution:
- "[|v \<in> redexes; u \<in> redexes; w \<in> redexes; n \<in> nat|]
- ==> subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)"
+ "\<lbrakk>v \<in> redexes; u \<in> redexes; w \<in> redexes; n \<in> nat\<rbrakk>
+ \<Longrightarrow> subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)"
by (simp add: subst_rec_subst_rec)
@@ -249,21 +249,21 @@
lemma lift_rec_preserve_comp [rule_format, simp]:
- "u \<sim> v ==> \<forall>m \<in> nat. lift_rec(u,m) \<sim> lift_rec(v,m)"
+ "u \<sim> v \<Longrightarrow> \<forall>m \<in> nat. lift_rec(u,m) \<sim> lift_rec(v,m)"
by (erule Scomp.induct, simp_all add: comp_refl)
lemma subst_rec_preserve_comp [rule_format, simp]:
- "u2 \<sim> v2 ==> \<forall>m \<in> nat. \<forall>u1 \<in> redexes. \<forall>v1 \<in> redexes.
+ "u2 \<sim> v2 \<Longrightarrow> \<forall>m \<in> nat. \<forall>u1 \<in> redexes. \<forall>v1 \<in> redexes.
u1 \<sim> v1\<longrightarrow> subst_rec(u1,u2,m) \<sim> subst_rec(v1,v2,m)"
by (erule Scomp.induct,
simp_all add: subst_Var lift_rec_preserve_comp comp_refl)
lemma lift_rec_preserve_reg [simp]:
- "regular(u) ==> \<forall>m \<in> nat. regular(lift_rec(u,m))"
+ "regular(u) \<Longrightarrow> \<forall>m \<in> nat. regular(lift_rec(u,m))"
by (erule Sreg.induct, simp_all add: lift_rec_Var)
lemma subst_rec_preserve_reg [simp]:
- "regular(v) ==>
+ "regular(v) \<Longrightarrow>
\<forall>m \<in> nat. \<forall>u \<in> redexes. regular(u)\<longrightarrow>regular(subst_rec(u,v,m))"
by (erule Sreg.induct, simp_all add: subst_Var lift_rec_preserve_reg)
--- a/src/ZF/Sum.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Sum.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,20 +10,20 @@
text\<open>And the "Part" primitive for simultaneous recursive type definitions\<close>
definition sum :: "[i,i]=>i" (infixr \<open>+\<close> 65) where
- "A+B == {0}*A \<union> {1}*B"
+ "A+B \<equiv> {0}*A \<union> {1}*B"
definition Inl :: "i=>i" where
- "Inl(a) == <0,a>"
+ "Inl(a) \<equiv> <0,a>"
definition Inr :: "i=>i" where
- "Inr(b) == <1,b>"
+ "Inr(b) \<equiv> <1,b>"
definition "case" :: "[i=>i, i=>i, i]=>i" where
- "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"
+ "case(c,d) \<equiv> (%<y,z>. cond(y, d(z), c(z)))"
(*operator for selecting out the various summands*)
definition Part :: "[i,i=>i] => i" where
- "Part(A,h) == {x \<in> A. \<exists>z. x = h(z)}"
+ "Part(A,h) \<equiv> {x \<in> A. \<exists>z. x = h(z)}"
subsection\<open>Rules for the \<^term>\<open>Part\<close> Primitive\<close>
@@ -34,14 +34,14 @@
done
lemma Part_eqI [intro]:
- "[| a \<in> A; a=h(b) |] ==> a \<in> Part(A,h)"
+ "\<lbrakk>a \<in> A; a=h(b)\<rbrakk> \<Longrightarrow> a \<in> Part(A,h)"
by (unfold Part_def, blast)
lemmas PartI = refl [THEN [2] Part_eqI]
lemma PartE [elim!]:
- "[| a \<in> Part(A,h); !!z. [| a \<in> A; a=h(z) |] ==> P
- |] ==> P"
+ "\<lbrakk>a \<in> Part(A,h); \<And>z. \<lbrakk>a \<in> A; a=h(z)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
apply (unfold Part_def, blast)
done
@@ -60,19 +60,19 @@
(** Introduction rules for the injections **)
-lemma InlI [intro!,simp,TC]: "a \<in> A ==> Inl(a) \<in> A+B"
+lemma InlI [intro!,simp,TC]: "a \<in> A \<Longrightarrow> Inl(a) \<in> A+B"
by (unfold sum_defs, blast)
-lemma InrI [intro!,simp,TC]: "b \<in> B ==> Inr(b) \<in> A+B"
+lemma InrI [intro!,simp,TC]: "b \<in> B \<Longrightarrow> Inr(b) \<in> A+B"
by (unfold sum_defs, blast)
(** Elimination rules **)
lemma sumE [elim!]:
- "[| u \<in> A+B;
- !!x. [| x \<in> A; u=Inl(x) |] ==> P;
- !!y. [| y \<in> B; u=Inr(y) |] ==> P
- |] ==> P"
+ "\<lbrakk>u \<in> A+B;
+ \<And>x. \<lbrakk>x \<in> A; u=Inl(x)\<rbrakk> \<Longrightarrow> P;
+ \<And>y. \<lbrakk>y \<in> B; u=Inr(y)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (unfold sum_defs, blast)
(** Injection and freeness equivalences, for rewriting **)
@@ -100,10 +100,10 @@
lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE, elim!]
-lemma InlD: "Inl(a): A+B ==> a \<in> A"
+lemma InlD: "Inl(a): A+B \<Longrightarrow> a \<in> A"
by blast
-lemma InrD: "Inr(b): A+B ==> b \<in> B"
+lemma InrD: "Inr(b): A+B \<Longrightarrow> b \<in> B"
by blast
lemma sum_iff: "u \<in> A+B \<longleftrightarrow> (\<exists>x. x \<in> A & u=Inl(x)) | (\<exists>y. y \<in> B & u=Inr(y))"
@@ -134,26 +134,26 @@
by (simp add: sum_defs)
lemma case_type [TC]:
- "[| u \<in> A+B;
- !!x. x \<in> A ==> c(x): C(Inl(x));
- !!y. y \<in> B ==> d(y): C(Inr(y))
- |] ==> case(c,d,u) \<in> C(u)"
+ "\<lbrakk>u \<in> A+B;
+ \<And>x. x \<in> A \<Longrightarrow> c(x): C(Inl(x));
+ \<And>y. y \<in> B \<Longrightarrow> d(y): C(Inr(y))
+\<rbrakk> \<Longrightarrow> case(c,d,u) \<in> C(u)"
by auto
-lemma expand_case: "u \<in> A+B ==>
+lemma expand_case: "u \<in> A+B \<Longrightarrow>
R(case(c,d,u)) \<longleftrightarrow>
((\<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:
- "[| z \<in> A+B;
- !!x. x \<in> A ==> c(x)=c'(x);
- !!y. y \<in> B ==> d(y)=d'(y)
- |] ==> case(c,d,z) = case(c',d',z)"
+ "\<lbrakk>z \<in> A+B;
+ \<And>x. x \<in> A \<Longrightarrow> c(x)=c'(x);
+ \<And>y. y \<in> B \<Longrightarrow> d(y)=d'(y)
+\<rbrakk> \<Longrightarrow> case(c,d,z) = case(c',d',z)"
by auto
-lemma case_case: "z \<in> A+B ==>
+lemma case_case: "z \<in> A+B \<Longrightarrow>
case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =
case(%x. c(c'(x)), %y. d(d'(y)), z)"
by auto
@@ -161,7 +161,7 @@
subsection\<open>More Rules for \<^term>\<open>Part(A,h)\<close>\<close>
-lemma Part_mono: "A<=B ==> Part(A,h)<=Part(B,h)"
+lemma Part_mono: "A<=B \<Longrightarrow> Part(A,h)<=Part(B,h)"
by blast
lemma Part_Collect: "Part(Collect(A,P), h) = Collect(Part(A,h), P)"
@@ -176,7 +176,7 @@
lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y \<in> B}"
by blast
-lemma PartD1: "a \<in> Part(A,h) ==> a \<in> A"
+lemma PartD1: "a \<in> Part(A,h) \<Longrightarrow> 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 \<in> Part(B,h)}"
by blast
-lemma Part_sum_equality: "C \<subseteq> A+B ==> Part(C,Inl) \<union> Part(C,Inr) = C"
+lemma Part_sum_equality: "C \<subseteq> A+B \<Longrightarrow> Part(C,Inl) \<union> Part(C,Inr) = C"
by blast
end
--- a/src/ZF/Trancl.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Trancl.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,44 +9,44 @@
definition
refl :: "[i,i]=>o" where
- "refl(A,r) == (\<forall>x\<in>A. <x,x> \<in> r)"
+ "refl(A,r) \<equiv> (\<forall>x\<in>A. <x,x> \<in> r)"
definition
irrefl :: "[i,i]=>o" where
- "irrefl(A,r) == \<forall>x\<in>A. <x,x> \<notin> r"
+ "irrefl(A,r) \<equiv> \<forall>x\<in>A. <x,x> \<notin> r"
definition
sym :: "i=>o" where
- "sym(r) == \<forall>x y. <x,y>: r \<longrightarrow> <y,x>: r"
+ "sym(r) \<equiv> \<forall>x y. <x,y>: r \<longrightarrow> <y,x>: r"
definition
asym :: "i=>o" where
- "asym(r) == \<forall>x y. <x,y>:r \<longrightarrow> ~ <y,x>:r"
+ "asym(r) \<equiv> \<forall>x y. <x,y>:r \<longrightarrow> \<not> <y,x>:r"
definition
antisym :: "i=>o" where
- "antisym(r) == \<forall>x y.<x,y>:r \<longrightarrow> <y,x>:r \<longrightarrow> x=y"
+ "antisym(r) \<equiv> \<forall>x y.<x,y>:r \<longrightarrow> <y,x>:r \<longrightarrow> x=y"
definition
trans :: "i=>o" where
- "trans(r) == \<forall>x y z. <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
+ "trans(r) \<equiv> \<forall>x y z. <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
definition
trans_on :: "[i,i]=>o" (\<open>trans[_]'(_')\<close>) where
- "trans[A](r) == \<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A.
+ "trans[A](r) \<equiv> \<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" (\<open>(_^*)\<close> [100] 100) (*refl/transitive closure*) where
- "r^* == lfp(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
+ "r^* \<equiv> lfp(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
definition
trancl :: "i=>i" (\<open>(_^+)\<close> [100] 100) (*transitive closure*) where
- "r^+ == r O r^*"
+ "r^+ \<equiv> r O r^*"
definition
equiv :: "[i,i]=>o" where
- "equiv(A,r) == r \<subseteq> A*A & refl(A,r) & sym(r) & trans(r)"
+ "equiv(A,r) \<equiv> r \<subseteq> A*A & refl(A,r) & sym(r) & trans(r)"
subsection\<open>General properties of relations\<close>
@@ -54,43 +54,43 @@
subsubsection\<open>irreflexivity\<close>
lemma irreflI:
- "[| !!x. x \<in> A ==> <x,x> \<notin> r |] ==> irrefl(A,r)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> <x,x> \<notin> r\<rbrakk> \<Longrightarrow> irrefl(A,r)"
by (simp add: irrefl_def)
-lemma irreflE: "[| irrefl(A,r); x \<in> A |] ==> <x,x> \<notin> r"
+lemma irreflE: "\<lbrakk>irrefl(A,r); x \<in> A\<rbrakk> \<Longrightarrow> <x,x> \<notin> r"
by (simp add: irrefl_def)
subsubsection\<open>symmetry\<close>
lemma symI:
- "[| !!x y.<x,y>: r ==> <y,x>: r |] ==> sym(r)"
+ "\<lbrakk>\<And>x y.<x,y>: r \<Longrightarrow> <y,x>: r\<rbrakk> \<Longrightarrow> sym(r)"
by (unfold sym_def, blast)
-lemma symE: "[| sym(r); <x,y>: r |] ==> <y,x>: r"
+lemma symE: "\<lbrakk>sym(r); <x,y>: r\<rbrakk> \<Longrightarrow> <y,x>: r"
by (unfold sym_def, blast)
subsubsection\<open>antisymmetry\<close>
lemma antisymI:
- "[| !!x y.[| <x,y>: r; <y,x>: r |] ==> x=y |] ==> antisym(r)"
+ "\<lbrakk>\<And>x y.\<lbrakk><x,y>: r; <y,x>: r\<rbrakk> \<Longrightarrow> x=y\<rbrakk> \<Longrightarrow> antisym(r)"
by (simp add: antisym_def, blast)
-lemma antisymE: "[| antisym(r); <x,y>: r; <y,x>: r |] ==> x=y"
+lemma antisymE: "\<lbrakk>antisym(r); <x,y>: r; <y,x>: r\<rbrakk> \<Longrightarrow> x=y"
by (simp add: antisym_def, blast)
subsubsection\<open>transitivity\<close>
-lemma transD: "[| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r"
+lemma transD: "\<lbrakk>trans(r); <a,b>:r; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c>:r"
by (unfold trans_def, blast)
lemma trans_onD:
- "[| trans[A](r); <a,b>:r; <b,c>:r; a \<in> A; b \<in> A; c \<in> A |] ==> <a,c>:r"
+ "\<lbrakk>trans[A](r); <a,b>:r; <b,c>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> <a,c>:r"
by (unfold trans_on_def, blast)
-lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)"
+lemma trans_imp_trans_on: "trans(r) \<Longrightarrow> trans[A](r)"
by (unfold trans_def trans_on_def, blast)
-lemma trans_on_imp_trans: "[|trans[A](r); r \<subseteq> A*A|] ==> trans(r)"
+lemma trans_on_imp_trans: "\<lbrakk>trans[A](r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> trans(r)"
by (simp add: trans_on_def trans_def, blast)
@@ -100,7 +100,7 @@
"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^* \<subseteq> s^*"
+lemma rtrancl_mono: "r<=s \<Longrightarrow> r^* \<subseteq> s^*"
apply (unfold rtrancl_def)
apply (rule lfp_mono)
apply (rule rtrancl_bnd_mono)+
@@ -122,23 +122,23 @@
done
(*Reflexivity of rtrancl*)
-lemma rtrancl_refl: "[| a \<in> field(r) |] ==> <a,a> \<in> r^*"
+lemma rtrancl_refl: "\<lbrakk>a \<in> field(r)\<rbrakk> \<Longrightarrow> <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> \<in> r^*; <b,c> \<in> r |] ==> <a,c> \<in> r^*"
+lemma rtrancl_into_rtrancl: "\<lbrakk><a,b> \<in> r^*; <b,c> \<in> r\<rbrakk> \<Longrightarrow> <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> \<in> r ==> <a,b> \<in> r^*"
+lemma r_into_rtrancl: "<a,b> \<in> r \<Longrightarrow> <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 \<subseteq> r^*"
+lemma r_subset_rtrancl: "relation(r) \<Longrightarrow> r \<subseteq> r^*"
by (simp add: relation_def, blast intro: r_into_rtrancl)
lemma rtrancl_field: "field(r^*) = field(r)"
@@ -148,20 +148,20 @@
(** standard induction rule **)
lemma rtrancl_full_induct [case_names initial step, consumes 1]:
- "[| <a,b> \<in> r^*;
- !!x. x \<in> field(r) ==> P(<x,x>);
- !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |]
- ==> P(<a,b>)"
+ "\<lbrakk><a,b> \<in> r^*;
+ \<And>x. x \<in> field(r) \<Longrightarrow> P(<x,x>);
+ \<And>x y z.\<lbrakk>P(<x,y>); <x,y>: r^*; <y,z>: r\<rbrakk> \<Longrightarrow> P(<x,z>)\<rbrakk>
+ \<Longrightarrow> P(<a,b>)"
by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast)
(*nice induction rule.
Tried adding the typing hypotheses y,z \<in> field(r), but these
caused expensive case splits!*)
lemma rtrancl_induct [case_names initial step, induct set: rtrancl]:
- "[| <a,b> \<in> r^*;
+ "\<lbrakk><a,b> \<in> r^*;
P(a);
- !!y z.[| <a,y> \<in> r^*; <y,z> \<in> r; P(y) |] ==> P(z)
- |] ==> P(b)"
+ \<And>y z.\<lbrakk><a,y> \<in> r^*; <y,z> \<in> r; P(y)\<rbrakk> \<Longrightarrow> P(z)
+\<rbrakk> \<Longrightarrow> P(b)"
(*by induction on this formula*)
apply (subgoal_tac "\<forall>y. <a,b> = <a,y> \<longrightarrow> P (y) ")
(*now solve first subgoal: this formula is sufficient*)
@@ -170,7 +170,7 @@
apply (erule rtrancl_full_induct, blast+)
done
-(*transitivity of transitive closure!! -- by induction.*)
+(*transitivity of transitive closure\<And>-- by induction.*)
lemma trans_rtrancl: "trans(r^*)"
apply (unfold trans_def)
apply (intro allI impI)
@@ -182,9 +182,9 @@
(*elimination of rtrancl -- by induction on a special formula*)
lemma rtranclE:
- "[| <a,b> \<in> r^*; (a=b) ==> P;
- !!y.[| <a,y> \<in> r^*; <y,b> \<in> r |] ==> P |]
- ==> P"
+ "\<lbrakk><a,b> \<in> r^*; (a=b) \<Longrightarrow> P;
+ \<And>y.\<lbrakk><a,y> \<in> r^*; <y,b> \<in> r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
apply (subgoal_tac "a = b | (\<exists>y. <a,y> \<in> r^* & <y,b> \<in> r) ")
(*see HOL/trancl*)
apply blast
@@ -207,29 +207,29 @@
(** Conversions between trancl and rtrancl **)
-lemma trancl_into_rtrancl: "<a,b> \<in> r^+ ==> <a,b> \<in> r^*"
+lemma trancl_into_rtrancl: "<a,b> \<in> r^+ \<Longrightarrow> <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> \<in> r ==> <a,b> \<in> r^+"
+lemma r_into_trancl: "<a,b> \<in> r \<Longrightarrow> <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 \<subseteq> r^+"
+lemma r_subset_trancl: "relation(r) \<Longrightarrow> 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> \<in> r^*; <b,c> \<in> r |] ==> <a,c> \<in> r^+"
+lemma rtrancl_into_trancl1: "\<lbrakk><a,b> \<in> r^*; <b,c> \<in> r\<rbrakk> \<Longrightarrow> <a,c> \<in> r^+"
by (unfold trancl_def, blast)
(*intro rule from r and r^* *)
lemma rtrancl_into_trancl2:
- "[| <a,b> \<in> r; <b,c> \<in> r^* |] ==> <a,c> \<in> r^+"
+ "\<lbrakk><a,b> \<in> r; <b,c> \<in> r^*\<rbrakk> \<Longrightarrow> <a,c> \<in> r^+"
apply (erule rtrancl_induct)
apply (erule r_into_trancl)
apply (blast intro: r_into_trancl trancl_trans)
@@ -237,10 +237,10 @@
(*Nice induction rule for trancl*)
lemma trancl_induct [case_names initial step, induct set: trancl]:
- "[| <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)"
+ "\<lbrakk><a,b> \<in> r^+;
+ \<And>y. \<lbrakk><a,y> \<in> r\<rbrakk> \<Longrightarrow> P(y);
+ \<And>y z.\<lbrakk><a,y> \<in> r^+; <y,z> \<in> r; P(y)\<rbrakk> \<Longrightarrow> P(z)
+\<rbrakk> \<Longrightarrow> P(b)"
apply (rule compEpair)
apply (unfold trancl_def, assumption)
(*by induction on this formula*)
@@ -253,10 +253,10 @@
(*elimination of r^+ -- NOT an induction rule*)
lemma tranclE:
- "[| <a,b> \<in> r^+;
- <a,b> \<in> r ==> P;
- !!y.[| <a,y> \<in> r^+; <y,b> \<in> r |] ==> P
- |] ==> P"
+ "\<lbrakk><a,b> \<in> r^+;
+ <a,b> \<in> r \<Longrightarrow> P;
+ \<And>y.\<lbrakk><a,y> \<in> r^+; <y,b> \<in> r\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
apply (subgoal_tac "<a,b> \<in> r | (\<exists>y. <a,y> \<in> r^+ & <y,b> \<in> r) ")
apply blast
apply (rule compEpair)
@@ -275,13 +275,13 @@
apply (blast dest: trancl_type [THEN subsetD])
done
-lemma trancl_subset_times: "r \<subseteq> A * A ==> r^+ \<subseteq> A * A"
+lemma trancl_subset_times: "r \<subseteq> A * A \<Longrightarrow> r^+ \<subseteq> A * A"
by (insert trancl_type [of r], blast)
-lemma trancl_mono: "r<=s ==> r^+ \<subseteq> s^+"
+lemma trancl_mono: "r<=s \<Longrightarrow> r^+ \<subseteq> s^+"
by (unfold trancl_def, intro comp_mono rtrancl_mono)
-lemma trancl_eq_r: "[|relation(r); trans(r)|] ==> r^+ = r"
+lemma trancl_eq_r: "\<lbrakk>relation(r); trans(r)\<rbrakk> \<Longrightarrow> r^+ = r"
apply (rule equalityI)
prefer 2 apply (erule r_subset_trancl, clarify)
apply (frule trancl_type [THEN subsetD], clarify)
@@ -304,13 +304,13 @@
apply (blast intro: rtrancl_trans)
done
-lemma rtrancl_subset: "[| R \<subseteq> S; S \<subseteq> R^* |] ==> S^* = R^*"
+lemma rtrancl_subset: "\<lbrakk>R \<subseteq> S; S \<subseteq> R^*\<rbrakk> \<Longrightarrow> S^* = R^*"
apply (drule rtrancl_mono)
apply (drule rtrancl_mono, simp_all, blast)
done
lemma rtrancl_Un_rtrancl:
- "[| relation(r); relation(s) |] ==> (r^* \<union> s^*)^* = (r \<union> s)^*"
+ "\<lbrakk>relation(r); relation(s)\<rbrakk> \<Longrightarrow> (r^* \<union> s^*)^* = (r \<union> s)^*"
apply (rule rtrancl_subset)
apply (blast dest: r_subset_rtrancl)
apply (blast intro: rtrancl_mono [THEN subsetD])
@@ -320,7 +320,7 @@
(** rtrancl **)
-lemma rtrancl_converseD: "<x,y>:converse(r)^* ==> <x,y>:converse(r^*)"
+lemma rtrancl_converseD: "<x,y>:converse(r)^* \<Longrightarrow> <x,y>:converse(r^*)"
apply (rule converseI)
apply (frule rtrancl_type [THEN subsetD])
apply (erule rtrancl_induct)
@@ -328,7 +328,7 @@
apply (blast intro: r_into_rtrancl rtrancl_trans)
done
-lemma rtrancl_converseI: "<x,y>:converse(r^*) ==> <x,y>:converse(r)^*"
+lemma rtrancl_converseI: "<x,y>:converse(r^*) \<Longrightarrow> <x,y>:converse(r)^*"
apply (drule converseD)
apply (frule rtrancl_type [THEN subsetD])
apply (erule rtrancl_induct)
@@ -344,12 +344,12 @@
(** trancl **)
-lemma trancl_converseD: "<a, b>:converse(r)^+ ==> <a, b>:converse(r^+)"
+lemma trancl_converseD: "<a, b>:converse(r)^+ \<Longrightarrow> <a, b>:converse(r^+)"
apply (erule trancl_induct)
apply (auto intro: r_into_trancl trancl_trans)
done
-lemma trancl_converseI: "<x,y>:converse(r^+) ==> <x,y>:converse(r)^+"
+lemma trancl_converseI: "<x,y>:converse(r^+) \<Longrightarrow> <x,y>:converse(r)^+"
apply (drule converseD)
apply (erule trancl_induct)
apply (auto intro: r_into_trancl trancl_trans)
@@ -362,9 +362,9 @@
done
lemma converse_trancl_induct [case_names initial step, consumes 1]:
-"[| <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)"
+"\<lbrakk><a, b>:r^+; \<And>y. <y, b> :r \<Longrightarrow> P(y);
+ \<And>y z. \<lbrakk><y, z> \<in> r; <z, b> \<in> r^+; P(z)\<rbrakk> \<Longrightarrow> P(y)\<rbrakk>
+ \<Longrightarrow> P(a)"
apply (drule converseI)
apply (simp (no_asm_use) add: trancl_converse [symmetric])
apply (erule trancl_induct)
--- a/src/ZF/UNITY/AllocBase.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/AllocBase.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,7 +10,7 @@
abbreviation (input)
tokbag :: i (* tokbags could be multisets...or any ordered type?*)
where
- "tokbag == nat"
+ "tokbag \<equiv> nat"
axiomatization
NbT :: i and (* Number of tokens in system *)
@@ -42,15 +42,15 @@
definition
all_distinct :: "i=>o" where
- "all_distinct(l) == all_distinct0(l)=1"
+ "all_distinct(l) \<equiv> all_distinct0(l)=1"
definition
state_of :: "i =>i" \<comment> \<open>coersion from anyting to state\<close> where
- "state_of(s) == if s \<in> state then s else st0"
+ "state_of(s) \<equiv> if s \<in> state then s else st0"
definition
lift :: "i =>(i=>i)" \<comment> \<open>simplifies the expression of programs\<close> where
- "lift(x) == %s. s`x"
+ "lift(x) \<equiv> %s. s`x"
text\<open>function to show that the set of variables is infinite\<close>
consts
@@ -66,7 +66,7 @@
definition
nat_var_inj :: "i=>i" where
- "nat_var_inj(n) == Var(nat_list_inj(n))"
+ "nat_var_inj(n) \<equiv> Var(nat_list_inj(n))"
subsection\<open>Various simple lemmas\<close>
@@ -90,7 +90,7 @@
by (force simp add: INT_iff)
lemma setsum_fun_mono [rule_format]:
- "n\<in>nat ==>
+ "n\<in>nat \<Longrightarrow>
(\<forall>i\<in>nat. i<n \<longrightarrow> f(i) $\<le> g(i)) \<longrightarrow>
setsum(f, n) $\<le> setsum(g,n)"
apply (induct_tac "n", simp_all)
@@ -102,17 +102,17 @@
apply (rule zadd_zle_mono, simp_all)
done
-lemma tokens_type [simp,TC]: "l\<in>list(A) ==> tokens(l)\<in>nat"
+lemma tokens_type [simp,TC]: "l\<in>list(A) \<Longrightarrow> tokens(l)\<in>nat"
by (erule list.induct, auto)
lemma tokens_mono_aux [rule_format]:
- "xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)
+ "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)
\<longrightarrow> tokens(xs) \<le> tokens(ys)"
apply (induct_tac "xs")
apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def)
done
-lemma tokens_mono: "<xs, ys>\<in>prefix(A) ==> tokens(xs) \<le> tokens(ys)"
+lemma tokens_mono: "<xs, ys>\<in>prefix(A) \<Longrightarrow> tokens(xs) \<le> tokens(ys)"
apply (cut_tac prefix_type)
apply (blast intro: tokens_mono_aux)
done
@@ -123,31 +123,31 @@
done
lemma tokens_append [simp]:
-"[| xs\<in>list(A); ys\<in>list(A) |] ==> tokens(xs@ys) = tokens(xs) #+ tokens(ys)"
+"\<lbrakk>xs\<in>list(A); ys\<in>list(A)\<rbrakk> \<Longrightarrow> tokens(xs@ys) = tokens(xs) #+ tokens(ys)"
apply (induct_tac "xs", auto)
done
subsection\<open>The function \<^term>\<open>bag_of\<close>\<close>
-lemma bag_of_type [simp,TC]: "l\<in>list(A) ==>bag_of(l)\<in>Mult(A)"
+lemma bag_of_type [simp,TC]: "l\<in>list(A) \<Longrightarrow>bag_of(l)\<in>Mult(A)"
apply (induct_tac "l")
apply (auto simp add: Mult_iff_multiset)
done
lemma bag_of_multiset:
- "l\<in>list(A) ==> multiset(bag_of(l)) & mset_of(bag_of(l))<=A"
+ "l\<in>list(A) \<Longrightarrow> multiset(bag_of(l)) & mset_of(bag_of(l))<=A"
apply (drule bag_of_type)
apply (auto simp add: Mult_iff_multiset)
done
lemma bag_of_append [simp]:
- "[| xs\<in>list(A); ys\<in>list(A)|] ==> bag_of(xs@ys) = bag_of(xs) +# bag_of(ys)"
+ "\<lbrakk>xs\<in>list(A); ys\<in>list(A)\<rbrakk> \<Longrightarrow> bag_of(xs@ys) = bag_of(xs) +# bag_of(ys)"
apply (induct_tac "xs")
apply (auto simp add: bag_of_multiset munion_assoc)
done
lemma bag_of_mono_aux [rule_format]:
- "xs\<in>list(A) ==> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)
+ "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)
\<longrightarrow> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
apply (induct_tac "xs", simp_all, clarify)
apply (frule_tac l = ys in bag_of_multiset)
@@ -158,8 +158,8 @@
done
lemma bag_of_mono [intro]:
- "[| <xs, ys>\<in>prefix(A); xs\<in>list(A); ys\<in>list(A) |]
- ==> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
+ "\<lbrakk><xs, ys>\<in>prefix(A); xs\<in>list(A); ys\<in>list(A)\<rbrakk>
+ \<Longrightarrow> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
apply (blast intro: bag_of_mono_aux)
done
@@ -172,7 +172,7 @@
lemmas nat_into_Fin = eqpoll_refl [THEN [2] Fin_lemma]
-lemma list_Int_length_Fin: "l \<in> list(A) ==> C \<inter> length(l) \<in> Fin(length(l))"
+lemma list_Int_length_Fin: "l \<in> list(A) \<Longrightarrow> C \<inter> length(l) \<in> Fin(length(l))"
apply (drule length_type)
apply (rule Fin_subset)
apply (rule Int_lower2)
@@ -182,7 +182,7 @@
lemma mem_Int_imp_lt_length:
- "[|xs \<in> list(A); k \<in> C \<inter> length(xs)|] ==> k < length(xs)"
+ "\<lbrakk>xs \<in> list(A); k \<in> C \<inter> length(xs)\<rbrakk> \<Longrightarrow> k < length(xs)"
by (simp add: ltI)
lemma Int_succ_right:
@@ -191,8 +191,8 @@
lemma bag_of_sublist_lemma:
- "[|C \<subseteq> nat; x \<in> A; xs \<in> list(A)|]
- ==> msetsum(\<lambda>i. {#nth(i, xs @ [x])#}, C \<inter> succ(length(xs)), A) =
+ "\<lbrakk>C \<subseteq> nat; x \<in> A; xs \<in> list(A)\<rbrakk>
+ \<Longrightarrow> msetsum(\<lambda>i. {#nth(i, xs @ [x])#}, C \<inter> succ(length(xs)), A) =
(if length(xs) \<in> C then
{#x#} +# msetsum(\<lambda>x. {#nth(x, xs)#}, C \<inter> length(xs), A)
else msetsum(\<lambda>x. {#nth(x, xs)#}, C \<inter> length(xs), A))"
@@ -209,8 +209,8 @@
done
lemma bag_of_sublist_lemma2:
- "l\<in>list(A) ==>
- C \<subseteq> nat ==>
+ "l\<in>list(A) \<Longrightarrow>
+ C \<subseteq> nat \<Longrightarrow>
bag_of(sublist(l, C)) =
msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
apply (erule list_append_induct)
@@ -219,14 +219,14 @@
done
-lemma nat_Int_length_eq: "l \<in> list(A) ==> nat \<inter> length(l) = length(l)"
+lemma nat_Int_length_eq: "l \<in> list(A) \<Longrightarrow> nat \<inter> length(l) = length(l)"
apply (rule Int_absorb1)
apply (rule OrdmemD, auto)
done
(*eliminating the assumption C<=nat*)
lemma bag_of_sublist:
- "l\<in>list(A) ==>
+ "l\<in>list(A) \<Longrightarrow>
bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
apply (subgoal_tac " bag_of (sublist (l, C \<inter> nat)) = msetsum (%i. {#nth (i, l) #}, C \<inter> length (l), A) ")
apply (simp add: sublist_Int_eq)
@@ -234,7 +234,7 @@
done
lemma bag_of_sublist_Un_Int:
-"l\<in>list(A) ==>
+"l\<in>list(A) \<Longrightarrow>
bag_of(sublist(l, B \<union> C)) +# bag_of(sublist(l, B \<inter> C)) =
bag_of(sublist(l, B)) +# bag_of(sublist(l, C))"
apply (subgoal_tac "B \<inter> C \<inter> length (l) = (B \<inter> length (l)) \<inter> (C \<inter> length (l))")
@@ -247,16 +247,16 @@
lemma bag_of_sublist_Un_disjoint:
- "[| l\<in>list(A); B \<inter> C = 0 |]
- ==> bag_of(sublist(l, B \<union> C)) =
+ "\<lbrakk>l\<in>list(A); B \<inter> C = 0\<rbrakk>
+ \<Longrightarrow> bag_of(sublist(l, B \<union> C)) =
bag_of(sublist(l, B)) +# bag_of(sublist(l, C))"
by (simp add: bag_of_sublist_Un_Int [symmetric] bag_of_multiset)
lemma bag_of_sublist_UN_disjoint [rule_format]:
- "[|Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> A(i) \<inter> A(j) = 0;
- l\<in>list(B) |]
- ==> bag_of(sublist(l, \<Union>i\<in>I. A(i))) =
+ "\<lbrakk>Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> A(i) \<inter> A(j) = 0;
+ l\<in>list(B)\<rbrakk>
+ \<Longrightarrow> bag_of(sublist(l, \<Union>i\<in>I. A(i))) =
(msetsum(%i. bag_of(sublist(l, A(i))), I, B)) "
apply (simp (no_asm_simp) del: UN_simps
add: UN_simps [symmetric] bag_of_sublist)
@@ -286,7 +286,7 @@
subsubsection\<open>The function \<^term>\<open>state_of\<close>\<close>
-lemma state_of_state: "s\<in>state ==> state_of(s)=s"
+lemma state_of_state: "s\<in>state \<Longrightarrow> state_of(s)=s"
by (unfold state_of_def, auto)
declare state_of_state [simp]
@@ -328,10 +328,10 @@
lemmas Follows_state_ofD2 =
Follows_state_of_eq [THEN equalityD2, THEN subsetD]
-lemma nat_list_inj_type: "n\<in>nat ==> nat_list_inj(n)\<in>list(nat)"
+lemma nat_list_inj_type: "n\<in>nat \<Longrightarrow> nat_list_inj(n)\<in>list(nat)"
by (induct_tac "n", auto)
-lemma length_nat_list_inj: "n\<in>nat ==> length(nat_list_inj(n)) = n"
+lemma length_nat_list_inj: "n\<in>nat \<Longrightarrow> length(nat_list_inj(n)) = n"
by (induct_tac "n", auto)
lemma var_infinite_lemma:
@@ -348,20 +348,20 @@
apply (rule var_infinite_lemma)
done
-lemma var_not_Finite: "~Finite(var)"
+lemma var_not_Finite: "\<not>Finite(var)"
apply (insert nat_not_Finite)
apply (blast intro: lepoll_Finite [OF nat_lepoll_var])
done
-lemma not_Finite_imp_exist: "~Finite(A) ==> \<exists>x. x\<in>A"
+lemma not_Finite_imp_exist: "\<not>Finite(A) \<Longrightarrow> \<exists>x. x\<in>A"
apply (subgoal_tac "A\<noteq>0")
apply (auto simp add: Finite_0)
done
lemma Inter_Diff_var_iff:
- "Finite(A) ==> b\<in>(\<Inter>(RepFun(var-A, B))) \<longleftrightarrow> (\<forall>x\<in>var-A. b\<in>B(x))"
+ "Finite(A) \<Longrightarrow> b\<in>(\<Inter>(RepFun(var-A, B))) \<longleftrightarrow> (\<forall>x\<in>var-A. b\<in>B(x))"
apply (subgoal_tac "\<exists>x. x\<in>var-A", auto)
-apply (subgoal_tac "~Finite (var-A) ")
+apply (subgoal_tac "\<not>Finite (var-A) ")
apply (drule not_Finite_imp_exist, auto)
apply (cut_tac var_not_Finite)
apply (erule swap)
@@ -369,10 +369,10 @@
done
lemma Inter_var_DiffD:
- "[| b\<in>\<Inter>(RepFun(var-A, B)); Finite(A); x\<in>var-A |] ==> b\<in>B(x)"
+ "\<lbrakk>b\<in>\<Inter>(RepFun(var-A, B)); Finite(A); x\<in>var-A\<rbrakk> \<Longrightarrow> b\<in>B(x)"
by (simp add: Inter_Diff_var_iff)
-(* [| Finite(A); (\<forall>x\<in>var-A. b\<in>B(x)) |] ==> b\<in>\<Inter>(RepFun(var-A, B)) *)
+(* \<lbrakk>Finite(A); (\<forall>x\<in>var-A. b\<in>B(x))\<rbrakk> \<Longrightarrow> b\<in>\<Inter>(RepFun(var-A, B)) *)
lemmas Inter_var_DiffI = Inter_Diff_var_iff [THEN iffD2]
declare Inter_var_DiffI [intro!]
@@ -382,15 +382,15 @@
by (insert Acts_type [of F], auto)
lemma setsum_nsetsum_eq:
- "[| Finite(A); \<forall>x\<in>A. g(x)\<in>nat |]
- ==> setsum(%x. $#(g(x)), A) = $# nsetsum(%x. g(x), A)"
+ "\<lbrakk>Finite(A); \<forall>x\<in>A. g(x)\<in>nat\<rbrakk>
+ \<Longrightarrow> setsum(%x. $#(g(x)), A) = $# nsetsum(%x. g(x), A)"
apply (erule Finite_induct)
apply (auto simp add: int_of_add)
done
lemma nsetsum_cong:
- "[| A=B; \<forall>x\<in>A. f(x)=g(x); \<forall>x\<in>A. g(x)\<in>nat; Finite(A) |]
- ==> nsetsum(f, A) = nsetsum(g, B)"
+ "\<lbrakk>A=B; \<forall>x\<in>A. f(x)=g(x); \<forall>x\<in>A. g(x)\<in>nat; Finite(A)\<rbrakk>
+ \<Longrightarrow> nsetsum(f, A) = nsetsum(g, B)"
apply (subgoal_tac "$# nsetsum (f, A) = $# nsetsum (g, B)", simp)
apply (simp add: setsum_nsetsum_eq [symmetric] cong: setsum_cong)
done
--- a/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,11 +10,11 @@
abbreviation
NbR :: i (*number of consumed messages*) where
- "NbR == Var([succ(2)])"
+ "NbR \<equiv> Var([succ(2)])"
abbreviation
available_tok :: i (*number of free tokens (T in paper)*) where
- "available_tok == Var([succ(succ(2))])"
+ "available_tok \<equiv> Var([succ(succ(2))])"
axiomatization where
alloc_type_assumes [simp]:
@@ -24,7 +24,7 @@
"default_val(NbR) = 0 & default_val(available_tok)=0"
definition
- "alloc_giv_act ==
+ "alloc_giv_act \<equiv>
{<s, t> \<in> state*state.
\<exists>k. k = length(s`giv) &
t = s(giv := s`giv @ [nth(k, s`ask)],
@@ -32,7 +32,7 @@
k < length(s`ask) & nth(k, s`ask) \<le> s`available_tok}"
definition
- "alloc_rel_act ==
+ "alloc_rel_act \<equiv>
{<s, t> \<in> state*state.
t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
NbR := succ(s`NbR)) &
@@ -41,7 +41,7 @@
definition
(*The initial condition s`giv=[] is missing from the
original definition: S. O. Ehmety *)
- "alloc_prog ==
+ "alloc_prog \<equiv>
mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil},
{alloc_giv_act, alloc_rel_act},
\<Union>G \<in> preserves(lift(available_tok)) \<inter>
@@ -49,12 +49,12 @@
preserves(lift(giv)). Acts(G))"
-lemma available_tok_value_type [simp,TC]: "s\<in>state ==> s`available_tok \<in> nat"
+lemma available_tok_value_type [simp,TC]: "s\<in>state \<Longrightarrow> s`available_tok \<in> nat"
apply (unfold state_def)
apply (drule_tac a = available_tok in apply_type, auto)
done
-lemma NbR_value_type [simp,TC]: "s\<in>state ==> s`NbR \<in> nat"
+lemma NbR_value_type [simp,TC]: "s\<in>state \<Longrightarrow> s`NbR \<in> nat"
apply (unfold state_def)
apply (drule_tac a = NbR in apply_type, auto)
done
@@ -141,8 +141,8 @@
done
lemma giv_Bounded_lemma2:
-"[| G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel)) |]
- ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state. s`NbR \<le> length(s`rel)} \<inter>
+"\<lbrakk>G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel))\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in> Stable({s\<in>state. s`NbR \<le> length(s`rel)} \<inter>
{s\<in>state. s`available_tok #+ tokens(s`giv) =
NbT #+ tokens(take(s`NbR, s`rel))})"
apply (cut_tac giv_Bounded_lamma1)
@@ -214,8 +214,8 @@
subsubsection\<open>First, we lead up to a proof of Lemma 49, page 28.\<close>
lemma alloc_prog_transient_lemma:
- "[|G \<in> program; k\<in>nat|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>G \<in> program; k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
transient({s\<in>state. k \<le> length(s`rel)} \<inter>
{s\<in>state. succ(s`NbR) = k})"
apply auto
@@ -232,8 +232,8 @@
done
lemma alloc_prog_rel_Stable_NbR_lemma:
- "[| G \<in> program; alloc_prog ok G; k\<in>nat |]
- ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> succ(s ` NbR)})"
+ "\<lbrakk>G \<in> program; alloc_prog ok G; k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> succ(s ` NbR)})"
apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety, auto)
apply (blast intro: le_trans leI)
apply (drule_tac f = "lift (NbR)" and A = nat in preserves_imp_increasing)
@@ -244,9 +244,9 @@
done
lemma alloc_prog_NbR_LeadsTo_lemma:
- "[| G \<in> program; alloc_prog ok G;
- alloc_prog \<squnion> G \<in> Incr(lift(rel)); k\<in>nat |]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>G \<in> program; alloc_prog ok G;
+ alloc_prog \<squnion> G \<in> Incr(lift(rel)); k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. k \<le> length(s`rel)} \<inter> {s\<in>state. succ(s`NbR) = k}
\<longmapsto>w {s\<in>state. k \<le> s`NbR}"
apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k \<le> length (s`rel)})")
@@ -264,9 +264,9 @@
done
lemma alloc_prog_NbR_LeadsTo_lemma2 [rule_format]:
- "[| G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
- k\<in>nat; n \<in> nat; n < k |]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
+ k\<in>nat; n \<in> nat; n < k\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state . k \<le> length(s ` rel)} \<inter> {s\<in>state . s ` NbR = n}
\<longmapsto>w {x \<in> state. k \<le> length(x`rel)} \<inter>
(\<Union>m \<in> greater_than(n). {x \<in> state. x ` NbR=m})"
@@ -292,15 +292,15 @@
apply (blast intro: lt_trans2)
done
-lemma Collect_vimage_eq: "u\<in>nat ==> {<s,f(s)>. s \<in> A} -`` u = {s\<in>A. f(s) < u}"
+lemma Collect_vimage_eq: "u\<in>nat \<Longrightarrow> {<s,f(s)>. s \<in> A} -`` u = {s\<in>A. f(s) < u}"
by (force simp add: lt_def)
(* Lemma 49, page 28 *)
lemma alloc_prog_NbR_LeadsTo_lemma3:
- "[|G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
- k\<in>nat|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>G \<in> program; alloc_prog ok G; alloc_prog \<squnion> G \<in> Incr(lift(rel));
+ k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. k \<le> length(s`rel)} \<longmapsto>w {s\<in>state. k \<le> s`NbR}"
(* Proof by induction over the difference between k and n *)
apply (rule_tac f = "\<lambda>s\<in>state. k #- s`NbR" in LessThan_induct)
@@ -322,12 +322,12 @@
subsubsection\<open>Towards proving lemma 50, page 29\<close>
lemma alloc_prog_giv_Ensures_lemma:
-"[| G \<in> program; k\<in>nat; alloc_prog ok G;
- alloc_prog \<squnion> G \<in> Incr(lift(ask)) |] ==>
+"\<lbrakk>G \<in> program; k\<in>nat; alloc_prog ok G;
+ alloc_prog \<squnion> G \<in> Incr(lift(ask))\<rbrakk> \<Longrightarrow>
alloc_prog \<squnion> G \<in>
{s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
{s\<in>state. k < length(s`ask)} \<inter> {s\<in>state. length(s`giv)=k}
- Ensures {s\<in>state. ~ k <length(s`ask)} \<union> {s\<in>state. length(s`giv) \<noteq> k}"
+ Ensures {s\<in>state. \<not> k <length(s`ask)} \<union> {s\<in>state. length(s`giv) \<noteq> k}"
apply (rule EnsuresI, auto)
apply (erule_tac [2] V = "G\<notin>u" for u in thin_rl)
apply (rule_tac [2] act = alloc_giv_act in transientI)
@@ -339,7 +339,7 @@
apply (rule_tac [2] ReplaceI)
apply (rule_tac [2] x = "x (giv := x ` giv @ [nth (length(x`giv), x ` ask) ], available_tok := x ` available_tok #- nth (length(x`giv), x ` ask))" in exI)
apply (auto intro!: state_update_type simp add: app_type)
-apply (rule_tac A = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<inter> {s\<in>state . k < length(s ` ask) } \<inter> {s\<in>state. length(s`giv) =k}" and A' = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<union> {s\<in>state. ~ k < length(s`ask) } \<union> {s\<in>state . length(s ` giv) \<noteq> k}" in Constrains_weaken)
+apply (rule_tac A = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<inter> {s\<in>state . k < length(s ` ask) } \<inter> {s\<in>state. length(s`giv) =k}" and A' = "{s\<in>state . nth (length(s ` giv), s ` ask) \<le> s ` available_tok} \<union> {s\<in>state. \<not> k < length(s`ask) } \<union> {s\<in>state . length(s ` giv) \<noteq> k}" in Constrains_weaken)
apply (auto dest: ActsD simp add: Constrains_def constrains_def alloc_prog_def [THEN def_prg_Acts] alloc_prog_ok_iff)
apply (subgoal_tac "length(xa ` giv @ [nth (length(xa ` giv), xa ` ask) ]) = length(xa ` giv) #+ 1")
apply (rule_tac [2] trans)
@@ -356,8 +356,8 @@
done
lemma alloc_prog_giv_Stable_lemma:
-"[| G \<in> program; alloc_prog ok G; k\<in>nat |]
- ==> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> length(s`giv)})"
+"\<lbrakk>G \<in> program; alloc_prog ok G; k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in> Stable({s\<in>state . k \<le> length(s`giv)})"
apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety)
apply (auto intro: leI)
apply (drule_tac f = "lift (giv)" and g = length in imp_preserves_comp)
@@ -370,14 +370,14 @@
(* Lemma 50, page 29 *)
lemma alloc_prog_giv_LeadsTo_lemma:
-"[| G \<in> program; alloc_prog ok G;
- alloc_prog \<squnion> G \<in> Incr(lift(ask)); k\<in>nat |]
- ==> alloc_prog \<squnion> G \<in>
+"\<lbrakk>G \<in> program; alloc_prog ok G;
+ alloc_prog \<squnion> G \<in> Incr(lift(ask)); k\<in>nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. nth(length(s`giv), s`ask) \<le> s`available_tok} \<inter>
{s\<in>state. k < length(s`ask)} \<inter>
{s\<in>state. length(s`giv) = k}
\<longmapsto>w {s\<in>state. k < length(s`giv)}"
-apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} \<longmapsto>w {s\<in>state. ~ k <length(s`ask) } \<union> {s\<in>state. length(s`giv) \<noteq> k}")
+apply (subgoal_tac "alloc_prog \<squnion> G \<in> {s\<in>state. nth (length(s`giv), s`ask) \<le> s`available_tok} \<inter> {s\<in>state. k < length(s`ask) } \<inter> {s\<in>state. length(s`giv) = k} \<longmapsto>w {s\<in>state. \<not> k <length(s`ask) } \<union> {s\<in>state. length(s`giv) \<noteq> k}")
prefer 2 apply (blast intro: alloc_prog_giv_Ensures_lemma [THEN LeadsTo_Basis])
apply (subgoal_tac "alloc_prog \<squnion> G \<in> Stable ({s\<in>state. k < length(s`ask) }) ")
apply (drule PSP_Stable, assumption)
@@ -397,10 +397,10 @@
tokens given does not exceed the number returned, then the upper limit
(\<^term>\<open>NbT\<close>) does not exceed the number currently available.\<close>
lemma alloc_prog_Always_lemma:
-"[| G \<in> program; alloc_prog ok G;
+"\<lbrakk>G \<in> program; alloc_prog ok G;
alloc_prog \<squnion> G \<in> Incr(lift(ask));
- alloc_prog \<squnion> G \<in> Incr(lift(rel)) |]
- ==> alloc_prog \<squnion> G \<in>
+ alloc_prog \<squnion> G \<in> Incr(lift(rel))\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
Always({s\<in>state. tokens(s`giv) \<le> tokens(take(s`NbR, s`rel)) \<longrightarrow>
NbT \<le> s`available_tok})"
apply (subgoal_tac
@@ -424,19 +424,19 @@
subsubsection\<open>Main lemmas towards proving property (31)\<close>
lemma LeadsTo_strength_R:
- "[| F \<in> C \<longmapsto>w B'; F \<in> A-C \<longmapsto>w B; B'<=B |] ==> F \<in> A \<longmapsto>w B"
+ "\<lbrakk>F \<in> C \<longmapsto>w B'; F \<in> A-C \<longmapsto>w B; B'<=B\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B"
by (blast intro: LeadsTo_weaken LeadsTo_Un_Un)
lemma PSP_StableI:
-"[| F \<in> Stable(C); F \<in> A - C \<longmapsto>w B;
- F \<in> A \<inter> C \<longmapsto>w B \<union> (state - C) |] ==> F \<in> A \<longmapsto>w B"
+"\<lbrakk>F \<in> Stable(C); F \<in> A - C \<longmapsto>w B;
+ F \<in> A \<inter> C \<longmapsto>w B \<union> (state - C)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B"
apply (rule_tac A = " (A-C) \<union> (A \<inter> C)" in LeadsTo_weaken_L)
prefer 2 apply blast
apply (rule LeadsTo_Un, assumption)
apply (blast intro: LeadsTo_weaken dest: PSP_Stable)
done
-lemma state_compl_eq [simp]: "state - {s\<in>state. P(s)} = {s\<in>state. ~P(s)}"
+lemma state_compl_eq [simp]: "state - {s\<in>state. P(s)} = {s\<in>state. \<not>P(s)}"
by auto
(*needed?*)
@@ -461,7 +461,7 @@
will eventually recognize that they've been released.*)
lemma (in alloc_progress) tokens_take_NbR_lemma:
"k \<in> tokbag
- ==> alloc_prog \<squnion> G \<in>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. k \<le> tokens(s`rel)}
\<longmapsto>w {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
apply (rule single_LeadsTo_I, safe)
@@ -480,7 +480,7 @@
\<longmapsto>w *)
lemma (in alloc_progress) tokens_take_NbR_lemma2:
"k \<in> tokbag
- ==> alloc_prog \<squnion> G \<in>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. tokens(s`giv) = k}
\<longmapsto>w {s\<in>state. k \<le> tokens(take(s`NbR, s`rel))}"
apply (rule LeadsTo_Trans)
@@ -492,8 +492,8 @@
(*Third step in proof of (31): by PSP with the fact that giv increases *)
lemma (in alloc_progress) length_giv_disj:
- "[| k \<in> tokbag; n \<in> nat |]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>k \<in> tokbag; n \<in> nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
\<longmapsto>w
{s\<in>state. (length(s`giv) = n & tokens(s`giv) = k &
@@ -512,8 +512,8 @@
(*Fourth step in proof of (31): we apply lemma (51) *)
lemma (in alloc_progress) length_giv_disj2:
- "[|k \<in> tokbag; n \<in> nat|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>k \<in> tokbag; n \<in> nat\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`giv) = n & tokens(s`giv) = k}
\<longmapsto>w
{s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
@@ -526,7 +526,7 @@
k\<in>nat *)
lemma (in alloc_progress) length_giv_disj3:
"n \<in> nat
- ==> alloc_prog \<squnion> G \<in>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`giv) = n}
\<longmapsto>w
{s\<in>state. (length(s`giv) = n & NbT \<le> s`available_tok) |
@@ -540,8 +540,8 @@
(*Sixth step in proof of (31): from the fifth step, by PSP with the
assumption that ask increases *)
lemma (in alloc_progress) length_ask_giv:
- "[|k \<in> nat; n < k|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>k \<in> nat; n < k\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`ask) = k & length(s`giv) = n}
\<longmapsto>w
{s\<in>state. (NbT \<le> s`available_tok & length(s`giv) < length(s`ask) &
@@ -563,8 +563,8 @@
(*Seventh step in proof of (31): no request (ask[k]) exceeds NbT *)
lemma (in alloc_progress) length_ask_giv2:
- "[|k \<in> nat; n < k|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>k \<in> nat; n < k\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`ask) = k & length(s`giv) = n}
\<longmapsto>w
{s\<in>state. (nth(length(s`giv), s`ask) \<le> s`available_tok &
@@ -580,8 +580,8 @@
(*Eighth step in proof of (31): by 50, we get |giv| > n. *)
lemma (in alloc_progress) extend_giv:
- "[| k \<in> nat; n < k|]
- ==> alloc_prog \<squnion> G \<in>
+ "\<lbrakk>k \<in> nat; n < k\<rbrakk>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. length(s`ask) = k & length(s`giv) = n}
\<longmapsto>w {s\<in>state. n < length(s`giv)}"
apply (rule LeadsTo_Un_duplicate)
@@ -597,7 +597,7 @@
we can't expect |ask| to remain fixed until |giv| increases.*)
lemma (in alloc_progress) alloc_prog_ask_LeadsTo_giv:
"k \<in> nat
- ==> alloc_prog \<squnion> G \<in>
+ \<Longrightarrow> alloc_prog \<squnion> G \<in>
{s\<in>state. k \<le> length(s`ask)} \<longmapsto>w {s\<in>state. k \<le> length(s`giv)}"
(* Proof by induction over the difference between k and n *)
apply (rule_tac f = "\<lambda>s\<in>state. k #- length(s`giv)" in LessThan_induct)
@@ -622,7 +622,7 @@
(*Final lemma: combine previous result with lemma (30)*)
lemma (in alloc_progress) final:
"h \<in> list(tokbag)
- ==> alloc_prog \<squnion> G
+ \<Longrightarrow> alloc_prog \<squnion> G
\<in> {s\<in>state. <h, s`ask> \<in> prefix(tokbag)} \<longmapsto>w
{s\<in>state. <h, s`giv> \<in> prefix(tokbag)}"
apply (rule single_LeadsTo_I)
--- a/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 16:51:35 2022 +0100
@@ -7,10 +7,10 @@
theory ClientImpl imports AllocBase Guar begin
-abbreviation "ask == Var(Nil)" (* input history: tokens requested *)
-abbreviation "giv == Var([0])" (* output history: tokens granted *)
-abbreviation "rel == Var([1])" (* input history: tokens released *)
-abbreviation "tok == Var([2])" (* the number of available tokens *)
+abbreviation "ask \<equiv> Var(Nil)" (* input history: tokens requested *)
+abbreviation "giv \<equiv> Var([0])" (* output history: tokens granted *)
+abbreviation "rel \<equiv> Var([1])" (* input history: tokens released *)
+abbreviation "tok \<equiv> Var([2])" (* the number of available tokens *)
axiomatization where
type_assumes:
@@ -21,11 +21,11 @@
default_val(rel) = Nil & default_val(tok) = 0"
-(*Array indexing is translated to list indexing as A[n] == nth(n-1,A). *)
+(*Array indexing is translated to list indexing as A[n] \<equiv> nth(n-1,A). *)
definition
(** Release some client_tokens **)
- "client_rel_act ==
+ "client_rel_act \<equiv>
{<s,t> \<in> state*state.
\<exists>nrel \<in> nat. nrel = length(s`rel) &
t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) &
@@ -36,14 +36,14 @@
(** Including t=s suppresses fairness, allowing the non-trivial part
of the action to be ignored **)
definition
- "client_tok_act == {<s,t> \<in> state*state. t=s |
+ "client_tok_act \<equiv> {<s,t> \<in> state*state. t=s |
t = s(tok:=succ(s`tok mod NbT))}"
definition
- "client_ask_act == {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
+ "client_ask_act \<equiv> {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
definition
- "client_prog ==
+ "client_prog \<equiv>
mk_program({s \<in> state. s`tok \<le> NbT & s`giv = Nil &
s`ask = Nil & s`rel = Nil},
{client_rel_act, client_tok_act, client_ask_act},
@@ -55,22 +55,22 @@
declare type_assumes [simp] default_val_assumes [simp]
(* This part should be automated *)
-lemma ask_value_type [simp,TC]: "s \<in> state ==> s`ask \<in> list(nat)"
+lemma ask_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`ask \<in> list(nat)"
apply (unfold state_def)
apply (drule_tac a = ask in apply_type, auto)
done
-lemma giv_value_type [simp,TC]: "s \<in> state ==> s`giv \<in> list(nat)"
+lemma giv_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`giv \<in> list(nat)"
apply (unfold state_def)
apply (drule_tac a = giv in apply_type, auto)
done
-lemma rel_value_type [simp,TC]: "s \<in> state ==> s`rel \<in> list(nat)"
+lemma rel_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`rel \<in> list(nat)"
apply (unfold state_def)
apply (drule_tac a = rel in apply_type, auto)
done
-lemma tok_value_type [simp,TC]: "s \<in> state ==> s`tok \<in> nat"
+lemma tok_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`tok \<in> nat"
apply (unfold state_def)
apply (drule_tac a = tok in apply_type, auto)
done
@@ -105,14 +105,14 @@
lemma preserves_lift_imp_stable:
- "G \<in> preserves(lift(ff)) ==> G \<in> stable({s \<in> state. P(s`ff)})"
+ "G \<in> preserves(lift(ff)) \<Longrightarrow> G \<in> stable({s \<in> state. P(s`ff)})"
apply (drule preserves_imp_stable)
apply (simp add: lift_def)
done
lemma preserves_imp_prefix:
"G \<in> preserves(lift(ff))
- ==> G \<in> stable({s \<in> state. \<langle>k, s`ff\<rangle> \<in> prefix(nat)})"
+ \<Longrightarrow> G \<in> stable({s \<in> state. \<langle>k, s`ff\<rangle> \<in> prefix(nat)})"
by (erule preserves_lift_imp_stable)
(*Safety property 1 \<in> ask, rel are increasing: (24) *)
@@ -135,8 +135,8 @@
With no Substitution Axiom, we must prove the two invariants simultaneously. *)
lemma ask_Bounded_lemma:
-"[| client_prog ok G; G \<in> program |]
- ==> client_prog \<squnion> G \<in>
+"\<lbrakk>client_prog ok G; G \<in> program\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G \<in>
Always({s \<in> state. s`tok \<le> NbT} \<inter>
{s \<in> state. \<forall>elt \<in> set_of_list(s`ask). elt \<le> NbT})"
apply (rotate_tac -1)
@@ -166,22 +166,22 @@
by (safety, auto)
lemma client_prog_Join_Stable_rel_le_giv:
-"[| client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel)) |]
- ==> client_prog \<squnion> G \<in> Stable({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})"
+"\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel))\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G \<in> Stable({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})"
apply (rule client_prog_stable_rel_le_giv [THEN Increasing_preserves_Stable])
apply (auto simp add: lift_def)
done
lemma client_prog_Join_Always_rel_le_giv:
- "[| client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel)) |]
- ==> client_prog \<squnion> G \<in> Always({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})"
+ "\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); G \<in> preserves(lift(rel))\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G \<in> Always({s \<in> state. <s`rel, s`giv> \<in> prefix(nat)})"
by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv)
lemma def_act_eq:
- "A == {<s, t> \<in> state*state. P(s, t)} ==> A={<s, t> \<in> state*state. P(s, t)}"
+ "A \<equiv> {<s, t> \<in> state*state. P(s, t)} \<Longrightarrow> A={<s, t> \<in> state*state. P(s, t)}"
by auto
-lemma act_subset: "A={<s,t> \<in> state*state. P(s, t)} ==> A<=state*state"
+lemma act_subset: "A={<s,t> \<in> state*state. P(s, t)} \<Longrightarrow> A<=state*state"
by auto
lemma transient_lemma:
@@ -215,8 +215,8 @@
done
lemma induct_lemma:
-"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |]
- ==> client_prog \<squnion> G \<in>
+"\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G \<in>
{s \<in> state. s`rel = k & <k,h> \<in> strict_prefix(nat)
& <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
\<longmapsto>w {s \<in> state. <k, s`rel> \<in> strict_prefix(nat)
@@ -244,8 +244,8 @@
done
lemma rel_progress_lemma:
-"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |]
- ==> client_prog \<squnion> G \<in>
+"\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G \<in>
{s \<in> state. <s`rel, h> \<in> strict_prefix(nat)
& <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
\<longmapsto>w {s \<in> state. <h, s`rel> \<in> prefix(nat)}"
@@ -273,8 +273,8 @@
done
lemma progress_lemma:
-"[| client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program |]
- ==> client_prog \<squnion> G
+"\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program\<rbrakk>
+ \<Longrightarrow> client_prog \<squnion> G
\<in> {s \<in> state. <h, s`giv> \<in> prefix(nat) & h pfixGe s`ask}
\<longmapsto>w {s \<in> state. <h, s`rel> \<in> prefix(nat)}"
apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI],
--- a/src/ZF/UNITY/Comp.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Comp.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,40 +19,40 @@
definition
component :: "[i,i]=>o" (infixl \<open>component\<close> 65) where
- "F component H == (\<exists>G. F \<squnion> G = H)"
+ "F component H \<equiv> (\<exists>G. F \<squnion> G = H)"
definition
strict_component :: "[i,i]=>o" (infixl \<open>strict'_component\<close> 65) where
- "F strict_component H == F component H & F\<noteq>H"
+ "F strict_component H \<equiv> F component H & F\<noteq>H"
definition
(* A stronger form of the component relation *)
component_of :: "[i,i]=>o" (infixl \<open>component'_of\<close> 65) where
- "F component_of H == (\<exists>G. F ok G & F \<squnion> G = H)"
+ "F component_of H \<equiv> (\<exists>G. F ok G & F \<squnion> G = H)"
definition
strict_component_of :: "[i,i]=>o" (infixl \<open>strict'_component'_of\<close> 65) where
- "F strict_component_of H == F component_of H & F\<noteq>H"
+ "F strict_component_of H \<equiv> F component_of H & F\<noteq>H"
definition
(* Preserves a state function f, in particular a variable *)
preserves :: "(i=>i)=>i" where
- "preserves(f) == {F:program. \<forall>z. F: stable({s:state. f(s) = z})}"
+ "preserves(f) \<equiv> {F:program. \<forall>z. F: stable({s:state. f(s) = z})}"
definition
fun_pair :: "[i=>i, i =>i] =>(i=>i)" where
- "fun_pair(f, g) == %x. <f(x), g(x)>"
+ "fun_pair(f, g) \<equiv> %x. <f(x), g(x)>"
definition
localize :: "[i=>i, i] => i" where
- "localize(f,F) == mk_program(Init(F), Acts(F),
+ "localize(f,F) \<equiv> mk_program(Init(F), Acts(F),
AllowedActs(F) \<inter> (\<Union>G\<in>preserves(f). Acts(G)))"
(*** component and strict_component relations ***)
lemma componentI:
- "H component F | H component G ==> H component (F \<squnion> G)"
+ "H component F | H component G \<Longrightarrow> H component (F \<squnion> G)"
apply (unfold component_def, auto)
apply (rule_tac x = "Ga \<squnion> G" in exI)
apply (rule_tac [2] x = "G \<squnion> F" in exI)
@@ -60,26 +60,26 @@
done
lemma component_eq_subset:
- "G \<in> program ==> (F component G) \<longleftrightarrow>
+ "G \<in> program \<Longrightarrow> (F component G) \<longleftrightarrow>
(Init(G) \<subseteq> Init(F) & Acts(F) \<subseteq> Acts(G) & AllowedActs(G) \<subseteq> AllowedActs(F))"
apply (unfold component_def, auto)
apply (rule exI)
apply (rule program_equalityI, auto)
done
-lemma component_SKIP [simp]: "F \<in> program ==> SKIP component F"
+lemma component_SKIP [simp]: "F \<in> program \<Longrightarrow> SKIP component F"
apply (unfold component_def)
apply (rule_tac x = F in exI)
apply (force intro: Join_SKIP_left)
done
-lemma component_refl [simp]: "F \<in> program ==> F component F"
+lemma component_refl [simp]: "F \<in> program \<Longrightarrow> F component F"
apply (unfold component_def)
apply (rule_tac x = F in exI)
apply (force intro: Join_SKIP_right)
done
-lemma SKIP_minimal: "F component SKIP ==> programify(F) = SKIP"
+lemma SKIP_minimal: "F component SKIP \<Longrightarrow> programify(F) = SKIP"
apply (rule program_equalityI)
apply (simp_all add: component_eq_subset, blast)
done
@@ -93,14 +93,14 @@
apply blast
done
-lemma Join_absorb1: "F component G ==> F \<squnion> G = G"
+lemma Join_absorb1: "F component G \<Longrightarrow> F \<squnion> G = G"
by (auto simp add: component_def Join_left_absorb)
-lemma Join_absorb2: "G component F ==> F \<squnion> G = F"
+lemma Join_absorb2: "G component F \<Longrightarrow> F \<squnion> G = F"
by (auto simp add: Join_ac component_def)
lemma JOIN_component_iff:
- "H \<in> program==>(JOIN(I,F) component H) \<longleftrightarrow> (\<forall>i \<in> I. F(i) component H)"
+ "H \<in> program\<Longrightarrow>(JOIN(I,F) component H) \<longleftrightarrow> (\<forall>i \<in> I. F(i) component H)"
apply (case_tac "I=0", force)
apply (simp (no_asm_simp) add: component_eq_subset)
apply auto
@@ -110,32 +110,32 @@
apply (blast elim!: not_emptyE)+
done
-lemma component_JOIN: "i \<in> I ==> F(i) component (\<Squnion>i \<in> I. (F(i)))"
+lemma component_JOIN: "i \<in> I \<Longrightarrow> F(i) component (\<Squnion>i \<in> I. (F(i)))"
apply (unfold component_def)
apply (blast intro: JOIN_absorb)
done
-lemma component_trans: "[| F component G; G component H |] ==> F component H"
+lemma component_trans: "\<lbrakk>F component G; G component H\<rbrakk> \<Longrightarrow> F component H"
apply (unfold component_def)
apply (blast intro: Join_assoc [symmetric])
done
lemma component_antisym:
- "[| F \<in> program; G \<in> program |] ==>(F component G & G component F) \<longrightarrow> F = G"
+ "\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F component G & G component F) \<longrightarrow> F = G"
apply (simp (no_asm_simp) add: component_eq_subset)
apply clarify
apply (rule program_equalityI, auto)
done
lemma Join_component_iff:
- "H \<in> program ==>
+ "H \<in> program \<Longrightarrow>
((F \<squnion> G) component H) \<longleftrightarrow> (F component H & G component H)"
apply (simp (no_asm_simp) add: component_eq_subset)
apply blast
done
lemma component_constrains:
- "[| F component G; G \<in> A co B; F \<in> program |] ==> F \<in> A co B"
+ "\<lbrakk>F component G; G \<in> A co B; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> A co B"
apply (frule constrainsD2)
apply (auto simp add: constrains_def component_eq_subset)
done
@@ -150,13 +150,13 @@
done
lemma preservesI [rule_format]:
- "\<forall>z. F \<in> stable({s \<in> state. f(s) = z}) ==> F \<in> preserves(f)"
+ "\<forall>z. F \<in> stable({s \<in> state. f(s) = z}) \<Longrightarrow> F \<in> preserves(f)"
apply (auto simp add: preserves_def)
apply (blast dest: stableD2)
done
lemma preserves_imp_eq:
- "[| F \<in> preserves(f); act \<in> Acts(F); <s,t> \<in> act |] ==> f(s) = f(t)"
+ "\<lbrakk>F \<in> preserves(f); act \<in> Acts(F); <s,t> \<in> act\<rbrakk> \<Longrightarrow> f(s) = f(t)"
apply (unfold preserves_def stable_def constrains_def)
apply (subgoal_tac "s \<in> state & t \<in> state")
prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force)
@@ -199,7 +199,7 @@
lemma preserves_type: "preserves(v)<=program"
by (unfold preserves_def, auto)
-lemma preserves_into_program [TC]: "F \<in> preserves(f) ==> F \<in> program"
+lemma preserves_into_program [TC]: "F \<in> preserves(f) \<Longrightarrow> F \<in> program"
by (blast intro: preserves_type [THEN subsetD])
lemma subset_preserves_comp: "preserves(f) \<subseteq> preserves(g comp f)"
@@ -208,7 +208,7 @@
apply (drule_tac x = act in bspec, auto)
done
-lemma imp_preserves_comp: "F \<in> preserves(f) ==> F \<in> preserves(g comp f)"
+lemma imp_preserves_comp: "F \<in> preserves(f) \<Longrightarrow> F \<in> preserves(g comp f)"
by (blast intro: subset_preserves_comp [THEN subsetD])
lemma preserves_subset_stable: "preserves(f) \<subseteq> stable({s \<in> state. P(f(s))})"
@@ -219,18 +219,18 @@
done
lemma preserves_imp_stable:
- "F \<in> preserves(f) ==> F \<in> stable({s \<in> state. P(f(s))})"
+ "F \<in> preserves(f) \<Longrightarrow> F \<in> stable({s \<in> state. P(f(s))})"
by (blast intro: preserves_subset_stable [THEN subsetD])
lemma preserves_imp_increasing:
- "[| F \<in> preserves(f); \<forall>x \<in> state. f(x):A |] ==> F \<in> Increasing.increasing(A, r, f)"
+ "\<lbrakk>F \<in> preserves(f); \<forall>x \<in> state. f(x):A\<rbrakk> \<Longrightarrow> F \<in> Increasing.increasing(A, r, f)"
apply (unfold Increasing.increasing_def)
apply (auto intro: preserves_into_program)
apply (rule_tac P = "%x. <k, x>:r" in preserves_imp_stable, auto)
done
lemma preserves_id_subset_stable:
- "st_set(A) ==> preserves(%x. x) \<subseteq> stable(A)"
+ "st_set(A) \<Longrightarrow> preserves(%x. x) \<subseteq> stable(A)"
apply (unfold preserves_def stable_def constrains_def, auto)
apply (drule_tac x = xb in spec)
apply (drule_tac x = act in bspec)
@@ -238,7 +238,7 @@
done
lemma preserves_id_imp_stable:
- "[| F \<in> preserves(%x. x); st_set(A) |] ==> F \<in> stable(A)"
+ "\<lbrakk>F \<in> preserves(%x. x); st_set(A)\<rbrakk> \<Longrightarrow> F \<in> stable(A)"
by (blast intro: preserves_id_subset_stable [THEN subsetD])
(** Added by Sidi **)
@@ -246,23 +246,23 @@
(* component_of is stronger than component *)
lemma component_of_imp_component:
-"F component_of H ==> F component H"
+"F component_of H \<Longrightarrow> F component H"
apply (unfold component_def component_of_def, blast)
done
(* component_of satisfies many of component's properties *)
-lemma component_of_refl [simp]: "F \<in> program ==> F component_of F"
+lemma component_of_refl [simp]: "F \<in> program \<Longrightarrow> F component_of F"
apply (unfold component_of_def)
apply (rule_tac x = SKIP in exI, auto)
done
-lemma component_of_SKIP [simp]: "F \<in> program ==>SKIP component_of F"
+lemma component_of_SKIP [simp]: "F \<in> program \<Longrightarrow>SKIP component_of F"
apply (unfold component_of_def, auto)
apply (rule_tac x = F in exI, auto)
done
lemma component_of_trans:
- "[| F component_of G; G component_of H |] ==> F component_of H"
+ "\<lbrakk>F component_of G; G component_of H\<rbrakk> \<Longrightarrow> F component_of H"
apply (unfold component_of_def)
apply (blast intro: Join_assoc [symmetric])
done
@@ -285,8 +285,8 @@
(** Theorems used in ClientImpl **)
lemma stable_localTo_stable2:
- "[| F \<in> stable({s \<in> state. P(f(s), g(s))}); G \<in> preserves(f); G \<in> preserves(g) |]
- ==> F \<squnion> G \<in> stable({s \<in> state. P(f(s), g(s))})"
+ "\<lbrakk>F \<in> stable({s \<in> state. P(f(s), g(s))}); G \<in> preserves(f); G \<in> preserves(g)\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> stable({s \<in> state. P(f(s), g(s))})"
apply (auto dest: ActsD preserves_into_program simp add: stable_def constrains_def)
apply (case_tac "act \<in> Acts (F) ")
apply auto
@@ -295,10 +295,10 @@
done
lemma Increasing_preserves_Stable:
- "[| F \<in> stable({s \<in> state. <f(s), g(s)>:r}); G \<in> preserves(f);
+ "\<lbrakk>F \<in> stable({s \<in> state. <f(s), g(s)>:r}); G \<in> preserves(f);
F \<squnion> G \<in> Increasing(A, r, g);
- \<forall>x \<in> state. f(x):A & g(x):A |]
- ==> F \<squnion> G \<in> Stable({s \<in> state. <f(s), g(s)>:r})"
+ \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> Stable({s \<in> state. <f(s), g(s)>:r})"
apply (auto simp add: stable_def Stable_def Increasing_def Constrains_def all_conj_distrib)
apply (simp_all add: constrains_type [THEN subsetD] preserves_type [THEN subsetD])
apply (blast intro: constrains_weaken)
@@ -315,15 +315,15 @@
(** Lemma used in AllocImpl **)
lemma Constrains_UN_left:
- "[| \<forall>x \<in> I. F \<in> A(x) Co B; F \<in> program |] ==> F:(\<Union>x \<in> I. A(x)) Co B"
+ "\<lbrakk>\<forall>x \<in> I. F \<in> A(x) Co B; F \<in> program\<rbrakk> \<Longrightarrow> F:(\<Union>x \<in> I. A(x)) Co B"
by (unfold Constrains_def constrains_def, auto)
lemma stable_Join_Stable:
- "[| F \<in> stable({s \<in> state. P(f(s), g(s))});
+ "\<lbrakk>F \<in> stable({s \<in> state. P(f(s), g(s))});
\<forall>k \<in> A. F \<squnion> G \<in> Stable({s \<in> state. P(k, g(s))});
- G \<in> preserves(f); \<forall>s \<in> state. f(s):A|]
- ==> F \<squnion> G \<in> Stable({s \<in> state. P(f(s), g(s))})"
+ G \<in> preserves(f); \<forall>s \<in> state. f(s):A\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> Stable({s \<in> state. P(f(s), g(s))})"
apply (unfold stable_def Stable_def preserves_def)
apply (rule_tac A = "(\<Union>k \<in> A. {s \<in> state. f(s)=k} \<inter> {s \<in> state. P (f (s), g (s))})" in Constrains_weaken_L)
prefer 2 apply blast
--- a/src/ZF/UNITY/Constrains.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Constrains.thy Tue Sep 27 16:51:35 2022 +0100
@@ -18,10 +18,10 @@
"(init \<union> (\<Union>act\<in>acts. field(act)))*list(\<Union>act\<in>acts. field(act))"
intros
(*Initial trace is empty*)
- Init: "s: init ==> <s,[]> \<in> traces(init,acts)"
+ Init: "s: init \<Longrightarrow> <s,[]> \<in> traces(init,acts)"
- Acts: "[| act:acts; <s,evs> \<in> traces(init,acts); <s,s'>: act |]
- ==> <s', Cons(s,evs)> \<in> traces(init, acts)"
+ Acts: "\<lbrakk>act:acts; <s,evs> \<in> traces(init,acts); <s,s'>: act\<rbrakk>
+ \<Longrightarrow> <s', Cons(s,evs)> \<in> traces(init, acts)"
type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1
@@ -31,30 +31,30 @@
domains
"reachable(F)" \<subseteq> "Init(F) \<union> (\<Union>act\<in>Acts(F). field(act))"
intros
- Init: "s:Init(F) ==> s:reachable(F)"
+ Init: "s:Init(F) \<Longrightarrow> s:reachable(F)"
- Acts: "[| act: Acts(F); s:reachable(F); <s,s'>: act |]
- ==> s':reachable(F)"
+ Acts: "\<lbrakk>act: Acts(F); s:reachable(F); <s,s'>: act\<rbrakk>
+ \<Longrightarrow> s':reachable(F)"
type_intros UnI1 UnI2 fieldI2 UN_I
definition
Constrains :: "[i,i] => i" (infixl \<open>Co\<close> 60) where
- "A Co B == {F:program. F:(reachable(F) \<inter> A) co B}"
+ "A Co B \<equiv> {F:program. F:(reachable(F) \<inter> A) co B}"
definition
op_Unless :: "[i, i] => i" (infixl \<open>Unless\<close> 60) where
- "A Unless B == (A-B) Co (A \<union> B)"
+ "A Unless B \<equiv> (A-B) Co (A \<union> B)"
definition
Stable :: "i => i" where
- "Stable(A) == A Co A"
+ "Stable(A) \<equiv> A Co A"
definition
(*Always is the weak form of "invariant"*)
Always :: "i => i" where
- "Always(A) == initially(A) \<inter> Stable(A)"
+ "Always(A) \<equiv> initially(A) \<inter> Stable(A)"
(*** traces and reachable ***)
@@ -80,7 +80,7 @@
declare state_Int_reachable [iff]
lemma reachable_equiv_traces:
-"F \<in> program ==> reachable(F)={s \<in> state. \<exists>evs. <s,evs>:traces(Init(F), Acts(F))}"
+"F \<in> program \<Longrightarrow> reachable(F)={s \<in> state. \<exists>evs. <s,evs>:traces(Init(F), Acts(F))}"
apply (rule equalityI, safe)
apply (blast dest: reachable_type [THEN subsetD])
apply (erule_tac [2] traces.induct)
@@ -91,8 +91,8 @@
lemma Init_into_reachable: "Init(F) \<subseteq> reachable(F)"
by (blast intro: reachable.intros)
-lemma stable_reachable: "[| F \<in> program; G \<in> program;
- Acts(G) \<subseteq> Acts(F) |] ==> G \<in> stable(reachable(F))"
+lemma stable_reachable: "\<lbrakk>F \<in> program; G \<in> program;
+ Acts(G) \<subseteq> Acts(F)\<rbrakk> \<Longrightarrow> G \<in> stable(reachable(F))"
apply (blast intro: stableI constrainsI st_setI
reachable_type [THEN subsetD] reachable.intros)
done
@@ -102,13 +102,13 @@
(*The set of all reachable states is an invariant...*)
lemma invariant_reachable:
- "F \<in> program ==> F \<in> invariant(reachable(F))"
+ "F \<in> program \<Longrightarrow> F \<in> invariant(reachable(F))"
apply (unfold invariant_def initially_def)
apply (blast intro: reachable_type [THEN subsetD] reachable.intros)
done
(*...in fact the strongest invariant!*)
-lemma invariant_includes_reachable: "F \<in> invariant(A) ==> reachable(F) \<subseteq> A"
+lemma invariant_includes_reachable: "F \<in> invariant(A) \<Longrightarrow> reachable(F) \<subseteq> A"
apply (cut_tac F = F in Acts_type)
apply (cut_tac F = F in Init_type)
apply (cut_tac F = F in reachable_type)
@@ -120,7 +120,7 @@
(*** Co ***)
-lemma constrains_reachable_Int: "F \<in> B co B'==>F:(reachable(F) \<inter> B) co (reachable(F) \<inter> B')"
+lemma constrains_reachable_Int: "F \<in> B co B'\<Longrightarrow>F:(reachable(F) \<inter> B) co (reachable(F) \<inter> B')"
apply (frule constrains_type [THEN subsetD])
apply (frule stable_reachable [OF _ _ subset_refl])
apply (simp_all add: stable_def constrains_Int)
@@ -136,15 +136,15 @@
lemmas Constrains_def2 = Constrains_eq_constrains [THEN eq_reflection]
-lemma constrains_imp_Constrains: "F \<in> A co A' ==> F \<in> A Co A'"
+lemma constrains_imp_Constrains: "F \<in> A co A' \<Longrightarrow> F \<in> A Co A'"
apply (unfold Constrains_def)
apply (blast intro: constrains_weaken_L dest: constrainsD2)
done
lemma ConstrainsI:
- "[|!!act s s'. [| act \<in> Acts(F); <s,s'>:act; s \<in> A |] ==> s':A';
- F \<in> program|]
- ==> F \<in> A Co A'"
+ "\<lbrakk>\<And>act s s'. \<lbrakk>act \<in> Acts(F); <s,s'>:act; s \<in> A\<rbrakk> \<Longrightarrow> s':A';
+ F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> A Co A'"
apply (auto simp add: Constrains_def constrains_def st_set_def)
apply (blast dest: reachable_type [THEN subsetD])
done
@@ -166,34 +166,34 @@
declare Constrains_state [iff]
lemma Constrains_weaken_R:
- "[| F \<in> A Co A'; A'<=B' |] ==> F \<in> A Co B'"
+ "\<lbrakk>F \<in> A Co A'; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> A Co B'"
apply (unfold Constrains_def2)
apply (blast intro: constrains_weaken_R)
done
lemma Constrains_weaken_L:
- "[| F \<in> A Co A'; B<=A |] ==> F \<in> B Co A'"
+ "\<lbrakk>F \<in> A Co A'; B<=A\<rbrakk> \<Longrightarrow> F \<in> B Co A'"
apply (unfold Constrains_def2)
apply (blast intro: constrains_weaken_L st_set_subset)
done
lemma Constrains_weaken:
- "[| F \<in> A Co A'; B<=A; A'<=B' |] ==> F \<in> B Co B'"
+ "\<lbrakk>F \<in> A Co A'; B<=A; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> B Co B'"
apply (unfold Constrains_def2)
apply (blast intro: constrains_weaken st_set_subset)
done
(** Union **)
lemma Constrains_Un:
- "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A \<union> B) Co (A' \<union> B')"
+ "\<lbrakk>F \<in> A Co A'; F \<in> B Co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) Co (A' \<union> B')"
apply (unfold Constrains_def2, auto)
apply (simp add: Int_Un_distrib)
apply (blast intro: constrains_Un)
done
lemma Constrains_UN:
- "[|(!!i. i \<in> I==>F \<in> A(i) Co A'(i)); F \<in> program|]
- ==> F:(\<Union>i \<in> I. A(i)) Co (\<Union>i \<in> I. A'(i))"
+ "\<lbrakk>(\<And>i. i \<in> I\<Longrightarrow>F \<in> A(i) Co A'(i)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F:(\<Union>i \<in> I. A(i)) Co (\<Union>i \<in> I. A'(i))"
by (auto intro: constrains_UN simp del: UN_simps
simp add: Constrains_def2 Int_UN_distrib)
@@ -201,33 +201,33 @@
(** Intersection **)
lemma Constrains_Int:
- "[| F \<in> A Co A'; F \<in> B Co B'|]==> F:(A \<inter> B) Co (A' \<inter> B')"
+ "\<lbrakk>F \<in> A Co A'; F \<in> B Co B'\<rbrakk>\<Longrightarrow> F:(A \<inter> B) Co (A' \<inter> B')"
apply (unfold Constrains_def)
apply (subgoal_tac "reachable (F) \<inter> (A \<inter> B) = (reachable (F) \<inter> A) \<inter> (reachable (F) \<inter> B) ")
apply (auto intro: constrains_Int)
done
lemma Constrains_INT:
- "[| (!!i. i \<in> I ==>F \<in> A(i) Co A'(i)); F \<in> program |]
- ==> F:(\<Inter>i \<in> I. A(i)) Co (\<Inter>i \<in> I. A'(i))"
+ "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow>F \<in> A(i) Co A'(i)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F:(\<Inter>i \<in> I. A(i)) Co (\<Inter>i \<in> I. A'(i))"
apply (simp (no_asm_simp) del: INT_simps add: Constrains_def INT_extend_simps)
apply (rule constrains_INT)
apply (auto simp add: Constrains_def)
done
-lemma Constrains_imp_subset: "F \<in> A Co A' ==> reachable(F) \<inter> A \<subseteq> A'"
+lemma Constrains_imp_subset: "F \<in> A Co A' \<Longrightarrow> reachable(F) \<inter> A \<subseteq> A'"
apply (unfold Constrains_def)
apply (blast dest: constrains_imp_subset)
done
lemma Constrains_trans:
- "[| F \<in> A Co B; F \<in> B Co C |] ==> F \<in> A Co C"
+ "\<lbrakk>F \<in> A Co B; F \<in> B Co C\<rbrakk> \<Longrightarrow> F \<in> A Co C"
apply (unfold Constrains_def2)
apply (blast intro: constrains_trans constrains_weaken)
done
lemma Constrains_cancel:
-"[| F \<in> A Co (A' \<union> B); F \<in> B Co B' |] ==> F \<in> A Co (A' \<union> B')"
+"\<lbrakk>F \<in> A Co (A' \<union> B); F \<in> B Co B'\<rbrakk> \<Longrightarrow> F \<in> A Co (A' \<union> B')"
apply (unfold Constrains_def2)
apply (simp (no_asm_use) add: Int_Un_distrib)
apply (blast intro: constrains_cancel)
@@ -237,13 +237,13 @@
(* Useful because there's no Stable_weaken. [Tanja Vos] *)
lemma stable_imp_Stable:
-"F \<in> stable(A) ==> F \<in> Stable(A)"
+"F \<in> stable(A) \<Longrightarrow> F \<in> Stable(A)"
apply (unfold stable_def Stable_def)
apply (erule constrains_imp_Constrains)
done
-lemma Stable_eq: "[| F \<in> Stable(A); A = B |] ==> F \<in> Stable(B)"
+lemma Stable_eq: "\<lbrakk>F \<in> Stable(A); A = B\<rbrakk> \<Longrightarrow> F \<in> Stable(B)"
by blast
lemma Stable_eq_stable:
@@ -251,53 +251,53 @@
apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2)
done
-lemma StableI: "F \<in> A Co A ==> F \<in> Stable(A)"
+lemma StableI: "F \<in> A Co A \<Longrightarrow> F \<in> Stable(A)"
by (unfold Stable_def, assumption)
-lemma StableD: "F \<in> Stable(A) ==> F \<in> A Co A"
+lemma StableD: "F \<in> Stable(A) \<Longrightarrow> F \<in> A Co A"
by (unfold Stable_def, assumption)
lemma Stable_Un:
- "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable(A \<union> A')"
+ "\<lbrakk>F \<in> Stable(A); F \<in> Stable(A')\<rbrakk> \<Longrightarrow> F \<in> Stable(A \<union> A')"
apply (unfold Stable_def)
apply (blast intro: Constrains_Un)
done
lemma Stable_Int:
- "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable (A \<inter> A')"
+ "\<lbrakk>F \<in> Stable(A); F \<in> Stable(A')\<rbrakk> \<Longrightarrow> F \<in> Stable (A \<inter> A')"
apply (unfold Stable_def)
apply (blast intro: Constrains_Int)
done
lemma Stable_Constrains_Un:
- "[| F \<in> Stable(C); F \<in> A Co (C \<union> A') |]
- ==> F \<in> (C \<union> A) Co (C \<union> A')"
+ "\<lbrakk>F \<in> Stable(C); F \<in> A Co (C \<union> A')\<rbrakk>
+ \<Longrightarrow> F \<in> (C \<union> A) Co (C \<union> A')"
apply (unfold Stable_def)
apply (blast intro: Constrains_Un [THEN Constrains_weaken_R])
done
lemma Stable_Constrains_Int:
- "[| F \<in> Stable(C); F \<in> (C \<inter> A) Co A' |]
- ==> F \<in> (C \<inter> A) Co (C \<inter> A')"
+ "\<lbrakk>F \<in> Stable(C); F \<in> (C \<inter> A) Co A'\<rbrakk>
+ \<Longrightarrow> F \<in> (C \<inter> A) Co (C \<inter> A')"
apply (unfold Stable_def)
apply (blast intro: Constrains_Int [THEN Constrains_weaken])
done
lemma Stable_UN:
- "[| (!!i. i \<in> I ==> F \<in> Stable(A(i))); F \<in> program |]
- ==> F \<in> Stable (\<Union>i \<in> I. A(i))"
+ "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> Stable(A(i))); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> Stable (\<Union>i \<in> I. A(i))"
apply (simp add: Stable_def)
apply (blast intro: Constrains_UN)
done
lemma Stable_INT:
- "[|(!!i. i \<in> I ==> F \<in> Stable(A(i))); F \<in> program |]
- ==> F \<in> Stable (\<Inter>i \<in> I. A(i))"
+ "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> Stable(A(i))); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> Stable (\<Inter>i \<in> I. A(i))"
apply (simp add: Stable_def)
apply (blast intro: Constrains_INT)
done
-lemma Stable_reachable: "F \<in> program ==>F \<in> Stable (reachable(F))"
+lemma Stable_reachable: "F \<in> program \<Longrightarrow>F \<in> Stable (reachable(F))"
apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb)
done
@@ -307,12 +307,12 @@
done
(*** The Elimination Theorem. The "free" m has become universally quantified!
- Should the premise be !!m instead of \<forall>m ? Would make it harder to use
+ Should the premise be \<And>m instead of \<forall>m ? Would make it harder to use
in forward proof. ***)
lemma Elimination:
- "[| \<forall>m \<in> M. F \<in> ({s \<in> A. x(s) = m}) Co (B(m)); F \<in> program |]
- ==> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
+ "\<lbrakk>\<forall>m \<in> M. F \<in> ({s \<in> A. x(s) = m}) Co (B(m)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
apply (unfold Constrains_def, auto)
apply (rule_tac A1 = "reachable (F) \<inter> A"
in UNITY.elimination [THEN constrains_weaken_L])
@@ -321,8 +321,8 @@
(* As above, but for the special case of A=state *)
lemma Elimination2:
- "[| \<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} Co B(m); F \<in> program |]
- ==> F \<in> {s \<in> state. x(s):M} Co (\<Union>m \<in> M. B(m))"
+ "\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} Co B(m); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> {s \<in> state. x(s):M} Co (\<Union>m \<in> M. B(m))"
apply (blast intro: Elimination)
done
@@ -338,20 +338,20 @@
(** Natural deduction rules for "Always A" **)
lemma AlwaysI:
-"[| Init(F)<=A; F \<in> Stable(A) |] ==> F \<in> Always(A)"
+"\<lbrakk>Init(F)<=A; F \<in> Stable(A)\<rbrakk> \<Longrightarrow> F \<in> Always(A)"
apply (unfold Always_def initially_def)
apply (frule Stable_type [THEN subsetD], auto)
done
-lemma AlwaysD: "F \<in> Always(A) ==> Init(F)<=A & F \<in> Stable(A)"
+lemma AlwaysD: "F \<in> Always(A) \<Longrightarrow> Init(F)<=A & F \<in> Stable(A)"
by (simp add: Always_def initially_def)
lemmas AlwaysE = AlwaysD [THEN conjE]
lemmas Always_imp_Stable = AlwaysD [THEN conjunct2]
(*The set of all reachable states is Always*)
-lemma Always_includes_reachable: "F \<in> Always(A) ==> reachable(F) \<subseteq> A"
+lemma Always_includes_reachable: "F \<in> Always(A) \<Longrightarrow> reachable(F) \<subseteq> A"
apply (simp (no_asm_use) add: Stable_def Constrains_def constrains_def Always_def initially_def)
apply (rule subsetI)
apply (erule reachable.induct)
@@ -359,7 +359,7 @@
done
lemma invariant_imp_Always:
- "F \<in> invariant(A) ==> F \<in> Always(A)"
+ "F \<in> invariant(A) \<Longrightarrow> F \<in> Always(A)"
apply (unfold Always_def invariant_def Stable_def stable_def)
apply (blast intro: constrains_imp_Constrains)
done
@@ -389,11 +389,11 @@
done
declare Always_state_eq [simp]
-lemma state_AlwaysI: "F \<in> program ==> F \<in> Always(state)"
+lemma state_AlwaysI: "F \<in> program \<Longrightarrow> F \<in> Always(state)"
by (auto dest: reachable_type [THEN subsetD]
simp add: Always_eq_includes_reachable)
-lemma Always_eq_UN_invariant: "st_set(A) ==> Always(A) = (\<Union>I \<in> Pow(A). invariant(I))"
+lemma Always_eq_UN_invariant: "st_set(A) \<Longrightarrow> Always(A) = (\<Union>I \<in> Pow(A). invariant(I))"
apply (simp (no_asm) add: Always_eq_includes_reachable)
apply (rule equalityI, auto)
apply (blast intro: invariantI rev_subsetD [OF _ Init_into_reachable]
@@ -401,30 +401,30 @@
dest: invariant_type [THEN subsetD])+
done
-lemma Always_weaken: "[| F \<in> Always(A); A \<subseteq> B |] ==> F \<in> Always(B)"
+lemma Always_weaken: "\<lbrakk>F \<in> Always(A); A \<subseteq> B\<rbrakk> \<Longrightarrow> F \<in> Always(B)"
by (auto simp add: Always_eq_includes_reachable)
(*** "Co" rules involving Always ***)
lemmas Int_absorb2 = subset_Int_iff [unfolded iff_def, THEN conjunct1, THEN mp]
-lemma Always_Constrains_pre: "F \<in> Always(I) ==> (F:(I \<inter> A) Co A') \<longleftrightarrow> (F \<in> A Co A')"
+lemma Always_Constrains_pre: "F \<in> Always(I) \<Longrightarrow> (F:(I \<inter> A) Co A') \<longleftrightarrow> (F \<in> A Co A')"
apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_def Int_assoc [symmetric])
done
-lemma Always_Constrains_post: "F \<in> Always(I) ==> (F \<in> A Co (I \<inter> A')) \<longleftrightarrow>(F \<in> A Co A')"
+lemma Always_Constrains_post: "F \<in> Always(I) \<Longrightarrow> (F \<in> A Co (I \<inter> A')) \<longleftrightarrow>(F \<in> A Co A')"
apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_eq_constrains Int_assoc [symmetric])
done
-lemma Always_ConstrainsI: "[| F \<in> Always(I); F \<in> (I \<inter> A) Co A' |] ==> F \<in> A Co A'"
+lemma Always_ConstrainsI: "\<lbrakk>F \<in> Always(I); F \<in> (I \<inter> A) Co A'\<rbrakk> \<Longrightarrow> F \<in> A Co A'"
by (blast intro: Always_Constrains_pre [THEN iffD1])
-(* [| F \<in> Always(I); F \<in> A Co A' |] ==> F \<in> A Co (I \<inter> A') *)
+(* \<lbrakk>F \<in> Always(I); F \<in> A Co A'\<rbrakk> \<Longrightarrow> F \<in> A Co (I \<inter> A') *)
lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2]
(*The analogous proof of Always_LeadsTo_weaken doesn't terminate*)
lemma Always_Constrains_weaken:
-"[|F \<in> Always(C); F \<in> A Co A'; C \<inter> B<=A; C \<inter> A'<=B'|]==>F \<in> B Co B'"
+"\<lbrakk>F \<in> Always(C); F \<in> A Co A'; C \<inter> B<=A; C \<inter> A'<=B'\<rbrakk>\<Longrightarrow>F \<in> B Co B'"
apply (rule Always_ConstrainsI)
apply (drule_tac [2] Always_ConstrainsD, simp_all)
apply (blast intro: Constrains_weaken)
@@ -435,18 +435,18 @@
by (auto simp add: Always_eq_includes_reachable)
(* the premise i \<in> I is need since \<Inter>is formally not defined for I=0 *)
-lemma Always_INT_distrib: "i \<in> I==>Always(\<Inter>i \<in> I. A(i)) = (\<Inter>i \<in> I. Always(A(i)))"
+lemma Always_INT_distrib: "i \<in> I\<Longrightarrow>Always(\<Inter>i \<in> I. A(i)) = (\<Inter>i \<in> I. Always(A(i)))"
apply (rule equalityI)
apply (auto simp add: Inter_iff Always_eq_includes_reachable)
done
-lemma Always_Int_I: "[| F \<in> Always(A); F \<in> Always(B) |] ==> F \<in> Always(A \<inter> B)"
+lemma Always_Int_I: "\<lbrakk>F \<in> Always(A); F \<in> Always(B)\<rbrakk> \<Longrightarrow> F \<in> Always(A \<inter> B)"
apply (simp (no_asm_simp) add: Always_Int_distrib)
done
(*Allows a kind of "implication introduction"*)
-lemma Always_Diff_Un_eq: "[| F \<in> Always(A) |] ==> (F \<in> Always(C-A \<union> B)) \<longleftrightarrow> (F \<in> Always(B))"
+lemma Always_Diff_Un_eq: "\<lbrakk>F \<in> Always(A)\<rbrakk> \<Longrightarrow> (F \<in> Always(C-A \<union> B)) \<longleftrightarrow> (F \<in> Always(B))"
by (auto simp add: Always_eq_includes_reachable)
(*Delete the nearest invariance assumption (which will be the second one
--- a/src/ZF/UNITY/Distributor.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Distributor.thy Tue Sep 27 16:51:35 2022 +0100
@@ -14,7 +14,7 @@
definition
distr_follows :: "[i, i, i, i =>i] =>i" where
- "distr_follows(A, In, iIn, Out) ==
+ "distr_follows(A, In, iIn, Out) \<equiv>
(lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
(lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients})
@@ -27,13 +27,13 @@
definition
distr_allowed_acts :: "[i=>i]=>i" where
- "distr_allowed_acts(Out) ==
+ "distr_allowed_acts(Out) \<equiv>
{D \<in> program. AllowedActs(D) =
cons(id(state), \<Union>G \<in> (\<Inter>n\<in>nat. preserves(lift(Out(n)))). Acts(G))}"
definition
distr_spec :: "[i, i, i, i =>i]=>i" where
- "distr_spec(A, In, iIn, Out) ==
+ "distr_spec(A, In, iIn, Out) \<equiv>
distr_follows(A, In, iIn, Out) \<inter> distr_allowed_acts(Out)"
locale distr =
@@ -54,19 +54,19 @@
and distr_spec: "D \<in> distr_spec(A, In, iIn, Out)"
-lemma (in distr) In_value_type [simp,TC]: "s \<in> state ==> s`In \<in> list(A)"
+lemma (in distr) In_value_type [simp,TC]: "s \<in> state \<Longrightarrow> s`In \<in> list(A)"
apply (unfold state_def)
apply (drule_tac a = In in apply_type, auto)
done
lemma (in distr) iIn_value_type [simp,TC]:
- "s \<in> state ==> s`iIn \<in> list(nat)"
+ "s \<in> state \<Longrightarrow> s`iIn \<in> list(nat)"
apply (unfold state_def)
apply (drule_tac a = iIn in apply_type, auto)
done
lemma (in distr) Out_value_type [simp,TC]:
- "s \<in> state ==> s`Out(n):list(A)"
+ "s \<in> state \<Longrightarrow> s`Out(n):list(A)"
apply (unfold state_def)
apply (drule_tac a = "Out (n)" in apply_type)
apply auto
@@ -78,7 +78,7 @@
done
lemma (in distr) D_ok_iff:
- "G \<in> program ==>
+ "G \<in> program \<Longrightarrow>
D ok G \<longleftrightarrow> ((\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)))) & D \<in> Allowed(G))"
apply (cut_tac distr_spec)
apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def
@@ -101,10 +101,10 @@
done
lemma (in distr) distr_bag_Follows_lemma:
-"[| \<forall>n \<in> nat. G \<in> preserves(lift(Out(n)));
+"\<lbrakk>\<forall>n \<in> nat. G \<in> preserves(lift(Out(n)));
D \<squnion> G \<in> Always({s \<in> state.
- \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients}) |]
- ==> D \<squnion> G \<in> Always
+ \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients})\<rbrakk>
+ \<Longrightarrow> D \<squnion> G \<in> Always
({s \<in> state. msetsum(%n. bag_of (sublist(s`In,
{k \<in> nat. k < length(s`iIn) &
nth(k, s`iIn)= n})),
--- a/src/ZF/UNITY/FP.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/FP.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,11 +13,11 @@
definition
FP_Orig :: "i=>i" where
- "FP_Orig(F) == \<Union>({A \<in> Pow(state). \<forall>B. F \<in> stable(A \<inter> B)})"
+ "FP_Orig(F) \<equiv> \<Union>({A \<in> Pow(state). \<forall>B. F \<in> stable(A \<inter> B)})"
definition
FP :: "i=>i" where
- "FP(F) == {s\<in>state. F \<in> stable({s})}"
+ "FP(F) \<equiv> {s\<in>state. F \<in> stable({s})}"
lemma FP_Orig_type: "FP_Orig(F) \<subseteq> state"
@@ -36,18 +36,18 @@
apply (rule FP_type)
done
-lemma stable_FP_Orig_Int: "F \<in> program ==> F \<in> stable(FP_Orig(F) \<inter> B)"
+lemma stable_FP_Orig_Int: "F \<in> program \<Longrightarrow> F \<in> stable(FP_Orig(F) \<inter> B)"
apply (simp only: FP_Orig_def stable_def Int_Union2)
apply (blast intro: constrains_UN)
done
lemma FP_Orig_weakest2:
- "[| \<forall>B. F \<in> stable (A \<inter> B); st_set(A) |] ==> A \<subseteq> FP_Orig(F)"
+ "\<lbrakk>\<forall>B. F \<in> stable (A \<inter> B); st_set(A)\<rbrakk> \<Longrightarrow> A \<subseteq> FP_Orig(F)"
by (unfold FP_Orig_def stable_def st_set_def, blast)
lemmas FP_Orig_weakest = allI [THEN FP_Orig_weakest2]
-lemma stable_FP_Int: "F \<in> program ==> F \<in> stable (FP(F) \<inter> B)"
+lemma stable_FP_Int: "F \<in> program \<Longrightarrow> F \<in> stable (FP(F) \<inter> B)"
apply (subgoal_tac "FP (F) \<inter> B = (\<Union>x\<in>B. FP (F) \<inter> {x}) ")
prefer 2 apply blast
apply (simp (no_asm_simp) add: Int_cons_right)
@@ -56,10 +56,10 @@
apply (auto simp add: cons_absorb)
done
-lemma FP_subset_FP_Orig: "F \<in> program ==> FP(F) \<subseteq> FP_Orig(F)"
+lemma FP_subset_FP_Orig: "F \<in> program \<Longrightarrow> FP(F) \<subseteq> FP_Orig(F)"
by (rule stable_FP_Int [THEN FP_Orig_weakest], auto)
-lemma FP_Orig_subset_FP: "F \<in> program ==> FP_Orig(F) \<subseteq> FP(F)"
+lemma FP_Orig_subset_FP: "F \<in> program \<Longrightarrow> FP_Orig(F) \<subseteq> FP(F)"
apply (unfold FP_Orig_def FP_def, clarify)
apply (drule_tac x = "{x}" in spec)
apply (simp add: Int_cons_right)
@@ -67,17 +67,17 @@
apply (auto simp add: cons_absorb st_set_def)
done
-lemma FP_equivalence: "F \<in> program ==> FP(F) = FP_Orig(F)"
+lemma FP_equivalence: "F \<in> program \<Longrightarrow> FP(F) = FP_Orig(F)"
by (blast intro!: FP_Orig_subset_FP FP_subset_FP_Orig)
lemma FP_weakest [rule_format]:
- "[| \<forall>B. F \<in> stable(A \<inter> B); F \<in> program; st_set(A) |] ==> A \<subseteq> FP(F)"
+ "\<lbrakk>\<forall>B. F \<in> stable(A \<inter> B); F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> A \<subseteq> FP(F)"
by (simp add: FP_equivalence FP_Orig_weakest)
lemma Diff_FP:
- "[| F \<in> program; st_set(A) |]
- ==> A-FP(F) = (\<Union>act \<in> Acts(F). A - {s \<in> state. act``{s} \<subseteq> {s}})"
+ "\<lbrakk>F \<in> program; st_set(A)\<rbrakk>
+ \<Longrightarrow> A-FP(F) = (\<Union>act \<in> Acts(F). A - {s \<in> state. act``{s} \<subseteq> {s}})"
by (unfold FP_def stable_def constrains_def st_set_def, blast)
end
--- a/src/ZF/UNITY/Follows.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Follows.thy Tue Sep 27 16:51:35 2022 +0100
@@ -11,7 +11,7 @@
definition
Follows :: "[i, i, i=>i, i=>i] => i" where
- "Follows(A, r, f, g) ==
+ "Follows(A, r, f, g) \<equiv>
Increasing(A, r, g) Int
Increasing(A, r,f) Int
Always({s \<in> state. <f(s), g(s)>:r}) Int
@@ -19,57 +19,57 @@
abbreviation
Incr :: "[i=>i]=>i" where
- "Incr(f) == Increasing(list(nat), prefix(nat), f)"
+ "Incr(f) \<equiv> Increasing(list(nat), prefix(nat), f)"
abbreviation
n_Incr :: "[i=>i]=>i" where
- "n_Incr(f) == Increasing(nat, Le, f)"
+ "n_Incr(f) \<equiv> Increasing(nat, Le, f)"
abbreviation
s_Incr :: "[i=>i]=>i" where
- "s_Incr(f) == Increasing(Pow(nat), SetLe(nat), f)"
+ "s_Incr(f) \<equiv> Increasing(Pow(nat), SetLe(nat), f)"
abbreviation
m_Incr :: "[i=>i]=>i" where
- "m_Incr(f) == Increasing(Mult(nat), MultLe(nat, Le), f)"
+ "m_Incr(f) \<equiv> Increasing(Mult(nat), MultLe(nat, Le), f)"
abbreviation
n_Fols :: "[i=>i, i=>i]=>i" (infixl \<open>n'_Fols\<close> 65) where
- "f n_Fols g == Follows(nat, Le, f, g)"
+ "f n_Fols g \<equiv> Follows(nat, Le, f, g)"
abbreviation
Follows' :: "[i=>i, i=>i, i, i] => i"
(\<open>(_ /Fols _ /Wrt (_ /'/ _))\<close> [60, 0, 0, 60] 60) where
- "f Fols g Wrt r/A == Follows(A,r,f,g)"
+ "f Fols g Wrt r/A \<equiv> Follows(A,r,f,g)"
(*Does this hold for "invariant"?*)
lemma Follows_cong:
- "[|A=A'; r=r'; !!x. x \<in> state ==> f(x)=f'(x); !!x. x \<in> state ==> g(x)=g'(x)|] ==> Follows(A, r, f, g) = Follows(A', r', f', g')"
+ "\<lbrakk>A=A'; r=r'; \<And>x. x \<in> state \<Longrightarrow> f(x)=f'(x); \<And>x. x \<in> state \<Longrightarrow> g(x)=g'(x)\<rbrakk> \<Longrightarrow> Follows(A, r, f, g) = Follows(A', r', f', g')"
by (simp add: Increasing_def Follows_def)
lemma subset_Always_comp:
-"[| mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A |] ==>
+"\<lbrakk>mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
Always({x \<in> state. <f(x), g(x)> \<in> r})<=Always({x \<in> state. <(h comp f)(x), (h comp g)(x)> \<in> s})"
apply (unfold mono1_def metacomp_def)
apply (auto simp add: Always_eq_includes_reachable)
done
lemma imp_Always_comp:
-"[| F \<in> Always({x \<in> state. <f(x), g(x)> \<in> r});
- mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A |] ==>
+"\<lbrakk>F \<in> Always({x \<in> state. <f(x), g(x)> \<in> r});
+ mono1(A, r, B, s, h); \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
F \<in> Always({x \<in> state. <(h comp f)(x), (h comp g)(x)> \<in> s})"
by (blast intro: subset_Always_comp [THEN subsetD])
lemma imp_Always_comp2:
-"[| F \<in> Always({x \<in> state. <f1(x), f(x)> \<in> r});
+"\<lbrakk>F \<in> Always({x \<in> state. <f1(x), f(x)> \<in> r});
F \<in> Always({x \<in> state. <g1(x), g(x)> \<in> s});
mono2(A, r, B, s, C, t, h);
- \<forall>x \<in> state. f1(x):A & f(x):A & g1(x):B & g(x):B |]
- ==> F \<in> Always({x \<in> state. <h(f1(x), g1(x)), h(f(x), g(x))> \<in> t})"
+ \<forall>x \<in> state. f1(x):A & f(x):A & g1(x):B & g(x):B\<rbrakk>
+ \<Longrightarrow> F \<in> Always({x \<in> state. <h(f1(x), g1(x)), h(f(x), g(x))> \<in> t})"
apply (auto simp add: Always_eq_includes_reachable mono2_def)
apply (auto dest!: subsetD)
done
@@ -77,8 +77,8 @@
(* comp LeadsTo *)
lemma subset_LeadsTo_comp:
-"[| mono1(A, r, B, s, h); refl(A,r); trans[B](s);
- \<forall>x \<in> state. f(x):A & g(x):A |] ==>
+"\<lbrakk>mono1(A, r, B, s, h); refl(A,r); trans[B](s);
+ \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
(\<Inter>j \<in> A. {s \<in> state. <j, g(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f(s)> \<in> r}) <=
(\<Inter>k \<in> B. {x \<in> state. <k, (h comp g)(x)> \<in> s} \<longmapsto>w {x \<in> state. <k, (h comp f)(x)> \<in> s})"
@@ -97,18 +97,18 @@
done
lemma imp_LeadsTo_comp:
-"[| F:(\<Inter>j \<in> A. {s \<in> state. <j, g(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f(s)> \<in> r});
+"\<lbrakk>F:(\<Inter>j \<in> A. {s \<in> state. <j, g(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f(s)> \<in> r});
mono1(A, r, B, s, h); refl(A,r); trans[B](s);
- \<forall>x \<in> state. f(x):A & g(x):A |] ==>
+ \<forall>x \<in> state. f(x):A & g(x):A\<rbrakk> \<Longrightarrow>
F:(\<Inter>k \<in> B. {x \<in> state. <k, (h comp g)(x)> \<in> s} \<longmapsto>w {x \<in> state. <k, (h comp f)(x)> \<in> s})"
apply (rule subset_LeadsTo_comp [THEN subsetD], auto)
done
lemma imp_LeadsTo_comp_right:
-"[| F \<in> Increasing(B, s, g);
+"\<lbrakk>F \<in> Increasing(B, s, g);
\<forall>j \<in> A. F: {s \<in> state. <j, f(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f1(s)> \<in> r};
mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t);
- \<forall>x \<in> state. f1(x):A & f(x):A & g(x):B; k \<in> C |] ==>
+ \<forall>x \<in> state. f1(x):A & f(x):A & g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f1(x), g(x))> \<in> t}"
apply (unfold mono2_def Increasing_def)
apply (rule single_LeadsTo_I, auto)
@@ -128,10 +128,10 @@
done
lemma imp_LeadsTo_comp_left:
-"[| F \<in> Increasing(A, r, f);
+"\<lbrakk>F \<in> Increasing(A, r, f);
\<forall>j \<in> B. F: {x \<in> state. <j, g(x)> \<in> s} \<longmapsto>w {x \<in> state. <j,g1(x)> \<in> s};
mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t);
- \<forall>x \<in> state. f(x):A & g1(x):B & g(x):B; k \<in> C |] ==>
+ \<forall>x \<in> state. f(x):A & g1(x):B & g(x):B; k \<in> C\<rbrakk> \<Longrightarrow>
F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f(x), g1(x))> \<in> t}"
apply (unfold mono2_def Increasing_def)
apply (rule single_LeadsTo_I, auto)
@@ -152,12 +152,12 @@
(** This general result is used to prove Follows Un, munion, etc. **)
lemma imp_LeadsTo_comp2:
-"[| F \<in> Increasing(A, r, f1) \<inter> Increasing(B, s, g);
+"\<lbrakk>F \<in> Increasing(A, r, f1) \<inter> Increasing(B, s, g);
\<forall>j \<in> A. F: {s \<in> state. <j, f(s)> \<in> r} \<longmapsto>w {s \<in> state. <j,f1(s)> \<in> r};
\<forall>j \<in> B. F: {x \<in> state. <j, g(x)> \<in> s} \<longmapsto>w {x \<in> state. <j,g1(x)> \<in> s};
mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t);
- \<forall>x \<in> state. f(x):A & g1(x):B & f1(x):A &g(x):B; k \<in> C |]
- ==> F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f1(x), g1(x))> \<in> t}"
+ \<forall>x \<in> state. f(x):A & g1(x):B & f1(x):A &g(x):B; k \<in> C\<rbrakk>
+ \<Longrightarrow> F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f1(x), g1(x))> \<in> t}"
apply (rule_tac B = "{x \<in> state. <k, h (f1 (x), g (x))> \<in> t}" in LeadsTo_Trans)
apply (blast intro: imp_LeadsTo_comp_right)
apply (blast intro: imp_LeadsTo_comp_left)
@@ -169,25 +169,25 @@
apply (blast dest: Increasing_type [THEN subsetD])
done
-lemma Follows_into_program [TC]: "F \<in> Follows(A, r, f, g) ==> F \<in> program"
+lemma Follows_into_program [TC]: "F \<in> Follows(A, r, f, g) \<Longrightarrow> F \<in> program"
by (blast dest: Follows_type [THEN subsetD])
lemma FollowsD:
-"F \<in> Follows(A, r, f, g)==>
+"F \<in> Follows(A, r, f, g)\<Longrightarrow>
F \<in> program & (\<exists>a. a \<in> A) & (\<forall>x \<in> state. f(x):A & g(x):A)"
apply (unfold Follows_def)
apply (blast dest: IncreasingD)
done
lemma Follows_constantI:
- "[| F \<in> program; c \<in> A; refl(A, r) |] ==> F \<in> Follows(A, r, %x. c, %x. c)"
+ "\<lbrakk>F \<in> program; c \<in> A; refl(A, r)\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, %x. c, %x. c)"
apply (unfold Follows_def, auto)
apply (auto simp add: refl_def)
done
lemma subset_Follows_comp:
-"[| mono1(A, r, B, s, h); refl(A, r); trans[B](s) |]
- ==> Follows(A, r, f, g) \<subseteq> Follows(B, s, h comp f, h comp g)"
+"\<lbrakk>mono1(A, r, B, s, h); refl(A, r); trans[B](s)\<rbrakk>
+ \<Longrightarrow> Follows(A, r, f, g) \<subseteq> Follows(B, s, h comp f, h comp g)"
apply (unfold Follows_def, clarify)
apply (frule_tac f = g in IncreasingD)
apply (frule_tac f = f in IncreasingD)
@@ -198,8 +198,8 @@
done
lemma imp_Follows_comp:
-"[| F \<in> Follows(A, r, f, g); mono1(A, r, B, s, h); refl(A, r); trans[B](s) |]
- ==> F \<in> Follows(B, s, h comp f, h comp g)"
+"\<lbrakk>F \<in> Follows(A, r, f, g); mono1(A, r, B, s, h); refl(A, r); trans[B](s)\<rbrakk>
+ \<Longrightarrow> F \<in> Follows(B, s, h comp f, h comp g)"
apply (blast intro: subset_Follows_comp [THEN subsetD])
done
@@ -208,9 +208,9 @@
(* 2-place monotone operation \<in> this general result is
used to prove Follows_Un, Follows_munion *)
lemma imp_Follows_comp2:
-"[| F \<in> Follows(A, r, f1, f); F \<in> Follows(B, s, g1, g);
- mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t) |]
- ==> F \<in> Follows(C, t, %x. h(f1(x), g1(x)), %x. h(f(x), g(x)))"
+"\<lbrakk>F \<in> Follows(A, r, f1, f); F \<in> Follows(B, s, g1, g);
+ mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t)\<rbrakk>
+ \<Longrightarrow> F \<in> Follows(C, t, %x. h(f1(x), g1(x)), %x. h(f(x), g(x)))"
apply (unfold Follows_def, clarify)
apply (frule_tac f = g in IncreasingD)
apply (frule_tac f = f in IncreasingD)
@@ -232,8 +232,8 @@
lemma Follows_trans:
- "[| F \<in> Follows(A, r, f, g); F \<in> Follows(A,r, g, h);
- trans[A](r) |] ==> F \<in> Follows(A, r, f, h)"
+ "\<lbrakk>F \<in> Follows(A, r, f, g); F \<in> Follows(A,r, g, h);
+ trans[A](r)\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, f, h)"
apply (frule_tac f = f in FollowsD)
apply (frule_tac f = g in FollowsD)
apply (simp add: Follows_def)
@@ -246,35 +246,35 @@
(** Destruction rules for Follows **)
lemma Follows_imp_Increasing_left:
- "F \<in> Follows(A, r, f,g) ==> F \<in> Increasing(A, r, f)"
+ "F \<in> Follows(A, r, f,g) \<Longrightarrow> F \<in> Increasing(A, r, f)"
by (unfold Follows_def, blast)
lemma Follows_imp_Increasing_right:
- "F \<in> Follows(A, r, f,g) ==> F \<in> Increasing(A, r, g)"
+ "F \<in> Follows(A, r, f,g) \<Longrightarrow> F \<in> Increasing(A, r, g)"
by (unfold Follows_def, blast)
lemma Follows_imp_Always:
- "F :Follows(A, r, f, g) ==> F \<in> Always({s \<in> state. <f(s),g(s)> \<in> r})"
+ "F :Follows(A, r, f, g) \<Longrightarrow> F \<in> Always({s \<in> state. <f(s),g(s)> \<in> r})"
by (unfold Follows_def, blast)
lemma Follows_imp_LeadsTo:
- "[| F \<in> Follows(A, r, f, g); k \<in> A |] ==>
+ "\<lbrakk>F \<in> Follows(A, r, f, g); k \<in> A\<rbrakk> \<Longrightarrow>
F: {s \<in> state. <k,g(s)> \<in> r } \<longmapsto>w {s \<in> state. <k,f(s)> \<in> r}"
by (unfold Follows_def, blast)
lemma Follows_LeadsTo_pfixLe:
- "[| F \<in> Follows(list(nat), gen_prefix(nat, Le), f, g); k \<in> list(nat) |]
- ==> F \<in> {s \<in> state. k pfixLe g(s)} \<longmapsto>w {s \<in> state. k pfixLe f(s)}"
+ "\<lbrakk>F \<in> Follows(list(nat), gen_prefix(nat, Le), f, g); k \<in> list(nat)\<rbrakk>
+ \<Longrightarrow> F \<in> {s \<in> state. k pfixLe g(s)} \<longmapsto>w {s \<in> state. k pfixLe f(s)}"
by (blast intro: Follows_imp_LeadsTo)
lemma Follows_LeadsTo_pfixGe:
- "[| F \<in> Follows(list(nat), gen_prefix(nat, Ge), f, g); k \<in> list(nat) |]
- ==> F \<in> {s \<in> state. k pfixGe g(s)} \<longmapsto>w {s \<in> state. k pfixGe f(s)}"
+ "\<lbrakk>F \<in> Follows(list(nat), gen_prefix(nat, Ge), f, g); k \<in> list(nat)\<rbrakk>
+ \<Longrightarrow> F \<in> {s \<in> state. k pfixGe g(s)} \<longmapsto>w {s \<in> state. k pfixGe f(s)}"
by (blast intro: Follows_imp_LeadsTo)
lemma Always_Follows1:
-"[| F \<in> Always({s \<in> state. f(s) = g(s)}); F \<in> Follows(A, r, f, h);
- \<forall>x \<in> state. g(x):A |] ==> F \<in> Follows(A, r, g, h)"
+"\<lbrakk>F \<in> Always({s \<in> state. f(s) = g(s)}); F \<in> Follows(A, r, f, h);
+ \<forall>x \<in> state. g(x):A\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, g, h)"
apply (unfold Follows_def Increasing_def Stable_def)
apply (simp add: INT_iff, auto)
apply (rule_tac [3] C = "{s \<in> state. f(s)=g(s)}"
@@ -291,8 +291,8 @@
lemma Always_Follows2:
-"[| F \<in> Always({s \<in> state. g(s) = h(s)});
- F \<in> Follows(A, r, f, g); \<forall>x \<in> state. h(x):A |] ==> F \<in> Follows(A, r, f, h)"
+"\<lbrakk>F \<in> Always({s \<in> state. g(s) = h(s)});
+ F \<in> Follows(A, r, f, g); \<forall>x \<in> state. h(x):A\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, f, h)"
apply (unfold Follows_def Increasing_def Stable_def)
apply (simp add: INT_iff, auto)
apply (rule_tac [3] C = "{s \<in> state. g (s) =h (s) }"
@@ -322,27 +322,27 @@
by (unfold part_order_def, auto)
lemma increasing_Un:
- "[| F \<in> Increasing.increasing(Pow(A), SetLe(A), f);
- F \<in> Increasing.increasing(Pow(A), SetLe(A), g) |]
- ==> F \<in> Increasing.increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
+ "\<lbrakk>F \<in> Increasing.increasing(Pow(A), SetLe(A), f);
+ F \<in> Increasing.increasing(Pow(A), SetLe(A), g)\<rbrakk>
+ \<Longrightarrow> F \<in> Increasing.increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
by (rule_tac h = "(Un)" in imp_increasing_comp2, auto)
lemma Increasing_Un:
- "[| F \<in> Increasing(Pow(A), SetLe(A), f);
- F \<in> Increasing(Pow(A), SetLe(A), g) |]
- ==> F \<in> Increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
+ "\<lbrakk>F \<in> Increasing(Pow(A), SetLe(A), f);
+ F \<in> Increasing(Pow(A), SetLe(A), g)\<rbrakk>
+ \<Longrightarrow> F \<in> Increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
by (rule_tac h = "(Un)" in imp_Increasing_comp2, auto)
lemma Always_Un:
- "[| F \<in> Always({s \<in> state. f1(s) \<subseteq> f(s)});
- F \<in> Always({s \<in> state. g1(s) \<subseteq> g(s)}) |]
- ==> F \<in> Always({s \<in> state. f1(s) \<union> g1(s) \<subseteq> f(s) \<union> g(s)})"
+ "\<lbrakk>F \<in> Always({s \<in> state. f1(s) \<subseteq> f(s)});
+ F \<in> Always({s \<in> state. g1(s) \<subseteq> g(s)})\<rbrakk>
+ \<Longrightarrow> F \<in> Always({s \<in> state. f1(s) \<union> g1(s) \<subseteq> f(s) \<union> g(s)})"
by (simp add: Always_eq_includes_reachable, blast)
lemma Follows_Un:
-"[| F \<in> Follows(Pow(A), SetLe(A), f1, f);
- F \<in> Follows(Pow(A), SetLe(A), g1, g) |]
- ==> F \<in> Follows(Pow(A), SetLe(A), %s. f1(s) \<union> g1(s), %s. f(s) \<union> g(s))"
+"\<lbrakk>F \<in> Follows(Pow(A), SetLe(A), f1, f);
+ F \<in> Follows(Pow(A), SetLe(A), g1, g)\<rbrakk>
+ \<Longrightarrow> F \<in> Follows(Pow(A), SetLe(A), %s. f1(s) \<union> g1(s), %s. f(s) \<union> g(s))"
by (rule_tac h = "(Un)" in imp_Follows_comp2, auto)
(** Multiset union properties (with the MultLe ordering) **)
@@ -351,12 +351,12 @@
by (unfold MultLe_def refl_def, auto)
lemma MultLe_refl1 [simp]:
- "[| multiset(M); mset_of(M)<=A |] ==> <M, M> \<in> MultLe(A, r)"
+ "\<lbrakk>multiset(M); mset_of(M)<=A\<rbrakk> \<Longrightarrow> <M, M> \<in> MultLe(A, r)"
apply (unfold MultLe_def id_def lam_def)
apply (auto simp add: Mult_iff_multiset)
done
-lemma MultLe_refl2 [simp]: "M \<in> Mult(A) ==> <M, M> \<in> MultLe(A, r)"
+lemma MultLe_refl2 [simp]: "M \<in> Mult(A) \<Longrightarrow> <M, M> \<in> MultLe(A, r)"
by (unfold MultLe_def id_def lam_def, auto)
@@ -371,14 +371,14 @@
done
lemma MultLe_trans:
- "[| <M,K> \<in> MultLe(A,r); <K,N> \<in> MultLe(A,r) |] ==> <M,N> \<in> MultLe(A,r)"
+ "\<lbrakk><M,K> \<in> MultLe(A,r); <K,N> \<in> MultLe(A,r)\<rbrakk> \<Longrightarrow> <M,N> \<in> MultLe(A,r)"
apply (cut_tac A=A in trans_on_MultLe)
apply (drule trans_onD, assumption)
apply (auto dest: MultLe_type [THEN subsetD])
done
lemma part_order_imp_part_ord:
- "part_order(A, r) ==> part_ord(A, r-id(A))"
+ "part_order(A, r) \<Longrightarrow> part_ord(A, r-id(A))"
apply (unfold part_order_def part_ord_def)
apply (simp add: refl_def id_def lam_def irrefl_def, auto)
apply (simp (no_asm) add: trans_on_def)
@@ -389,7 +389,7 @@
done
lemma antisym_MultLe [simp]:
- "part_order(A, r) ==> antisym(MultLe(A,r))"
+ "part_order(A, r) \<Longrightarrow> antisym(MultLe(A,r))"
apply (unfold MultLe_def antisym_def)
apply (drule part_order_imp_part_ord, auto)
apply (drule irrefl_on_multirel)
@@ -399,13 +399,13 @@
done
lemma part_order_MultLe [simp]:
- "part_order(A, r) ==> part_order(Mult(A), MultLe(A, r))"
+ "part_order(A, r) \<Longrightarrow> part_order(Mult(A), MultLe(A, r))"
apply (frule antisym_MultLe)
apply (auto simp add: part_order_def)
done
lemma empty_le_MultLe [simp]:
-"[| multiset(M); mset_of(M)<= A|] ==> <0, M> \<in> MultLe(A, r)"
+"\<lbrakk>multiset(M); mset_of(M)<= A\<rbrakk> \<Longrightarrow> <0, M> \<in> MultLe(A, r)"
apply (unfold MultLe_def)
apply (case_tac "M=0")
apply (auto simp add: FiniteFun.intros)
@@ -414,11 +414,11 @@
apply (auto simp add: Mult_iff_multiset)
done
-lemma empty_le_MultLe2 [simp]: "M \<in> Mult(A) ==> <0, M> \<in> MultLe(A, r)"
+lemma empty_le_MultLe2 [simp]: "M \<in> Mult(A) \<Longrightarrow> <0, M> \<in> MultLe(A, r)"
by (simp add: Mult_iff_multiset)
lemma munion_mono:
-"[| <M, N> \<in> MultLe(A, r); <K, L> \<in> MultLe(A, r) |] ==>
+"\<lbrakk><M, N> \<in> MultLe(A, r); <K, L> \<in> MultLe(A, r)\<rbrakk> \<Longrightarrow>
<M +# K, N +# L> \<in> MultLe(A, r)"
apply (unfold MultLe_def)
apply (auto intro: munion_multirel_mono1 munion_multirel_mono2
@@ -426,40 +426,40 @@
done
lemma increasing_munion:
- "[| F \<in> Increasing.increasing(Mult(A), MultLe(A,r), f);
- F \<in> Increasing.increasing(Mult(A), MultLe(A,r), g) |]
- ==> F \<in> Increasing.increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
+ "\<lbrakk>F \<in> Increasing.increasing(Mult(A), MultLe(A,r), f);
+ F \<in> Increasing.increasing(Mult(A), MultLe(A,r), g)\<rbrakk>
+ \<Longrightarrow> F \<in> Increasing.increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
by (rule_tac h = munion in imp_increasing_comp2, auto)
lemma Increasing_munion:
- "[| F \<in> Increasing(Mult(A), MultLe(A,r), f);
- F \<in> Increasing(Mult(A), MultLe(A,r), g)|]
- ==> F \<in> Increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
+ "\<lbrakk>F \<in> Increasing(Mult(A), MultLe(A,r), f);
+ F \<in> Increasing(Mult(A), MultLe(A,r), g)\<rbrakk>
+ \<Longrightarrow> F \<in> Increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
by (rule_tac h = munion in imp_Increasing_comp2, auto)
lemma Always_munion:
-"[| F \<in> Always({s \<in> state. <f1(s),f(s)> \<in> MultLe(A,r)});
+"\<lbrakk>F \<in> Always({s \<in> state. <f1(s),f(s)> \<in> MultLe(A,r)});
F \<in> Always({s \<in> state. <g1(s), g(s)> \<in> MultLe(A,r)});
- \<forall>x \<in> state. f1(x):Mult(A)&f(x):Mult(A) & g1(x):Mult(A) & g(x):Mult(A)|]
- ==> F \<in> Always({s \<in> state. <f1(s) +# g1(s), f(s) +# g(s)> \<in> MultLe(A,r)})"
+ \<forall>x \<in> state. f1(x):Mult(A)&f(x):Mult(A) & g1(x):Mult(A) & g(x):Mult(A)\<rbrakk>
+ \<Longrightarrow> F \<in> Always({s \<in> state. <f1(s) +# g1(s), f(s) +# g(s)> \<in> MultLe(A,r)})"
apply (rule_tac h = munion in imp_Always_comp2, simp_all)
apply (blast intro: munion_mono, simp_all)
done
lemma Follows_munion:
-"[| F \<in> Follows(Mult(A), MultLe(A, r), f1, f);
- F \<in> Follows(Mult(A), MultLe(A, r), g1, g) |]
- ==> F \<in> Follows(Mult(A), MultLe(A, r), %s. f1(s) +# g1(s), %s. f(s) +# g(s))"
+"\<lbrakk>F \<in> Follows(Mult(A), MultLe(A, r), f1, f);
+ F \<in> Follows(Mult(A), MultLe(A, r), g1, g)\<rbrakk>
+ \<Longrightarrow> F \<in> Follows(Mult(A), MultLe(A, r), %s. f1(s) +# g1(s), %s. f(s) +# g(s))"
by (rule_tac h = munion in imp_Follows_comp2, auto)
(** Used in ClientImp **)
lemma Follows_msetsum_UN:
-"!!f. [| \<forall>i \<in> I. F \<in> Follows(Mult(A), MultLe(A, r), f'(i), f(i));
+"\<And>f. \<lbrakk>\<forall>i \<in> I. F \<in> Follows(Mult(A), MultLe(A, r), f'(i), f(i));
\<forall>s. \<forall>i \<in> I. multiset(f'(i, s)) & mset_of(f'(i, s))<=A &
multiset(f(i, s)) & mset_of(f(i, s))<=A ;
- Finite(I); F \<in> program |]
- ==> F \<in> Follows(Mult(A),
+ Finite(I); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> Follows(Mult(A),
MultLe(A, r), %x. msetsum(%i. f'(i, x), I, A),
%x. msetsum(%i. f(i, x), I, A))"
apply (erule rev_mp)
--- a/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,7 +17,7 @@
definition (*really belongs in ZF/Trancl*)
part_order :: "[i, i] => o" where
- "part_order(A, r) == refl(A,r) & trans[A](r) & antisym(r)"
+ "part_order(A, r) \<equiv> refl(A,r) & trans[A](r) & antisym(r)"
consts
gen_prefix :: "[i, i] => i"
@@ -30,46 +30,46 @@
intros
Nil: "<[],[]>:gen_prefix(A, r)"
- prepend: "[| <xs,ys>:gen_prefix(A, r); <x,y>:r; x \<in> A; y \<in> A |]
- ==> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
+ prepend: "\<lbrakk><xs,ys>:gen_prefix(A, r); <x,y>:r; x \<in> A; y \<in> A\<rbrakk>
+ \<Longrightarrow> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
- append: "[| <xs,ys>:gen_prefix(A, r); zs:list(A) |]
- ==> <xs, ys@zs>:gen_prefix(A, r)"
+ append: "\<lbrakk><xs,ys>:gen_prefix(A, r); zs:list(A)\<rbrakk>
+ \<Longrightarrow> <xs, ys@zs>:gen_prefix(A, r)"
type_intros app_type list.Nil list.Cons
definition
prefix :: "i=>i" where
- "prefix(A) == gen_prefix(A, id(A))"
+ "prefix(A) \<equiv> gen_prefix(A, id(A))"
definition
strict_prefix :: "i=>i" where
- "strict_prefix(A) == prefix(A) - id(list(A))"
+ "strict_prefix(A) \<equiv> prefix(A) - id(list(A))"
(* less or equal and greater or equal over prefixes *)
abbreviation
pfixLe :: "[i, i] => o" (infixl \<open>pfixLe\<close> 50) where
- "xs pfixLe ys == <xs, ys>:gen_prefix(nat, Le)"
+ "xs pfixLe ys \<equiv> <xs, ys>:gen_prefix(nat, Le)"
abbreviation
pfixGe :: "[i, i] => o" (infixl \<open>pfixGe\<close> 50) where
- "xs pfixGe ys == <xs, ys>:gen_prefix(nat, Ge)"
+ "xs pfixGe ys \<equiv> <xs, ys>:gen_prefix(nat, Ge)"
lemma reflD:
- "[| refl(A, r); x \<in> A |] ==> <x,x>:r"
+ "\<lbrakk>refl(A, r); x \<in> A\<rbrakk> \<Longrightarrow> <x,x>:r"
apply (unfold refl_def, auto)
done
(*** preliminary lemmas ***)
-lemma Nil_gen_prefix: "xs \<in> list(A) ==> <[], xs> \<in> gen_prefix(A, r)"
+lemma Nil_gen_prefix: "xs \<in> list(A) \<Longrightarrow> <[], xs> \<in> gen_prefix(A, r)"
by (drule gen_prefix.append [OF gen_prefix.Nil], simp)
declare Nil_gen_prefix [simp]
-lemma gen_prefix_length_le: "<xs,ys> \<in> gen_prefix(A, r) ==> length(xs) \<le> length(ys)"
+lemma gen_prefix_length_le: "<xs,ys> \<in> gen_prefix(A, r) \<Longrightarrow> length(xs) \<le> length(ys)"
apply (erule gen_prefix.induct)
apply (subgoal_tac [3] "ys \<in> list (A) ")
apply (auto dest: gen_prefix.dom_subset [THEN subsetD]
@@ -78,8 +78,8 @@
lemma Cons_gen_prefix_aux:
- "[| <xs', ys'> \<in> gen_prefix(A, r) |]
- ==> (\<forall>x xs. x \<in> A \<longrightarrow> xs'= Cons(x,xs) \<longrightarrow>
+ "\<lbrakk><xs', ys'> \<in> gen_prefix(A, r)\<rbrakk>
+ \<Longrightarrow> (\<forall>x xs. x \<in> A \<longrightarrow> xs'= Cons(x,xs) \<longrightarrow>
(\<exists>y ys. y \<in> A & ys' = Cons(y,ys) &
<x,y>:r & <xs, ys> \<in> gen_prefix(A, r)))"
apply (erule gen_prefix.induct)
@@ -87,9 +87,9 @@
done
lemma Cons_gen_prefixE:
- "[| <Cons(x,xs), zs> \<in> gen_prefix(A, r);
- !!y ys. [|zs = Cons(y, ys); y \<in> A; x \<in> A; <x,y>:r;
- <xs,ys> \<in> gen_prefix(A, r) |] ==> P |] ==> P"
+ "\<lbrakk><Cons(x,xs), zs> \<in> gen_prefix(A, r);
+ \<And>y ys. \<lbrakk>zs = Cons(y, ys); y \<in> A; x \<in> A; <x,y>:r;
+ <xs,ys> \<in> gen_prefix(A, r)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
apply (frule gen_prefix.dom_subset [THEN subsetD], auto)
apply (blast dest: Cons_gen_prefix_aux)
done
@@ -104,7 +104,7 @@
(** Monotonicity of gen_prefix **)
-lemma gen_prefix_mono2: "r<=s ==> gen_prefix(A, r) \<subseteq> gen_prefix(A, s)"
+lemma gen_prefix_mono2: "r<=s \<Longrightarrow> gen_prefix(A, r) \<subseteq> gen_prefix(A, s)"
apply clarify
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
apply (erule rev_mp)
@@ -112,7 +112,7 @@
apply (auto intro: gen_prefix.append)
done
-lemma gen_prefix_mono1: "A<=B ==>gen_prefix(A, r) \<subseteq> gen_prefix(B, r)"
+lemma gen_prefix_mono1: "A<=B \<Longrightarrow>gen_prefix(A, r) \<subseteq> gen_prefix(B, r)"
apply clarify
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
apply (erule rev_mp)
@@ -126,7 +126,7 @@
intro: gen_prefix.append list_mono [THEN subsetD])
done
-lemma gen_prefix_mono: "[| A \<subseteq> B; r \<subseteq> s |] ==> gen_prefix(A, r) \<subseteq> gen_prefix(B, s)"
+lemma gen_prefix_mono: "\<lbrakk>A \<subseteq> B; r \<subseteq> s\<rbrakk> \<Longrightarrow> gen_prefix(A, r) \<subseteq> gen_prefix(B, s)"
apply (rule subset_trans)
apply (rule gen_prefix_mono1)
apply (rule_tac [2] gen_prefix_mono2, auto)
@@ -135,7 +135,7 @@
(*** gen_prefix order ***)
(* reflexivity *)
-lemma refl_gen_prefix: "refl(A, r) ==> refl(list(A), gen_prefix(A, r))"
+lemma refl_gen_prefix: "refl(A, r) \<Longrightarrow> refl(list(A), gen_prefix(A, r))"
apply (unfold refl_def, auto)
apply (induct_tac "x", auto)
done
@@ -144,7 +144,7 @@
(* Transitivity *)
(* A lemma for proving gen_prefix_trans_comp *)
-lemma append_gen_prefix [rule_format (no_asm)]: "xs \<in> list(A) ==>
+lemma append_gen_prefix [rule_format (no_asm)]: "xs \<in> list(A) \<Longrightarrow>
\<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) \<longrightarrow> <xs, zs>: gen_prefix(A, r)"
apply (erule list.induct)
apply (auto dest: gen_prefix.dom_subset [THEN subsetD])
@@ -153,7 +153,7 @@
(* Lemma proving transitivity and more*)
lemma gen_prefix_trans_comp [rule_format (no_asm)]:
- "<x, y>: gen_prefix(A, r) ==>
+ "<x, y>: gen_prefix(A, r) \<Longrightarrow>
(\<forall>z \<in> list(A). <y,z> \<in> gen_prefix(A, s)\<longrightarrow><x, z> \<in> gen_prefix(A, s O r))"
apply (erule gen_prefix.induct)
apply (auto elim: ConsE simp add: Nil_gen_prefix)
@@ -162,10 +162,10 @@
apply (drule_tac xs = ys and r = s in append_gen_prefix, auto)
done
-lemma trans_comp_subset: "trans(r) ==> r O r \<subseteq> r"
+lemma trans_comp_subset: "trans(r) \<Longrightarrow> r O r \<subseteq> r"
by (auto dest: transD)
-lemma trans_gen_prefix: "trans(r) ==> trans(gen_prefix(A,r))"
+lemma trans_gen_prefix: "trans(r) \<Longrightarrow> trans(gen_prefix(A,r))"
apply (simp (no_asm) add: trans_def)
apply clarify
apply (rule trans_comp_subset [THEN gen_prefix_mono2, THEN subsetD], assumption)
@@ -174,14 +174,14 @@
done
lemma trans_on_gen_prefix:
- "trans(r) ==> trans[list(A)](gen_prefix(A, r))"
+ "trans(r) \<Longrightarrow> trans[list(A)](gen_prefix(A, r))"
apply (drule_tac A = A in trans_gen_prefix)
apply (unfold trans_def trans_on_def, blast)
done
lemma prefix_gen_prefix_trans:
- "[| <x,y> \<in> prefix(A); <y, z> \<in> gen_prefix(A, r); r<=A*A |]
- ==> <x, z> \<in> gen_prefix(A, r)"
+ "\<lbrakk><x,y> \<in> prefix(A); <y, z> \<in> gen_prefix(A, r); r<=A*A\<rbrakk>
+ \<Longrightarrow> <x, z> \<in> gen_prefix(A, r)"
apply (unfold prefix_def)
apply (rule_tac P = "%r. <x,z> \<in> gen_prefix (A, r) " in right_comp_id [THEN subst])
apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+
@@ -189,8 +189,8 @@
lemma gen_prefix_prefix_trans:
-"[| <x,y> \<in> gen_prefix(A,r); <y, z> \<in> prefix(A); r<=A*A |]
- ==> <x, z> \<in> gen_prefix(A, r)"
+"\<lbrakk><x,y> \<in> gen_prefix(A,r); <y, z> \<in> prefix(A); r<=A*A\<rbrakk>
+ \<Longrightarrow> <x, z> \<in> gen_prefix(A, r)"
apply (unfold prefix_def)
apply (rule_tac P = "%r. <x,z> \<in> gen_prefix (A, r) " in left_comp_id [THEN subst])
apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+
@@ -198,10 +198,10 @@
(** Antisymmetry **)
-lemma nat_le_lemma [rule_format]: "n \<in> nat ==> \<forall>b \<in> nat. n #+ b \<le> n \<longrightarrow> b = 0"
+lemma nat_le_lemma [rule_format]: "n \<in> nat \<Longrightarrow> \<forall>b \<in> nat. n #+ b \<le> n \<longrightarrow> b = 0"
by (induct_tac "n", auto)
-lemma antisym_gen_prefix: "antisym(r) ==> antisym(gen_prefix(A, r))"
+lemma antisym_gen_prefix: "antisym(r) \<Longrightarrow> antisym(gen_prefix(A, r))"
apply (simp (no_asm) add: antisym_def)
apply (rule impI [THEN allI, THEN allI])
apply (erule gen_prefix.induct, blast)
@@ -225,12 +225,12 @@
(*** recursion equations ***)
-lemma gen_prefix_Nil: "xs \<in> list(A) ==> <xs, []> \<in> gen_prefix(A,r) \<longleftrightarrow> (xs = [])"
+lemma gen_prefix_Nil: "xs \<in> list(A) \<Longrightarrow> <xs, []> \<in> gen_prefix(A,r) \<longleftrightarrow> (xs = [])"
by (induct_tac "xs", auto)
declare gen_prefix_Nil [simp]
lemma same_gen_prefix_gen_prefix:
- "[| refl(A, r); xs \<in> list(A) |] ==>
+ "\<lbrakk>refl(A, r); xs \<in> list(A)\<rbrakk> \<Longrightarrow>
<xs@ys, xs@zs>: gen_prefix(A, r) \<longleftrightarrow> <ys,zs> \<in> gen_prefix(A, r)"
apply (unfold refl_def)
apply (induct_tac "xs")
@@ -238,14 +238,14 @@
done
declare same_gen_prefix_gen_prefix [simp]
-lemma gen_prefix_Cons: "[| xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
+lemma gen_prefix_Cons: "\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
<xs, Cons(y,ys)> \<in> gen_prefix(A,r) \<longleftrightarrow>
(xs=[] | (\<exists>z zs. xs=Cons(z,zs) & z \<in> A & <z,y>:r & <zs,ys> \<in> gen_prefix(A,r)))"
apply (induct_tac "xs", auto)
done
-lemma gen_prefix_take_append: "[| refl(A,r); <xs,ys> \<in> gen_prefix(A, r); zs \<in> list(A) |]
- ==> <xs@zs, take(length(xs), ys) @ zs> \<in> gen_prefix(A, r)"
+lemma gen_prefix_take_append: "\<lbrakk>refl(A,r); <xs,ys> \<in> gen_prefix(A, r); zs \<in> list(A)\<rbrakk>
+ \<Longrightarrow> <xs@zs, take(length(xs), ys) @ zs> \<in> gen_prefix(A, r)"
apply (erule gen_prefix.induct)
apply (simp (no_asm_simp))
apply (frule_tac [!] gen_prefix.dom_subset [THEN subsetD], auto)
@@ -254,21 +254,21 @@
apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type)
done
-lemma gen_prefix_append_both: "[| refl(A, r); <xs,ys> \<in> gen_prefix(A,r);
- length(xs) = length(ys); zs \<in> list(A) |]
- ==> <xs@zs, ys @ zs> \<in> gen_prefix(A, r)"
+lemma gen_prefix_append_both: "\<lbrakk>refl(A, r); <xs,ys> \<in> gen_prefix(A,r);
+ length(xs) = length(ys); zs \<in> list(A)\<rbrakk>
+ \<Longrightarrow> <xs@zs, ys @ zs> \<in> gen_prefix(A, r)"
apply (drule_tac zs = zs in gen_prefix_take_append, assumption+)
apply (subgoal_tac "take (length (xs), ys) =ys")
apply (auto intro!: take_all dest: gen_prefix.dom_subset [THEN subsetD])
done
(*NOT suitable for rewriting since [y] has the form y#ys*)
-lemma append_cons_conv: "xs \<in> list(A) ==> xs @ Cons(y, ys) = (xs @ [y]) @ ys"
+lemma append_cons_conv: "xs \<in> list(A) \<Longrightarrow> xs @ Cons(y, ys) = (xs @ [y]) @ ys"
by (auto simp add: app_assoc)
lemma append_one_gen_prefix_lemma [rule_format]:
- "[| <xs,ys> \<in> gen_prefix(A, r); refl(A, r) |]
- ==> length(xs) < length(ys) \<longrightarrow>
+ "\<lbrakk><xs,ys> \<in> gen_prefix(A, r); refl(A, r)\<rbrakk>
+ \<Longrightarrow> length(xs) < length(ys) \<longrightarrow>
<xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
apply (erule gen_prefix.induct, blast)
apply (frule gen_prefix.dom_subset [THEN subsetD], clarify)
@@ -290,15 +290,15 @@
apply (auto elim: ConsE simp add: gen_prefix_append_both)
done
-lemma append_one_gen_prefix: "[| <xs,ys>: gen_prefix(A, r); length(xs) < length(ys); refl(A, r) |]
- ==> <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
+lemma append_one_gen_prefix: "\<lbrakk><xs,ys>: gen_prefix(A, r); length(xs) < length(ys); refl(A, r)\<rbrakk>
+ \<Longrightarrow> <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
apply (blast intro: append_one_gen_prefix_lemma)
done
(** Proving the equivalence with Charpentier's definition **)
-lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \<in> list(A) ==>
+lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \<in> list(A) \<Longrightarrow>
\<forall>ys \<in> list(A). \<forall>i \<in> nat. i < length(xs)
\<longrightarrow> <xs, ys>: gen_prefix(A, r) \<longrightarrow> <nth(i, xs), nth(i, ys)>:r"
apply (induct_tac "xs", simp, clarify)
@@ -306,14 +306,14 @@
apply (erule natE, auto)
done
-lemma gen_prefix_imp_nth: "[| <xs,ys> \<in> gen_prefix(A,r); i < length(xs)|]
- ==> <nth(i, xs), nth(i, ys)>:r"
+lemma gen_prefix_imp_nth: "\<lbrakk><xs,ys> \<in> gen_prefix(A,r); i < length(xs)\<rbrakk>
+ \<Longrightarrow> <nth(i, xs), nth(i, ys)>:r"
apply (cut_tac A = A in gen_prefix.dom_subset)
apply (rule gen_prefix_imp_nth_lemma)
apply (auto simp add: lt_nat_in_nat)
done
-lemma nth_imp_gen_prefix [rule_format]: "xs \<in> list(A) ==>
+lemma nth_imp_gen_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow>
\<forall>ys \<in> list(A). length(xs) \<le> length(ys)
\<longrightarrow> (\<forall>i. i < length(xs) \<longrightarrow> <nth(i, xs), nth(i,ys)>:r)
\<longrightarrow> <xs, ys> \<in> gen_prefix(A, r)"
@@ -363,7 +363,7 @@
(* Monotonicity of "set" operator WRT prefix *)
lemma set_of_list_prefix_mono:
-"<xs,ys> \<in> prefix(A) ==> set_of_list(xs) \<subseteq> set_of_list(ys)"
+"<xs,ys> \<in> prefix(A) \<Longrightarrow> set_of_list(xs) \<subseteq> set_of_list(ys)"
apply (unfold prefix_def)
apply (erule gen_prefix.induct)
@@ -374,7 +374,7 @@
(** recursion equations **)
-lemma Nil_prefix: "xs \<in> list(A) ==> <[],xs> \<in> prefix(A)"
+lemma Nil_prefix: "xs \<in> list(A) \<Longrightarrow> <[],xs> \<in> prefix(A)"
apply (unfold prefix_def)
apply (simp (no_asm_simp) add: Nil_gen_prefix)
@@ -398,7 +398,7 @@
declare Cons_prefix_Cons [iff]
lemma same_prefix_prefix:
-"xs \<in> list(A)==> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (<ys,zs> \<in> prefix(A))"
+"xs \<in> list(A)\<Longrightarrow> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (<ys,zs> \<in> prefix(A))"
apply (unfold prefix_def)
apply (subgoal_tac "refl (A,id (A))")
apply (simp (no_asm_simp))
@@ -406,21 +406,21 @@
done
declare same_prefix_prefix [simp]
-lemma same_prefix_prefix_Nil: "xs \<in> list(A) ==> <xs@ys,xs> \<in> prefix(A) \<longleftrightarrow> (<ys,[]> \<in> prefix(A))"
+lemma same_prefix_prefix_Nil: "xs \<in> list(A) \<Longrightarrow> <xs@ys,xs> \<in> prefix(A) \<longleftrightarrow> (<ys,[]> \<in> prefix(A))"
apply (rule_tac P = "%x. <u, x>:v \<longleftrightarrow> w(x)" for u v w in app_right_Nil [THEN subst])
apply (rule_tac [2] same_prefix_prefix, auto)
done
declare same_prefix_prefix_Nil [simp]
lemma prefix_appendI:
-"[| <xs,ys> \<in> prefix(A); zs \<in> list(A) |] ==> <xs,ys@zs> \<in> prefix(A)"
+"\<lbrakk><xs,ys> \<in> prefix(A); zs \<in> list(A)\<rbrakk> \<Longrightarrow> <xs,ys@zs> \<in> prefix(A)"
apply (unfold prefix_def)
apply (erule gen_prefix.append, assumption)
done
declare prefix_appendI [simp]
lemma prefix_Cons:
-"[| xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
+"\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
<xs,Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow>
(xs=[] | (\<exists>zs. xs=Cons(y,zs) & <zs,ys> \<in> prefix(A)))"
apply (unfold prefix_def)
@@ -428,8 +428,8 @@
done
lemma append_one_prefix:
- "[| <xs,ys> \<in> prefix(A); length(xs) < length(ys) |]
- ==> <xs @ [nth(length(xs),ys)], ys> \<in> prefix(A)"
+ "\<lbrakk><xs,ys> \<in> prefix(A); length(xs) < length(ys)\<rbrakk>
+ \<Longrightarrow> <xs @ [nth(length(xs),ys)], ys> \<in> prefix(A)"
apply (unfold prefix_def)
apply (subgoal_tac "refl (A, id (A))")
apply (simp (no_asm_simp) add: append_one_gen_prefix)
@@ -437,7 +437,7 @@
done
lemma prefix_length_le:
-"<xs,ys> \<in> prefix(A) ==> length(xs) \<le> length(ys)"
+"<xs,ys> \<in> prefix(A) \<Longrightarrow> length(xs) \<le> length(ys)"
apply (unfold prefix_def)
apply (blast dest: gen_prefix_length_le)
done
@@ -454,7 +454,7 @@
done
lemma strict_prefix_length_lt_aux:
- "<xs,ys> \<in> prefix(A) ==> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)"
+ "<xs,ys> \<in> prefix(A) \<Longrightarrow> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)"
apply (unfold prefix_def)
apply (erule gen_prefix.induct, clarify)
apply (subgoal_tac [!] "ys \<in> list(A) & xs \<in> list(A)")
@@ -467,7 +467,7 @@
done
lemma strict_prefix_length_lt:
- "<xs,ys>:strict_prefix(A) ==> length(xs) < length(ys)"
+ "<xs,ys>:strict_prefix(A) \<Longrightarrow> length(xs) < length(ys)"
apply (unfold strict_prefix_def)
apply (rule strict_prefix_length_lt_aux [THEN mp])
apply (auto dest: prefix_type [THEN subsetD])
@@ -493,7 +493,7 @@
done
lemma prefix_snoc:
-"[|xs \<in> list(A); ys \<in> list(A); y \<in> A |] ==>
+"\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
<xs, ys@[y]> \<in> prefix(A) \<longleftrightarrow> (xs = ys@[y] | <xs,ys> \<in> prefix(A))"
apply (simp (no_asm) add: prefix_iff)
apply (rule iffI, clarify)
@@ -503,7 +503,7 @@
done
declare prefix_snoc [simp]
-lemma prefix_append_iff [rule_format]: "zs \<in> list(A) ==> \<forall>xs \<in> list(A). \<forall>ys \<in> list(A).
+lemma prefix_append_iff [rule_format]: "zs \<in> list(A) \<Longrightarrow> \<forall>xs \<in> list(A). \<forall>ys \<in> list(A).
(<xs, ys@zs> \<in> prefix(A)) \<longleftrightarrow>
(<xs,ys> \<in> prefix(A) | (\<exists>us. xs = ys@us & <us,zs> \<in> prefix(A)))"
apply (erule list_append_induct, force, clarify)
@@ -518,14 +518,14 @@
(*Although the prefix ordering is not linear, the prefixes of a list
are linearly ordered.*)
-lemma common_prefix_linear_lemma [rule_format]: "[| zs \<in> list(A); xs \<in> list(A); ys \<in> list(A) |]
- ==> <xs, zs> \<in> prefix(A) \<longrightarrow> <ys,zs> \<in> prefix(A)
+lemma common_prefix_linear_lemma [rule_format]: "\<lbrakk>zs \<in> list(A); xs \<in> list(A); ys \<in> list(A)\<rbrakk>
+ \<Longrightarrow> <xs, zs> \<in> prefix(A) \<longrightarrow> <ys,zs> \<in> prefix(A)
\<longrightarrow><xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
apply (erule list_append_induct, auto)
done
-lemma common_prefix_linear: "[|<xs, zs> \<in> prefix(A); <ys,zs> \<in> prefix(A) |]
- ==> <xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
+lemma common_prefix_linear: "\<lbrakk><xs, zs> \<in> prefix(A); <ys,zs> \<in> prefix(A)\<rbrakk>
+ \<Longrightarrow> <xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
apply (cut_tac prefix_type)
apply (blast del: disjCI intro: common_prefix_linear_lemma)
done
@@ -565,19 +565,19 @@
by (unfold part_order_def, auto)
declare part_order_Le [simp]
-lemma pfixLe_refl: "x \<in> list(nat) ==> x pfixLe x"
+lemma pfixLe_refl: "x \<in> list(nat) \<Longrightarrow> x pfixLe x"
by (blast intro: refl_gen_prefix [THEN reflD] refl_Le)
declare pfixLe_refl [simp]
-lemma pfixLe_trans: "[| x pfixLe y; y pfixLe z |] ==> x pfixLe z"
+lemma pfixLe_trans: "\<lbrakk>x pfixLe y; y pfixLe z\<rbrakk> \<Longrightarrow> x pfixLe z"
by (blast intro: trans_gen_prefix [THEN transD] trans_Le)
-lemma pfixLe_antisym: "[| x pfixLe y; y pfixLe x |] ==> x = y"
+lemma pfixLe_antisym: "\<lbrakk>x pfixLe y; y pfixLe x\<rbrakk> \<Longrightarrow> x = y"
by (blast intro: antisym_gen_prefix [THEN antisymE] antisym_Le)
lemma prefix_imp_pfixLe:
-"<xs,ys>:prefix(nat)==> xs pfixLe ys"
+"<xs,ys>:prefix(nat)\<Longrightarrow> xs pfixLe ys"
apply (unfold prefix_def)
apply (rule gen_prefix_mono [THEN subsetD], auto)
@@ -599,25 +599,25 @@
done
declare trans_Ge [iff]
-lemma pfixGe_refl: "x \<in> list(nat) ==> x pfixGe x"
+lemma pfixGe_refl: "x \<in> list(nat) \<Longrightarrow> x pfixGe x"
by (blast intro: refl_gen_prefix [THEN reflD])
declare pfixGe_refl [simp]
-lemma pfixGe_trans: "[| x pfixGe y; y pfixGe z |] ==> x pfixGe z"
+lemma pfixGe_trans: "\<lbrakk>x pfixGe y; y pfixGe z\<rbrakk> \<Longrightarrow> x pfixGe z"
by (blast intro: trans_gen_prefix [THEN transD])
-lemma pfixGe_antisym: "[| x pfixGe y; y pfixGe x |] ==> x = y"
+lemma pfixGe_antisym: "\<lbrakk>x pfixGe y; y pfixGe x\<rbrakk> \<Longrightarrow> x = y"
by (blast intro: antisym_gen_prefix [THEN antisymE])
lemma prefix_imp_pfixGe:
- "<xs,ys>:prefix(nat) ==> xs pfixGe ys"
+ "<xs,ys>:prefix(nat) \<Longrightarrow> xs pfixGe ys"
apply (unfold prefix_def Ge_def)
apply (rule gen_prefix_mono [THEN subsetD], auto)
done
(* Added by Sidi \<in> prefix and take *)
lemma prefix_imp_take:
-"<xs, ys> \<in> prefix(A) ==> xs = take(length(xs), ys)"
+"<xs, ys> \<in> prefix(A) \<Longrightarrow> xs = take(length(xs), ys)"
apply (unfold prefix_def)
apply (erule gen_prefix.induct)
@@ -630,17 +630,17 @@
apply (simp_all (no_asm_simp) add: diff_is_0_iff)
done
-lemma prefix_length_equal: "[|<xs,ys> \<in> prefix(A); length(xs)=length(ys)|] ==> xs = ys"
+lemma prefix_length_equal: "\<lbrakk><xs,ys> \<in> prefix(A); length(xs)=length(ys)\<rbrakk> \<Longrightarrow> xs = ys"
apply (cut_tac A = A in prefix_type)
apply (drule subsetD, auto)
apply (drule prefix_imp_take)
apply (erule trans, simp)
done
-lemma prefix_length_le_equal: "[|<xs,ys> \<in> prefix(A); length(ys) \<le> length(xs)|] ==> xs = ys"
+lemma prefix_length_le_equal: "\<lbrakk><xs,ys> \<in> prefix(A); length(ys) \<le> length(xs)\<rbrakk> \<Longrightarrow> xs = ys"
by (blast intro: prefix_length_equal le_anti_sym prefix_length_le)
-lemma take_prefix [rule_format]: "xs \<in> list(A) ==> \<forall>n \<in> nat. <take(n, xs), xs> \<in> prefix(A)"
+lemma take_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow> \<forall>n \<in> nat. <take(n, xs), xs> \<in> prefix(A)"
apply (unfold prefix_def)
apply (erule list.induct, simp, clarify)
apply (erule natE, auto)
@@ -654,21 +654,21 @@
apply (blast intro: take_prefix length_type)
done
-lemma prefix_imp_nth: "[| <xs,ys> \<in> prefix(A); i < length(xs)|] ==> nth(i,xs) = nth(i,ys)"
+lemma prefix_imp_nth: "\<lbrakk><xs,ys> \<in> prefix(A); i < length(xs)\<rbrakk> \<Longrightarrow> nth(i,xs) = nth(i,ys)"
by (auto dest!: gen_prefix_imp_nth simp add: prefix_def)
lemma nth_imp_prefix:
- "[|xs \<in> list(A); ys \<in> list(A); length(xs) \<le> length(ys);
- !!i. i < length(xs) ==> nth(i, xs) = nth(i,ys)|]
- ==> <xs,ys> \<in> prefix(A)"
+ "\<lbrakk>xs \<in> list(A); ys \<in> list(A); length(xs) \<le> length(ys);
+ \<And>i. i < length(xs) \<Longrightarrow> nth(i, xs) = nth(i,ys)\<rbrakk>
+ \<Longrightarrow> <xs,ys> \<in> prefix(A)"
apply (auto simp add: prefix_def nth_imp_gen_prefix)
apply (auto intro!: nth_imp_gen_prefix simp add: prefix_def)
apply (blast intro: nth_type lt_trans2)
done
-lemma length_le_prefix_imp_prefix: "[|length(xs) \<le> length(ys);
- <xs,zs> \<in> prefix(A); <ys,zs> \<in> prefix(A)|] ==> <xs,ys> \<in> prefix(A)"
+lemma length_le_prefix_imp_prefix: "\<lbrakk>length(xs) \<le> length(ys);
+ <xs,zs> \<in> prefix(A); <ys,zs> \<in> prefix(A)\<rbrakk> \<Longrightarrow> <xs,ys> \<in> prefix(A)"
apply (cut_tac A = A in prefix_type)
apply (rule nth_imp_prefix, blast, blast)
apply assumption
--- a/src/ZF/UNITY/Guar.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Guar.thy Tue Sep 27 16:51:35 2022 +0100
@@ -25,7 +25,7 @@
(* To be moved to theory WFair???? *)
-lemma leadsTo_Basis': "[| F \<in> A co A \<union> B; F \<in> transient(A); st_set(B) |] ==> F \<in> A \<longmapsto> B"
+lemma leadsTo_Basis': "\<lbrakk>F \<in> A co A \<union> B; F \<in> transient(A); st_set(B)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (frule constrainsD2)
apply (drule_tac B = "A-B" in constrains_weaken_L, blast)
apply (drule_tac B = "A-B" in transient_strengthen, blast)
@@ -38,63 +38,63 @@
definition
ex_prop :: "i => o" where
- "ex_prop(X) == X<=program &
+ "ex_prop(X) \<equiv> X<=program &
(\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X | G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X)"
definition
strict_ex_prop :: "i => o" where
- "strict_ex_prop(X) == X<=program &
+ "strict_ex_prop(X) \<equiv> X<=program &
(\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F \<in> X | G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X))"
definition
uv_prop :: "i => o" where
- "uv_prop(X) == X<=program &
+ "uv_prop(X) \<equiv> X<=program &
(SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X & G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X))"
definition
strict_uv_prop :: "i => o" where
- "strict_uv_prop(X) == X<=program &
+ "strict_uv_prop(X) \<equiv> X<=program &
(SKIP \<in> X & (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow>(F \<in> X & G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X)))"
definition
guar :: "[i, i] => i" (infixl \<open>guarantees\<close> 55) where
(*higher than membership, lower than Co*)
- "X guarantees Y == {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X \<longrightarrow> F \<squnion> G \<in> Y}"
+ "X guarantees Y \<equiv> {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X \<longrightarrow> F \<squnion> G \<in> Y}"
definition
(* Weakest guarantees *)
wg :: "[i,i] => i" where
- "wg(F,Y) == \<Union>({X \<in> Pow(program). F:(X guarantees Y)})"
+ "wg(F,Y) \<equiv> \<Union>({X \<in> Pow(program). F:(X guarantees Y)})"
definition
(* Weakest existential property stronger than X *)
wx :: "i =>i" where
- "wx(X) == \<Union>({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
+ "wx(X) \<equiv> \<Union>({Y \<in> Pow(program). Y<=X & ex_prop(Y)})"
definition
(*Ill-defined programs can arise through "\<squnion>"*)
welldef :: i where
- "welldef == {F \<in> program. Init(F) \<noteq> 0}"
+ "welldef \<equiv> {F \<in> program. Init(F) \<noteq> 0}"
definition
refines :: "[i, i, i] => o" (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10) where
- "G refines F wrt X ==
+ "G refines F wrt X \<equiv>
\<forall>H \<in> program. (F ok H & G ok H & F \<squnion> H \<in> welldef \<inter> X)
\<longrightarrow> (G \<squnion> H \<in> welldef \<inter> X)"
definition
iso_refines :: "[i,i, i] => o" (\<open>(3_ iso'_refines _ wrt _)\<close> [10,10,10] 10) where
- "G iso_refines F wrt X == F \<in> welldef \<inter> X \<longrightarrow> G \<in> welldef \<inter> X"
+ "G iso_refines F wrt X \<equiv> F \<in> welldef \<inter> X \<longrightarrow> G \<in> welldef \<inter> X"
(*** existential properties ***)
-lemma ex_imp_subset_program: "ex_prop(X) ==> X\<subseteq>program"
+lemma ex_imp_subset_program: "ex_prop(X) \<Longrightarrow> X\<subseteq>program"
by (simp add: ex_prop_def)
lemma ex1 [rule_format]:
"GG \<in> Fin(program)
- ==> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (%G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X"
+ \<Longrightarrow> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (%G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X"
apply (unfold ex_prop_def)
apply (erule Fin_induct)
apply (simp_all add: OK_cons_iff)
@@ -102,7 +102,7 @@
done
lemma ex2 [rule_format]:
-"X \<subseteq> program ==>
+"X \<subseteq> program \<Longrightarrow>
(\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<longrightarrow> OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X)
\<longrightarrow> ex_prop(X)"
apply (unfold ex_prop_def, clarify)
@@ -131,13 +131,13 @@
(*** universal properties ***)
-lemma uv_imp_subset_program: "uv_prop(X)==> X\<subseteq>program"
+lemma uv_imp_subset_program: "uv_prop(X)\<Longrightarrow> X\<subseteq>program"
apply (unfold uv_prop_def)
apply (simp (no_asm_simp))
done
lemma uv1 [rule_format]:
- "GG \<in> Fin(program) ==>
+ "GG \<in> Fin(program) \<Longrightarrow>
(uv_prop(X)\<longrightarrow> GG \<subseteq> X & OK(GG, (%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)"
apply (unfold uv_prop_def)
apply (erule Fin_induct)
@@ -145,7 +145,7 @@
done
lemma uv2 [rule_format]:
- "X\<subseteq>program ==>
+ "X\<subseteq>program \<Longrightarrow>
(\<forall>GG \<in> Fin(program). GG \<subseteq> X & OK(GG,(%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)
\<longrightarrow> uv_prop(X)"
apply (unfold uv_prop_def, auto)
@@ -166,14 +166,14 @@
(*** guarantees ***)
lemma guaranteesI:
- "[| (!!G. [| F ok G; F \<squnion> G \<in> X; G \<in> program |] ==> F \<squnion> G \<in> Y);
- F \<in> program |]
- ==> F \<in> X guarantees Y"
+ "\<lbrakk>(\<And>G. \<lbrakk>F ok G; F \<squnion> G \<in> X; G \<in> program\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> Y);
+ F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> X guarantees Y"
by (simp add: guar_def component_def)
lemma guaranteesD:
- "[| F \<in> X guarantees Y; F ok G; F \<squnion> G \<in> X; G \<in> program |]
- ==> F \<squnion> G \<in> Y"
+ "\<lbrakk>F \<in> X guarantees Y; F ok G; F \<squnion> G \<in> X; G \<in> program\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> Y"
by (simp add: guar_def component_def)
@@ -181,41 +181,41 @@
The major premise can no longer be F\<subseteq>H since we need to reason about G*)
lemma component_guaranteesD:
- "[| F \<in> X guarantees Y; F \<squnion> G = H; H \<in> X; F ok G; G \<in> program |]
- ==> H \<in> Y"
+ "\<lbrakk>F \<in> X guarantees Y; F \<squnion> G = H; H \<in> X; F ok G; G \<in> program\<rbrakk>
+ \<Longrightarrow> H \<in> Y"
by (simp add: guar_def, blast)
lemma guarantees_weaken:
- "[| F \<in> X guarantees X'; Y \<subseteq> X; X' \<subseteq> Y' |] ==> F \<in> Y guarantees Y'"
+ "\<lbrakk>F \<in> X guarantees X'; Y \<subseteq> X; X' \<subseteq> Y'\<rbrakk> \<Longrightarrow> F \<in> Y guarantees Y'"
by (simp add: guar_def, auto)
lemma subset_imp_guarantees_program:
- "X \<subseteq> Y ==> X guarantees Y = program"
+ "X \<subseteq> Y \<Longrightarrow> X guarantees Y = program"
by (unfold guar_def, blast)
(*Equivalent to subset_imp_guarantees_UNIV but more intuitive*)
lemma subset_imp_guarantees:
- "[| X \<subseteq> Y; F \<in> program |] ==> F \<in> X guarantees Y"
+ "\<lbrakk>X \<subseteq> Y; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> X guarantees Y"
by (unfold guar_def, blast)
-lemma component_of_Join1: "F ok G ==> F component_of (F \<squnion> G)"
+lemma component_of_Join1: "F ok G \<Longrightarrow> F component_of (F \<squnion> G)"
by (unfold component_of_def, blast)
-lemma component_of_Join2: "F ok G ==> G component_of (F \<squnion> G)"
+lemma component_of_Join2: "F ok G \<Longrightarrow> G component_of (F \<squnion> G)"
apply (subst Join_commute)
apply (blast intro: ok_sym component_of_Join1)
done
(*Remark at end of section 4.1 *)
lemma ex_prop_imp:
- "ex_prop(Y) ==> (Y = (program guarantees Y))"
+ "ex_prop(Y) \<Longrightarrow> (Y = (program guarantees Y))"
apply (simp (no_asm_use) add: ex_prop_equiv guar_def component_of_def)
apply clarify
apply (rule equalityI, blast, safe)
apply (drule_tac x = x in bspec, assumption, force)
done
-lemma guarantees_imp: "(Y = program guarantees Y) ==> ex_prop(Y)"
+lemma guarantees_imp: "(Y = program guarantees Y) \<Longrightarrow> ex_prop(Y)"
apply (unfold guar_def)
apply (simp (no_asm_simp) add: ex_prop_equiv)
apply safe
@@ -230,7 +230,7 @@
(** Distributive laws. Re-orient to perform miniscoping **)
lemma guarantees_UN_left:
- "i \<in> I ==>(\<Union>i \<in> I. X(i)) guarantees Y = (\<Inter>i \<in> I. X(i) guarantees Y)"
+ "i \<in> I \<Longrightarrow>(\<Union>i \<in> I. X(i)) guarantees Y = (\<Inter>i \<in> I. X(i) guarantees Y)"
apply (unfold guar_def)
apply (rule equalityI, safe)
prefer 2 apply force
@@ -244,7 +244,7 @@
done
lemma guarantees_INT_right:
- "i \<in> I ==> X guarantees (\<Inter>i \<in> I. Y(i)) = (\<Inter>i \<in> I. X guarantees Y(i))"
+ "i \<in> I \<Longrightarrow> X guarantees (\<Inter>i \<in> I. Y(i)) = (\<Inter>i \<in> I. X guarantees Y(i))"
apply (unfold guar_def)
apply (rule equalityI, safe, blast+)
done
@@ -254,12 +254,12 @@
by (unfold guar_def, blast)
lemma guarantees_Int_right_I:
- "[| F \<in> Z guarantees X; F \<in> Z guarantees Y |]
- ==> F \<in> Z guarantees (X \<inter> Y)"
+ "\<lbrakk>F \<in> Z guarantees X; F \<in> Z guarantees Y\<rbrakk>
+ \<Longrightarrow> F \<in> Z guarantees (X \<inter> Y)"
by (simp (no_asm_simp) add: guarantees_Int_right)
lemma guarantees_INT_right_iff:
- "i \<in> I==> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) \<longleftrightarrow>
+ "i \<in> I\<Longrightarrow> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) \<longleftrightarrow>
(\<forall>i \<in> I. F \<in> X guarantees Y(i))"
by (simp add: guarantees_INT_right INT_iff, blast)
@@ -276,36 +276,36 @@
**)
lemma combining1:
- "[| F \<in> V guarantees X; F \<in> (X \<inter> Y) guarantees Z |]
- ==> F \<in> (V \<inter> Y) guarantees Z"
+ "\<lbrakk>F \<in> V guarantees X; F \<in> (X \<inter> Y) guarantees Z\<rbrakk>
+ \<Longrightarrow> F \<in> (V \<inter> Y) guarantees Z"
by (unfold guar_def, blast)
lemma combining2:
- "[| F \<in> V guarantees (X \<union> Y); F \<in> Y guarantees Z |]
- ==> F \<in> V guarantees (X \<union> Z)"
+ "\<lbrakk>F \<in> V guarantees (X \<union> Y); F \<in> Y guarantees Z\<rbrakk>
+ \<Longrightarrow> F \<in> V guarantees (X \<union> Z)"
by (unfold guar_def, blast)
(** The following two follow Chandy-Sanders, but the use of object-quantifiers
does not suit Isabelle... **)
-(*Premise should be (!!i. i \<in> I ==> F \<in> X guarantees Y i) *)
+(*Premise should be (\<And>i. i \<in> I \<Longrightarrow> F \<in> X guarantees Y i) *)
lemma all_guarantees:
- "[| \<forall>i \<in> I. F \<in> X guarantees Y(i); i \<in> I |]
- ==> F \<in> X guarantees (\<Inter>i \<in> I. Y(i))"
+ "\<lbrakk>\<forall>i \<in> I. F \<in> X guarantees Y(i); i \<in> I\<rbrakk>
+ \<Longrightarrow> F \<in> X guarantees (\<Inter>i \<in> I. Y(i))"
by (unfold guar_def, blast)
-(*Premises should be [| F \<in> X guarantees Y i; i \<in> I |] *)
+(*Premises should be \<lbrakk>F \<in> X guarantees Y i; i \<in> I\<rbrakk> *)
lemma ex_guarantees:
- "\<exists>i \<in> I. F \<in> X guarantees Y(i) ==> F \<in> X guarantees (\<Union>i \<in> I. Y(i))"
+ "\<exists>i \<in> I. F \<in> X guarantees Y(i) \<Longrightarrow> F \<in> X guarantees (\<Union>i \<in> I. Y(i))"
by (unfold guar_def, blast)
(*** Additional guarantees laws, by lcp ***)
lemma guarantees_Join_Int:
- "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |]
- ==> F \<squnion> G: (U \<inter> X) guarantees (V \<inter> Y)"
+ "\<lbrakk>F \<in> U guarantees V; G \<in> X guarantees Y; F ok G\<rbrakk>
+ \<Longrightarrow> F \<squnion> G: (U \<inter> X) guarantees (V \<inter> Y)"
apply (unfold guar_def)
apply (simp (no_asm))
@@ -317,8 +317,8 @@
done
lemma guarantees_Join_Un:
- "[| F \<in> U guarantees V; G \<in> X guarantees Y; F ok G |]
- ==> F \<squnion> G: (U \<union> X) guarantees (V \<union> Y)"
+ "\<lbrakk>F \<in> U guarantees V; G \<in> X guarantees Y; F ok G\<rbrakk>
+ \<Longrightarrow> F \<squnion> G: (U \<union> X) guarantees (V \<union> Y)"
apply (unfold guar_def)
apply (simp (no_asm))
apply safe
@@ -332,8 +332,8 @@
done
lemma guarantees_JOIN_INT:
- "[| \<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F); i \<in> I |]
- ==> (\<Squnion>i \<in> I. F(i)) \<in> (\<Inter>i \<in> I. X(i)) guarantees (\<Inter>i \<in> I. Y(i))"
+ "\<lbrakk>\<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F); i \<in> I\<rbrakk>
+ \<Longrightarrow> (\<Squnion>i \<in> I. F(i)) \<in> (\<Inter>i \<in> I. X(i)) guarantees (\<Inter>i \<in> I. Y(i))"
apply (unfold guar_def, safe)
prefer 2 apply blast
apply (drule_tac x = xa in bspec)
@@ -343,8 +343,8 @@
done
lemma guarantees_JOIN_UN:
- "[| \<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F) |]
- ==> JOIN(I,F) \<in> (\<Union>i \<in> I. X(i)) guarantees (\<Union>i \<in> I. Y(i))"
+ "\<lbrakk>\<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F)\<rbrakk>
+ \<Longrightarrow> JOIN(I,F) \<in> (\<Union>i \<in> I. X(i)) guarantees (\<Union>i \<in> I. Y(i))"
apply (unfold guar_def, auto)
apply (drule_tac x = y in bspec, simp_all, safe)
apply (rename_tac G y)
@@ -355,20 +355,20 @@
(*** guarantees laws for breaking down the program, by lcp ***)
lemma guarantees_Join_I1:
- "[| F \<in> X guarantees Y; F ok G |] ==> F \<squnion> G \<in> X guarantees Y"
+ "\<lbrakk>F \<in> X guarantees Y; F ok G\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> X guarantees Y"
apply (simp add: guar_def, safe)
apply (simp add: Join_assoc)
done
lemma guarantees_Join_I2:
- "[| G \<in> X guarantees Y; F ok G |] ==> F \<squnion> G \<in> X guarantees Y"
+ "\<lbrakk>G \<in> X guarantees Y; F ok G\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> X guarantees Y"
apply (simp add: Join_commute [of _ G] ok_commute [of _ G])
apply (blast intro: guarantees_Join_I1)
done
lemma guarantees_JOIN_I:
- "[| i \<in> I; F(i) \<in> X guarantees Y; OK(I,F) |]
- ==> (\<Squnion>i \<in> I. F(i)) \<in> X guarantees Y"
+ "\<lbrakk>i \<in> I; F(i) \<in> X guarantees Y; OK(I,F)\<rbrakk>
+ \<Longrightarrow> (\<Squnion>i \<in> I. F(i)) \<in> X guarantees Y"
apply (unfold guar_def, safe)
apply (drule_tac x = "JOIN (I-{i},F) \<squnion> G" in bspec)
apply (simp (no_asm))
@@ -377,10 +377,10 @@
(*** well-definedness ***)
-lemma Join_welldef_D1: "F \<squnion> G \<in> welldef ==> programify(F) \<in> welldef"
+lemma Join_welldef_D1: "F \<squnion> G \<in> welldef \<Longrightarrow> programify(F) \<in> welldef"
by (unfold welldef_def, auto)
-lemma Join_welldef_D2: "F \<squnion> G \<in> welldef ==> programify(G) \<in> welldef"
+lemma Join_welldef_D2: "F \<squnion> G \<in> welldef \<Longrightarrow> programify(G) \<in> welldef"
by (unfold welldef_def, auto)
(*** refinement ***)
@@ -396,7 +396,7 @@
lemma guarantees_type: "X guarantees Y \<subseteq> program"
by (unfold guar_def, auto)
-lemma wgD2: "G \<in> wg(F, X) ==> G \<in> program & F \<in> program"
+lemma wgD2: "G \<in> wg(F, X) \<Longrightarrow> G \<in> program & F \<in> program"
apply (unfold wg_def, auto)
apply (blast dest: guarantees_type [THEN subsetD])
done
@@ -407,10 +407,10 @@
by (unfold guar_def component_of_def, force)
lemma wg_weakest:
- "!!X. [| F \<in> (X guarantees Y); X \<subseteq> program |] ==> X \<subseteq> wg(F,Y)"
+ "\<And>X. \<lbrakk>F \<in> (X guarantees Y); X \<subseteq> program\<rbrakk> \<Longrightarrow> X \<subseteq> wg(F,Y)"
by (unfold wg_def, auto)
-lemma wg_guarantees: "F \<in> program ==> F \<in> wg(F,Y) guarantees Y"
+lemma wg_guarantees: "F \<in> program \<Longrightarrow> F \<in> wg(F,Y) guarantees Y"
by (unfold wg_def guar_def, blast)
lemma wg_equiv:
@@ -423,7 +423,7 @@
done
lemma component_of_wg:
- "F component_of H ==> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X & F \<in> program & H \<in> program)"
+ "F component_of H \<Longrightarrow> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X & F \<in> program & H \<in> program)"
by (simp (no_asm_simp) add: wg_equiv)
lemma wg_finite [rule_format]:
@@ -439,7 +439,7 @@
done
lemma wg_ex_prop:
- "ex_prop(X) ==> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)"
+ "ex_prop(X) \<Longrightarrow> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) & F \<in> program)"
apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv)
apply blast
done
@@ -510,7 +510,7 @@
Reasoning About Program composition paper *)
lemma stable_guarantees_Always:
- "[| Init(F) \<subseteq> A; F \<in> program |] ==> F \<in> stable(A) guarantees Always(A)"
+ "\<lbrakk>Init(F) \<subseteq> A; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> stable(A) guarantees Always(A)"
apply (rule guaranteesI)
prefer 2 apply assumption
apply (simp (no_asm) add: Join_commute)
@@ -520,8 +520,8 @@
done
lemma constrains_guarantees_leadsTo:
- "[| F \<in> transient(A); st_set(B) |]
- ==> F: (A co A \<union> B) guarantees (A \<longmapsto> (B-A))"
+ "\<lbrakk>F \<in> transient(A); st_set(B)\<rbrakk>
+ \<Longrightarrow> F: (A co A \<union> B) guarantees (A \<longmapsto> (B-A))"
apply (rule guaranteesI)
prefer 2 apply (blast dest: transient_type [THEN subsetD])
apply (rule leadsTo_Basis')
--- a/src/ZF/UNITY/Increasing.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Increasing.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,19 +12,19 @@
definition
increasing :: "[i, i, i=>i] => i" (\<open>increasing[_]'(_, _')\<close>) where
- "increasing[A](r, f) ==
+ "increasing[A](r, f) \<equiv>
{F \<in> program. (\<forall>k \<in> A. F \<in> stable({s \<in> state. <k, f(s)> \<in> r})) &
(\<forall>x \<in> state. f(x):A)}"
definition
Increasing :: "[i, i, i=>i] => i" (\<open>Increasing[_]'(_, _')\<close>) where
- "Increasing[A](r, f) ==
+ "Increasing[A](r, f) \<equiv>
{F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) &
(\<forall>x \<in> state. f(x):A)}"
abbreviation (input)
IncWrt :: "[i=>i, i, i] => i" (\<open>(_ IncreasingWrt _ '/ _)\<close> [60, 0, 60] 60) where
- "f IncreasingWrt r/A == Increasing[A](r,f)"
+ "f IncreasingWrt r/A \<equiv> Increasing[A](r,f)"
(** increasing **)
@@ -32,15 +32,15 @@
lemma increasing_type: "increasing[A](r, f) \<subseteq> program"
by (unfold increasing_def, blast)
-lemma increasing_into_program: "F \<in> increasing[A](r, f) ==> F \<in> program"
+lemma increasing_into_program: "F \<in> increasing[A](r, f) \<Longrightarrow> F \<in> program"
by (unfold increasing_def, blast)
lemma increasing_imp_stable:
-"[| F \<in> increasing[A](r, f); x \<in> A |] ==>F \<in> stable({s \<in> state. <x, f(s)>:r})"
+"\<lbrakk>F \<in> increasing[A](r, f); x \<in> A\<rbrakk> \<Longrightarrow>F \<in> stable({s \<in> state. <x, f(s)>:r})"
by (unfold increasing_def, blast)
lemma increasingD:
-"F \<in> increasing[A](r,f) ==> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
+"F \<in> increasing[A](r,f) \<Longrightarrow> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
apply (unfold increasing_def)
apply (subgoal_tac "\<exists>x. x \<in> state")
apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state)
@@ -54,7 +54,7 @@
done
lemma subset_increasing_comp:
-"[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==>
+"\<lbrakk>mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow>
increasing[A](r, f) \<subseteq> increasing[B](s, g comp f)"
apply (unfold increasing_def stable_def part_order_def
constrains_def mono1_def metacomp_def, clarify, simp)
@@ -74,8 +74,8 @@
done
lemma imp_increasing_comp:
- "[| F \<in> increasing[A](r, f); mono1(A, r, B, s, g);
- refl(A, r); trans[B](s) |] ==> F \<in> increasing[B](s, g comp f)"
+ "\<lbrakk>F \<in> increasing[A](r, f); mono1(A, r, B, s, g);
+ refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow> F \<in> increasing[B](s, g comp f)"
by (rule subset_increasing_comp [THEN subsetD], auto)
lemma strict_increasing:
@@ -92,7 +92,7 @@
(** Increasing **)
lemma increasing_imp_Increasing:
- "F \<in> increasing[A](r, f) ==> F \<in> Increasing[A](r, f)"
+ "F \<in> increasing[A](r, f) \<Longrightarrow> F \<in> Increasing[A](r, f)"
apply (unfold increasing_def Increasing_def)
apply (auto intro: stable_imp_Stable)
@@ -101,15 +101,15 @@
lemma Increasing_type: "Increasing[A](r, f) \<subseteq> program"
by (unfold Increasing_def, auto)
-lemma Increasing_into_program: "F \<in> Increasing[A](r, f) ==> F \<in> program"
+lemma Increasing_into_program: "F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program"
by (unfold Increasing_def, auto)
lemma Increasing_imp_Stable:
-"[| F \<in> Increasing[A](r, f); a \<in> A |] ==> F \<in> Stable({s \<in> state. <a,f(s)>:r})"
+"\<lbrakk>F \<in> Increasing[A](r, f); a \<in> A\<rbrakk> \<Longrightarrow> F \<in> Stable({s \<in> state. <a,f(s)>:r})"
by (unfold Increasing_def, blast)
lemma IncreasingD:
-"F \<in> Increasing[A](r, f) ==> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
+"F \<in> Increasing[A](r, f) \<Longrightarrow> F \<in> program & (\<exists>a. a \<in> A) & (\<forall>s \<in> state. f(s):A)"
apply (unfold Increasing_def)
apply (subgoal_tac "\<exists>x. x \<in> state")
apply (auto intro: st0_in_state)
@@ -122,7 +122,7 @@
done
lemma subset_Increasing_comp:
-"[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==>
+"\<lbrakk>mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk> \<Longrightarrow>
Increasing[A](r, f) \<subseteq> Increasing[B](s, g comp f)"
apply (unfold Increasing_def Stable_def Constrains_def part_order_def
constrains_def mono1_def metacomp_def, safe)
@@ -144,8 +144,8 @@
done
lemma imp_Increasing_comp:
- "[| F \<in> Increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s) |]
- ==> F \<in> Increasing[B](s, g comp f)"
+ "\<lbrakk>F \<in> Increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s)\<rbrakk>
+ \<Longrightarrow> F \<in> Increasing[B](s, g comp f)"
apply (rule subset_Increasing_comp [THEN subsetD], auto)
done
@@ -161,9 +161,9 @@
(** Two-place monotone operations **)
lemma imp_increasing_comp2:
-"[| F \<in> increasing[A](r, f); F \<in> increasing[B](s, g);
- mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t) |]
- ==> F \<in> increasing[C](t, %x. h(f(x), g(x)))"
+"\<lbrakk>F \<in> increasing[A](r, f); F \<in> increasing[B](s, g);
+ mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk>
+ \<Longrightarrow> F \<in> increasing[C](t, %x. h(f(x), g(x)))"
apply (unfold increasing_def stable_def
part_order_def constrains_def mono2_def, clarify, simp)
apply clarify
@@ -193,8 +193,8 @@
done
lemma imp_Increasing_comp2:
-"[| F \<in> Increasing[A](r, f); F \<in> Increasing[B](s, g);
- mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t) |] ==>
+"\<lbrakk>F \<in> Increasing[A](r, f); F \<in> Increasing[B](s, g);
+ mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk> \<Longrightarrow>
F \<in> Increasing[C](t, %x. h(f(x), g(x)))"
apply (unfold Increasing_def stable_def
part_order_def constrains_def mono2_def Stable_def Constrains_def, safe)
--- a/src/ZF/UNITY/Merge.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Merge.thy Tue Sep 27 16:51:35 2022 +0100
@@ -14,27 +14,27 @@
definition
(*spec (10)*)
merge_increasing :: "[i, i, i] =>i" where
- "merge_increasing(A, Out, iOut) == program guarantees
+ "merge_increasing(A, Out, iOut) \<equiv> program guarantees
(lift(Out) IncreasingWrt prefix(A)/list(A)) Int
(lift(iOut) IncreasingWrt prefix(nat)/list(nat))"
definition
(*spec (11)*)
merge_eq_Out :: "[i, i] =>i" where
- "merge_eq_Out(Out, iOut) == program guarantees
+ "merge_eq_Out(Out, iOut) \<equiv> program guarantees
Always({s \<in> state. length(s`Out) = length(s`iOut)})"
definition
(*spec (12)*)
merge_bounded :: "i=>i" where
- "merge_bounded(iOut) == program guarantees
+ "merge_bounded(iOut) \<equiv> program guarantees
Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})"
definition
(*spec (13)*)
(* Parameter A represents the type of tokens *)
merge_follows :: "[i, i=>i, i, i] =>i" where
- "merge_follows(A, In, Out, iOut) ==
+ "merge_follows(A, In, Out, iOut) \<equiv>
(\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))
guarantees
(\<Inter>n \<in> Nclients.
@@ -45,19 +45,19 @@
definition
(*spec: preserves part*)
merge_preserves :: "[i=>i] =>i" where
- "merge_preserves(In) == \<Inter>n \<in> nat. preserves(lift(In(n)))"
+ "merge_preserves(In) \<equiv> \<Inter>n \<in> nat. preserves(lift(In(n)))"
definition
(* environmental constraints*)
merge_allowed_acts :: "[i, i] =>i" where
- "merge_allowed_acts(Out, iOut) ==
+ "merge_allowed_acts(Out, iOut) \<equiv>
{F \<in> program. AllowedActs(F) =
cons(id(state), (\<Union>G \<in> preserves(lift(Out)) \<inter>
preserves(lift(iOut)). Acts(G)))}"
definition
merge_spec :: "[i, i =>i, i, i]=>i" where
- "merge_spec(A, In, Out, iOut) ==
+ "merge_spec(A, In, Out, iOut) \<equiv>
merge_increasing(A, Out, iOut) \<inter> merge_eq_Out(Out, iOut) \<inter>
merge_bounded(iOut) \<inter> merge_follows(A, In, Out, iOut)
\<inter> merge_allowed_acts(Out, iOut) \<inter> merge_preserves(In)"
@@ -84,18 +84,18 @@
and merge_spec: "M \<in> merge_spec(A, In, Out, iOut)"
-lemma (in merge) In_value_type [TC,simp]: "s \<in> state ==> s`In(n) \<in> list(A)"
+lemma (in merge) In_value_type [TC,simp]: "s \<in> state \<Longrightarrow> s`In(n) \<in> list(A)"
apply (unfold state_def)
apply (drule_tac a = "In (n)" in apply_type)
apply auto
done
-lemma (in merge) Out_value_type [TC,simp]: "s \<in> state ==> s`Out \<in> list(A)"
+lemma (in merge) Out_value_type [TC,simp]: "s \<in> state \<Longrightarrow> s`Out \<in> list(A)"
apply (unfold state_def)
apply (drule_tac a = Out in apply_type, auto)
done
-lemma (in merge) iOut_value_type [TC,simp]: "s \<in> state ==> s`iOut \<in> list(nat)"
+lemma (in merge) iOut_value_type [TC,simp]: "s \<in> state \<Longrightarrow> s`iOut \<in> list(nat)"
apply (unfold state_def)
apply (drule_tac a = iOut in apply_type, auto)
done
@@ -113,7 +113,7 @@
done
lemma (in merge) M_ok_iff:
- "G \<in> program ==>
+ "G \<in> program \<Longrightarrow>
M ok G \<longleftrightarrow> (G \<in> preserves(lift(Out)) &
G \<in> preserves(lift(iOut)) & M \<in> Allowed(G))"
apply (cut_tac merge_spec)
@@ -121,9 +121,9 @@
done
lemma (in merge) merge_Always_Out_eq_iOut:
- "[| G \<in> preserves(lift(Out)); G \<in> preserves(lift(iOut));
- M \<in> Allowed(G) |]
- ==> M \<squnion> G \<in> Always({s \<in> state. length(s`Out)=length(s`iOut)})"
+ "\<lbrakk>G \<in> preserves(lift(Out)); G \<in> preserves(lift(iOut));
+ M \<in> Allowed(G)\<rbrakk>
+ \<Longrightarrow> M \<squnion> G \<in> Always({s \<in> state. length(s`Out)=length(s`iOut)})"
apply (frule preserves_type [THEN subsetD])
apply (subgoal_tac "G \<in> program")
prefer 2 apply assumption
@@ -133,8 +133,8 @@
done
lemma (in merge) merge_Bounded:
- "[| G \<in> preserves(lift(iOut)); G \<in> preserves(lift(Out));
- M \<in> Allowed(G) |] ==>
+ "\<lbrakk>G \<in> preserves(lift(iOut)); G \<in> preserves(lift(Out));
+ M \<in> Allowed(G)\<rbrakk> \<Longrightarrow>
M \<squnion> G: Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})"
apply (frule preserves_type [THEN subsetD])
apply (frule M_ok_iff)
@@ -143,9 +143,9 @@
done
lemma (in merge) merge_bag_Follows_lemma:
-"[| G \<in> preserves(lift(iOut));
- G: preserves(lift(Out)); M \<in> Allowed(G) |]
- ==> M \<squnion> G \<in> Always
+"\<lbrakk>G \<in> preserves(lift(iOut));
+ G: preserves(lift(Out)); M \<in> Allowed(G)\<rbrakk>
+ \<Longrightarrow> M \<squnion> G \<in> Always
({s \<in> state. msetsum(%i. bag_of(sublist(s`Out,
{k \<in> nat. k < length(s`iOut) & nth(k, s`iOut)=i})),
Nclients, A) = bag_of(s`Out)})"
--- a/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 16:51:35 2022 +0100
@@ -13,14 +13,14 @@
definition
mono1 :: "[i, i, i, i, i=>i] => o" where
- "mono1(A, r, B, s, f) ==
+ "mono1(A, r, B, s, f) \<equiv>
(\<forall>x \<in> A. \<forall>y \<in> A. <x,y> \<in> r \<longrightarrow> <f(x), f(y)> \<in> s) & (\<forall>x \<in> A. f(x) \<in> B)"
(* monotonicity of a 2-place meta-function f *)
definition
mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o" where
- "mono2(A, r, B, s, C, t, f) ==
+ "mono2(A, r, B, s, C, t, f) \<equiv>
(\<forall>x \<in> A. \<forall>y \<in> A. \<forall>u \<in> B. \<forall>v \<in> B.
<x,y> \<in> r & <u,v> \<in> s \<longrightarrow> <f(x,u), f(y,v)> \<in> t) &
(\<forall>x \<in> A. \<forall>y \<in> B. f(x,y) \<in> C)"
@@ -29,29 +29,29 @@
definition
SetLe :: "i =>i" where
- "SetLe(A) == {<x,y> \<in> Pow(A)*Pow(A). x \<subseteq> y}"
+ "SetLe(A) \<equiv> {<x,y> \<in> Pow(A)*Pow(A). x \<subseteq> y}"
definition
MultLe :: "[i,i] =>i" where
- "MultLe(A, r) == multirel(A, r - id(A)) \<union> id(Mult(A))"
+ "MultLe(A, r) \<equiv> multirel(A, r - id(A)) \<union> id(Mult(A))"
lemma mono1D:
- "[| mono1(A, r, B, s, f); <x, y> \<in> r; x \<in> A; y \<in> A |] ==> <f(x), f(y)> \<in> s"
+ "\<lbrakk>mono1(A, r, B, s, f); <x, y> \<in> r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <f(x), f(y)> \<in> s"
by (unfold mono1_def, auto)
lemma mono2D:
- "[| mono2(A, r, B, s, C, t, f);
- <x, y> \<in> r; <u,v> \<in> s; x \<in> A; y \<in> A; u \<in> B; v \<in> B |]
- ==> <f(x, u), f(y,v)> \<in> t"
+ "\<lbrakk>mono2(A, r, B, s, C, t, f);
+ <x, y> \<in> r; <u,v> \<in> s; x \<in> A; y \<in> A; u \<in> B; v \<in> B\<rbrakk>
+ \<Longrightarrow> <f(x, u), f(y,v)> \<in> t"
by (unfold mono2_def, auto)
(** Monotonicity of take **)
lemma take_mono_left_lemma:
- "[| i \<le> j; xs \<in> list(A); i \<in> nat; j \<in> nat |]
- ==> <take(i, xs), take(j, xs)> \<in> prefix(A)"
+ "\<lbrakk>i \<le> j; xs \<in> list(A); i \<in> nat; j \<in> nat\<rbrakk>
+ \<Longrightarrow> <take(i, xs), take(j, xs)> \<in> prefix(A)"
apply (case_tac "length (xs) \<le> i")
apply (subgoal_tac "length (xs) \<le> j")
apply (simp)
@@ -67,18 +67,18 @@
done
lemma take_mono_left:
- "[| i \<le> j; xs \<in> list(A); j \<in> nat |]
- ==> <take(i, xs), take(j, xs)> \<in> prefix(A)"
+ "\<lbrakk>i \<le> j; xs \<in> list(A); j \<in> nat\<rbrakk>
+ \<Longrightarrow> <take(i, xs), take(j, xs)> \<in> prefix(A)"
by (blast intro: le_in_nat take_mono_left_lemma)
lemma take_mono_right:
- "[| <xs,ys> \<in> prefix(A); i \<in> nat |]
- ==> <take(i, xs), take(i, ys)> \<in> prefix(A)"
+ "\<lbrakk><xs,ys> \<in> prefix(A); i \<in> nat\<rbrakk>
+ \<Longrightarrow> <take(i, xs), take(i, ys)> \<in> prefix(A)"
by (auto simp add: prefix_iff)
lemma take_mono:
- "[| i \<le> j; <xs, ys> \<in> prefix(A); j \<in> nat |]
- ==> <take(i, xs), take(j, ys)> \<in> prefix(A)"
+ "\<lbrakk>i \<le> j; <xs, ys> \<in> prefix(A); j \<in> nat\<rbrakk>
+ \<Longrightarrow> <take(i, xs), take(j, ys)> \<in> prefix(A)"
apply (rule_tac b = "take (j, xs) " in prefix_trans)
apply (auto dest: prefix_type [THEN subsetD] intro: take_mono_left take_mono_right)
done
--- a/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,23 +10,23 @@
definition
lcomm :: "[i, i, [i,i]=>i]=>o" where
- "lcomm(A, B, f) ==
+ "lcomm(A, B, f) \<equiv>
(\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> B. f(x, f(y, z))= f(y, f(x, z))) &
(\<forall>x \<in> A. \<forall>y \<in> B. f(x, y) \<in> B)"
definition
general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i" where
- "general_setsum(C, B, e, f, g) ==
+ "general_setsum(C, B, e, f, g) \<equiv>
if Finite(C) then fold[cons(e, B)](%x y. f(g(x), y), e, C) else e"
definition
msetsum :: "[i=>i, i, i]=>i" where
- "msetsum(g, C, B) == normalize(general_setsum(C, Mult(B), 0, (+#), g))"
+ "msetsum(g, C, B) \<equiv> normalize(general_setsum(C, Mult(B), 0, (+#), g))"
definition (* sum for naturals *)
nsetsum :: "[i=>i, i] =>i" where
- "nsetsum(g, C) == general_setsum(C, nat, 0, (#+), g)"
+ "nsetsum(g, C) \<equiv> general_setsum(C, nat, 0, (#+), g)"
lemma mset_of_normalize: "mset_of(normalize(M)) \<subseteq> mset_of(M)"
@@ -36,7 +36,7 @@
by (simp add: general_setsum_def)
lemma general_setsum_cons [simp]:
-"[| C \<in> Fin(A); a \<in> A; a\<notin>C; e \<in> B; \<forall>x \<in> A. g(x) \<in> B; lcomm(B, B, f) |] ==>
+"\<lbrakk>C \<in> Fin(A); a \<in> A; a\<notin>C; e \<in> B; \<forall>x \<in> A. g(x) \<in> B; lcomm(B, B, f)\<rbrakk> \<Longrightarrow>
general_setsum(cons(a, C), B, e, f, g) =
f(g(a), general_setsum(C, B, e, f,g))"
apply (simp add: general_setsum_def)
@@ -48,7 +48,7 @@
(** lcomm **)
-lemma lcomm_mono: "[| C\<subseteq>A; lcomm(A, B, f) |] ==> lcomm(C, B,f)"
+lemma lcomm_mono: "\<lbrakk>C\<subseteq>A; lcomm(A, B, f)\<rbrakk> \<Longrightarrow> lcomm(C, B,f)"
by (auto simp add: lcomm_def, blast)
lemma munion_lcomm [simp]: "lcomm(Mult(A), Mult(A), (+#))"
@@ -57,8 +57,8 @@
(** msetsum **)
lemma multiset_general_setsum:
- "[| C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x))& mset_of(g(x))\<subseteq>B |]
- ==> multiset(general_setsum(C, B -||> nat - {0}, 0, \<lambda>u v. u +# v, g))"
+ "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x))& mset_of(g(x))\<subseteq>B\<rbrakk>
+ \<Longrightarrow> multiset(general_setsum(C, B -||> nat - {0}, 0, \<lambda>u v. u +# v, g))"
apply (erule Fin_induct, auto)
apply (subst general_setsum_cons)
apply (auto simp add: Mult_iff_multiset)
@@ -68,8 +68,8 @@
by (simp add: msetsum_def)
lemma msetsum_cons [simp]:
- "[| C \<in> Fin(A); a\<notin>C; a \<in> A; \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]
- ==> msetsum(g, cons(a, C), B) = g(a) +# msetsum(g, C, B)"
+ "\<lbrakk>C \<in> Fin(A); a\<notin>C; a \<in> A; \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>
+ \<Longrightarrow> msetsum(g, cons(a, C), B) = g(a) +# msetsum(g, C, B)"
apply (simp add: msetsum_def)
apply (subst general_setsum_cons)
apply (auto simp add: multiset_general_setsum Mult_iff_multiset)
@@ -81,15 +81,15 @@
by (simp add: msetsum_def)
lemma mset_of_msetsum:
- "[| C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]
- ==> mset_of(msetsum(g, C, B))\<subseteq>B"
+ "\<lbrakk>C \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>
+ \<Longrightarrow> mset_of(msetsum(g, C, B))\<subseteq>B"
by (erule Fin_induct, auto)
(*The reversed orientation looks more natural, but LOOPS as a simprule!*)
lemma msetsum_Un_Int:
- "[| C \<in> Fin(A); D \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]
- ==> msetsum(g, C \<union> D, B) +# msetsum(g, C \<inter> D, B)
+ "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A); \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>
+ \<Longrightarrow> msetsum(g, C \<union> D, B) +# msetsum(g, C \<inter> D, B)
= msetsum(g, C, B) +# msetsum(g, D, B)"
apply (erule Fin_induct)
apply (subgoal_tac [2] "cons (x, y) \<union> D = cons (x, y \<union> D) ")
@@ -106,19 +106,19 @@
lemma msetsum_Un_disjoint:
- "[| C \<in> Fin(A); D \<in> Fin(A); C \<inter> D = 0;
- \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |]
- ==> msetsum(g, C \<union> D, B) = msetsum(g, C, B) +# msetsum(g,D, B)"
+ "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A); C \<inter> D = 0;
+ \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk>
+ \<Longrightarrow> msetsum(g, C \<union> D, B) = msetsum(g, C, B) +# msetsum(g,D, B)"
apply (subst msetsum_Un_Int [symmetric])
apply (auto simp add: msetsum_multiset)
done
lemma UN_Fin_lemma [rule_format (no_asm)]:
- "I \<in> Fin(A) ==> (\<forall>i \<in> I. C(i) \<in> Fin(B)) \<longrightarrow> (\<Union>i \<in> I. C(i)):Fin(B)"
+ "I \<in> Fin(A) \<Longrightarrow> (\<forall>i \<in> I. C(i) \<in> Fin(B)) \<longrightarrow> (\<Union>i \<in> I. C(i)):Fin(B)"
by (erule Fin_induct, auto)
lemma msetsum_UN_disjoint [rule_format (no_asm)]:
- "[| I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A) |] ==>
+ "\<lbrakk>I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A)\<rbrakk> \<Longrightarrow>
(\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B) \<longrightarrow>
(\<forall>i \<in> I. \<forall>j \<in> I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow>
msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)"
@@ -138,9 +138,9 @@
done
lemma msetsum_addf:
- "[| C \<in> Fin(A);
+ "\<lbrakk>C \<in> Fin(A);
\<forall>x \<in> A. multiset(f(x)) & mset_of(f(x))\<subseteq>B;
- \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B |] ==>
+ \<forall>x \<in> A. multiset(g(x)) & mset_of(g(x))\<subseteq>B\<rbrakk> \<Longrightarrow>
msetsum(%x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)"
apply (subgoal_tac "\<forall>x \<in> A. multiset (f(x) +# g(x)) & mset_of (f(x) +# g(x)) \<subseteq> B")
apply (erule Fin_induct)
@@ -148,16 +148,16 @@
done
lemma msetsum_cong:
- "[| C=D; !!x. x \<in> D ==> f(x) = g(x) |]
- ==> msetsum(f, C, B) = msetsum(g, D, B)"
+ "\<lbrakk>C=D; \<And>x. x \<in> D \<Longrightarrow> f(x) = g(x)\<rbrakk>
+ \<Longrightarrow> msetsum(f, C, B) = msetsum(g, D, B)"
by (simp add: msetsum_def general_setsum_def cong add: fold_cong)
-lemma multiset_union_diff: "[| multiset(M); multiset(N) |] ==> M +# N -# N = M"
+lemma multiset_union_diff: "\<lbrakk>multiset(M); multiset(N)\<rbrakk> \<Longrightarrow> M +# N -# N = M"
by (simp add: multiset_equality)
-lemma msetsum_Un: "[| C \<in> Fin(A); D \<in> Fin(A);
- \<forall>x \<in> A. multiset(f(x)) & mset_of(f(x)) \<subseteq> B |]
- ==> msetsum(f, C \<union> D, B) =
+lemma msetsum_Un: "\<lbrakk>C \<in> Fin(A); D \<in> Fin(A);
+ \<forall>x \<in> A. multiset(f(x)) & mset_of(f(x)) \<subseteq> B\<rbrakk>
+ \<Longrightarrow> msetsum(f, C \<union> D, B) =
msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C \<inter> D, B)"
apply (subgoal_tac "C \<union> D \<in> Fin (A) & C \<inter> D \<in> Fin (A) ")
apply clarify
@@ -170,7 +170,7 @@
by (simp add: nsetsum_def)
lemma nsetsum_cons [simp]:
- "[| Finite(C); x\<notin>C |] ==> nsetsum(g, cons(x, C))= g(x) #+ nsetsum(g, C)"
+ "\<lbrakk>Finite(C); x\<notin>C\<rbrakk> \<Longrightarrow> nsetsum(g, cons(x, C))= g(x) #+ nsetsum(g, C)"
apply (simp add: nsetsum_def general_setsum_def)
apply (rule_tac A = "cons (x, C)" in fold_typing.fold_cons)
apply (auto intro: Finite_cons_lemma simp add: fold_typing_def)
--- a/src/ZF/UNITY/Mutex.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Mutex.thy Tue Sep 27 16:51:35 2022 +0100
@@ -21,11 +21,11 @@
the usual ZF typechecking: an ill-tyed expressions reduce to the empty set.
\<close>
-abbreviation "p == Var([0])"
-abbreviation "m == Var([1])"
-abbreviation "n == Var([0,0])"
-abbreviation "u == Var([0,1])"
-abbreviation "v == Var([1,0])"
+abbreviation "p \<equiv> Var([0])"
+abbreviation "m \<equiv> Var([1])"
+abbreviation "n \<equiv> Var([0,0])"
+abbreviation "u \<equiv> Var([0,1])"
+abbreviation "v \<equiv> Var([1,0])"
axiomatization where \<comment> \<open>Type declarations\<close>
p_type: "type_of(p)=bool & default_val(p)=0" and
@@ -36,56 +36,56 @@
definition
(** The program for process U **)
- "U0 == {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
+ "U0 \<equiv> {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
definition
- "U1 == {<s,t>:state*state. t = s(p:= s`v, m:=#2) & s`m = #1}"
+ "U1 \<equiv> {<s,t>:state*state. t = s(p:= s`v, m:=#2) & s`m = #1}"
definition
- "U2 == {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
+ "U2 \<equiv> {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
definition
- "U3 == {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
+ "U3 \<equiv> {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
definition
- "U4 == {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
+ "U4 \<equiv> {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
(** The program for process V **)
definition
- "V0 == {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
+ "V0 \<equiv> {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
definition
- "V1 == {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
+ "V1 \<equiv> {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
definition
- "V2 == {<s,t>:state*state. t = s(n:=#3) & s`p=1 & s`n = #2}"
+ "V2 \<equiv> {<s,t>:state*state. t = s(n:=#3) & s`p=1 & s`n = #2}"
definition
- "V3 == {<s,t>:state*state. t = s (v:=0, n:=#4) & s`n = #3}"
+ "V3 \<equiv> {<s,t>:state*state. t = s (v:=0, n:=#4) & s`n = #3}"
definition
- "V4 == {<s,t>:state*state. t = s (p:=0, n:=#0) & s`n = #4}"
+ "V4 \<equiv> {<s,t>:state*state. t = s (p:=0, n:=#0) & s`n = #4}"
definition
- "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
+ "Mutex \<equiv> mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
{U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))"
(** The correct invariants **)
definition
- "IU == {s:state. (s`u = 1\<longleftrightarrow>(#1 $\<le> s`m & s`m $\<le> #3))
+ "IU \<equiv> {s:state. (s`u = 1\<longleftrightarrow>(#1 $\<le> s`m & s`m $\<le> #3))
& (s`m = #3 \<longrightarrow> s`p=0)}"
definition
- "IV == {s:state. (s`v = 1 \<longleftrightarrow> (#1 $\<le> s`n & s`n $\<le> #3))
+ "IV \<equiv> {s:state. (s`v = 1 \<longleftrightarrow> (#1 $\<le> s`n & s`n $\<le> #3))
& (s`n = #3 \<longrightarrow> s`p=1)}"
(** The faulty invariant (for U alone) **)
definition
- "bad_IU == {s:state. (s`u = 1 \<longleftrightarrow> (#1 $\<le> s`m & s`m $\<le> #3))&
+ "bad_IU \<equiv> {s:state. (s`u = 1 \<longleftrightarrow> (#1 $\<le> s`m & s`m $\<le> #3))&
(#3 $\<le> s`m & s`m $\<le> #4 \<longrightarrow> s`p=0)}"
@@ -93,27 +93,27 @@
declare p_type [simp] u_type [simp] v_type [simp] m_type [simp] n_type [simp]
-lemma u_value_type: "s \<in> state ==>s`u \<in> bool"
+lemma u_value_type: "s \<in> state \<Longrightarrow>s`u \<in> bool"
apply (unfold state_def)
apply (drule_tac a = u in apply_type, auto)
done
-lemma v_value_type: "s \<in> state ==> s`v \<in> bool"
+lemma v_value_type: "s \<in> state \<Longrightarrow> s`v \<in> bool"
apply (unfold state_def)
apply (drule_tac a = v in apply_type, auto)
done
-lemma p_value_type: "s \<in> state ==> s`p \<in> bool"
+lemma p_value_type: "s \<in> state \<Longrightarrow> s`p \<in> bool"
apply (unfold state_def)
apply (drule_tac a = p in apply_type, auto)
done
-lemma m_value_type: "s \<in> state ==> s`m \<in> int"
+lemma m_value_type: "s \<in> state \<Longrightarrow> s`m \<in> int"
apply (unfold state_def)
apply (drule_tac a = m in apply_type, auto)
done
-lemma n_value_type: "s \<in> state ==>s`n \<in> int"
+lemma n_value_type: "s \<in> state \<Longrightarrow>s`n \<in> int"
apply (unfold state_def)
apply (drule_tac a = n in apply_type, auto)
done
@@ -174,14 +174,14 @@
done
(*The safety property: mutual exclusion*)
-lemma mutual_exclusion: "Mutex \<in> Always({s \<in> state. ~(s`m = #3 & s`n = #3)})"
+lemma mutual_exclusion: "Mutex \<in> Always({s \<in> state. \<not>(s`m = #3 & s`n = #3)})"
apply (rule Always_weaken)
apply (rule Always_Int_I [OF IU IV], auto)
done
(*The bad invariant FAILS in V1*)
-lemma less_lemma: "[| x$<#1; #3 $\<le> x |] ==> P"
+lemma less_lemma: "\<lbrakk>x$<#1; #3 $\<le> x\<rbrakk> \<Longrightarrow> P"
apply (drule_tac j = "#1" and k = "#3" in zless_zle_trans)
apply (drule_tac [2] j = x in zle_zless_trans, auto)
done
@@ -232,7 +232,7 @@
lemma U_lemma1: "Mutex \<in> {s \<in> state. s`m = #1} \<longmapsto>w {s \<in> state. s`p =1}"
by (rule LeadsTo_Trans [OF U_F1 [THEN LeadsTo_weaken_R] U_lemma2], blast)
-lemma eq_123: "i \<in> int ==> (#1 $\<le> i & i $\<le> #3) \<longleftrightarrow> (i=#1 | i=#2 | i=#3)"
+lemma eq_123: "i \<in> int \<Longrightarrow> (#1 $\<le> i & i $\<le> #3) \<longleftrightarrow> (i=#1 | i=#2 | i=#3)"
apply auto
apply (auto simp add: neq_iff_zless)
apply (drule_tac [4] j = "#3" and i = i in zle_zless_trans)
--- a/src/ZF/UNITY/State.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/State.thy Tue Sep 27 16:51:35 2022 +0100
@@ -21,20 +21,20 @@
default_val :: "i=>i"
definition
- "state == \<Prod>x \<in> var. cons(default_val(x), type_of(x))"
+ "state \<equiv> \<Prod>x \<in> var. cons(default_val(x), type_of(x))"
definition
- "st0 == \<lambda>x \<in> var. default_val(x)"
+ "st0 \<equiv> \<lambda>x \<in> var. default_val(x)"
definition
st_set :: "i=>o" where
(* To prevent typing conditions like `A<=state' from
being used in combination with the rules `constrains_weaken', etc. *)
- "st_set(A) == A<=state"
+ "st_set(A) \<equiv> A<=state"
definition
st_compl :: "i=>i" where
- "st_compl(A) == state-A"
+ "st_compl(A) \<equiv> state-A"
lemma st0_in_state [simp,TC]: "st0 \<in> state"
@@ -59,16 +59,16 @@
(* Intersection *)
-lemma st_set_Int [intro!]: "st_set(A) | st_set(B) ==> st_set(A \<inter> B)"
+lemma st_set_Int [intro!]: "st_set(A) | st_set(B) \<Longrightarrow> st_set(A \<inter> B)"
by (simp add: st_set_def, auto)
lemma st_set_Inter [intro!]:
- "(S=0) | (\<exists>A \<in> S. st_set(A)) ==> st_set(\<Inter>(S))"
+ "(S=0) | (\<exists>A \<in> S. st_set(A)) \<Longrightarrow> st_set(\<Inter>(S))"
apply (simp add: st_set_def Inter_def, auto)
done
(* Diff *)
-lemma st_set_DiffI [intro!]: "st_set(A) ==> st_set(A - B)"
+lemma st_set_DiffI [intro!]: "st_set(A) \<Longrightarrow> st_set(A - B)"
by (simp add: st_set_def, auto)
lemma Collect_Int_state [simp]: "Collect(state,P) \<inter> state = Collect(state,P)"
@@ -80,18 +80,18 @@
(* Introduction and destruction rules for st_set *)
-lemma st_setI: "A \<subseteq> state ==> st_set(A)"
+lemma st_setI: "A \<subseteq> state \<Longrightarrow> st_set(A)"
by (simp add: st_set_def)
-lemma st_setD: "st_set(A) ==> A<=state"
+lemma st_setD: "st_set(A) \<Longrightarrow> A<=state"
by (simp add: st_set_def)
-lemma st_set_subset: "[| st_set(A); B<=A |] ==> st_set(B)"
+lemma st_set_subset: "\<lbrakk>st_set(A); B<=A\<rbrakk> \<Longrightarrow> st_set(B)"
by (simp add: st_set_def, auto)
lemma state_update_type:
- "[| s \<in> state; x \<in> var; y \<in> type_of(x) |] ==> s(x:=y):state"
+ "\<lbrakk>s \<in> state; x \<in> var; y \<in> type_of(x)\<rbrakk> \<Longrightarrow> s(x:=y):state"
apply (simp add: state_def)
apply (blast intro: update_type)
done
@@ -103,12 +103,12 @@
by (simp add: st_compl_def)
lemma st_compl_Collect [simp]:
- "st_compl({s \<in> state. P(s)}) = {s \<in> state. ~P(s)}"
+ "st_compl({s \<in> state. P(s)}) = {s \<in> state. \<not>P(s)}"
by (simp add: st_compl_def, auto)
(*For using "disjunction" (union over an index set) to eliminate a variable.*)
lemma UN_conj_eq:
- "\<forall>d\<in>D. f(d) \<in> A ==> (\<Union>k\<in>A. {d\<in>D. P(d) & f(d) = k}) = {d\<in>D. P(d)}"
+ "\<forall>d\<in>D. f(d) \<in> A \<Longrightarrow> (\<Union>k\<in>A. {d\<in>D. P(d) & f(d) = k}) = {d\<in>D. P(d)}"
by blast
end
--- a/src/ZF/UNITY/SubstAx.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/SubstAx.thy Tue Sep 27 16:51:35 2022 +0100
@@ -14,18 +14,18 @@
definition
(* The definitions below are not `conventional', but yield simpler rules *)
Ensures :: "[i,i] => i" (infixl \<open>Ensures\<close> 60) where
- "A Ensures B == {F \<in> program. F \<in> (reachable(F) \<inter> A) ensures (reachable(F) \<inter> B) }"
+ "A Ensures B \<equiv> {F \<in> program. F \<in> (reachable(F) \<inter> A) ensures (reachable(F) \<inter> B) }"
definition
LeadsTo :: "[i, i] => i" (infixl \<open>\<longmapsto>w\<close> 60) where
- "A \<longmapsto>w B == {F \<in> program. F:(reachable(F) \<inter> A) \<longmapsto> (reachable(F) \<inter> B)}"
+ "A \<longmapsto>w B \<equiv> {F \<in> program. F:(reachable(F) \<inter> A) \<longmapsto> (reachable(F) \<inter> B)}"
(*Resembles the previous definition of LeadsTo*)
(* Equivalence with the HOL-like definition *)
lemma LeadsTo_eq:
-"st_set(B)==> A \<longmapsto>w B = {F \<in> program. F:(reachable(F) \<inter> A) \<longmapsto> B}"
+"st_set(B)\<Longrightarrow> A \<longmapsto>w B = {F \<in> program. F:(reachable(F) \<inter> A) \<longmapsto> B}"
apply (unfold LeadsTo_def)
apply (blast dest: psp_stable2 leadsToD2 constrainsD2 intro: leadsTo_weaken)
done
@@ -36,35 +36,35 @@
(*** Specialized laws for handling invariants ***)
(** Conjoining an Always property **)
-lemma Always_LeadsTo_pre: "F \<in> Always(I) ==> (F:(I \<inter> A) \<longmapsto>w A') \<longleftrightarrow> (F \<in> A \<longmapsto>w A')"
+lemma Always_LeadsTo_pre: "F \<in> Always(I) \<Longrightarrow> (F:(I \<inter> A) \<longmapsto>w A') \<longleftrightarrow> (F \<in> A \<longmapsto>w A')"
by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2)
-lemma Always_LeadsTo_post: "F \<in> Always(I) ==> (F \<in> A \<longmapsto>w (I \<inter> A')) \<longleftrightarrow> (F \<in> A \<longmapsto>w A')"
+lemma Always_LeadsTo_post: "F \<in> Always(I) \<Longrightarrow> (F \<in> A \<longmapsto>w (I \<inter> A')) \<longleftrightarrow> (F \<in> A \<longmapsto>w A')"
apply (unfold LeadsTo_def)
apply (simp add: Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2)
done
(* Like 'Always_LeadsTo_pre RS iffD1', but with premises in the good order *)
-lemma Always_LeadsToI: "[| F \<in> Always(C); F \<in> (C \<inter> A) \<longmapsto>w A' |] ==> F \<in> A \<longmapsto>w A'"
+lemma Always_LeadsToI: "\<lbrakk>F \<in> Always(C); F \<in> (C \<inter> A) \<longmapsto>w A'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w A'"
by (blast intro: Always_LeadsTo_pre [THEN iffD1])
(* Like 'Always_LeadsTo_post RS iffD2', but with premises in the good order *)
-lemma Always_LeadsToD: "[| F \<in> Always(C); F \<in> A \<longmapsto>w A' |] ==> F \<in> A \<longmapsto>w (C \<inter> A')"
+lemma Always_LeadsToD: "\<lbrakk>F \<in> Always(C); F \<in> A \<longmapsto>w A'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w (C \<inter> A')"
by (blast intro: Always_LeadsTo_post [THEN iffD2])
(*** Introduction rules \<in> Basis, Trans, Union ***)
-lemma LeadsTo_Basis: "F \<in> A Ensures B ==> F \<in> A \<longmapsto>w B"
+lemma LeadsTo_Basis: "F \<in> A Ensures B \<Longrightarrow> F \<in> A \<longmapsto>w B"
by (auto simp add: Ensures_def LeadsTo_def)
lemma LeadsTo_Trans:
- "[| F \<in> A \<longmapsto>w B; F \<in> B \<longmapsto>w C |] ==> F \<in> A \<longmapsto>w C"
+ "\<lbrakk>F \<in> A \<longmapsto>w B; F \<in> B \<longmapsto>w C\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w C"
apply (simp (no_asm_use) add: LeadsTo_def)
apply (blast intro: leadsTo_Trans)
done
lemma LeadsTo_Union:
-"[|(!!A. A \<in> S ==> F \<in> A \<longmapsto>w B); F \<in> program|]==>F \<in> \<Union>(S) \<longmapsto>w B"
+"\<lbrakk>(\<And>A. A \<in> S \<Longrightarrow> F \<in> A \<longmapsto>w B); F \<in> program\<rbrakk>\<Longrightarrow>F \<in> \<Union>(S) \<longmapsto>w B"
apply (simp add: LeadsTo_def)
apply (subst Int_Union_Union2)
apply (rule leadsTo_UN, auto)
@@ -72,23 +72,23 @@
(*** Derived rules ***)
-lemma leadsTo_imp_LeadsTo: "F \<in> A \<longmapsto> B ==> F \<in> A \<longmapsto>w B"
+lemma leadsTo_imp_LeadsTo: "F \<in> A \<longmapsto> B \<Longrightarrow> F \<in> A \<longmapsto>w B"
apply (frule leadsToD2, clarify)
apply (simp (no_asm_simp) add: LeadsTo_eq)
apply (blast intro: leadsTo_weaken_L)
done
(*Useful with cancellation, disjunction*)
-lemma LeadsTo_Un_duplicate: "F \<in> A \<longmapsto>w (A' \<union> A') ==> F \<in> A \<longmapsto>w A'"
+lemma LeadsTo_Un_duplicate: "F \<in> A \<longmapsto>w (A' \<union> A') \<Longrightarrow> F \<in> A \<longmapsto>w A'"
by (simp add: Un_ac)
lemma LeadsTo_Un_duplicate2:
- "F \<in> A \<longmapsto>w (A' \<union> C \<union> C) ==> F \<in> A \<longmapsto>w (A' \<union> C)"
+ "F \<in> A \<longmapsto>w (A' \<union> C \<union> C) \<Longrightarrow> F \<in> A \<longmapsto>w (A' \<union> C)"
by (simp add: Un_ac)
lemma LeadsTo_UN:
- "[|(!!i. i \<in> I ==> F \<in> A(i) \<longmapsto>w B); F \<in> program|]
- ==>F:(\<Union>i \<in> I. A(i)) \<longmapsto>w B"
+ "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto>w B); F \<in> program\<rbrakk>
+ \<Longrightarrow>F:(\<Union>i \<in> I. A(i)) \<longmapsto>w B"
apply (simp add: LeadsTo_def)
apply (simp (no_asm_simp) del: UN_simps add: Int_UN_distrib)
apply (rule leadsTo_UN, auto)
@@ -96,7 +96,7 @@
(*Binary union introduction rule*)
lemma LeadsTo_Un:
- "[| F \<in> A \<longmapsto>w C; F \<in> B \<longmapsto>w C |] ==> F \<in> (A \<union> B) \<longmapsto>w C"
+ "\<lbrakk>F \<in> A \<longmapsto>w C; F \<in> B \<longmapsto>w C\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) \<longmapsto>w C"
apply (subst Un_eq_Union)
apply (rule LeadsTo_Union)
apply (auto dest: LeadsTo_type [THEN subsetD])
@@ -104,11 +104,11 @@
(*Lets us look at the starting state*)
lemma single_LeadsTo_I:
- "[|(!!s. s \<in> A ==> F:{s} \<longmapsto>w B); F \<in> program|]==>F \<in> A \<longmapsto>w B"
+ "\<lbrakk>(\<And>s. s \<in> A \<Longrightarrow> F:{s} \<longmapsto>w B); F \<in> program\<rbrakk>\<Longrightarrow>F \<in> A \<longmapsto>w B"
apply (subst UN_singleton [symmetric], rule LeadsTo_UN, auto)
done
-lemma subset_imp_LeadsTo: "[| A \<subseteq> B; F \<in> program |] ==> F \<in> A \<longmapsto>w B"
+lemma subset_imp_LeadsTo: "\<lbrakk>A \<subseteq> B; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B"
apply (simp (no_asm_simp) add: LeadsTo_def)
apply (blast intro: subset_imp_leadsTo)
done
@@ -122,33 +122,33 @@
by (auto dest: LeadsTo_type [THEN subsetD] simp add: LeadsTo_eq)
declare LeadsTo_state [iff]
-lemma LeadsTo_weaken_R: "[| F \<in> A \<longmapsto>w A'; A'<=B'|] ==> F \<in> A \<longmapsto>w B'"
+lemma LeadsTo_weaken_R: "\<lbrakk>F \<in> A \<longmapsto>w A'; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B'"
apply (unfold LeadsTo_def)
apply (auto intro: leadsTo_weaken_R)
done
-lemma LeadsTo_weaken_L: "[| F \<in> A \<longmapsto>w A'; B \<subseteq> A |] ==> F \<in> B \<longmapsto>w A'"
+lemma LeadsTo_weaken_L: "\<lbrakk>F \<in> A \<longmapsto>w A'; B \<subseteq> A\<rbrakk> \<Longrightarrow> F \<in> B \<longmapsto>w A'"
apply (unfold LeadsTo_def)
apply (auto intro: leadsTo_weaken_L)
done
-lemma LeadsTo_weaken: "[| F \<in> A \<longmapsto>w A'; B<=A; A'<=B' |] ==> F \<in> B \<longmapsto>w B'"
+lemma LeadsTo_weaken: "\<lbrakk>F \<in> A \<longmapsto>w A'; B<=A; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> B \<longmapsto>w B'"
by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans)
lemma Always_LeadsTo_weaken:
-"[| F \<in> Always(C); F \<in> A \<longmapsto>w A'; C \<inter> B \<subseteq> A; C \<inter> A' \<subseteq> B' |]
- ==> F \<in> B \<longmapsto>w B'"
+"\<lbrakk>F \<in> Always(C); F \<in> A \<longmapsto>w A'; C \<inter> B \<subseteq> A; C \<inter> A' \<subseteq> B'\<rbrakk>
+ \<Longrightarrow> F \<in> B \<longmapsto>w B'"
apply (blast dest: Always_LeadsToI intro: LeadsTo_weaken Always_LeadsToD)
done
(** Two theorems for "proof lattices" **)
-lemma LeadsTo_Un_post: "F \<in> A \<longmapsto>w B ==> F:(A \<union> B) \<longmapsto>w B"
+lemma LeadsTo_Un_post: "F \<in> A \<longmapsto>w B \<Longrightarrow> F:(A \<union> B) \<longmapsto>w B"
by (blast dest: LeadsTo_type [THEN subsetD]
intro: LeadsTo_Un subset_imp_LeadsTo)
-lemma LeadsTo_Trans_Un: "[| F \<in> A \<longmapsto>w B; F \<in> B \<longmapsto>w C |]
- ==> F \<in> (A \<union> B) \<longmapsto>w C"
+lemma LeadsTo_Trans_Un: "\<lbrakk>F \<in> A \<longmapsto>w B; F \<in> B \<longmapsto>w C\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<union> B) \<longmapsto>w C"
apply (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans dest: LeadsTo_type [THEN subsetD])
done
@@ -166,14 +166,14 @@
(** More rules using the premise "Always(I)" **)
-lemma EnsuresI: "[| F:(A-B) Co (A \<union> B); F \<in> transient (A-B) |] ==> F \<in> A Ensures B"
+lemma EnsuresI: "\<lbrakk>F:(A-B) Co (A \<union> B); F \<in> transient (A-B)\<rbrakk> \<Longrightarrow> F \<in> A Ensures B"
apply (simp add: Ensures_def Constrains_eq_constrains)
apply (blast intro: ensuresI constrains_weaken transient_strengthen dest: constrainsD2)
done
-lemma Always_LeadsTo_Basis: "[| F \<in> Always(I); F \<in> (I \<inter> (A-A')) Co (A \<union> A');
- F \<in> transient (I \<inter> (A-A')) |]
- ==> F \<in> A \<longmapsto>w A'"
+lemma Always_LeadsTo_Basis: "\<lbrakk>F \<in> Always(I); F \<in> (I \<inter> (A-A')) Co (A \<union> A');
+ F \<in> transient (I \<inter> (A-A'))\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto>w A'"
apply (rule Always_LeadsToI, assumption)
apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen)
done
@@ -181,36 +181,36 @@
(*Set difference: maybe combine with leadsTo_weaken_L??
This is the most useful form of the "disjunction" rule*)
lemma LeadsTo_Diff:
- "[| F \<in> (A-B) \<longmapsto>w C; F \<in> (A \<inter> B) \<longmapsto>w C |] ==> F \<in> A \<longmapsto>w C"
+ "\<lbrakk>F \<in> (A-B) \<longmapsto>w C; F \<in> (A \<inter> B) \<longmapsto>w C\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w C"
by (blast intro: LeadsTo_Un LeadsTo_weaken)
lemma LeadsTo_UN_UN:
- "[|(!!i. i \<in> I ==> F \<in> A(i) \<longmapsto>w A'(i)); F \<in> program |]
- ==> F \<in> (\<Union>i \<in> I. A(i)) \<longmapsto>w (\<Union>i \<in> I. A'(i))"
+ "\<lbrakk>(\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto>w A'(i)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Union>i \<in> I. A(i)) \<longmapsto>w (\<Union>i \<in> I. A'(i))"
apply (rule LeadsTo_Union, auto)
apply (blast intro: LeadsTo_weaken_R)
done
(*Binary union version*)
lemma LeadsTo_Un_Un:
- "[| F \<in> A \<longmapsto>w A'; F \<in> B \<longmapsto>w B' |] ==> F:(A \<union> B) \<longmapsto>w (A' \<union> B')"
+ "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> B \<longmapsto>w B'\<rbrakk> \<Longrightarrow> F:(A \<union> B) \<longmapsto>w (A' \<union> B')"
by (blast intro: LeadsTo_Un LeadsTo_weaken_R)
(** The cancellation law **)
-lemma LeadsTo_cancel2: "[| F \<in> A \<longmapsto>w(A' \<union> B); F \<in> B \<longmapsto>w B' |] ==> F \<in> A \<longmapsto>w (A' \<union> B')"
+lemma LeadsTo_cancel2: "\<lbrakk>F \<in> A \<longmapsto>w(A' \<union> B); F \<in> B \<longmapsto>w B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w (A' \<union> B')"
by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans dest: LeadsTo_type [THEN subsetD])
lemma Un_Diff: "A \<union> (B - A) = A \<union> B"
by auto
-lemma LeadsTo_cancel_Diff2: "[| F \<in> A \<longmapsto>w (A' \<union> B); F \<in> (B-A') \<longmapsto>w B' |] ==> F \<in> A \<longmapsto>w (A' \<union> B')"
+lemma LeadsTo_cancel_Diff2: "\<lbrakk>F \<in> A \<longmapsto>w (A' \<union> B); F \<in> (B-A') \<longmapsto>w B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w (A' \<union> B')"
apply (rule LeadsTo_cancel2)
prefer 2 apply assumption
apply (simp (no_asm_simp) add: Un_Diff)
done
-lemma LeadsTo_cancel1: "[| F \<in> A \<longmapsto>w (B \<union> A'); F \<in> B \<longmapsto>w B' |] ==> F \<in> A \<longmapsto>w (B' \<union> A')"
+lemma LeadsTo_cancel1: "\<lbrakk>F \<in> A \<longmapsto>w (B \<union> A'); F \<in> B \<longmapsto>w B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w (B' \<union> A')"
apply (simp add: Un_commute)
apply (blast intro!: LeadsTo_cancel2)
done
@@ -218,7 +218,7 @@
lemma Diff_Un2: "(B - A) \<union> A = B \<union> A"
by auto
-lemma LeadsTo_cancel_Diff1: "[| F \<in> A \<longmapsto>w (B \<union> A'); F \<in> (B-A') \<longmapsto>w B' |] ==> F \<in> A \<longmapsto>w (B' \<union> A')"
+lemma LeadsTo_cancel_Diff1: "\<lbrakk>F \<in> A \<longmapsto>w (B \<union> A'); F \<in> (B-A') \<longmapsto>w B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w (B' \<union> A')"
apply (rule LeadsTo_cancel1)
prefer 2 apply assumption
apply (simp (no_asm_simp) add: Diff_Un2)
@@ -227,7 +227,7 @@
(** The impossibility law **)
(*The set "A" may be non-empty, but it contains no reachable states*)
-lemma LeadsTo_empty: "F \<in> A \<longmapsto>w 0 ==> F \<in> Always (state -A)"
+lemma LeadsTo_empty: "F \<in> A \<longmapsto>w 0 \<Longrightarrow> F \<in> Always (state -A)"
apply (simp (no_asm_use) add: LeadsTo_def Always_eq_includes_reachable)
apply (cut_tac reachable_type)
apply (auto dest!: leadsTo_empty)
@@ -236,26 +236,26 @@
(** PSP \<in> Progress-Safety-Progress **)
(*Special case of PSP \<in> Misra's "stable conjunction"*)
-lemma PSP_Stable: "[| F \<in> A \<longmapsto>w A'; F \<in> Stable(B) |]==> F:(A \<inter> B) \<longmapsto>w (A' \<inter> B)"
+lemma PSP_Stable: "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> Stable(B)\<rbrakk>\<Longrightarrow> F:(A \<inter> B) \<longmapsto>w (A' \<inter> B)"
apply (simp add: LeadsTo_def Stable_eq_stable, clarify)
apply (drule psp_stable, assumption)
apply (simp add: Int_ac)
done
-lemma PSP_Stable2: "[| F \<in> A \<longmapsto>w A'; F \<in> Stable(B) |] ==> F \<in> (B \<inter> A) \<longmapsto>w (B \<inter> A')"
+lemma PSP_Stable2: "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> Stable(B)\<rbrakk> \<Longrightarrow> F \<in> (B \<inter> A) \<longmapsto>w (B \<inter> A')"
apply (simp (no_asm_simp) add: PSP_Stable Int_ac)
done
-lemma PSP: "[| F \<in> A \<longmapsto>w A'; F \<in> B Co B'|]==> F \<in> (A \<inter> B') \<longmapsto>w ((A' \<inter> B) \<union> (B' - B))"
+lemma PSP: "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> B Co B'\<rbrakk>\<Longrightarrow> F \<in> (A \<inter> B') \<longmapsto>w ((A' \<inter> B) \<union> (B' - B))"
apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains)
apply (blast dest: psp intro: leadsTo_weaken)
done
-lemma PSP2: "[| F \<in> A \<longmapsto>w A'; F \<in> B Co B' |]==> F:(B' \<inter> A) \<longmapsto>w ((B \<inter> A') \<union> (B' - B))"
+lemma PSP2: "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> B Co B'\<rbrakk>\<Longrightarrow> F:(B' \<inter> A) \<longmapsto>w ((B \<inter> A') \<union> (B' - B))"
by (simp (no_asm_simp) add: PSP Int_ac)
lemma PSP_Unless:
-"[| F \<in> A \<longmapsto>w A'; F \<in> B Unless B'|]==> F:(A \<inter> B) \<longmapsto>w ((A' \<inter> B) \<union> B')"
+"\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> B Unless B'\<rbrakk>\<Longrightarrow> F:(A \<inter> B) \<longmapsto>w ((A' \<inter> B) \<union> B')"
apply (unfold op_Unless_def)
apply (drule PSP, assumption)
apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo)
@@ -264,11 +264,11 @@
(*** Induction rules ***)
(** Meta or object quantifier ????? **)
-lemma LeadsTo_wf_induct: "[| wf(r);
+lemma LeadsTo_wf_induct: "\<lbrakk>wf(r);
\<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) \<longmapsto>w
((A \<inter> f-``(converse(r) `` {m})) \<union> B);
- field(r)<=I; A<=f-``I; F \<in> program |]
- ==> F \<in> A \<longmapsto>w B"
+ field(r)<=I; A<=f-``I; F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto>w B"
apply (simp (no_asm_use) add: LeadsTo_def)
apply auto
apply (erule_tac I = I and f = f in leadsTo_wf_induct, safe)
@@ -278,8 +278,8 @@
done
-lemma LessThan_induct: "[| \<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto>w ((A \<inter> f-``m) \<union> B);
- A<=f-``nat; F \<in> program |] ==> F \<in> A \<longmapsto>w B"
+lemma LessThan_induct: "\<lbrakk>\<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto>w ((A \<inter> f-``m) \<union> B);
+ A<=f-``nat; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B"
apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct])
apply (simp_all add: nat_measure_field)
apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
@@ -297,16 +297,16 @@
(*** Completion \<in> Binary and General Finite versions ***)
-lemma Completion: "[| F \<in> A \<longmapsto>w (A' \<union> C); F \<in> A' Co (A' \<union> C);
- F \<in> B \<longmapsto>w (B' \<union> C); F \<in> B' Co (B' \<union> C) |]
- ==> F \<in> (A \<inter> B) \<longmapsto>w ((A' \<inter> B') \<union> C)"
+lemma Completion: "\<lbrakk>F \<in> A \<longmapsto>w (A' \<union> C); F \<in> A' Co (A' \<union> C);
+ F \<in> B \<longmapsto>w (B' \<union> C); F \<in> B' Co (B' \<union> C)\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto>w ((A' \<inter> B') \<union> C)"
apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains Int_Un_distrib)
apply (blast intro: completion leadsTo_weaken)
done
lemma Finite_completion_aux:
- "[| I \<in> Fin(X);F \<in> program |]
- ==> (\<forall>i \<in> I. F \<in> (A(i)) \<longmapsto>w (A'(i) \<union> C)) \<longrightarrow>
+ "\<lbrakk>I \<in> Fin(X);F \<in> program\<rbrakk>
+ \<Longrightarrow> (\<forall>i \<in> I. F \<in> (A(i)) \<longmapsto>w (A'(i) \<union> C)) \<longrightarrow>
(\<forall>i \<in> I. F \<in> (A'(i)) Co (A'(i) \<union> C)) \<longrightarrow>
F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto>w ((\<Inter>i \<in> I. A'(i)) \<union> C)"
apply (erule Fin_induct)
@@ -317,16 +317,16 @@
done
lemma Finite_completion:
- "[| I \<in> Fin(X); !!i. i \<in> I ==> F \<in> A(i) \<longmapsto>w (A'(i) \<union> C);
- !!i. i \<in> I ==> F \<in> A'(i) Co (A'(i) \<union> C);
- F \<in> program |]
- ==> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto>w ((\<Inter>i \<in> I. A'(i)) \<union> C)"
+ "\<lbrakk>I \<in> Fin(X); \<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto>w (A'(i) \<union> C);
+ \<And>i. i \<in> I \<Longrightarrow> F \<in> A'(i) Co (A'(i) \<union> C);
+ F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto>w ((\<Inter>i \<in> I. A'(i)) \<union> C)"
by (blast intro: Finite_completion_aux [THEN mp, THEN mp])
lemma Stable_completion:
- "[| F \<in> A \<longmapsto>w A'; F \<in> Stable(A');
- F \<in> B \<longmapsto>w B'; F \<in> Stable(B') |]
- ==> F \<in> (A \<inter> B) \<longmapsto>w (A' \<inter> B')"
+ "\<lbrakk>F \<in> A \<longmapsto>w A'; F \<in> Stable(A');
+ F \<in> B \<longmapsto>w B'; F \<in> Stable(B')\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto>w (A' \<inter> B')"
apply (unfold Stable_def)
apply (rule_tac C1 = 0 in Completion [THEN LeadsTo_weaken_R])
prefer 5
@@ -335,10 +335,10 @@
done
lemma Finite_stable_completion:
- "[| I \<in> Fin(X);
- (!!i. i \<in> I ==> F \<in> A(i) \<longmapsto>w A'(i));
- (!!i. i \<in> I ==>F \<in> Stable(A'(i))); F \<in> program |]
- ==> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto>w (\<Inter>i \<in> I. A'(i))"
+ "\<lbrakk>I \<in> Fin(X);
+ (\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto>w A'(i));
+ (\<And>i. i \<in> I \<Longrightarrow>F \<in> Stable(A'(i))); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto>w (\<Inter>i \<in> I. A'(i))"
apply (unfold Stable_def)
apply (rule_tac C1 = 0 in Finite_completion [THEN LeadsTo_weaken_R], simp_all)
apply (rule_tac [3] subset_refl, auto)
--- a/src/ZF/UNITY/UNITY.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/UNITY.thy Tue Sep 27 16:51:35 2022 +0100
@@ -14,7 +14,7 @@
definition
program :: i where
- "program == {<init, acts, allowed>:
+ "program \<equiv> {<init, acts, allowed>:
Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)).
id(state) \<in> acts & id(state) \<in> allowed}"
@@ -22,80 +22,80 @@
mk_program :: "[i,i,i]=>i" where
\<comment> \<open>The definition yields a program thanks to the coercions
init \<inter> state, acts \<inter> Pow(state*state), etc.\<close>
- "mk_program(init, acts, allowed) ==
+ "mk_program(init, acts, allowed) \<equiv>
<init \<inter> state, cons(id(state), acts \<inter> Pow(state*state)),
cons(id(state), allowed \<inter> Pow(state*state))>"
definition
SKIP :: i (\<open>\<bottom>\<close>) where
- "SKIP == mk_program(state, 0, Pow(state*state))"
+ "SKIP \<equiv> mk_program(state, 0, Pow(state*state))"
(* Coercion from anything to program *)
definition
programify :: "i=>i" where
- "programify(F) == if F \<in> program then F else SKIP"
+ "programify(F) \<equiv> if F \<in> program then F else SKIP"
definition
RawInit :: "i=>i" where
- "RawInit(F) == fst(F)"
+ "RawInit(F) \<equiv> fst(F)"
definition
Init :: "i=>i" where
- "Init(F) == RawInit(programify(F))"
+ "Init(F) \<equiv> RawInit(programify(F))"
definition
RawActs :: "i=>i" where
- "RawActs(F) == cons(id(state), fst(snd(F)))"
+ "RawActs(F) \<equiv> cons(id(state), fst(snd(F)))"
definition
Acts :: "i=>i" where
- "Acts(F) == RawActs(programify(F))"
+ "Acts(F) \<equiv> RawActs(programify(F))"
definition
RawAllowedActs :: "i=>i" where
- "RawAllowedActs(F) == cons(id(state), snd(snd(F)))"
+ "RawAllowedActs(F) \<equiv> cons(id(state), snd(snd(F)))"
definition
AllowedActs :: "i=>i" where
- "AllowedActs(F) == RawAllowedActs(programify(F))"
+ "AllowedActs(F) \<equiv> RawAllowedActs(programify(F))"
definition
Allowed :: "i =>i" where
- "Allowed(F) == {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}"
+ "Allowed(F) \<equiv> {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}"
definition
initially :: "i=>i" where
- "initially(A) == {F \<in> program. Init(F)\<subseteq>A}"
+ "initially(A) \<equiv> {F \<in> program. Init(F)\<subseteq>A}"
definition "constrains" :: "[i, i] => i" (infixl \<open>co\<close> 60) where
- "A co B == {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) & st_set(A)}"
+ "A co B \<equiv> {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) & st_set(A)}"
\<comment> \<open>the condition \<^term>\<open>st_set(A)\<close> makes the definition slightly
stronger than the HOL one\<close>
definition unless :: "[i, i] => i" (infixl \<open>unless\<close> 60) where
- "A unless B == (A - B) co (A \<union> B)"
+ "A unless B \<equiv> (A - B) co (A \<union> B)"
definition
stable :: "i=>i" where
- "stable(A) == A co A"
+ "stable(A) \<equiv> A co A"
definition
strongest_rhs :: "[i, i] => i" where
- "strongest_rhs(F, A) == \<Inter>({B \<in> Pow(state). F \<in> A co B})"
+ "strongest_rhs(F, A) \<equiv> \<Inter>({B \<in> Pow(state). F \<in> A co B})"
definition
invariant :: "i => i" where
- "invariant(A) == initially(A) \<inter> stable(A)"
+ "invariant(A) \<equiv> initially(A) \<inter> stable(A)"
(* meta-function composition *)
definition
metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl \<open>comp\<close> 65) where
- "f comp g == %x. f(g(x))"
+ "f comp g \<equiv> %x. f(g(x))"
definition
pg_compl :: "i=>i" where
- "pg_compl(X)== program - X"
+ "pg_compl(X)\<equiv> program - X"
text\<open>SKIP\<close>
lemma SKIP_in_program [iff,TC]: "SKIP \<in> program"
@@ -105,7 +105,7 @@
subsection\<open>The function \<^term>\<open>programify\<close>, the coercion from anything to
program\<close>
-lemma programify_program [simp]: "F \<in> program ==> programify(F)=F"
+lemma programify_program [simp]: "F \<in> program \<Longrightarrow> programify(F)=F"
by (force simp add: programify_def)
lemma programify_in_program [iff,TC]: "programify(F) \<in> program"
@@ -127,13 +127,13 @@
subsection\<open>The Inspectors for Programs\<close>
-lemma id_in_RawActs: "F \<in> program ==>id(state) \<in> RawActs(F)"
+lemma id_in_RawActs: "F \<in> program \<Longrightarrow>id(state) \<in> RawActs(F)"
by (auto simp add: program_def RawActs_def)
lemma id_in_Acts [iff,TC]: "id(state) \<in> Acts(F)"
by (simp add: id_in_RawActs Acts_def)
-lemma id_in_RawAllowedActs: "F \<in> program ==>id(state) \<in> RawAllowedActs(F)"
+lemma id_in_RawAllowedActs: "F \<in> program \<Longrightarrow>id(state) \<in> RawAllowedActs(F)"
by (auto simp add: program_def RawAllowedActs_def)
lemma id_in_AllowedActs [iff,TC]: "id(state) \<in> AllowedActs(F)"
@@ -149,14 +149,14 @@
subsection\<open>Types of the Inspectors\<close>
-lemma RawInit_type: "F \<in> program ==> RawInit(F)\<subseteq>state"
+lemma RawInit_type: "F \<in> program \<Longrightarrow> RawInit(F)\<subseteq>state"
by (auto simp add: program_def RawInit_def)
-lemma RawActs_type: "F \<in> program ==> RawActs(F)\<subseteq>Pow(state*state)"
+lemma RawActs_type: "F \<in> program \<Longrightarrow> RawActs(F)\<subseteq>Pow(state*state)"
by (auto simp add: program_def RawActs_def)
lemma RawAllowedActs_type:
- "F \<in> program ==> RawAllowedActs(F)\<subseteq>Pow(state*state)"
+ "F \<in> program \<Longrightarrow> RawAllowedActs(F)\<subseteq>Pow(state*state)"
by (auto simp add: program_def RawAllowedActs_def)
lemma Init_type: "Init(F)\<subseteq>state"
@@ -176,11 +176,11 @@
by (simp add: RawAllowedActs_type AllowedActs_def)
text\<open>Needed in Behaviors\<close>
-lemma ActsD: "[| act \<in> Acts(F); <s,s'> \<in> act |] ==> s \<in> state & s' \<in> state"
+lemma ActsD: "\<lbrakk>act \<in> Acts(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state & s' \<in> state"
by (blast dest: Acts_type [THEN subsetD])
lemma AllowedActsD:
- "[| act \<in> AllowedActs(F); <s,s'> \<in> act |] ==> s \<in> state & s' \<in> state"
+ "\<lbrakk>act \<in> AllowedActs(F); <s,s'> \<in> act\<rbrakk> \<Longrightarrow> s \<in> state & s' \<in> state"
by (blast dest: AllowedActs_type [THEN subsetD])
subsection\<open>Simplification rules involving \<^term>\<open>state\<close>, \<^term>\<open>Init\<close>,
@@ -280,7 +280,7 @@
text\<open>Equality of UNITY programs\<close>
lemma raw_surjective_mk_program:
- "F \<in> program ==> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F"
+ "F \<in> program \<Longrightarrow> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F"
apply (auto simp add: program_def mk_program_def RawInit_def RawActs_def
RawAllowedActs_def, blast+)
done
@@ -290,45 +290,45 @@
by (auto simp add: raw_surjective_mk_program Init_def Acts_def AllowedActs_def)
lemma program_equalityI:
- "[|Init(F) = Init(G); Acts(F) = Acts(G);
- AllowedActs(F) = AllowedActs(G); F \<in> program; G \<in> program |] ==> F = G"
+ "\<lbrakk>Init(F) = Init(G); Acts(F) = Acts(G);
+ AllowedActs(F) = AllowedActs(G); F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow> F = G"
apply (subgoal_tac "programify(F) = programify(G)")
apply simp
apply (simp only: surjective_mk_program [symmetric])
done
lemma program_equalityE:
- "[|F = G;
- [|Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G) |]
- ==> P |]
- ==> P"
+ "\<lbrakk>F = G;
+ \<lbrakk>Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G)\<rbrakk>
+ \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by force
lemma program_equality_iff:
- "[| F \<in> program; G \<in> program |] ==>(F=G) \<longleftrightarrow>
+ "\<lbrakk>F \<in> program; G \<in> program\<rbrakk> \<Longrightarrow>(F=G) \<longleftrightarrow>
(Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(G))"
by (blast intro: program_equalityI program_equalityE)
subsection\<open>These rules allow "lazy" definition expansion\<close>
lemma def_prg_Init:
- "F == mk_program (init,acts,allowed) ==> Init(F) = init \<inter> state"
+ "F \<equiv> mk_program (init,acts,allowed) \<Longrightarrow> Init(F) = init \<inter> state"
by auto
lemma def_prg_Acts:
- "F == mk_program (init,acts,allowed)
- ==> Acts(F) = cons(id(state), acts \<inter> Pow(state*state))"
+ "F \<equiv> mk_program (init,acts,allowed)
+ \<Longrightarrow> Acts(F) = cons(id(state), acts \<inter> Pow(state*state))"
by auto
lemma def_prg_AllowedActs:
- "F == mk_program (init,acts,allowed)
- ==> AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
+ "F \<equiv> mk_program (init,acts,allowed)
+ \<Longrightarrow> AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
by auto
lemma def_prg_simps:
- "[| F == mk_program (init,acts,allowed) |]
- ==> Init(F) = init \<inter> state &
+ "\<lbrakk>F \<equiv> mk_program (init,acts,allowed)\<rbrakk>
+ \<Longrightarrow> Init(F) = init \<inter> state &
Acts(F) = cons(id(state), acts \<inter> Pow(state*state)) &
AllowedActs(F) = cons(id(state), allowed \<inter> Pow(state*state))"
by auto
@@ -336,12 +336,12 @@
text\<open>An action is expanded only if a pair of states is being tested against it\<close>
lemma def_act_simp:
- "[| act == {<s,s'> \<in> A*B. P(s, s')} |]
- ==> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B & P(s, s'))"
+ "\<lbrakk>act \<equiv> {<s,s'> \<in> A*B. P(s, s')}\<rbrakk>
+ \<Longrightarrow> (<s,s'> \<in> act) \<longleftrightarrow> (<s,s'> \<in> A*B & P(s, s'))"
by auto
text\<open>A set is expanded only if an element is being tested against it\<close>
-lemma def_set_simp: "A == B ==> (x \<in> A) \<longleftrightarrow> (x \<in> B)"
+lemma def_set_simp: "A \<equiv> B \<Longrightarrow> (x \<in> A) \<longleftrightarrow> (x \<in> B)"
by auto
@@ -351,15 +351,15 @@
by (force simp add: constrains_def)
lemma constrainsI:
- "[|(!!act s s'. [| act: Acts(F); <s,s'> \<in> act; s \<in> A|] ==> s' \<in> A');
- F \<in> program; st_set(A) |] ==> F \<in> A co A'"
+ "\<lbrakk>(\<And>act s s'. \<lbrakk>act: Acts(F); <s,s'> \<in> act; s \<in> A\<rbrakk> \<Longrightarrow> s' \<in> A');
+ F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A co A'"
by (force simp add: constrains_def)
lemma constrainsD:
- "F \<in> A co B ==> \<forall>act \<in> Acts(F). act``A\<subseteq>B"
+ "F \<in> A co B \<Longrightarrow> \<forall>act \<in> Acts(F). act``A\<subseteq>B"
by (force simp add: constrains_def)
-lemma constrainsD2: "F \<in> A co B ==> F \<in> program & st_set(A)"
+lemma constrainsD2: "F \<in> A co B \<Longrightarrow> F \<in> program & st_set(A)"
by (force simp add: constrains_def)
lemma constrains_empty [iff]: "F \<in> 0 co B \<longleftrightarrow> F \<in> program"
@@ -380,18 +380,18 @@
text\<open>monotonic in 2nd argument\<close>
lemma constrains_weaken_R:
- "[| F \<in> A co A'; A'\<subseteq>B' |] ==> F \<in> A co B'"
+ "\<lbrakk>F \<in> A co A'; A'\<subseteq>B'\<rbrakk> \<Longrightarrow> F \<in> A co B'"
apply (unfold constrains_def, blast)
done
text\<open>anti-monotonic in 1st argument\<close>
lemma constrains_weaken_L:
- "[| F \<in> A co A'; B\<subseteq>A |] ==> F \<in> B co A'"
+ "\<lbrakk>F \<in> A co A'; B\<subseteq>A\<rbrakk> \<Longrightarrow> F \<in> B co A'"
apply (unfold constrains_def st_set_def, blast)
done
lemma constrains_weaken:
- "[| F \<in> A co A'; B\<subseteq>A; A'\<subseteq>B' |] ==> F \<in> B co B'"
+ "\<lbrakk>F \<in> A co A'; B\<subseteq>A; A'\<subseteq>B'\<rbrakk> \<Longrightarrow> F \<in> B co B'"
apply (drule constrains_weaken_R)
apply (drule_tac [2] constrains_weaken_L, blast+)
done
@@ -400,12 +400,12 @@
subsection\<open>Constrains and Union\<close>
lemma constrains_Un:
- "[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<union> B) co (A' \<union> B')"
+ "\<lbrakk>F \<in> A co A'; F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) co (A' \<union> B')"
by (auto simp add: constrains_def st_set_def, force)
lemma constrains_UN:
- "[|!!i. i \<in> I ==> F \<in> A(i) co A'(i); F \<in> program |]
- ==> F \<in> (\<Union>i \<in> I. A(i)) co (\<Union>i \<in> I. A'(i))"
+ "\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) co A'(i); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Union>i \<in> I. A(i)) co (\<Union>i \<in> I. A'(i))"
by (force simp add: constrains_def st_set_def)
lemma constrains_Un_distrib:
@@ -413,7 +413,7 @@
by (force simp add: constrains_def st_set_def)
lemma constrains_UN_distrib:
- "i \<in> I ==> (\<Union>i \<in> I. A(i)) co B = (\<Inter>i \<in> I. A(i) co B)"
+ "i \<in> I \<Longrightarrow> (\<Union>i \<in> I. A(i)) co B = (\<Inter>i \<in> I. A(i) co B)"
by (force simp add: constrains_def st_set_def)
@@ -423,16 +423,16 @@
by (force simp add: constrains_def st_set_def)
lemma constrains_INT_distrib:
- "x \<in> I ==> A co (\<Inter>i \<in> I. B(i)) = (\<Inter>i \<in> I. A co B(i))"
+ "x \<in> I \<Longrightarrow> A co (\<Inter>i \<in> I. B(i)) = (\<Inter>i \<in> I. A co B(i))"
by (force simp add: constrains_def st_set_def)
lemma constrains_Int:
- "[| F \<in> A co A'; F \<in> B co B' |] ==> F \<in> (A \<inter> B) co (A' \<inter> B')"
+ "\<lbrakk>F \<in> A co A'; F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> (A \<inter> B) co (A' \<inter> B')"
by (force simp add: constrains_def st_set_def)
lemma constrains_INT [rule_format]:
- "[| \<forall>i \<in> I. F \<in> A(i) co A'(i); F \<in> program|]
- ==> F \<in> (\<Inter>i \<in> I. A(i)) co (\<Inter>i \<in> I. A'(i))"
+ "\<lbrakk>\<forall>i \<in> I. F \<in> A(i) co A'(i); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) co (\<Inter>i \<in> I. A'(i))"
apply (case_tac "I=0")
apply (simp add: Inter_def)
apply (erule not_emptyE)
@@ -442,22 +442,22 @@
(* The rule below simulates the HOL's one for (\<Inter>z. A i) co (\<Inter>z. B i) *)
lemma constrains_All:
-"[| \<forall>z. F:{s \<in> state. P(s, z)} co {s \<in> state. Q(s, z)}; F \<in> program |]==>
+"\<lbrakk>\<forall>z. F:{s \<in> state. P(s, z)} co {s \<in> state. Q(s, z)}; F \<in> program\<rbrakk>\<Longrightarrow>
F:{s \<in> state. \<forall>z. P(s, z)} co {s \<in> state. \<forall>z. Q(s, z)}"
by (unfold constrains_def, blast)
lemma constrains_imp_subset:
- "[| F \<in> A co A' |] ==> A \<subseteq> A'"
+ "\<lbrakk>F \<in> A co A'\<rbrakk> \<Longrightarrow> A \<subseteq> A'"
by (unfold constrains_def st_set_def, force)
text\<open>The reasoning is by subsets since "co" refers to single actions
only. So this rule isn't that useful.\<close>
-lemma constrains_trans: "[| F \<in> A co B; F \<in> B co C |] ==> F \<in> A co C"
+lemma constrains_trans: "\<lbrakk>F \<in> A co B; F \<in> B co C\<rbrakk> \<Longrightarrow> F \<in> A co C"
by (unfold constrains_def st_set_def, auto, blast)
lemma constrains_cancel:
-"[| F \<in> A co (A' \<union> B); F \<in> B co B' |] ==> F \<in> A co (A' \<union> B')"
+"\<lbrakk>F \<in> A co (A' \<union> B); F \<in> B co B'\<rbrakk> \<Longrightarrow> F \<in> A co (A' \<union> B')"
apply (drule_tac A = B in constrains_imp_subset)
apply (blast intro: constrains_weaken_R)
done
@@ -468,12 +468,12 @@
lemma unless_type: "A unless B \<subseteq> program"
by (force simp add: unless_def constrains_def)
-lemma unlessI: "[| F \<in> (A-B) co (A \<union> B) |] ==> F \<in> A unless B"
+lemma unlessI: "\<lbrakk>F \<in> (A-B) co (A \<union> B)\<rbrakk> \<Longrightarrow> F \<in> A unless B"
apply (unfold unless_def)
apply (blast dest: constrainsD2)
done
-lemma unlessD: "F :A unless B ==> F \<in> (A-B) co (A \<union> B)"
+lemma unlessD: "F :A unless B \<Longrightarrow> F \<in> (A-B) co (A \<union> B)"
by (unfold unless_def, auto)
@@ -482,10 +482,10 @@
lemma initially_type: "initially(A) \<subseteq> program"
by (unfold initially_def, blast)
-lemma initiallyI: "[| F \<in> program; Init(F)\<subseteq>A |] ==> F \<in> initially(A)"
+lemma initiallyI: "\<lbrakk>F \<in> program; Init(F)\<subseteq>A\<rbrakk> \<Longrightarrow> F \<in> initially(A)"
by (unfold initially_def, blast)
-lemma initiallyD: "F \<in> initially(A) ==> Init(F)\<subseteq>A"
+lemma initiallyD: "F \<in> initially(A) \<Longrightarrow> Init(F)\<subseteq>A"
by (unfold initially_def, blast)
@@ -494,13 +494,13 @@
lemma stable_type: "stable(A)\<subseteq>program"
by (unfold stable_def constrains_def, blast)
-lemma stableI: "F \<in> A co A ==> F \<in> stable(A)"
+lemma stableI: "F \<in> A co A \<Longrightarrow> F \<in> stable(A)"
by (unfold stable_def, assumption)
-lemma stableD: "F \<in> stable(A) ==> F \<in> A co A"
+lemma stableD: "F \<in> stable(A) \<Longrightarrow> F \<in> A co A"
by (unfold stable_def, assumption)
-lemma stableD2: "F \<in> stable(A) ==> F \<in> program & st_set(A)"
+lemma stableD2: "F \<in> stable(A) \<Longrightarrow> F \<in> program & st_set(A)"
by (unfold stable_def constrains_def, auto)
lemma stable_state [simp]: "stable(state) = program"
@@ -514,49 +514,49 @@
subsection\<open>Union and Intersection with \<^term>\<open>stable\<close>\<close>
lemma stable_Un:
- "[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable(A \<union> A')"
+ "\<lbrakk>F \<in> stable(A); F \<in> stable(A')\<rbrakk> \<Longrightarrow> F \<in> stable(A \<union> A')"
apply (unfold stable_def)
apply (blast intro: constrains_Un)
done
lemma stable_UN:
- "[|!!i. i\<in>I ==> F \<in> stable(A(i)); F \<in> program |]
- ==> F \<in> stable (\<Union>i \<in> I. A(i))"
+ "\<lbrakk>\<And>i. i\<in>I \<Longrightarrow> F \<in> stable(A(i)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> stable (\<Union>i \<in> I. A(i))"
apply (unfold stable_def)
apply (blast intro: constrains_UN)
done
lemma stable_Int:
- "[| F \<in> stable(A); F \<in> stable(A') |] ==> F \<in> stable (A \<inter> A')"
+ "\<lbrakk>F \<in> stable(A); F \<in> stable(A')\<rbrakk> \<Longrightarrow> F \<in> stable (A \<inter> A')"
apply (unfold stable_def)
apply (blast intro: constrains_Int)
done
lemma stable_INT:
- "[| !!i. i \<in> I ==> F \<in> stable(A(i)); F \<in> program |]
- ==> F \<in> stable (\<Inter>i \<in> I. A(i))"
+ "\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> F \<in> stable(A(i)); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> stable (\<Inter>i \<in> I. A(i))"
apply (unfold stable_def)
apply (blast intro: constrains_INT)
done
lemma stable_All:
- "[|\<forall>z. F \<in> stable({s \<in> state. P(s, z)}); F \<in> program|]
- ==> F \<in> stable({s \<in> state. \<forall>z. P(s, z)})"
+ "\<lbrakk>\<forall>z. F \<in> stable({s \<in> state. P(s, z)}); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> stable({s \<in> state. \<forall>z. P(s, z)})"
apply (unfold stable_def)
apply (rule constrains_All, auto)
done
lemma stable_constrains_Un:
- "[| F \<in> stable(C); F \<in> A co (C \<union> A') |] ==> F \<in> (C \<union> A) co (C \<union> A')"
+ "\<lbrakk>F \<in> stable(C); F \<in> A co (C \<union> A')\<rbrakk> \<Longrightarrow> F \<in> (C \<union> A) co (C \<union> A')"
apply (unfold stable_def constrains_def st_set_def, auto)
apply (blast dest!: bspec)
done
lemma stable_constrains_Int:
- "[| F \<in> stable(C); F \<in> (C \<inter> A) co A' |] ==> F \<in> (C \<inter> A) co (C \<inter> A')"
+ "\<lbrakk>F \<in> stable(C); F \<in> (C \<inter> A) co A'\<rbrakk> \<Longrightarrow> F \<in> (C \<inter> A) co (C \<inter> A')"
by (unfold stable_def constrains_def st_set_def, blast)
-(* [| F \<in> stable(C); F \<in> (C \<inter> A) co A |] ==> F \<in> stable(C \<inter> A) *)
+(* \<lbrakk>F \<in> stable(C); F \<in> (C \<inter> A) co A\<rbrakk> \<Longrightarrow> F \<in> stable(C \<inter> A) *)
lemmas stable_constrains_stable = stable_constrains_Int [THEN stableI]
subsection\<open>The Operator \<^term>\<open>invariant\<close>\<close>
@@ -566,15 +566,15 @@
apply (blast dest: stable_type [THEN subsetD])
done
-lemma invariantI: "[| Init(F)\<subseteq>A; F \<in> stable(A) |] ==> F \<in> invariant(A)"
+lemma invariantI: "\<lbrakk>Init(F)\<subseteq>A; F \<in> stable(A)\<rbrakk> \<Longrightarrow> F \<in> invariant(A)"
apply (unfold invariant_def initially_def)
apply (frule stable_type [THEN subsetD], auto)
done
-lemma invariantD: "F \<in> invariant(A) ==> Init(F)\<subseteq>A & F \<in> stable(A)"
+lemma invariantD: "F \<in> invariant(A) \<Longrightarrow> Init(F)\<subseteq>A & F \<in> stable(A)"
by (unfold invariant_def initially_def, auto)
-lemma invariantD2: "F \<in> invariant(A) ==> F \<in> program & st_set(A)"
+lemma invariantD2: "F \<in> invariant(A) \<Longrightarrow> F \<in> program & st_set(A)"
apply (unfold invariant_def)
apply (blast dest: stableD2)
done
@@ -582,7 +582,7 @@
text\<open>Could also say
\<^term>\<open>invariant(A) \<inter> invariant(B) \<subseteq> invariant (A \<inter> B)\<close>\<close>
lemma invariant_Int:
- "[| F \<in> invariant(A); F \<in> invariant(B) |] ==> F \<in> invariant(A \<inter> B)"
+ "\<lbrakk>F \<in> invariant(A); F \<in> invariant(B)\<rbrakk> \<Longrightarrow> F \<in> invariant(A \<inter> B)"
apply (unfold invariant_def initially_def)
apply (simp add: stable_Int, blast)
done
@@ -591,30 +591,30 @@
subsection\<open>The Elimination Theorem\<close>
(** The "free" m has become universally quantified!
- Should the premise be !!m instead of \<forall>m ? Would make it harder
+ Should the premise be \<And>m instead of \<forall>m ? Would make it harder
to use in forward proof. **)
text\<open>The general case is easier to prove than the special case!\<close>
lemma "elimination":
- "[| \<forall>m \<in> M. F \<in> {s \<in> A. x(s) = m} co B(m); F \<in> program |]
- ==> F \<in> {s \<in> A. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
+ "\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> A. x(s) = m} co B(m); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> {s \<in> A. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
by (auto simp add: constrains_def st_set_def, blast)
text\<open>As above, but for the special case of A=state\<close>
lemma elimination2:
- "[| \<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} co B(m); F \<in> program |]
- ==> F:{s \<in> state. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
+ "\<lbrakk>\<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} co B(m); F \<in> program\<rbrakk>
+ \<Longrightarrow> F:{s \<in> state. x(s) \<in> M} co (\<Union>m \<in> M. B(m))"
by (rule UNITY.elimination, auto)
subsection\<open>The Operator \<^term>\<open>strongest_rhs\<close>\<close>
lemma constrains_strongest_rhs:
- "[| F \<in> program; st_set(A) |] ==> F \<in> A co (strongest_rhs(F,A))"
+ "\<lbrakk>F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A co (strongest_rhs(F,A))"
by (auto simp add: constrains_def strongest_rhs_def st_set_def
dest: Acts_type [THEN subsetD])
lemma strongest_rhs_is_strongest:
- "[| F \<in> A co B; st_set(B) |] ==> strongest_rhs(F,A) \<subseteq> B"
+ "\<lbrakk>F \<in> A co B; st_set(B)\<rbrakk> \<Longrightarrow> strongest_rhs(F,A) \<subseteq> B"
by (auto simp add: constrains_def strongest_rhs_def st_set_def)
ML \<open>
--- a/src/ZF/UNITY/Union.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/Union.thy Tue Sep 27 16:51:35 2022 +0100
@@ -16,28 +16,28 @@
definition
(*FIXME: conjoin Init(F) \<inter> Init(G) \<noteq> 0 *)
ok :: "[i, i] => o" (infixl \<open>ok\<close> 65) where
- "F ok G == Acts(F) \<subseteq> AllowedActs(G) &
+ "F ok G \<equiv> Acts(F) \<subseteq> AllowedActs(G) &
Acts(G) \<subseteq> AllowedActs(F)"
definition
(*FIXME: conjoin (\<Inter>i \<in> I. Init(F(i))) \<noteq> 0 *)
OK :: "[i, i=>i] => o" where
- "OK(I,F) == (\<forall>i \<in> I. \<forall>j \<in> I-{i}. Acts(F(i)) \<subseteq> AllowedActs(F(j)))"
+ "OK(I,F) \<equiv> (\<forall>i \<in> I. \<forall>j \<in> I-{i}. Acts(F(i)) \<subseteq> AllowedActs(F(j)))"
definition
JOIN :: "[i, i=>i] => i" where
- "JOIN(I,F) == if I = 0 then SKIP else
+ "JOIN(I,F) \<equiv> if I = 0 then SKIP else
mk_program(\<Inter>i \<in> I. Init(F(i)), \<Union>i \<in> I. Acts(F(i)),
\<Inter>i \<in> I. AllowedActs(F(i)))"
definition
Join :: "[i, i] => i" (infixl \<open>\<squnion>\<close> 65) where
- "F \<squnion> G == mk_program (Init(F) \<inter> Init(G), Acts(F) \<union> Acts(G),
+ "F \<squnion> G \<equiv> mk_program (Init(F) \<inter> Init(G), Acts(F) \<union> Acts(G),
AllowedActs(F) \<inter> AllowedActs(G))"
definition
(*Characterizes safety properties. Used with specifying AllowedActs*)
safety_prop :: "i => o" where
- "safety_prop(X) == X\<subseteq>program &
+ "safety_prop(X) \<equiv> X\<subseteq>program &
SKIP \<in> X & (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) \<longrightarrow> G \<in> X)"
syntax
@@ -45,7 +45,7 @@
"_JOIN" :: "[pttrn, i, i] => i" (\<open>(3\<Squnion>_ \<in> _./ _)\<close> 10)
translations
- "\<Squnion>x \<in> A. B" == "CONST JOIN(A, (\<lambda>x. B))"
+ "\<Squnion>x \<in> A. B" == "CONST JOIN(A, (\<lambda>x. B))"
"\<Squnion>x y. B" == "\<Squnion>x. \<Squnion>y. B"
"\<Squnion>x. B" == "CONST JOIN(CONST state, (\<lambda>x. B))"
@@ -177,14 +177,14 @@
by (rule program_equalityI, auto)
lemma JOIN_cong [cong]:
- "[| I=J; !!i. i \<in> J ==> F(i) = G(i) |] ==>
+ "\<lbrakk>I=J; \<And>i. i \<in> J \<Longrightarrow> F(i) = G(i)\<rbrakk> \<Longrightarrow>
(\<Squnion>i \<in> I. F(i)) = (\<Squnion>i \<in> J. G(i))"
by (simp add: JOIN_def)
subsection\<open>JOIN laws\<close>
-lemma JOIN_absorb: "k \<in> I ==>F(k) \<squnion> (\<Squnion>i \<in> I. F(i)) = (\<Squnion>i \<in> I. F(i))"
+lemma JOIN_absorb: "k \<in> I \<Longrightarrow>F(k) \<squnion> (\<Squnion>i \<in> I. F(i)) = (\<Squnion>i \<in> I. F(i))"
apply (subst JOIN_cons [symmetric])
apply (auto simp add: cons_absorb)
done
@@ -208,7 +208,7 @@
by (simp add: JOIN_Join_distrib JOIN_constant)
text\<open>Used to prove guarantees_JOIN_I\<close>
-lemma JOIN_Join_diff: "i \<in> I==>F(i) \<squnion> JOIN(I - {i}, F) = JOIN(I, F)"
+lemma JOIN_Join_diff: "i \<in> I\<Longrightarrow>F(i) \<squnion> JOIN(I - {i}, F) = JOIN(I, F)"
apply (rule program_equalityI)
apply (auto elim!: not_emptyE)
done
@@ -221,7 +221,7 @@
I to be nonempty for other reasons anyway.*)
lemma JOIN_constrains:
- "i \<in> I==>(\<Squnion>i \<in> I. F(i)) \<in> A co B \<longleftrightarrow> (\<forall>i \<in> I. programify(F(i)) \<in> A co B)"
+ "i \<in> I\<Longrightarrow>(\<Squnion>i \<in> I. F(i)) \<in> A co B \<longleftrightarrow> (\<forall>i \<in> I. programify(F(i)) \<in> A co B)"
apply (unfold constrains_def JOIN_def st_set_def, auto)
prefer 2 apply blast
@@ -245,8 +245,8 @@
*)
lemma Join_constrains_weaken:
- "[| F \<in> A co A'; G \<in> B co B' |]
- ==> F \<squnion> G \<in> (A \<inter> B) co (A' \<union> B')"
+ "\<lbrakk>F \<in> A co A'; G \<in> B co B'\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> (A \<inter> B) co (A' \<union> B')"
apply (subgoal_tac "st_set (A) & st_set (B) & F \<in> program & G \<in> program")
prefer 2 apply (blast dest: constrainsD2, simp)
apply (blast intro: constrains_weaken)
@@ -254,7 +254,7 @@
(*If I=0, it degenerates to SKIP \<in> state co 0, which is false.*)
lemma JOIN_constrains_weaken:
- assumes major: "(!!i. i \<in> I ==> F(i) \<in> A(i) co A'(i))"
+ assumes major: "(\<And>i. i \<in> I \<Longrightarrow> F(i) \<in> A(i) co A'(i))"
and minor: "i \<in> I"
shows "(\<Squnion>i \<in> I. F(i)) \<in> (\<Inter>i \<in> I. A(i)) co (\<Union>i \<in> I. A'(i))"
apply (cut_tac minor)
@@ -274,7 +274,7 @@
done
lemma initially_JOIN_I:
- assumes major: "(!!i. i \<in> I ==>F(i) \<in> initially(A))"
+ assumes major: "(\<And>i. i \<in> I \<Longrightarrow>F(i) \<in> initially(A))"
and minor: "i \<in> I"
shows "(\<Squnion>i \<in> I. F(i)) \<in> initially(A)"
apply (cut_tac minor)
@@ -284,7 +284,7 @@
done
lemma invariant_JOIN_I:
- assumes major: "(!!i. i \<in> I ==> F(i) \<in> invariant(A))"
+ assumes major: "(\<And>i. i \<in> I \<Longrightarrow> F(i) \<in> invariant(A))"
and minor: "i \<in> I"
shows "(\<Squnion>i \<in> I. F(i)) \<in> invariant(A)"
apply (cut_tac minor)
@@ -302,12 +302,12 @@
by (simp add: stable_def)
lemma initially_JoinI [intro!]:
- "[| F \<in> initially(A); G \<in> initially(A) |] ==> F \<squnion> G \<in> initially(A)"
+ "\<lbrakk>F \<in> initially(A); G \<in> initially(A)\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> initially(A)"
by (unfold initially_def, auto)
lemma invariant_JoinI:
- "[| F \<in> invariant(A); G \<in> invariant(A) |]
- ==> F \<squnion> G \<in> invariant(A)"
+ "\<lbrakk>F \<in> invariant(A); G \<in> invariant(A)\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> invariant(A)"
apply (subgoal_tac "F \<in> program&G \<in> program")
prefer 2 apply (blast dest: invariantD2)
apply (simp add: invariant_def)
@@ -316,13 +316,13 @@
(* Fails if I=0 because \<Inter>i \<in> 0. A(i) = 0 *)
-lemma FP_JOIN: "i \<in> I ==> FP(\<Squnion>i \<in> I. F(i)) = (\<Inter>i \<in> I. FP (programify(F(i))))"
+lemma FP_JOIN: "i \<in> I \<Longrightarrow> FP(\<Squnion>i \<in> I. F(i)) = (\<Inter>i \<in> I. FP (programify(F(i))))"
by (auto simp add: FP_def Inter_def st_set_def JOIN_stable)
subsection\<open>Progress: transient, ensures\<close>
lemma JOIN_transient:
- "i \<in> I ==>
+ "i \<in> I \<Longrightarrow>
(\<Squnion>i \<in> I. F(i)) \<in> transient(A) \<longleftrightarrow> (\<exists>i \<in> I. programify(F(i)) \<in> transient(A))"
apply (auto simp add: transient_def JOIN_def)
apply (unfold st_set_def)
@@ -336,16 +336,16 @@
apply (auto simp add: transient_def Join_def Int_Un_distrib2)
done
-lemma Join_transient_I1: "F \<in> transient(A) ==> F \<squnion> G \<in> transient(A)"
+lemma Join_transient_I1: "F \<in> transient(A) \<Longrightarrow> F \<squnion> G \<in> transient(A)"
by (simp add: Join_transient transientD2)
-lemma Join_transient_I2: "G \<in> transient(A) ==> F \<squnion> G \<in> transient(A)"
+lemma Join_transient_I2: "G \<in> transient(A) \<Longrightarrow> F \<squnion> G \<in> transient(A)"
by (simp add: Join_transient transientD2)
-(*If I=0 it degenerates to (SKIP \<in> A ensures B) = False, i.e. to ~(A\<subseteq>B) *)
+(*If I=0 it degenerates to (SKIP \<in> A ensures B) = False, i.e. to \<not>(A\<subseteq>B) *)
lemma JOIN_ensures:
- "i \<in> I ==>
+ "i \<in> I \<Longrightarrow>
(\<Squnion>i \<in> I. F(i)) \<in> A ensures B \<longleftrightarrow>
((\<forall>i \<in> I. programify(F(i)) \<in> (A-B) co (A \<union> B)) &
(\<exists>i \<in> I. programify(F(i)) \<in> A ensures B))"
@@ -362,8 +362,8 @@
done
lemma stable_Join_constrains:
- "[| F \<in> stable(A); G \<in> A co A' |]
- ==> F \<squnion> G \<in> A co A'"
+ "\<lbrakk>F \<in> stable(A); G \<in> A co A'\<rbrakk>
+ \<Longrightarrow> F \<squnion> G \<in> A co A'"
apply (unfold stable_def constrains_def Join_def st_set_def)
apply (cut_tac F = F in Acts_type)
apply (cut_tac F = G in Acts_type, force)
@@ -372,7 +372,7 @@
(*Premise for G cannot use Always because F \<in> Stable A is
weaker than G \<in> stable A *)
lemma stable_Join_Always1:
- "[| F \<in> stable(A); G \<in> invariant(A) |] ==> F \<squnion> G \<in> Always(A)"
+ "\<lbrakk>F \<in> stable(A); G \<in> invariant(A)\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> Always(A)"
apply (subgoal_tac "F \<in> program & G \<in> program & st_set (A) ")
prefer 2 apply (blast dest: invariantD2 stableD2)
apply (simp add: Always_def invariant_def initially_def Stable_eq_stable)
@@ -381,7 +381,7 @@
(*As above, but exchanging the roles of F and G*)
lemma stable_Join_Always2:
- "[| F \<in> invariant(A); G \<in> stable(A) |] ==> F \<squnion> G \<in> Always(A)"
+ "\<lbrakk>F \<in> invariant(A); G \<in> stable(A)\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> Always(A)"
apply (subst Join_commute)
apply (blast intro: stable_Join_Always1)
done
@@ -389,7 +389,7 @@
lemma stable_Join_ensures1:
- "[| F \<in> stable(A); G \<in> A ensures B |] ==> F \<squnion> G \<in> A ensures B"
+ "\<lbrakk>F \<in> stable(A); G \<in> A ensures B\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> A ensures B"
apply (subgoal_tac "F \<in> program & G \<in> program & st_set (A) ")
prefer 2 apply (blast dest: stableD2 ensures_type [THEN subsetD])
apply (simp (no_asm_simp) add: Join_ensures)
@@ -400,7 +400,7 @@
(*As above, but exchanging the roles of F and G*)
lemma stable_Join_ensures2:
- "[| F \<in> A ensures B; G \<in> stable(A) |] ==> F \<squnion> G \<in> A ensures B"
+ "\<lbrakk>F \<in> A ensures B; G \<in> stable(A)\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> A ensures B"
apply (subst Join_commute)
apply (blast intro: stable_Join_ensures1)
done
@@ -433,7 +433,7 @@
by (auto simp add: ok_def)
(*useful? Not with the previous two around*)
-lemma ok_Join_commute_I: "[| F ok G; (F \<squnion> G) ok H |] ==> F ok (G \<squnion> H)"
+lemma ok_Join_commute_I: "\<lbrakk>F ok G; (F \<squnion> G) ok H\<rbrakk> \<Longrightarrow> F ok (G \<squnion> H)"
by (auto simp add: ok_def)
lemma ok_JOIN_iff1 [iff]: "F ok JOIN(I,G) \<longleftrightarrow> (\<forall>i \<in> I. F ok G(i))"
@@ -447,7 +447,7 @@
lemma OK_iff_ok: "OK(I,F) \<longleftrightarrow> (\<forall>i \<in> I. \<forall>j \<in> I-{i}. F(i) ok (F(j)))"
by (auto simp add: ok_def OK_def)
-lemma OK_imp_ok: "[| OK(I,F); i \<in> I; j \<in> I; i\<noteq>j|] ==> F(i) ok F(j)"
+lemma OK_imp_ok: "\<lbrakk>OK(I,F); i \<in> I; j \<in> I; i\<noteq>j\<rbrakk> \<Longrightarrow> F(i) ok F(j)"
by (auto simp add: OK_iff_ok)
@@ -474,7 +474,7 @@
done
lemma Allowed_JOIN [simp]:
- "i \<in> I ==>
+ "i \<in> I \<Longrightarrow>
Allowed(JOIN(I,F)) = (\<Inter>i \<in> I. Allowed(programify(F(i))))"
apply (auto simp add: Allowed_def, blast)
done
@@ -494,7 +494,7 @@
subsection\<open>safety_prop, for reasoning about given instances of "ok"\<close>
lemma safety_prop_Acts_iff:
- "safety_prop(X) ==> (Acts(G) \<subseteq> cons(id(state), (\<Union>F \<in> X. Acts(F)))) \<longleftrightarrow> (programify(G) \<in> X)"
+ "safety_prop(X) \<Longrightarrow> (Acts(G) \<subseteq> cons(id(state), (\<Union>F \<in> X. Acts(F)))) \<longleftrightarrow> (programify(G) \<in> X)"
apply (simp (no_asm_use) add: safety_prop_def)
apply clarify
apply (case_tac "G \<in> program", simp_all, blast, safe)
@@ -503,7 +503,7 @@
done
lemma safety_prop_AllowedActs_iff_Allowed:
- "safety_prop(X) ==>
+ "safety_prop(X) \<Longrightarrow>
(\<Union>G \<in> X. Acts(G)) \<subseteq> AllowedActs(F) \<longleftrightarrow> (X \<subseteq> Allowed(programify(F)))"
apply (simp add: Allowed_def safety_prop_Acts_iff [THEN iff_sym]
safety_prop_def, blast)
@@ -511,7 +511,7 @@
lemma Allowed_eq:
- "safety_prop(X) ==> Allowed(mk_program(init, acts, \<Union>F \<in> X. Acts(F))) = X"
+ "safety_prop(X) \<Longrightarrow> Allowed(mk_program(init, acts, \<Union>F \<in> X. Acts(F))) = X"
apply (subgoal_tac "cons (id (state), \<Union>(RepFun (X, Acts)) \<inter> Pow (state * state)) = \<Union>(RepFun (X, Acts))")
apply (rule_tac [2] equalityI)
apply (simp del: UN_simps add: Allowed_def safety_prop_Acts_iff safety_prop_def, auto)
@@ -519,8 +519,8 @@
done
lemma def_prg_Allowed:
- "[| F == mk_program (init, acts, \<Union>F \<in> X. Acts(F)); safety_prop(X) |]
- ==> Allowed(F) = X"
+ "\<lbrakk>F \<equiv> mk_program (init, acts, \<Union>F \<in> X. Acts(F)); safety_prop(X)\<rbrakk>
+ \<Longrightarrow> Allowed(F) = X"
by (simp add: Allowed_eq)
(*For safety_prop to hold, the property must be satisfiable!*)
@@ -530,17 +530,17 @@
(* To be used with resolution *)
lemma safety_prop_constrainsI [iff]:
- "[| A\<subseteq>B; st_set(A) |] ==>safety_prop(A co B)"
+ "\<lbrakk>A\<subseteq>B; st_set(A)\<rbrakk> \<Longrightarrow>safety_prop(A co B)"
by auto
lemma safety_prop_stable [iff]: "safety_prop(stable(A)) \<longleftrightarrow> st_set(A)"
by (simp add: stable_def)
-lemma safety_prop_stableI: "st_set(A) ==> safety_prop(stable(A))"
+lemma safety_prop_stableI: "st_set(A) \<Longrightarrow> safety_prop(stable(A))"
by auto
lemma safety_prop_Int [simp]:
- "[| safety_prop(X) ; safety_prop(Y) |] ==> safety_prop(X \<inter> Y)"
+ "\<lbrakk>safety_prop(X) ; safety_prop(Y)\<rbrakk> \<Longrightarrow> safety_prop(X \<inter> Y)"
apply (simp add: safety_prop_def, safe, blast)
apply (drule_tac [2] B = "\<Union>(RepFun (X \<inter> Y, Acts))" and C = "\<Union>(RepFun (Y, Acts))" in subset_trans)
apply (drule_tac B = "\<Union>(RepFun (X \<inter> Y, Acts))" and C = "\<Union>(RepFun (X, Acts))" in subset_trans)
@@ -549,7 +549,7 @@
(* If I=0 the conclusion becomes safety_prop(0) which is false *)
lemma safety_prop_Inter:
- assumes major: "(!!i. i \<in> I ==>safety_prop(X(i)))"
+ assumes major: "(\<And>i. i \<in> I \<Longrightarrow>safety_prop(X(i)))"
and minor: "i \<in> I"
shows "safety_prop(\<Inter>i \<in> I. X(i))"
apply (simp add: safety_prop_def)
@@ -565,8 +565,8 @@
done
lemma def_UNION_ok_iff:
-"[| F == mk_program(init,acts, \<Union>G \<in> X. Acts(G)); safety_prop(X) |]
- ==> F ok G \<longleftrightarrow> (programify(G) \<in> X & acts \<inter> Pow(state*state) \<subseteq> AllowedActs(G))"
+"\<lbrakk>F \<equiv> mk_program(init,acts, \<Union>G \<in> X. Acts(G)); safety_prop(X)\<rbrakk>
+ \<Longrightarrow> F ok G \<longleftrightarrow> (programify(G) \<in> X & acts \<inter> Pow(state*state) \<subseteq> AllowedActs(G))"
apply (unfold ok_def)
apply (drule_tac G = G in safety_prop_Acts_iff)
apply (cut_tac F = G in AllowedActs_type)
--- a/src/ZF/UNITY/WFair.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/UNITY/WFair.thy Tue Sep 27 16:51:35 2022 +0100
@@ -17,12 +17,12 @@
(* This definition specifies weak fairness. The rest of the theory
is generic to all forms of fairness.*)
transient :: "i=>i" where
- "transient(A) =={F \<in> program. (\<exists>act\<in>Acts(F). A<=domain(act) &
+ "transient(A) \<equiv>{F \<in> program. (\<exists>act\<in>Acts(F). A<=domain(act) &
act``A \<subseteq> state-A) & st_set(A)}"
definition
ensures :: "[i,i] => i" (infixl \<open>ensures\<close> 60) where
- "A ensures B == ((A-B) co (A \<union> B)) \<inter> transient(A-B)"
+ "A ensures B \<equiv> ((A-B) co (A \<union> B)) \<inter> transient(A-B)"
consts
@@ -33,9 +33,9 @@
domains
"leads(D, F)" \<subseteq> "Pow(D)*Pow(D)"
intros
- Basis: "[| F \<in> A ensures B; A \<in> Pow(D); B \<in> Pow(D) |] ==> <A,B>:leads(D, F)"
- Trans: "[| <A,B> \<in> leads(D, F); <B,C> \<in> leads(D, F) |] ==> <A,C>:leads(D, F)"
- Union: "[| S \<in> Pow({A \<in> S. <A, B>:leads(D, F)}); B \<in> Pow(D); S \<in> Pow(Pow(D)) |] ==>
+ Basis: "\<lbrakk>F \<in> A ensures B; A \<in> Pow(D); B \<in> Pow(D)\<rbrakk> \<Longrightarrow> <A,B>:leads(D, F)"
+ Trans: "\<lbrakk><A,B> \<in> leads(D, F); <B,C> \<in> leads(D, F)\<rbrakk> \<Longrightarrow> <A,C>:leads(D, F)"
+ Union: "\<lbrakk>S \<in> Pow({A \<in> S. <A, B>:leads(D, F)}); B \<in> Pow(D); S \<in> Pow(Pow(D))\<rbrakk> \<Longrightarrow>
<\<Union>(S),B>:leads(D, F)"
monos Pow_mono
@@ -44,12 +44,12 @@
definition
(* The Visible version of the LEADS-TO relation*)
leadsTo :: "[i, i] => i" (infixl \<open>\<longmapsto>\<close> 60) where
- "A \<longmapsto> B == {F \<in> program. <A,B>:leads(state, F)}"
+ "A \<longmapsto> B \<equiv> {F \<in> program. <A,B>:leads(state, F)}"
definition
(* wlt(F, B) is the largest set that leads to B*)
wlt :: "[i, i] => i" where
- "wlt(F, B) == \<Union>({A \<in> Pow(state). F \<in> A \<longmapsto> B})"
+ "wlt(F, B) \<equiv> \<Union>({A \<in> Pow(state). F \<in> A \<longmapsto> B})"
(** Ad-hoc set-theory rules **)
@@ -65,27 +65,27 @@
by (unfold transient_def, auto)
lemma transientD2:
-"F \<in> transient(A) ==> F \<in> program & st_set(A)"
+"F \<in> transient(A) \<Longrightarrow> F \<in> program & st_set(A)"
apply (unfold transient_def, auto)
done
-lemma stable_transient_empty: "[| F \<in> stable(A); F \<in> transient(A) |] ==> A = 0"
+lemma stable_transient_empty: "\<lbrakk>F \<in> stable(A); F \<in> transient(A)\<rbrakk> \<Longrightarrow> A = 0"
by (simp add: stable_def constrains_def transient_def, fast)
-lemma transient_strengthen: "[|F \<in> transient(A); B<=A|] ==> F \<in> transient(B)"
+lemma transient_strengthen: "\<lbrakk>F \<in> transient(A); B<=A\<rbrakk> \<Longrightarrow> F \<in> transient(B)"
apply (simp add: transient_def st_set_def, clarify)
apply (blast intro!: rev_bexI)
done
lemma transientI:
-"[|act \<in> Acts(F); A \<subseteq> domain(act); act``A \<subseteq> state-A;
- F \<in> program; st_set(A)|] ==> F \<in> transient(A)"
+"\<lbrakk>act \<in> Acts(F); A \<subseteq> domain(act); act``A \<subseteq> state-A;
+ F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> transient(A)"
by (simp add: transient_def, blast)
lemma transientE:
- "[| F \<in> transient(A);
- !!act. [| act \<in> Acts(F); A \<subseteq> domain(act); act``A \<subseteq> state-A|]==>P|]
- ==>P"
+ "\<lbrakk>F \<in> transient(A);
+ \<And>act. \<lbrakk>act \<in> Acts(F); A \<subseteq> domain(act); act``A \<subseteq> state-A\<rbrakk>\<Longrightarrow>P\<rbrakk>
+ \<Longrightarrow>P"
by (simp add: transient_def, blast)
lemma transient_state: "transient(state) = 0"
@@ -96,7 +96,7 @@
apply (auto intro: st0_in_state)
done
-lemma transient_state2: "state<=B ==> transient(B) = 0"
+lemma transient_state2: "state<=B \<Longrightarrow> transient(B) = 0"
apply (simp add: transient_def st_set_def)
apply (rule equalityI, auto)
apply (cut_tac F = x in Acts_type)
@@ -115,36 +115,36 @@
by (simp add: ensures_def constrains_def, auto)
lemma ensuresI:
-"[|F:(A-B) co (A \<union> B); F \<in> transient(A-B)|]==>F \<in> A ensures B"
+"\<lbrakk>F:(A-B) co (A \<union> B); F \<in> transient(A-B)\<rbrakk>\<Longrightarrow>F \<in> A ensures B"
apply (unfold ensures_def)
apply (auto simp add: transient_type [THEN subsetD])
done
(* Added by Sidi, from Misra's notes, Progress chapter, exercise 4 *)
-lemma ensuresI2: "[| F \<in> A co A \<union> B; F \<in> transient(A) |] ==> F \<in> A ensures B"
+lemma ensuresI2: "\<lbrakk>F \<in> A co A \<union> B; F \<in> transient(A)\<rbrakk> \<Longrightarrow> F \<in> A ensures B"
apply (drule_tac B = "A-B" in constrains_weaken_L)
apply (drule_tac [2] B = "A-B" in transient_strengthen)
apply (auto simp add: ensures_def transient_type [THEN subsetD])
done
-lemma ensuresD: "F \<in> A ensures B ==> F:(A-B) co (A \<union> B) & F \<in> transient (A-B)"
+lemma ensuresD: "F \<in> A ensures B \<Longrightarrow> F:(A-B) co (A \<union> B) & F \<in> transient (A-B)"
by (unfold ensures_def, auto)
-lemma ensures_weaken_R: "[|F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'"
+lemma ensures_weaken_R: "\<lbrakk>F \<in> A ensures A'; A'<=B'\<rbrakk> \<Longrightarrow> F \<in> A ensures B'"
apply (unfold ensures_def)
apply (blast intro: transient_strengthen constrains_weaken)
done
(*The L-version (precondition strengthening) fails, but we have this*)
lemma stable_ensures_Int:
- "[| F \<in> stable(C); F \<in> A ensures B |] ==> F:(C \<inter> A) ensures (C \<inter> B)"
+ "\<lbrakk>F \<in> stable(C); F \<in> A ensures B\<rbrakk> \<Longrightarrow> F:(C \<inter> A) ensures (C \<inter> B)"
apply (unfold ensures_def)
apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric])
apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken)
done
-lemma stable_transient_ensures: "[|F \<in> stable(A); F \<in> transient(C); A<=B \<union> C|] ==> F \<in> A ensures B"
+lemma stable_transient_ensures: "\<lbrakk>F \<in> stable(A); F \<in> transient(C); A<=B \<union> C\<rbrakk> \<Longrightarrow> F \<in> A ensures B"
apply (frule stable_type [THEN subsetD])
apply (simp add: ensures_def stable_def)
apply (blast intro: transient_strengthen constrains_weaken)
@@ -153,7 +153,7 @@
lemma ensures_eq: "(A ensures B) = (A unless B) \<inter> transient (A-B)"
by (auto simp add: ensures_def unless_def)
-lemma subset_imp_ensures: "[| F \<in> program; A<=B |] ==> F \<in> A ensures B"
+lemma subset_imp_ensures: "\<lbrakk>F \<in> program; A<=B\<rbrakk> \<Longrightarrow> F \<in> A ensures B"
by (auto simp add: ensures_def constrains_def transient_def st_set_def)
(*** leadsTo ***)
@@ -164,13 +164,13 @@
by (unfold leadsTo_def, auto)
lemma leadsToD2:
-"F \<in> A \<longmapsto> B ==> F \<in> program & st_set(A) & st_set(B)"
+"F \<in> A \<longmapsto> B \<Longrightarrow> F \<in> program & st_set(A) & st_set(B)"
apply (unfold leadsTo_def st_set_def)
apply (blast dest: leads_left leads_right)
done
lemma leadsTo_Basis:
- "[|F \<in> A ensures B; st_set(A); st_set(B)|] ==> F \<in> A \<longmapsto> B"
+ "\<lbrakk>F \<in> A ensures B; st_set(A); st_set(B)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (unfold leadsTo_def st_set_def)
apply (cut_tac ensures_type)
apply (auto intro: leads.Basis)
@@ -178,66 +178,66 @@
declare leadsTo_Basis [intro]
(* Added by Sidi, from Misra's notes, Progress chapter, exercise number 4 *)
-(* [| F \<in> program; A<=B; st_set(A); st_set(B) |] ==> A \<longmapsto> B *)
+(* \<lbrakk>F \<in> program; A<=B; st_set(A); st_set(B)\<rbrakk> \<Longrightarrow> A \<longmapsto> B *)
lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis]
-lemma leadsTo_Trans: "[|F \<in> A \<longmapsto> B; F \<in> B \<longmapsto> C |]==>F \<in> A \<longmapsto> C"
+lemma leadsTo_Trans: "\<lbrakk>F \<in> A \<longmapsto> B; F \<in> B \<longmapsto> C\<rbrakk>\<Longrightarrow>F \<in> A \<longmapsto> C"
apply (unfold leadsTo_def)
apply (auto intro: leads.Trans)
done
(* Better when used in association with leadsTo_weaken_R *)
-lemma transient_imp_leadsTo: "F \<in> transient(A) ==> F \<in> A \<longmapsto> (state-A)"
+lemma transient_imp_leadsTo: "F \<in> transient(A) \<Longrightarrow> F \<in> A \<longmapsto> (state-A)"
apply (unfold transient_def)
apply (blast intro: ensuresI [THEN leadsTo_Basis] constrains_weaken transientI)
done
(*Useful with cancellation, disjunction*)
-lemma leadsTo_Un_duplicate: "F \<in> A \<longmapsto> (A' \<union> A') ==> F \<in> A \<longmapsto> A'"
+lemma leadsTo_Un_duplicate: "F \<in> A \<longmapsto> (A' \<union> A') \<Longrightarrow> F \<in> A \<longmapsto> A'"
by simp
lemma leadsTo_Un_duplicate2:
- "F \<in> A \<longmapsto> (A' \<union> C \<union> C) ==> F \<in> A \<longmapsto> (A' \<union> C)"
+ "F \<in> A \<longmapsto> (A' \<union> C \<union> C) \<Longrightarrow> F \<in> A \<longmapsto> (A' \<union> C)"
by (simp add: Un_ac)
(*The Union introduction rule as we should have liked to state it*)
lemma leadsTo_Union:
- "[|!!A. A \<in> S ==> F \<in> A \<longmapsto> B; F \<in> program; st_set(B)|]
- ==> F \<in> \<Union>(S) \<longmapsto> B"
+ "\<lbrakk>\<And>A. A \<in> S \<Longrightarrow> F \<in> A \<longmapsto> B; F \<in> program; st_set(B)\<rbrakk>
+ \<Longrightarrow> F \<in> \<Union>(S) \<longmapsto> B"
apply (unfold leadsTo_def st_set_def)
apply (blast intro: leads.Union dest: leads_left)
done
lemma leadsTo_Union_Int:
- "[|!!A. A \<in> S ==>F \<in> (A \<inter> C) \<longmapsto> B; F \<in> program; st_set(B)|]
- ==> F \<in> (\<Union>(S)Int C)\<longmapsto> B"
+ "\<lbrakk>\<And>A. A \<in> S \<Longrightarrow>F \<in> (A \<inter> C) \<longmapsto> B; F \<in> program; st_set(B)\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Union>(S)Int C)\<longmapsto> B"
apply (unfold leadsTo_def st_set_def)
apply (simp only: Int_Union_Union)
apply (blast dest: leads_left intro: leads.Union)
done
lemma leadsTo_UN:
- "[| !!i. i \<in> I ==> F \<in> A(i) \<longmapsto> B; F \<in> program; st_set(B)|]
- ==> F:(\<Union>i \<in> I. A(i)) \<longmapsto> B"
+ "\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto> B; F \<in> program; st_set(B)\<rbrakk>
+ \<Longrightarrow> F:(\<Union>i \<in> I. A(i)) \<longmapsto> B"
apply (simp add: Int_Union_Union leadsTo_def st_set_def)
apply (blast dest: leads_left intro: leads.Union)
done
(* Binary union introduction rule *)
lemma leadsTo_Un:
- "[| F \<in> A \<longmapsto> C; F \<in> B \<longmapsto> C |] ==> F \<in> (A \<union> B) \<longmapsto> C"
+ "\<lbrakk>F \<in> A \<longmapsto> C; F \<in> B \<longmapsto> C\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) \<longmapsto> C"
apply (subst Un_eq_Union)
apply (blast intro: leadsTo_Union dest: leadsToD2)
done
lemma single_leadsTo_I:
- "[|!!x. x \<in> A==> F:{x} \<longmapsto> B; F \<in> program; st_set(B) |]
- ==> F \<in> A \<longmapsto> B"
+ "\<lbrakk>\<And>x. x \<in> A\<Longrightarrow> F:{x} \<longmapsto> B; F \<in> program; st_set(B)\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (rule_tac b = A in UN_singleton [THEN subst])
apply (rule leadsTo_UN, auto)
done
-lemma leadsTo_refl: "[| F \<in> program; st_set(A) |] ==> F \<in> A \<longmapsto> A"
+lemma leadsTo_refl: "\<lbrakk>F \<in> program; st_set(A)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> A"
by (blast intro: subset_imp_leadsTo)
lemma leadsTo_refl_iff: "F \<in> A \<longmapsto> A \<longleftrightarrow> F \<in> program & st_set(A)"
@@ -251,21 +251,21 @@
by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD)
declare leadsTo_state [iff]
-lemma leadsTo_weaken_R: "[| F \<in> A \<longmapsto> A'; A'<=B'; st_set(B') |] ==> F \<in> A \<longmapsto> B'"
+lemma leadsTo_weaken_R: "\<lbrakk>F \<in> A \<longmapsto> A'; A'<=B'; st_set(B')\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> B'"
by (blast dest: leadsToD2 intro: subset_imp_leadsTo leadsTo_Trans)
-lemma leadsTo_weaken_L: "[| F \<in> A \<longmapsto> A'; B<=A |] ==> F \<in> B \<longmapsto> A'"
+lemma leadsTo_weaken_L: "\<lbrakk>F \<in> A \<longmapsto> A'; B<=A\<rbrakk> \<Longrightarrow> F \<in> B \<longmapsto> A'"
apply (frule leadsToD2)
apply (blast intro: leadsTo_Trans subset_imp_leadsTo st_set_subset)
done
-lemma leadsTo_weaken: "[| F \<in> A \<longmapsto> A'; B<=A; A'<=B'; st_set(B')|]==> F \<in> B \<longmapsto> B'"
+lemma leadsTo_weaken: "\<lbrakk>F \<in> A \<longmapsto> A'; B<=A; A'<=B'; st_set(B')\<rbrakk>\<Longrightarrow> F \<in> B \<longmapsto> B'"
apply (frule leadsToD2)
apply (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans leadsTo_refl)
done
(* This rule has a nicer conclusion *)
-lemma transient_imp_leadsTo2: "[| F \<in> transient(A); state-A<=B; st_set(B)|] ==> F \<in> A \<longmapsto> B"
+lemma transient_imp_leadsTo2: "\<lbrakk>F \<in> transient(A); state-A<=B; st_set(B)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (frule transientD2)
apply (rule leadsTo_weaken_R)
apply (auto simp add: transient_imp_leadsTo)
@@ -285,45 +285,45 @@
text\<open>Set difference: maybe combine with \<open>leadsTo_weaken_L\<close>??\<close>
lemma leadsTo_Diff:
- "[| F: (A-B) \<longmapsto> C; F \<in> B \<longmapsto> C; st_set(C) |]
- ==> F \<in> A \<longmapsto> C"
+ "\<lbrakk>F: (A-B) \<longmapsto> C; F \<in> B \<longmapsto> C; st_set(C)\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto> C"
by (blast intro: leadsTo_Un leadsTo_weaken dest: leadsToD2)
lemma leadsTo_UN_UN:
- "[|!!i. i \<in> I ==> F \<in> A(i) \<longmapsto> A'(i); F \<in> program |]
- ==> F: (\<Union>i \<in> I. A(i)) \<longmapsto> (\<Union>i \<in> I. A'(i))"
+ "\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto> A'(i); F \<in> program\<rbrakk>
+ \<Longrightarrow> F: (\<Union>i \<in> I. A(i)) \<longmapsto> (\<Union>i \<in> I. A'(i))"
apply (rule leadsTo_Union)
apply (auto intro: leadsTo_weaken_R dest: leadsToD2)
done
(*Binary union version*)
-lemma leadsTo_Un_Un: "[| F \<in> A \<longmapsto> A'; F \<in> B \<longmapsto> B' |] ==> F \<in> (A \<union> B) \<longmapsto> (A' \<union> B')"
+lemma leadsTo_Un_Un: "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B \<longmapsto> B'\<rbrakk> \<Longrightarrow> F \<in> (A \<union> B) \<longmapsto> (A' \<union> B')"
apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ")
prefer 2 apply (blast dest: leadsToD2)
apply (blast intro: leadsTo_Un leadsTo_weaken_R)
done
(** The cancellation law **)
-lemma leadsTo_cancel2: "[|F \<in> A \<longmapsto> (A' \<union> B); F \<in> B \<longmapsto> B'|] ==> F \<in> A \<longmapsto> (A' \<union> B')"
+lemma leadsTo_cancel2: "\<lbrakk>F \<in> A \<longmapsto> (A' \<union> B); F \<in> B \<longmapsto> B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> (A' \<union> B')"
apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \<in> program")
prefer 2 apply (blast dest: leadsToD2)
apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl)
done
-lemma leadsTo_cancel_Diff2: "[|F \<in> A \<longmapsto> (A' \<union> B); F \<in> (B-A') \<longmapsto> B'|]==> F \<in> A \<longmapsto> (A' \<union> B')"
+lemma leadsTo_cancel_Diff2: "\<lbrakk>F \<in> A \<longmapsto> (A' \<union> B); F \<in> (B-A') \<longmapsto> B'\<rbrakk>\<Longrightarrow> F \<in> A \<longmapsto> (A' \<union> B')"
apply (rule leadsTo_cancel2)
prefer 2 apply assumption
apply (blast dest: leadsToD2 intro: leadsTo_weaken_R)
done
-lemma leadsTo_cancel1: "[| F \<in> A \<longmapsto> (B \<union> A'); F \<in> B \<longmapsto> B' |] ==> F \<in> A \<longmapsto> (B' \<union> A')"
+lemma leadsTo_cancel1: "\<lbrakk>F \<in> A \<longmapsto> (B \<union> A'); F \<in> B \<longmapsto> B'\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> (B' \<union> A')"
apply (simp add: Un_commute)
apply (blast intro!: leadsTo_cancel2)
done
lemma leadsTo_cancel_Diff1:
- "[|F \<in> A \<longmapsto> (B \<union> A'); F: (B-A') \<longmapsto> B'|]==> F \<in> A \<longmapsto> (B' \<union> A')"
+ "\<lbrakk>F \<in> A \<longmapsto> (B \<union> A'); F: (B-A') \<longmapsto> B'\<rbrakk>\<Longrightarrow> F \<in> A \<longmapsto> (B' \<union> A')"
apply (rule leadsTo_cancel1)
prefer 2 apply assumption
apply (blast intro: leadsTo_weaken_R dest: leadsToD2)
@@ -332,11 +332,11 @@
(*The INDUCTION rule as we should have liked to state it*)
lemma leadsTo_induct:
assumes major: "F \<in> za \<longmapsto> zb"
- and basis: "!!A B. [|F \<in> A ensures B; st_set(A); st_set(B)|] ==> P(A,B)"
- and trans: "!!A B C. [| F \<in> A \<longmapsto> B; P(A, B);
- F \<in> B \<longmapsto> C; P(B, C) |] ==> P(A,C)"
- and union: "!!B S. [| \<forall>A \<in> S. F \<in> A \<longmapsto> B; \<forall>A \<in> S. P(A,B);
- st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(\<Union>(S), B)"
+ and basis: "\<And>A B. \<lbrakk>F \<in> A ensures B; st_set(A); st_set(B)\<rbrakk> \<Longrightarrow> P(A,B)"
+ and trans: "\<And>A B C. \<lbrakk>F \<in> A \<longmapsto> B; P(A, B);
+ F \<in> B \<longmapsto> C; P(B, C)\<rbrakk> \<Longrightarrow> P(A,C)"
+ and union: "\<And>B S. \<lbrakk>\<forall>A \<in> S. F \<in> A \<longmapsto> B; \<forall>A \<in> S. P(A,B);
+ st_set(B); \<forall>A \<in> S. st_set(A)\<rbrakk> \<Longrightarrow> P(\<Union>(S), B)"
shows "P(za, zb)"
apply (cut_tac major)
apply (unfold leadsTo_def, clarify)
@@ -350,13 +350,13 @@
(* Added by Sidi, an induction rule without ensures *)
lemma leadsTo_induct2:
assumes major: "F \<in> za \<longmapsto> zb"
- and basis1: "!!A B. [| A<=B; st_set(B) |] ==> P(A, B)"
- and basis2: "!!A B. [| F \<in> A co A \<union> B; F \<in> transient(A); st_set(B) |]
- ==> P(A, B)"
- and trans: "!!A B C. [| F \<in> A \<longmapsto> B; P(A, B);
- F \<in> B \<longmapsto> C; P(B, C) |] ==> P(A,C)"
- and union: "!!B S. [| \<forall>A \<in> S. F \<in> A \<longmapsto> B; \<forall>A \<in> S. P(A,B);
- st_set(B); \<forall>A \<in> S. st_set(A)|] ==> P(\<Union>(S), B)"
+ and basis1: "\<And>A B. \<lbrakk>A<=B; st_set(B)\<rbrakk> \<Longrightarrow> P(A, B)"
+ and basis2: "\<And>A B. \<lbrakk>F \<in> A co A \<union> B; F \<in> transient(A); st_set(B)\<rbrakk>
+ \<Longrightarrow> P(A, B)"
+ and trans: "\<And>A B C. \<lbrakk>F \<in> A \<longmapsto> B; P(A, B);
+ F \<in> B \<longmapsto> C; P(B, C)\<rbrakk> \<Longrightarrow> P(A,C)"
+ and union: "\<And>B S. \<lbrakk>\<forall>A \<in> S. F \<in> A \<longmapsto> B; \<forall>A \<in> S. P(A,B);
+ st_set(B); \<forall>A \<in> S. st_set(A)\<rbrakk> \<Longrightarrow> P(\<Union>(S), B)"
shows "P(za, zb)"
apply (cut_tac major)
apply (erule leadsTo_induct)
@@ -379,11 +379,11 @@
(** Variant induction rule: on the preconditions for B **)
(*Lemma is the weak version: can't see how to do it in one step*)
lemma leadsTo_induct_pre_aux:
- "[| F \<in> za \<longmapsto> zb;
+ "\<lbrakk>F \<in> za \<longmapsto> zb;
P(zb);
- !!A B. [| F \<in> A ensures B; P(B); st_set(A); st_set(B) |] ==> P(A);
- !!S. [| \<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A) |] ==> P(\<Union>(S))
- |] ==> P(za)"
+ \<And>A B. \<lbrakk>F \<in> A ensures B; P(B); st_set(A); st_set(B)\<rbrakk> \<Longrightarrow> P(A);
+ \<And>S. \<lbrakk>\<forall>A \<in> S. P(A); \<forall>A \<in> S. st_set(A)\<rbrakk> \<Longrightarrow> P(\<Union>(S))
+\<rbrakk> \<Longrightarrow> P(za)"
txt\<open>by induction on this formula\<close>
apply (subgoal_tac "P (zb) \<longrightarrow> P (za) ")
txt\<open>now solve first subgoal: this formula is sufficient\<close>
@@ -394,11 +394,11 @@
lemma leadsTo_induct_pre:
- "[| F \<in> za \<longmapsto> zb;
+ "\<lbrakk>F \<in> za \<longmapsto> zb;
P(zb);
- !!A B. [| F \<in> A ensures B; F \<in> B \<longmapsto> zb; P(B); st_set(A) |] ==> P(A);
- !!S. \<forall>A \<in> S. F \<in> A \<longmapsto> zb & P(A) & st_set(A) ==> P(\<Union>(S))
- |] ==> P(za)"
+ \<And>A B. \<lbrakk>F \<in> A ensures B; F \<in> B \<longmapsto> zb; P(B); st_set(A)\<rbrakk> \<Longrightarrow> P(A);
+ \<And>S. \<forall>A \<in> S. F \<in> A \<longmapsto> zb & P(A) & st_set(A) \<Longrightarrow> P(\<Union>(S))
+\<rbrakk> \<Longrightarrow> P(za)"
apply (subgoal_tac " (F \<in> za \<longmapsto> zb) & P (za) ")
apply (erule conjunct2)
apply (frule leadsToD2)
@@ -410,7 +410,7 @@
(** The impossibility law **)
lemma leadsTo_empty:
- "F \<in> A \<longmapsto> 0 ==> A=0"
+ "F \<in> A \<longmapsto> 0 \<Longrightarrow> A=0"
apply (erule leadsTo_induct_pre)
apply (auto simp add: ensures_def constrains_def transient_def st_set_def)
apply (drule bspec, assumption)+
@@ -423,7 +423,7 @@
text\<open>Special case of PSP: Misra's "stable conjunction"\<close>
lemma psp_stable:
- "[| F \<in> A \<longmapsto> A'; F \<in> stable(B) |] ==> F:(A \<inter> B) \<longmapsto> (A' \<inter> B)"
+ "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> stable(B)\<rbrakk> \<Longrightarrow> F:(A \<inter> B) \<longmapsto> (A' \<inter> B)"
apply (unfold stable_def)
apply (frule leadsToD2)
apply (erule leadsTo_induct)
@@ -435,12 +435,12 @@
done
-lemma psp_stable2: "[|F \<in> A \<longmapsto> A'; F \<in> stable(B) |]==>F: (B \<inter> A) \<longmapsto> (B \<inter> A')"
+lemma psp_stable2: "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> stable(B)\<rbrakk>\<Longrightarrow>F: (B \<inter> A) \<longmapsto> (B \<inter> A')"
apply (simp (no_asm_simp) add: psp_stable Int_ac)
done
lemma psp_ensures:
-"[| F \<in> A ensures A'; F \<in> B co B' |]==> F: (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))"
+"\<lbrakk>F \<in> A ensures A'; F \<in> B co B'\<rbrakk>\<Longrightarrow> F: (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))"
apply (unfold ensures_def constrains_def st_set_def)
(*speeds up the proof*)
apply clarify
@@ -448,7 +448,7 @@
done
lemma psp:
-"[|F \<in> A \<longmapsto> A'; F \<in> B co B'; st_set(B')|]==> F:(A \<inter> B') \<longmapsto> ((A' \<inter> B) \<union> (B' - B))"
+"\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B co B'; st_set(B')\<rbrakk>\<Longrightarrow> F:(A \<inter> B') \<longmapsto> ((A' \<inter> B) \<union> (B' - B))"
apply (subgoal_tac "F \<in> program & st_set (A) & st_set (A') & st_set (B) ")
prefer 2 apply (blast dest!: constrainsD2 leadsToD2)
apply (erule leadsTo_induct)
@@ -463,13 +463,13 @@
done
-lemma psp2: "[| F \<in> A \<longmapsto> A'; F \<in> B co B'; st_set(B') |]
- ==> F \<in> (B' \<inter> A) \<longmapsto> ((B \<inter> A') \<union> (B' - B))"
+lemma psp2: "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B co B'; st_set(B')\<rbrakk>
+ \<Longrightarrow> F \<in> (B' \<inter> A) \<longmapsto> ((B \<inter> A') \<union> (B' - B))"
by (simp (no_asm_simp) add: psp Int_ac)
lemma psp_unless:
- "[| F \<in> A \<longmapsto> A'; F \<in> B unless B'; st_set(B); st_set(B') |]
- ==> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B) \<union> B')"
+ "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> B unless B'; st_set(B); st_set(B')\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B) \<union> B')"
apply (unfold unless_def)
apply (subgoal_tac "st_set (A) &st_set (A') ")
prefer 2 apply (blast dest: leadsToD2)
@@ -481,13 +481,13 @@
subsection\<open>Proving the induction rules\<close>
(** The most general rule \<in> r is any wf relation; f is any variant function **)
-lemma leadsTo_wf_induct_aux: "[| wf(r);
+lemma leadsTo_wf_induct_aux: "\<lbrakk>wf(r);
m \<in> I;
field(r)<=I;
F \<in> program; st_set(B);
\<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) \<longmapsto>
- ((A \<inter> f-``(converse(r)``{m})) \<union> B) |]
- ==> F \<in> (A \<inter> f-``{m}) \<longmapsto> B"
+ ((A \<inter> f-``(converse(r)``{m})) \<union> B)\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> f-``{m}) \<longmapsto> B"
apply (erule_tac a = m in wf_induct2, simp_all)
apply (subgoal_tac "F \<in> (A \<inter> (f-`` (converse (r) ``{x}))) \<longmapsto> B")
apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate)
@@ -497,13 +497,13 @@
done
(** Meta or object quantifier ? **)
-lemma leadsTo_wf_induct: "[| wf(r);
+lemma leadsTo_wf_induct: "\<lbrakk>wf(r);
field(r)<=I;
A<=f-``I;
F \<in> program; st_set(A); st_set(B);
\<forall>m \<in> I. F \<in> (A \<inter> f-``{m}) \<longmapsto>
- ((A \<inter> f-``(converse(r)``{m})) \<union> B) |]
- ==> F \<in> A \<longmapsto> B"
+ ((A \<inter> f-``(converse(r)``{m})) \<union> B)\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (rule_tac b = A in subst)
defer 1
apply (rule_tac I = I in leadsTo_UN)
@@ -522,7 +522,7 @@
done
-lemma Image_inverse_lessThan: "k<A ==> measure(A, %x. x) -`` {k} = k"
+lemma Image_inverse_lessThan: "k<A \<Longrightarrow> measure(A, %x. x) -`` {k} = k"
apply (rule equalityI)
apply (auto simp add: measure_def)
apply (blast intro: ltD)
@@ -534,10 +534,10 @@
(*Alternative proof is via the lemma F \<in> (A \<inter> f-`(lessThan m)) \<longmapsto> B*)
lemma lessThan_induct:
- "[| A<=f-``nat;
+ "\<lbrakk>A<=f-``nat;
F \<in> program; st_set(A); st_set(B);
- \<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto> ((A \<inter> f -`` m) \<union> B) |]
- ==> F \<in> A \<longmapsto> B"
+ \<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto> ((A \<inter> f -`` m) \<union> B)\<rbrakk>
+ \<Longrightarrow> F \<in> A \<longmapsto> B"
apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct])
apply (simp_all add: nat_measure_field)
apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
@@ -561,10 +561,10 @@
apply (blast dest: leadsToD2 intro!: leadsTo_Union)
done
-(* [| F \<in> program; st_set(B) |] ==> F \<in> wlt(F, B) \<longmapsto> B *)
+(* \<lbrakk>F \<in> program; st_set(B)\<rbrakk> \<Longrightarrow> F \<in> wlt(F, B) \<longmapsto> B *)
lemmas wlt_leadsTo = conjI [THEN wlt_leadsTo_iff [THEN iffD2]]
-lemma leadsTo_subset: "F \<in> A \<longmapsto> B ==> A \<subseteq> wlt(F, B)"
+lemma leadsTo_subset: "F \<in> A \<longmapsto> B \<Longrightarrow> A \<subseteq> wlt(F, B)"
apply (unfold wlt_def)
apply (frule leadsToD2)
apply (auto simp add: st_set_def)
@@ -577,24 +577,24 @@
done
(*Misra's property W4*)
-lemma wlt_increasing: "[| F \<in> program; st_set(B) |] ==> B \<subseteq> wlt(F,B)"
+lemma wlt_increasing: "\<lbrakk>F \<in> program; st_set(B)\<rbrakk> \<Longrightarrow> B \<subseteq> wlt(F,B)"
apply (rule leadsTo_subset)
apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [THEN iff_sym] subset_imp_leadsTo)
done
(*Used in the Trans case below*)
lemma leadsTo_123_aux:
- "[| B \<subseteq> A2;
+ "\<lbrakk>B \<subseteq> A2;
F \<in> (A1 - B) co (A1 \<union> B);
- F \<in> (A2 - C) co (A2 \<union> C) |]
- ==> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)"
+ F \<in> (A2 - C) co (A2 \<union> C)\<rbrakk>
+ \<Longrightarrow> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)"
apply (unfold constrains_def st_set_def, blast)
done
(*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
(* slightly different from the HOL one \<in> B here is bounded *)
lemma leadsTo_123: "F \<in> A \<longmapsto> A'
- ==> \<exists>B \<in> Pow(state). A<=B & F \<in> B \<longmapsto> A' & F \<in> (B-A') co (B \<union> A')"
+ \<Longrightarrow> \<exists>B \<in> Pow(state). A<=B & F \<in> B \<longmapsto> A' & F \<in> (B-A') co (B \<union> A')"
apply (frule leadsToD2)
apply (erule leadsTo_induct)
txt\<open>Basis\<close>
@@ -619,7 +619,7 @@
(*Misra's property W5*)
-lemma wlt_constrains_wlt: "[| F \<in> program; st_set(B) |] ==>F \<in> (wlt(F, B) - B) co (wlt(F,B))"
+lemma wlt_constrains_wlt: "\<lbrakk>F \<in> program; st_set(B)\<rbrakk> \<Longrightarrow>F \<in> (wlt(F, B) - B) co (wlt(F,B))"
apply (cut_tac F = F in wlt_leadsTo [THEN leadsTo_123], assumption, blast)
apply clarify
apply (subgoal_tac "Ba = wlt (F,B) ")
@@ -630,10 +630,10 @@
subsection\<open>Completion: Binary and General Finite versions\<close>
-lemma completion_aux: "[| W = wlt(F, (B' \<union> C));
+lemma completion_aux: "\<lbrakk>W = wlt(F, (B' \<union> C));
F \<in> A \<longmapsto> (A' \<union> C); F \<in> A' co (A' \<union> C);
- F \<in> B \<longmapsto> (B' \<union> C); F \<in> B' co (B' \<union> C) |]
- ==> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B') \<union> C)"
+ F \<in> B \<longmapsto> (B' \<union> C); F \<in> B' co (B' \<union> C)\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto> ((A' \<inter> B') \<union> C)"
apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \<in> program")
prefer 2
apply simp
@@ -665,7 +665,7 @@
lemmas completion = refl [THEN completion_aux]
lemma finite_completion_aux:
- "[| I \<in> Fin(X); F \<in> program; st_set(C) |] ==>
+ "\<lbrakk>I \<in> Fin(X); F \<in> program; st_set(C)\<rbrakk> \<Longrightarrow>
(\<forall>i \<in> I. F \<in> (A(i)) \<longmapsto> (A'(i) \<union> C)) \<longrightarrow>
(\<forall>i \<in> I. F \<in> (A'(i)) co (A'(i) \<union> C)) \<longrightarrow>
F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto> ((\<Inter>i \<in> I. A'(i)) \<union> C)"
@@ -677,16 +677,16 @@
done
lemma finite_completion:
- "[| I \<in> Fin(X);
- !!i. i \<in> I ==> F \<in> A(i) \<longmapsto> (A'(i) \<union> C);
- !!i. i \<in> I ==> F \<in> A'(i) co (A'(i) \<union> C); F \<in> program; st_set(C)|]
- ==> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto> ((\<Inter>i \<in> I. A'(i)) \<union> C)"
+ "\<lbrakk>I \<in> Fin(X);
+ \<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto> (A'(i) \<union> C);
+ \<And>i. i \<in> I \<Longrightarrow> F \<in> A'(i) co (A'(i) \<union> C); F \<in> program; st_set(C)\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto> ((\<Inter>i \<in> I. A'(i)) \<union> C)"
by (blast intro: finite_completion_aux [THEN mp, THEN mp])
lemma stable_completion:
- "[| F \<in> A \<longmapsto> A'; F \<in> stable(A');
- F \<in> B \<longmapsto> B'; F \<in> stable(B') |]
- ==> F \<in> (A \<inter> B) \<longmapsto> (A' \<inter> B')"
+ "\<lbrakk>F \<in> A \<longmapsto> A'; F \<in> stable(A');
+ F \<in> B \<longmapsto> B'; F \<in> stable(B')\<rbrakk>
+ \<Longrightarrow> F \<in> (A \<inter> B) \<longmapsto> (A' \<inter> B')"
apply (unfold stable_def)
apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+)
apply (blast dest: leadsToD2)
@@ -694,10 +694,10 @@
lemma finite_stable_completion:
- "[| I \<in> Fin(X);
- (!!i. i \<in> I ==> F \<in> A(i) \<longmapsto> A'(i));
- (!!i. i \<in> I ==> F \<in> stable(A'(i))); F \<in> program |]
- ==> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto> (\<Inter>i \<in> I. A'(i))"
+ "\<lbrakk>I \<in> Fin(X);
+ (\<And>i. i \<in> I \<Longrightarrow> F \<in> A(i) \<longmapsto> A'(i));
+ (\<And>i. i \<in> I \<Longrightarrow> F \<in> stable(A'(i))); F \<in> program\<rbrakk>
+ \<Longrightarrow> F \<in> (\<Inter>i \<in> I. A(i)) \<longmapsto> (\<Inter>i \<in> I. A'(i))"
apply (unfold stable_def)
apply (subgoal_tac "st_set (\<Inter>i \<in> I. A' (i))")
prefer 2 apply (blast dest: leadsToD2)
--- a/src/ZF/Univ.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Univ.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,26 +15,26 @@
definition
Vfrom :: "[i,i]=>i" where
- "Vfrom(A,i) == transrec(i, %x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))"
+ "Vfrom(A,i) \<equiv> transrec(i, %x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))"
abbreviation
Vset :: "i=>i" where
- "Vset(x) == Vfrom(0,x)"
+ "Vset(x) \<equiv> Vfrom(0,x)"
definition
Vrec :: "[i, [i,i]=>i] =>i" where
- "Vrec(a,H) == transrec(rank(a), %x g. \<lambda>z\<in>Vset(succ(x)).
+ "Vrec(a,H) \<equiv> 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. \<lambda>z\<in>Vset(succ(x)).
+ "Vrecursor(H,a) \<equiv> 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
- "univ(A) == Vfrom(A,nat)"
+ "univ(A) \<equiv> Vfrom(A,nat)"
subsection\<open>Immediate Consequences of the Definition of \<^term>\<open>Vfrom(A,i)\<close>\<close>
@@ -46,7 +46,7 @@
subsubsection\<open>Monotonicity\<close>
lemma Vfrom_mono [rule_format]:
- "A<=B ==> \<forall>j. i<=j \<longrightarrow> Vfrom(A,i) \<subseteq> Vfrom(B,j)"
+ "A<=B \<Longrightarrow> \<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])
@@ -55,7 +55,7 @@
apply (erule UN_mono, blast)
done
-lemma VfromI: "[| a \<in> Vfrom(A,j); j<i |] ==> a \<in> Vfrom(A,i)"
+lemma VfromI: "\<lbrakk>a \<in> Vfrom(A,j); j<i\<rbrakk> \<Longrightarrow> a \<in> Vfrom(A,i)"
by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]])
@@ -96,7 +96,7 @@
subsection\<open>Basic Closure Properties\<close>
-lemma zero_in_Vfrom: "y:x ==> 0 \<in> Vfrom(A,x)"
+lemma zero_in_Vfrom: "y:x \<Longrightarrow> 0 \<in> Vfrom(A,x)"
by (subst Vfrom, blast)
lemma i_subset_Vfrom: "i \<subseteq> Vfrom(A,i)"
@@ -111,25 +111,25 @@
lemmas A_into_Vfrom = A_subset_Vfrom [THEN subsetD]
-lemma subset_mem_Vfrom: "a \<subseteq> Vfrom(A,i) ==> a \<in> Vfrom(A,succ(i))"
+lemma subset_mem_Vfrom: "a \<subseteq> Vfrom(A,i) \<Longrightarrow> a \<in> Vfrom(A,succ(i))"
by (subst Vfrom, blast)
subsubsection\<open>Finite sets and ordered pairs\<close>
-lemma singleton_in_Vfrom: "a \<in> Vfrom(A,i) ==> {a} \<in> Vfrom(A,succ(i))"
+lemma singleton_in_Vfrom: "a \<in> Vfrom(A,i) \<Longrightarrow> {a} \<in> Vfrom(A,succ(i))"
by (rule subset_mem_Vfrom, safe)
lemma doubleton_in_Vfrom:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i) |] ==> {a,b} \<in> Vfrom(A,succ(i))"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Vfrom(A,succ(i))"
by (rule subset_mem_Vfrom, safe)
lemma Pair_in_Vfrom:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i) |] ==> <a,b> \<in> Vfrom(A,succ(succ(i)))"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Vfrom(A,succ(succ(i)))"
apply (unfold Pair_def)
apply (blast intro: doubleton_in_Vfrom)
done
-lemma succ_in_Vfrom: "a \<subseteq> Vfrom(A,i) ==> succ(a) \<in> Vfrom(A,succ(succ(i)))"
+lemma succ_in_Vfrom: "a \<subseteq> Vfrom(A,i) \<Longrightarrow> 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])
@@ -140,7 +140,7 @@
lemma Vfrom_0: "Vfrom(A,0) = A"
by (subst Vfrom, blast)
-lemma Vfrom_succ_lemma: "Ord(i) ==> Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))"
+lemma Vfrom_succ_lemma: "Ord(i) \<Longrightarrow> Vfrom(A,succ(i)) = A \<union> Pow(Vfrom(A,i))"
apply (rule Vfrom [THEN trans])
apply (rule equalityI [THEN subst_context,
OF _ succI1 [THEN RepFunI, THEN Union_upper]])
@@ -159,7 +159,7 @@
(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces
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))"
+lemma Vfrom_Union: "y:X \<Longrightarrow> Vfrom(A,\<Union>(X)) = (\<Union>y\<in>X. Vfrom(A,y))"
apply (subst Vfrom)
apply (rule equalityI)
txt\<open>first inclusion\<close>
@@ -181,15 +181,15 @@
(*NB. limit ordinals are non-empty:
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))"
+ "Limit(i) \<Longrightarrow> 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)
done
lemma Limit_VfromE:
- "[| a \<in> Vfrom(A,i); ~R ==> Limit(i);
- !!x. [| x<i; a \<in> Vfrom(A,x) |] ==> R
- |] ==> R"
+ "\<lbrakk>a \<in> Vfrom(A,i); \<not>R \<Longrightarrow> Limit(i);
+ \<And>x. \<lbrakk>x<i; a \<in> Vfrom(A,x)\<rbrakk> \<Longrightarrow> R
+\<rbrakk> \<Longrightarrow> R"
apply (rule classical)
apply (rule Limit_Vfrom_eq [THEN equalityD1, THEN subsetD, THEN UN_E])
prefer 2 apply assumption
@@ -198,7 +198,7 @@
done
lemma singleton_in_VLimit:
- "[| a \<in> Vfrom(A,i); Limit(i) |] ==> {a} \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> {a} \<in> Vfrom(A,i)"
apply (erule Limit_VfromE, assumption)
apply (erule singleton_in_Vfrom [THEN VfromI])
apply (blast intro: Limit_has_succ)
@@ -211,7 +211,7 @@
text\<open>Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)\<close>
lemma doubleton_in_VLimit:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i) |] ==> {a,b} \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> {a,b} \<in> Vfrom(A,i)"
apply (erule Limit_VfromE, assumption)
apply (erule Limit_VfromE, assumption)
apply (blast intro: VfromI [OF doubleton_in_Vfrom]
@@ -219,7 +219,7 @@
done
lemma Pair_in_VLimit:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i) |] ==> <a,b> \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Vfrom(A,i)"
txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
apply (erule Limit_VfromE, assumption)
apply (erule Limit_VfromE, assumption)
@@ -228,7 +228,7 @@
Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt)
done
-lemma product_VLimit: "Limit(i) ==> Vfrom(A,i) * Vfrom(A,i) \<subseteq> Vfrom(A,i)"
+lemma product_VLimit: "Limit(i) \<Longrightarrow> Vfrom(A,i) * Vfrom(A,i) \<subseteq> Vfrom(A,i)"
by (blast intro: Pair_in_VLimit)
lemmas Sigma_subset_VLimit =
@@ -237,29 +237,29 @@
lemmas nat_subset_VLimit =
subset_trans [OF nat_le_Limit [THEN le_imp_subset] i_subset_Vfrom]
-lemma nat_into_VLimit: "[| n: nat; Limit(i) |] ==> n \<in> Vfrom(A,i)"
+lemma nat_into_VLimit: "\<lbrakk>n: nat; Limit(i)\<rbrakk> \<Longrightarrow> n \<in> Vfrom(A,i)"
by (blast intro: nat_subset_VLimit [THEN subsetD])
subsubsection\<open>Closure under Disjoint Union\<close>
lemmas zero_in_VLimit = Limit_has_0 [THEN ltD, THEN zero_in_Vfrom]
-lemma one_in_VLimit: "Limit(i) ==> 1 \<in> Vfrom(A,i)"
+lemma one_in_VLimit: "Limit(i) \<Longrightarrow> 1 \<in> Vfrom(A,i)"
by (blast intro: nat_into_VLimit)
lemma Inl_in_VLimit:
- "[| a \<in> Vfrom(A,i); Limit(i) |] ==> Inl(a) \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> Inl(a) \<in> Vfrom(A,i)"
apply (unfold Inl_def)
apply (blast intro: zero_in_VLimit Pair_in_VLimit)
done
lemma Inr_in_VLimit:
- "[| b \<in> Vfrom(A,i); Limit(i) |] ==> Inr(b) \<in> Vfrom(A,i)"
+ "\<lbrakk>b \<in> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> Inr(b) \<in> Vfrom(A,i)"
apply (unfold Inr_def)
apply (blast intro: one_in_VLimit Pair_in_VLimit)
done
-lemma sum_VLimit: "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) \<subseteq> Vfrom(C,i)"
+lemma sum_VLimit: "Limit(i) \<Longrightarrow> 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]
@@ -268,14 +268,14 @@
subsection\<open>Properties assuming \<^term>\<open>Transset(A)\<close>\<close>
-lemma Transset_Vfrom: "Transset(A) ==> Transset(Vfrom(A,i))"
+lemma Transset_Vfrom: "Transset(A) \<Longrightarrow> Transset(Vfrom(A,i))"
apply (rule_tac a=i in eps_induct)
apply (subst Vfrom)
apply (blast intro!: Transset_Union_family Transset_Un Transset_Pow)
done
lemma Transset_Vfrom_succ:
- "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"
+ "Transset(A) \<Longrightarrow> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"
apply (rule Vfrom_succ [THEN trans])
apply (rule equalityI [OF _ Un_upper2])
apply (rule Un_least [OF _ subset_refl])
@@ -283,26 +283,26 @@
apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]])
done
-lemma Transset_Pair_subset: "[| <a,b> \<subseteq> C; Transset(C) |] ==> a: C & b: C"
+lemma Transset_Pair_subset: "\<lbrakk><a,b> \<subseteq> C; Transset(C)\<rbrakk> \<Longrightarrow> a: C & b: C"
by (unfold Pair_def Transset_def, blast)
lemma Transset_Pair_subset_VLimit:
- "[| <a,b> \<subseteq> Vfrom(A,i); Transset(A); Limit(i) |]
- ==> <a,b> \<in> Vfrom(A,i)"
+ "\<lbrakk><a,b> \<subseteq> Vfrom(A,i); Transset(A); Limit(i)\<rbrakk>
+ \<Longrightarrow> <a,b> \<in> Vfrom(A,i)"
apply (erule Transset_Pair_subset [THEN conjE])
apply (erule Transset_Vfrom)
apply (blast intro: Pair_in_VLimit)
done
lemma Union_in_Vfrom:
- "[| X \<in> Vfrom(A,j); Transset(A) |] ==> \<Union>(X) \<in> Vfrom(A, succ(j))"
+ "\<lbrakk>X \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> \<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)"
+ "\<lbrakk>X \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> \<Longrightarrow> \<Union>(X) \<in> Vfrom(A,i)"
apply (rule Limit_VfromE, assumption+)
apply (blast intro: Limit_has_succ VfromI Union_in_Vfrom)
done
@@ -315,10 +315,10 @@
text\<open>General theorem for membership in Vfrom(A,i) when i is a limit ordinal\<close>
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) |]
- ==> \<exists>k. h(x,y) \<in> Vfrom(A,k) & k<i |]
- ==> h(a,b) \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i);
+ \<And>x y j. \<lbrakk>j<i; 1:j; x \<in> Vfrom(A,j); y \<in> Vfrom(A,j)\<rbrakk>
+ \<Longrightarrow> \<exists>k. h(x,y) \<in> Vfrom(A,k) & k<i\<rbrakk>
+ \<Longrightarrow> h(a,b) \<in> Vfrom(A,i)"
txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
apply (erule Limit_VfromE, assumption)
apply (erule Limit_VfromE, assumption, atomize)
@@ -332,8 +332,8 @@
subsubsection\<open>Products\<close>
lemma prod_in_Vfrom:
- "[| a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A) |]
- ==> a*b \<in> Vfrom(A, succ(succ(succ(j))))"
+ "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A)\<rbrakk>
+ \<Longrightarrow> a*b \<in> Vfrom(A, succ(succ(succ(j))))"
apply (drule Transset_Vfrom)
apply (rule subset_mem_Vfrom)
apply (unfold Transset_def)
@@ -341,8 +341,8 @@
done
lemma prod_in_VLimit:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A) |]
- ==> a*b \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk>
+ \<Longrightarrow> a*b \<in> Vfrom(A,i)"
apply (erule in_VLimit, assumption+)
apply (blast intro: prod_in_Vfrom Limit_has_succ)
done
@@ -350,8 +350,8 @@
subsubsection\<open>Disjoint Sums, or Quine Ordered Pairs\<close>
lemma sum_in_Vfrom:
- "[| a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A); 1:j |]
- ==> a+b \<in> Vfrom(A, succ(succ(succ(j))))"
+ "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A); 1:j\<rbrakk>
+ \<Longrightarrow> a+b \<in> Vfrom(A, succ(succ(succ(j))))"
apply (unfold sum_def)
apply (drule Transset_Vfrom)
apply (rule subset_mem_Vfrom)
@@ -360,8 +360,8 @@
done
lemma sum_in_VLimit:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A) |]
- ==> a+b \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk>
+ \<Longrightarrow> a+b \<in> Vfrom(A,i)"
apply (erule in_VLimit, assumption+)
apply (blast intro: sum_in_Vfrom Limit_has_succ)
done
@@ -369,7 +369,7 @@
subsubsection\<open>Function Space!\<close>
lemma fun_in_Vfrom:
- "[| a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A) |] ==>
+ "\<lbrakk>a \<in> Vfrom(A,j); b \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow>
a->b \<in> Vfrom(A, succ(succ(succ(succ(j)))))"
apply (unfold Pi_def)
apply (drule Transset_Vfrom)
@@ -385,14 +385,14 @@
done
lemma fun_in_VLimit:
- "[| a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A) |]
- ==> a->b \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); b \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk>
+ \<Longrightarrow> a->b \<in> Vfrom(A,i)"
apply (erule in_VLimit, assumption+)
apply (blast intro: fun_in_Vfrom Limit_has_succ)
done
lemma Pow_in_Vfrom:
- "[| a \<in> Vfrom(A,j); Transset(A) |] ==> Pow(a) \<in> Vfrom(A, succ(succ(j)))"
+ "\<lbrakk>a \<in> Vfrom(A,j); Transset(A)\<rbrakk> \<Longrightarrow> Pow(a) \<in> Vfrom(A, succ(succ(j)))"
apply (drule Transset_Vfrom)
apply (rule subset_mem_Vfrom)
apply (unfold Transset_def)
@@ -400,7 +400,7 @@
done
lemma Pow_in_VLimit:
- "[| a \<in> Vfrom(A,i); Limit(i); Transset(A) |] ==> Pow(a) \<in> Vfrom(A,i)"
+ "\<lbrakk>a \<in> Vfrom(A,i); Limit(i); Transset(A)\<rbrakk> \<Longrightarrow> Pow(a) \<in> Vfrom(A,i)"
by (blast elim: Limit_VfromE intro: Limit_has_succ Pow_in_Vfrom VfromI)
@@ -414,7 +414,7 @@
subsubsection\<open>Characterisation of the elements of \<^term>\<open>Vset(i)\<close>\<close>
-lemma VsetD [rule_format]: "Ord(i) ==> \<forall>b. b \<in> Vset(i) \<longrightarrow> rank(b) < i"
+lemma VsetD [rule_format]: "Ord(i) \<Longrightarrow> \<forall>b. b \<in> Vset(i) \<longrightarrow> rank(b) < i"
apply (erule trans_induct)
apply (subst Vset, safe)
apply (subst rank)
@@ -422,18 +422,18 @@
done
lemma VsetI_lemma [rule_format]:
- "Ord(i) ==> \<forall>b. rank(b) \<in> i \<longrightarrow> b \<in> Vset(i)"
+ "Ord(i) \<Longrightarrow> \<forall>b. rank(b) \<in> i \<longrightarrow> b \<in> Vset(i)"
apply (erule trans_induct)
apply (rule allI)
apply (subst Vset)
apply (blast intro!: rank_lt [THEN ltD])
done
-lemma VsetI: "rank(x)<i ==> x \<in> Vset(i)"
+lemma VsetI: "rank(x)<i \<Longrightarrow> x \<in> Vset(i)"
by (blast intro: VsetI_lemma elim: ltE)
text\<open>Merely a lemma for the next result\<close>
-lemma Vset_Ord_rank_iff: "Ord(i) ==> b \<in> Vset(i) \<longleftrightarrow> rank(b) < i"
+lemma Vset_Ord_rank_iff: "Ord(i) \<Longrightarrow> b \<in> Vset(i) \<longleftrightarrow> rank(b) < i"
by (blast intro: VsetD VsetI)
lemma Vset_rank_iff [simp]: "b \<in> Vset(a) \<longleftrightarrow> rank(b) < rank(a)"
@@ -444,7 +444,7 @@
text\<open>This is rank(rank(a)) = rank(a)\<close>
declare Ord_rank [THEN rank_of_Ord, simp]
-lemma rank_Vset: "Ord(i) ==> rank(Vset(i)) = i"
+lemma rank_Vset: "Ord(i) \<Longrightarrow> rank(Vset(i)) = i"
apply (subst rank)
apply (rule equalityI, safe)
apply (blast intro: VsetD [THEN ltD])
@@ -453,7 +453,7 @@
Ord_in_Ord [THEN rank_of_Ord, THEN ssubst])
done
-lemma Finite_Vset: "i \<in> nat ==> Finite(Vset(i))"
+lemma Finite_Vset: "i \<in> nat \<Longrightarrow> Finite(Vset(i))"
apply (erule nat_induct)
apply (simp add: Vfrom_0)
apply (simp add: Vset_succ)
@@ -467,7 +467,7 @@
done
lemma Int_Vset_subset:
- "[| !!i. Ord(i) ==> a \<inter> Vset(i) \<subseteq> b |] ==> a \<subseteq> b"
+ "\<lbrakk>\<And>i. Ord(i) \<Longrightarrow> a \<inter> Vset(i) \<subseteq> b\<rbrakk> \<Longrightarrow> a \<subseteq> b"
apply (rule subset_trans)
apply (rule Int_greatest [OF subset_refl arg_subset_Vset_rank])
apply (blast intro: Ord_rank)
@@ -496,9 +496,9 @@
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
done
-text\<open>This form avoids giant explosions in proofs. NOTE USE OF ==\<close>
+text\<open>This form avoids giant explosions in proofs. NOTE USE OF \<equiv>\<close>
lemma def_Vrec:
- "[| !!x. h(x)==Vrec(x,H) |] ==>
+ "\<lbrakk>\<And>x. h(x)\<equiv>Vrec(x,H)\<rbrakk> \<Longrightarrow>
h(a) = H(a, \<lambda>x\<in>Vset(rank(a)). h(x))"
apply simp
apply (rule Vrec)
@@ -512,9 +512,9 @@
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
done
-text\<open>This form avoids giant explosions in proofs. NOTE USE OF ==\<close>
+text\<open>This form avoids giant explosions in proofs. NOTE USE OF \<equiv>\<close>
lemma def_Vrecursor:
- "h == Vrecursor(H) ==> h(a) = H(\<lambda>x\<in>Vset(rank(a)). h(x), a)"
+ "h \<equiv> Vrecursor(H) \<Longrightarrow> h(a) = H(\<lambda>x\<in>Vset(rank(a)). h(x), a)"
apply simp
apply (rule Vrecursor)
done
@@ -522,13 +522,13 @@
subsection\<open>The Datatype Universe: \<^term>\<open>univ(A)\<close>\<close>
-lemma univ_mono: "A<=B ==> univ(A) \<subseteq> univ(B)"
+lemma univ_mono: "A<=B \<Longrightarrow> univ(A) \<subseteq> univ(B)"
apply (unfold univ_def)
apply (erule Vfrom_mono)
apply (rule subset_refl)
done
-lemma Transset_univ: "Transset(A) ==> Transset(univ(A))"
+lemma Transset_univ: "Transset(A) \<Longrightarrow> Transset(univ(A))"
apply (unfold univ_def)
apply (erule Transset_Vfrom)
done
@@ -540,23 +540,23 @@
apply (rule Limit_nat [THEN Limit_Vfrom_eq])
done
-lemma subset_univ_eq_Int: "c \<subseteq> univ(A) ==> c = (\<Union>i\<in>nat. c \<inter> Vfrom(A,i))"
+lemma subset_univ_eq_Int: "c \<subseteq> univ(A) \<Longrightarrow> 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 \<subseteq> univ(X);
- !!i. i:nat ==> a \<inter> Vfrom(X,i) \<subseteq> b |]
- ==> a \<subseteq> b"
+ "\<lbrakk>a \<subseteq> univ(X);
+ \<And>i. i:nat \<Longrightarrow> a \<inter> Vfrom(X,i) \<subseteq> b\<rbrakk>
+ \<Longrightarrow> a \<subseteq> b"
apply (subst subset_univ_eq_Int, assumption)
apply (rule UN_least, simp)
done
lemma univ_Int_Vfrom_eq:
- "[| a \<subseteq> univ(X); b \<subseteq> univ(X);
- !!i. i:nat ==> a \<inter> Vfrom(X,i) = b \<inter> Vfrom(X,i)
- |] ==> a = b"
+ "\<lbrakk>a \<subseteq> univ(X); b \<subseteq> univ(X);
+ \<And>i. i:nat \<Longrightarrow> a \<inter> Vfrom(X,i) = b \<inter> Vfrom(X,i)
+\<rbrakk> \<Longrightarrow> a = b"
apply (rule equalityI)
apply (rule univ_Int_Vfrom_subset, assumption)
apply (blast elim: equalityCE)
@@ -583,25 +583,25 @@
subsubsection\<open>Closure under Unordered and Ordered Pairs\<close>
-lemma singleton_in_univ: "a: univ(A) ==> {a} \<in> univ(A)"
+lemma singleton_in_univ: "a: univ(A) \<Longrightarrow> {a} \<in> univ(A)"
apply (unfold univ_def)
apply (blast intro: singleton_in_VLimit Limit_nat)
done
lemma doubleton_in_univ:
- "[| a: univ(A); b: univ(A) |] ==> {a,b} \<in> univ(A)"
+ "\<lbrakk>a: univ(A); b: univ(A)\<rbrakk> \<Longrightarrow> {a,b} \<in> univ(A)"
apply (unfold univ_def)
apply (blast intro: doubleton_in_VLimit Limit_nat)
done
lemma Pair_in_univ:
- "[| a: univ(A); b: univ(A) |] ==> <a,b> \<in> univ(A)"
+ "\<lbrakk>a: univ(A); b: univ(A)\<rbrakk> \<Longrightarrow> <a,b> \<in> univ(A)"
apply (unfold univ_def)
apply (blast intro: Pair_in_VLimit Limit_nat)
done
lemma Union_in_univ:
- "[| X: univ(A); Transset(A) |] ==> \<Union>(X) \<in> univ(A)"
+ "\<lbrakk>X: univ(A); Transset(A)\<rbrakk> \<Longrightarrow> \<Union>(X) \<in> univ(A)"
apply (unfold univ_def)
apply (blast intro: Union_in_VLimit Limit_nat)
done
@@ -619,7 +619,7 @@
apply (rule i_subset_Vfrom)
done
-text\<open>n:nat ==> n:univ(A)\<close>
+text\<open>n:nat \<Longrightarrow> n:univ(A)\<close>
lemmas nat_into_univ = nat_subset_univ [THEN subsetD]
subsubsection\<open>Instances for 1 and 2\<close>
@@ -643,12 +643,12 @@
subsubsection\<open>Closure under Disjoint Union\<close>
-lemma Inl_in_univ: "a: univ(A) ==> Inl(a) \<in> univ(A)"
+lemma Inl_in_univ: "a: univ(A) \<Longrightarrow> Inl(a) \<in> univ(A)"
apply (unfold univ_def)
apply (erule Inl_in_VLimit [OF _ Limit_nat])
done
-lemma Inr_in_univ: "b: univ(A) ==> Inr(b) \<in> univ(A)"
+lemma Inr_in_univ: "b: univ(A) \<Longrightarrow> Inr(b) \<in> univ(A)"
apply (unfold univ_def)
apply (erule Inr_in_VLimit [OF _ Limit_nat])
done
@@ -661,7 +661,7 @@
lemmas sum_subset_univ = subset_trans [OF sum_mono sum_univ]
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)"
+ "\<lbrakk>A \<subseteq> univ(D); \<And>x. x \<in> A \<Longrightarrow> B(x) \<subseteq> univ(D)\<rbrakk> \<Longrightarrow> Sigma(A,B) \<subseteq> univ(D)"
apply (simp add: univ_def)
apply (blast intro: Sigma_subset_VLimit del: subsetI)
done
@@ -677,14 +677,14 @@
subsubsection\<open>Closure under Finite Powerset\<close>
lemma Fin_Vfrom_lemma:
- "[| b: Fin(Vfrom(A,i)); Limit(i) |] ==> \<exists>j. b \<subseteq> Vfrom(A,j) & j<i"
+ "\<lbrakk>b: Fin(Vfrom(A,i)); Limit(i)\<rbrakk> \<Longrightarrow> \<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)) \<subseteq> Vfrom(A,i)"
+lemma Fin_VLimit: "Limit(i) \<Longrightarrow> Fin(Vfrom(A,i)) \<subseteq> Vfrom(A,i)"
apply (rule subsetI)
apply (drule Fin_Vfrom_lemma, safe)
apply (rule Vfrom [THEN ssubst])
@@ -701,7 +701,7 @@
subsubsection\<open>Closure under Finite Powers: Functions from a Natural Number\<close>
lemma nat_fun_VLimit:
- "[| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
+ "\<lbrakk>n: nat; Limit(i)\<rbrakk> \<Longrightarrow> 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) \<subseteq> univ(A)"
+lemma nat_fun_univ: "n: nat \<Longrightarrow> n -> univ(A) \<subseteq> univ(A)"
apply (unfold univ_def)
apply (erule nat_fun_VLimit [OF _ Limit_nat])
done
@@ -719,7 +719,7 @@
text\<open>General but seldom-used version; normally the domain is fixed\<close>
lemma FiniteFun_VLimit1:
- "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
+ "Limit(i) \<Longrightarrow> 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)
@@ -732,20 +732,20 @@
text\<open>Version for a fixed domain\<close>
lemma FiniteFun_VLimit:
- "[| W \<subseteq> Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) \<subseteq> Vfrom(A,i)"
+ "\<lbrakk>W \<subseteq> Vfrom(A,i); Limit(i)\<rbrakk> \<Longrightarrow> 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 \<subseteq> univ(A) ==> W -||> univ(A) \<subseteq> univ(A)"
+ "W \<subseteq> univ(A) \<Longrightarrow> 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 \<subseteq> univ(A) |] ==> f \<in> univ(A)"
+ "\<lbrakk>f: W -||> univ(A); W \<subseteq> univ(A)\<rbrakk> \<Longrightarrow> f \<in> univ(A)"
by (erule FiniteFun_univ [THEN subsetD], assumption)
text\<open>Remove \<open>\<subseteq>\<close> from the rule above\<close>
@@ -758,8 +758,8 @@
text\<open>This version says a, b exist one level down, in the smaller set Vfrom(X,i)\<close>
lemma doubleton_in_Vfrom_D:
- "[| {a,b} \<in> Vfrom(X,succ(i)); Transset(X) |]
- ==> a \<in> Vfrom(X,i) & b \<in> Vfrom(X,i)"
+ "\<lbrakk>{a,b} \<in> Vfrom(X,succ(i)); Transset(X)\<rbrakk>
+ \<Longrightarrow> a \<in> Vfrom(X,i) & b \<in> Vfrom(X,i)"
by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD],
assumption, fast)
@@ -775,14 +775,14 @@
**)
lemma Pair_in_Vfrom_D:
- "[| <a,b> \<in> Vfrom(X,succ(i)); Transset(X) |]
- ==> a \<in> Vfrom(X,i) & b \<in> Vfrom(X,i)"
+ "\<lbrakk><a,b> \<in> Vfrom(X,succ(i)); Transset(X)\<rbrakk>
+ \<Longrightarrow> a \<in> Vfrom(X,i) & b \<in> Vfrom(X,i)"
apply (unfold Pair_def)
apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D)
done
lemma product_Int_Vfrom_subset:
- "Transset(X) ==>
+ "Transset(X) \<Longrightarrow>
(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)
--- a/src/ZF/WF.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/WF.thy Tue Sep 27 16:51:35 2022 +0100
@@ -21,58 +21,58 @@
definition
wf :: "i=>o" where
(*r is a well-founded relation*)
- "wf(r) == \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. <y,x>:r \<longrightarrow> ~ y \<in> Z)"
+ "wf(r) \<equiv> \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. <y,x>:r \<longrightarrow> \<not> y \<in> Z)"
definition
wf_on :: "[i,i]=>o" (\<open>wf[_]'(_')\<close>) where
(*r is well-founded on A*)
- "wf_on(A,r) == wf(r \<inter> A*A)"
+ "wf_on(A,r) \<equiv> wf(r \<inter> A*A)"
definition
is_recfun :: "[i, i, [i,i]=>i, i] =>o" where
- "is_recfun(r,a,H,f) == (f = (\<lambda>x\<in>r-``{a}. H(x, restrict(f, r-``{x}))))"
+ "is_recfun(r,a,H,f) \<equiv> (f = (\<lambda>x\<in>r-``{a}. H(x, restrict(f, r-``{x}))))"
definition
the_recfun :: "[i, i, [i,i]=>i] =>i" where
- "the_recfun(r,a,H) == (THE f. is_recfun(r,a,H,f))"
+ "the_recfun(r,a,H) \<equiv> (THE f. is_recfun(r,a,H,f))"
definition
wftrec :: "[i, i, [i,i]=>i] =>i" where
- "wftrec(r,a,H) == H(a, the_recfun(r,a,H))"
+ "wftrec(r,a,H) \<equiv> H(a, the_recfun(r,a,H))"
definition
wfrec :: "[i, i, [i,i]=>i] =>i" where
(*public version. Does not require r to be transitive*)
- "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))"
+ "wfrec(r,a,H) \<equiv> wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))"
definition
wfrec_on :: "[i, i, i, [i,i]=>i] =>i" (\<open>wfrec[_]'(_,_,_')\<close>) where
- "wfrec[A](r,a,H) == wfrec(r \<inter> A*A, a, H)"
+ "wfrec[A](r,a,H) \<equiv> wfrec(r \<inter> A*A, a, H)"
subsection\<open>Well-Founded Relations\<close>
subsubsection\<open>Equivalences between \<^term>\<open>wf\<close> and \<^term>\<open>wf_on\<close>\<close>
-lemma wf_imp_wf_on: "wf(r) ==> wf[A](r)"
+lemma wf_imp_wf_on: "wf(r) \<Longrightarrow> wf[A](r)"
by (unfold wf_def wf_on_def, force)
-lemma wf_on_imp_wf: "[|wf[A](r); r \<subseteq> A*A|] ==> wf(r)"
+lemma wf_on_imp_wf: "\<lbrakk>wf[A](r); r \<subseteq> A*A\<rbrakk> \<Longrightarrow> wf(r)"
by (simp add: wf_on_def subset_Int_iff)
-lemma wf_on_field_imp_wf: "wf[field(r)](r) ==> wf(r)"
+lemma wf_on_field_imp_wf: "wf[field(r)](r) \<Longrightarrow> wf(r)"
by (unfold wf_def wf_on_def, fast)
lemma wf_iff_wf_on_field: "wf(r) \<longleftrightarrow> wf[field(r)](r)"
by (blast intro: wf_imp_wf_on wf_on_field_imp_wf)
-lemma wf_on_subset_A: "[| wf[A](r); B<=A |] ==> wf[B](r)"
+lemma wf_on_subset_A: "\<lbrakk>wf[A](r); B<=A\<rbrakk> \<Longrightarrow> wf[B](r)"
by (unfold wf_on_def wf_def, fast)
-lemma wf_on_subset_r: "[| wf[A](r); s<=r |] ==> wf[A](s)"
+lemma wf_on_subset_r: "\<lbrakk>wf[A](r); s<=r\<rbrakk> \<Longrightarrow> wf[A](s)"
by (unfold wf_on_def wf_def, fast)
-lemma wf_subset: "[|wf(s); r<=s|] ==> wf(r)"
+lemma wf_subset: "\<lbrakk>wf(s); r<=s\<rbrakk> \<Longrightarrow> wf(r)"
by (simp add: wf_def, fast)
subsubsection\<open>Introduction Rules for \<^term>\<open>wf_on\<close>\<close>
@@ -80,7 +80,7 @@
text\<open>If every non-empty subset of \<^term>\<open>A\<close> has an \<^term>\<open>r\<close>-minimal element
then we have \<^term>\<open>wf[A](r)\<close>.\<close>
lemma wf_onI:
- assumes prem: "!!Z u. [| Z<=A; u \<in> Z; \<forall>x\<in>Z. \<exists>y\<in>Z. <y,x>:r |] ==> False"
+ assumes prem: "\<And>Z u. \<lbrakk>Z<=A; u \<in> Z; \<forall>x\<in>Z. \<exists>y\<in>Z. <y,x>:r\<rbrakk> \<Longrightarrow> False"
shows "wf[A](r)"
apply (unfold wf_on_def wf_def)
apply (rule equals0I [THEN disjCI, THEN allI])
@@ -89,10 +89,10 @@
text\<open>If \<^term>\<open>r\<close> allows well-founded induction over \<^term>\<open>A\<close>
then we have \<^term>\<open>wf[A](r)\<close>. Premise is equivalent to
- \<^prop>\<open>!!B. \<forall>x\<in>A. (\<forall>y. <y,x>: r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B ==> A<=B\<close>\<close>
+ \<^prop>\<open>\<And>B. \<forall>x\<in>A. (\<forall>y. <y,x>: r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B \<Longrightarrow> A<=B\<close>\<close>
lemma wf_onI2:
- assumes prem: "!!y B. [| \<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A |]
- ==> y \<in> B"
+ assumes prem: "\<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A\<rbrakk>
+ \<Longrightarrow> y \<in> B"
shows "wf[A](r)"
apply (rule wf_onI)
apply (rule_tac c=u in prem [THEN DiffE])
@@ -106,11 +106,11 @@
text\<open>Consider the least \<^term>\<open>z\<close> in \<^term>\<open>domain(r)\<close> such that
\<^term>\<open>P(z)\<close> does not hold...\<close>
lemma wf_induct_raw:
- "[| wf(r);
- !!x.[| \<forall>y. <y,x>: r \<longrightarrow> P(y) |] ==> P(x) |]
- ==> P(a)"
+ "\<lbrakk>wf(r);
+ \<And>x.\<lbrakk>\<forall>y. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> P(a)"
apply (unfold wf_def)
-apply (erule_tac x = "{z \<in> domain(r). ~ P(z)}" in allE)
+apply (erule_tac x = "{z \<in> domain(r). \<not> P(z)}" in allE)
apply blast
done
@@ -118,9 +118,9 @@
text\<open>The form of this rule is designed to match \<open>wfI\<close>\<close>
lemma wf_induct2:
- "[| wf(r); a \<in> A; field(r)<=A;
- !!x.[| x \<in> A; \<forall>y. <y,x>: r \<longrightarrow> P(y) |] ==> P(x) |]
- ==> P(a)"
+ "\<lbrakk>wf(r); a \<in> A; field(r)<=A;
+ \<And>x.\<lbrakk>x \<in> A; \<forall>y. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
+ \<Longrightarrow> P(a)"
apply (erule_tac P="a \<in> A" in rev_mp)
apply (erule_tac a=a in wf_induct, blast)
done
@@ -129,9 +129,9 @@
by blast
lemma wf_on_induct_raw [consumes 2, induct set: wf_on]:
- "[| wf[A](r); a \<in> A;
- !!x.[| x \<in> A; \<forall>y\<in>A. <y,x>: r \<longrightarrow> P(y) |] ==> P(x)
- |] ==> P(a)"
+ "\<lbrakk>wf[A](r); a \<in> A;
+ \<And>x.\<lbrakk>x \<in> A; \<forall>y\<in>A. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
+\<rbrakk> \<Longrightarrow> P(a)"
apply (unfold wf_on_def)
apply (erule wf_induct2, assumption)
apply (rule field_Int_square, blast)
@@ -144,10 +144,10 @@
text\<open>If \<^term>\<open>r\<close> allows well-founded induction
then we have \<^term>\<open>wf(r)\<close>.\<close>
lemma wfI:
- "[| field(r)<=A;
- !!y B. [| \<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A|]
- ==> y \<in> B |]
- ==> wf(r)"
+ "\<lbrakk>field(r)<=A;
+ \<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B; y \<in> A\<rbrakk>
+ \<Longrightarrow> y \<in> B\<rbrakk>
+ \<Longrightarrow> wf(r)"
apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf])
apply (rule wf_onI2)
prefer 2 apply blast
@@ -157,34 +157,34 @@
subsection\<open>Basic Properties of Well-Founded Relations\<close>
-lemma wf_not_refl: "wf(r) ==> <a,a> \<notin> r"
+lemma wf_not_refl: "wf(r) \<Longrightarrow> <a,a> \<notin> r"
by (erule_tac a=a in wf_induct, blast)
-lemma wf_not_sym [rule_format]: "wf(r) ==> \<forall>x. <a,x>:r \<longrightarrow> <x,a> \<notin> r"
+lemma wf_not_sym [rule_format]: "wf(r) \<Longrightarrow> \<forall>x. <a,x>:r \<longrightarrow> <x,a> \<notin> r"
by (erule_tac a=a in wf_induct, blast)
-(* @{term"[| wf(r); <a,x> \<in> r; ~P ==> <x,a> \<in> r |] ==> P"} *)
+(* @{term"\<lbrakk>wf(r); <a,x> \<in> r; \<not>P \<Longrightarrow> <x,a> \<in> r\<rbrakk> \<Longrightarrow> P"} *)
lemmas wf_asym = wf_not_sym [THEN swap]
-lemma wf_on_not_refl: "[| wf[A](r); a \<in> A |] ==> <a,a> \<notin> r"
+lemma wf_on_not_refl: "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> <a,a> \<notin> r"
by (erule_tac a=a in wf_on_induct, assumption, blast)
lemma wf_on_not_sym:
- "[| wf[A](r); a \<in> A |] ==> (\<And>b. b\<in>A \<Longrightarrow> <a,b>:r \<Longrightarrow> <b,a>\<notin>r)"
+ "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> (\<And>b. b\<in>A \<Longrightarrow> <a,b>:r \<Longrightarrow> <b,a>\<notin>r)"
apply (atomize (full), intro impI)
apply (erule_tac a=a in wf_on_induct, assumption, blast)
done
lemma wf_on_asym:
- "[| wf[A](r); ~Z ==> <a,b> \<in> r;
- <b,a> \<notin> r ==> Z; ~Z ==> a \<in> A; ~Z ==> b \<in> A |] ==> Z"
+ "\<lbrakk>wf[A](r); \<not>Z \<Longrightarrow> <a,b> \<in> r;
+ <b,a> \<notin> r \<Longrightarrow> Z; \<not>Z \<Longrightarrow> a \<in> A; \<not>Z \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> Z"
by (blast dest: wf_on_not_sym)
(*Needed to prove well_ordI. Could also reason that wf[A](r) means
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 \<in> A; b \<in> A; c \<in> A |] ==> P"
+ "\<lbrakk>wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> P"
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)
@@ -195,14 +195,14 @@
text\<open>transitive closure of a WF relation is WF provided
\<^term>\<open>A\<close> is downward closed\<close>
lemma wf_on_trancl:
- "[| wf[A](r); r-``A \<subseteq> A |] ==> wf[A](r^+)"
+ "\<lbrakk>wf[A](r); r-``A \<subseteq> A\<rbrakk> \<Longrightarrow> 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)
done
-lemma wf_trancl: "wf(r) ==> wf(r^+)"
+lemma wf_trancl: "wf(r) \<Longrightarrow> wf(r^+)"
apply (simp add: wf_iff_wf_on_field)
apply (rule wf_on_subset_A)
apply (erule wf_on_trancl)
@@ -219,7 +219,7 @@
subsection\<open>The Predicate \<^term>\<open>is_recfun\<close>\<close>
-lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f \<in> r-``{a} -> range(f)"
+lemma is_recfun_type: "is_recfun(r,a,H,f) \<Longrightarrow> f \<in> r-``{a} -> range(f)"
apply (unfold is_recfun_def)
apply (erule ssubst)
apply (rule lamI [THEN rangeI, THEN lam_type], assumption)
@@ -228,7 +228,7 @@
lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function]
lemma apply_recfun:
- "[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))"
+ "\<lbrakk>is_recfun(r,a,H,f); <x,a>:r\<rbrakk> \<Longrightarrow> f`x = H(x, restrict(f,r-``{x}))"
apply (unfold is_recfun_def)
txt\<open>replace f only on the left-hand side\<close>
apply (erule_tac P = "%x. t(x) = u" for t u in ssubst)
@@ -236,8 +236,8 @@
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 \<longrightarrow> <x,b>:r \<longrightarrow> f`x=g`x"
+ "\<lbrakk>wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g)\<rbrakk>
+ \<Longrightarrow> <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)
@@ -253,9 +253,9 @@
done
lemma is_recfun_cut:
- "[| wf(r); trans(r);
- is_recfun(r,a,H,f); is_recfun(r,b,H,g); <b,a>:r |]
- ==> restrict(f, r-``{b}) = g"
+ "\<lbrakk>wf(r); trans(r);
+ is_recfun(r,a,H,f); is_recfun(r,b,H,g); <b,a>:r\<rbrakk>
+ \<Longrightarrow> restrict(f, r-``{b}) = g"
apply (frule_tac f = f in is_recfun_type)
apply (rule fun_extension)
apply (blast dest: transD intro: restrict_type2)
@@ -266,23 +266,23 @@
subsection\<open>Recursion: Main Existence Lemma\<close>
lemma is_recfun_functional:
- "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g"
+ "\<lbrakk>wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g)\<rbrakk> \<Longrightarrow> f=g"
by (blast intro: fun_extension is_recfun_type is_recfun_equal)
lemma the_recfun_eq:
- "[| is_recfun(r,a,H,f); wf(r); trans(r) |] ==> the_recfun(r,a,H) = f"
+ "\<lbrakk>is_recfun(r,a,H,f); wf(r); trans(r)\<rbrakk> \<Longrightarrow> the_recfun(r,a,H) = f"
apply (unfold the_recfun_def)
apply (blast intro: is_recfun_functional)
done
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *)
lemma is_the_recfun:
- "[| is_recfun(r,a,H,f); wf(r); trans(r) |]
- ==> is_recfun(r, a, H, the_recfun(r,a,H))"
+ "\<lbrakk>is_recfun(r,a,H,f); wf(r); trans(r)\<rbrakk>
+ \<Longrightarrow> is_recfun(r, a, H, the_recfun(r,a,H))"
by (simp add: the_recfun_eq)
lemma unfold_the_recfun:
- "[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))"
+ "\<lbrakk>wf(r); trans(r)\<rbrakk> \<Longrightarrow> 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 = "\<lambda>y\<in>r-``{a1}. wftrec (r,y,H)" in is_the_recfun)
@@ -311,13 +311,13 @@
subsection\<open>Unfolding \<^term>\<open>wftrec(r,a,H)\<close>\<close>
lemma the_recfun_cut:
- "[| wf(r); trans(r); <b,a>:r |]
- ==> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"
+ "\<lbrakk>wf(r); trans(r); <b,a>:r\<rbrakk>
+ \<Longrightarrow> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"
by (blast intro: is_recfun_cut unfold_the_recfun)
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
lemma wftrec:
- "[| wf(r); trans(r) |] ==>
+ "\<lbrakk>wf(r); trans(r)\<rbrakk> \<Longrightarrow>
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])
@@ -329,7 +329,7 @@
(*NOT SUITABLE FOR REWRITING: it is recursive!*)
lemma wfrec:
- "wf(r) ==> wfrec(r,a,H) = H(a, \<lambda>x\<in>r-``{a}. wfrec(r,x,H))"
+ "wf(r) \<Longrightarrow> 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)
@@ -338,18 +338,18 @@
apply (rule subset_refl)
done
-(*This form avoids giant explosions in proofs. NOTE USE OF == *)
+(*This form avoids giant explosions in proofs. NOTE USE OF \<equiv> *)
lemma def_wfrec:
- "[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==>
+ "\<lbrakk>\<And>x. h(x)\<equiv>wfrec(r,x,H); wf(r)\<rbrakk> \<Longrightarrow>
h(a) = H(a, \<lambda>x\<in>r-``{a}. h(x))"
apply simp
apply (elim wfrec)
done
lemma wfrec_type:
- "[| wf(r); a \<in> A; field(r)<=A;
- !!x u. [| x \<in> A; u \<in> Pi(r-``{x}, B) |] ==> H(x,u) \<in> B(x)
- |] ==> wfrec(r,a,H) \<in> B(a)"
+ "\<lbrakk>wf(r); a \<in> A; field(r)<=A;
+ \<And>x u. \<lbrakk>x \<in> A; u \<in> Pi(r-``{x}, B)\<rbrakk> \<Longrightarrow> H(x,u) \<in> B(x)
+\<rbrakk> \<Longrightarrow> 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)
@@ -357,7 +357,7 @@
lemma wfrec_on:
- "[| wf[A](r); a \<in> A |] ==>
+ "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow>
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])
--- a/src/ZF/ZF.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ZF.thy Tue Sep 27 16:51:35 2022 +0100
@@ -15,27 +15,27 @@
definition
iterates_omega :: "[i=>i,i] => i" (\<open>(_^\<omega> '(_'))\<close> [60,1000] 60) where
- "F^\<omega> (x) == \<Union>n\<in>nat. F^n (x)"
+ "F^\<omega> (x) \<equiv> \<Union>n\<in>nat. F^n (x)"
lemma iterates_triv:
- "[| n\<in>nat; F(x) = x |] ==> F^n (x) = x"
+ "\<lbrakk>n\<in>nat; F(x) = x\<rbrakk> \<Longrightarrow> F^n (x) = x"
by (induct n rule: nat_induct, simp_all)
lemma iterates_type [TC]:
- "[| n \<in> nat; a \<in> A; !!x. x \<in> A ==> F(x) \<in> A |]
- ==> F^n (a) \<in> A"
+ "\<lbrakk>n \<in> nat; a \<in> A; \<And>x. x \<in> A \<Longrightarrow> F(x) \<in> A\<rbrakk>
+ \<Longrightarrow> F^n (a) \<in> A"
by (induct n rule: nat_induct, simp_all)
lemma iterates_omega_triv:
- "F(x) = x ==> F^\<omega> (x) = x"
+ "F(x) = x \<Longrightarrow> F^\<omega> (x) = x"
by (simp add: iterates_omega_def iterates_triv)
lemma Ord_iterates [simp]:
- "[| n\<in>nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |]
- ==> Ord(F^n (x))"
+ "\<lbrakk>n\<in>nat; \<And>i. Ord(i) \<Longrightarrow> Ord(F(i)); Ord(x)\<rbrakk>
+ \<Longrightarrow> Ord(F^n (x))"
by (induct n rule: nat_induct, simp_all)
-lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))"
+lemma iterates_commute: "n \<in> nat \<Longrightarrow> F(F^n (x)) = F^n (F(x))"
by (induct_tac n, simp_all)
@@ -46,7 +46,7 @@
definition
transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
- "transrec3(k, a, b, c) ==
+ "transrec3(k, a, b, c) \<equiv>
transrec(k, \<lambda>x r.
if x=0 then a
else if Limit(x) then c(x, \<lambda>y\<in>x. r`y)
@@ -60,7 +60,7 @@
by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
lemma transrec3_Limit:
- "Limit(i) ==>
+ "Limit(i) \<Longrightarrow>
transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))"
by (rule transrec3_def [THEN def_transrec, THEN trans], force)
--- a/src/ZF/ZF_Base.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ZF_Base.thy Tue Sep 27 16:51:35 2022 +0100
@@ -49,7 +49,7 @@
The resulting set (for functional P) is the same as with
PrimReplace, but the rules are simpler. *)
definition Replace :: "[i, [i, i] \<Rightarrow> o] \<Rightarrow> i"
- where "Replace(A,P) == PrimReplace(A, %x y. (\<exists>!z. P(x,z)) & P(x,y))"
+ where "Replace(A,P) \<equiv> PrimReplace(A, %x y. (\<exists>!z. P(x,z)) & P(x,y))"
syntax
"_Replace" :: "[pttrn, pttrn, i, o] => i" (\<open>(1{_ ./ _ \<in> _, _})\<close>)
@@ -60,7 +60,7 @@
(* Functional form of replacement -- analgous to ML's map functional *)
definition RepFun :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
- where "RepFun(A,f) == {y . x\<in>A, y=f(x)}"
+ where "RepFun(A,f) \<equiv> {y . x\<in>A, y=f(x)}"
syntax
"_RepFun" :: "[i, pttrn, i] => i" (\<open>(1{_ ./ _ \<in> _})\<close> [51,0,51])
@@ -71,7 +71,7 @@
(* Separation and Pairing can be derived from the Replacement
and Powerset Axioms using the following definitions. *)
definition Collect :: "[i, i \<Rightarrow> o] \<Rightarrow> i"
- where "Collect(A,P) == {y . x\<in>A, x=y & P(x)}"
+ where "Collect(A,P) \<equiv> {y . x\<in>A, x=y & P(x)}"
syntax
"_Collect" :: "[pttrn, i, o] \<Rightarrow> i" (\<open>(1{_ \<in> _ ./ _})\<close>)
@@ -82,7 +82,7 @@
subsection \<open>General union and intersection\<close>
definition Inter :: "i => i" (\<open>\<Inter>_\<close> [90] 90)
- where "\<Inter>(A) == { x\<in>\<Union>(A) . \<forall>y\<in>A. x\<in>y}"
+ where "\<Inter>(A) \<equiv> { x\<in>\<Union>(A) . \<forall>y\<in>A. x\<in>y}"
syntax
"_UNION" :: "[pttrn, i, i] => i" (\<open>(3\<Union>_\<in>_./ _)\<close> 10)
@@ -98,25 +98,25 @@
set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*)
definition Upair :: "[i, i] => i"
- where "Upair(a,b) == {y. x\<in>Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
+ where "Upair(a,b) \<equiv> {y. x\<in>Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
definition Subset :: "[i, i] \<Rightarrow> o" (infixl \<open>\<subseteq>\<close> 50) \<comment> \<open>subset relation\<close>
where subset_def: "A \<subseteq> B \<equiv> \<forall>x\<in>A. x\<in>B"
definition Diff :: "[i, i] \<Rightarrow> i" (infixl \<open>-\<close> 65) \<comment> \<open>set difference\<close>
- where "A - B == { x\<in>A . ~(x\<in>B) }"
+ where "A - B \<equiv> { x\<in>A . \<not>(x\<in>B) }"
definition Un :: "[i, i] \<Rightarrow> i" (infixl \<open>\<union>\<close> 65) \<comment> \<open>binary union\<close>
- where "A \<union> B == \<Union>(Upair(A,B))"
+ where "A \<union> B \<equiv> \<Union>(Upair(A,B))"
definition Int :: "[i, i] \<Rightarrow> i" (infixl \<open>\<inter>\<close> 70) \<comment> \<open>binary intersection\<close>
- where "A \<inter> B == \<Inter>(Upair(A,B))"
+ where "A \<inter> B \<equiv> \<Inter>(Upair(A,B))"
definition cons :: "[i, i] => i"
- where "cons(a,A) == Upair(a,a) \<union> A"
+ where "cons(a,A) \<equiv> Upair(a,a) \<union> A"
definition succ :: "i => i"
- where "succ(i) == cons(i, i)"
+ where "succ(i) \<equiv> cons(i, i)"
nonterminal "is"
syntax
@@ -154,30 +154,30 @@
subsection \<open>Definite descriptions -- via Replace over the set "1"\<close>
definition The :: "(i \<Rightarrow> o) \<Rightarrow> i" (binder \<open>THE \<close> 10)
- where the_def: "The(P) == \<Union>({y . x \<in> {0}, P(y)})"
+ where the_def: "The(P) \<equiv> \<Union>({y . x \<in> {0}, P(y)})"
definition If :: "[o, i, i] \<Rightarrow> i" (\<open>(if (_)/ then (_)/ else (_))\<close> [10] 10)
- where if_def: "if P then a else b == THE z. P & z=a | ~P & z=b"
+ where if_def: "if P then a else b \<equiv> THE z. P & z=a | \<not>P & z=b"
abbreviation (input)
old_if :: "[o, i, i] => i" (\<open>if '(_,_,_')\<close>)
- where "if(P,a,b) == If(P,a,b)"
+ where "if(P,a,b) \<equiv> If(P,a,b)"
subsection \<open>Ordered Pairing\<close>
(* this "symmetric" definition works better than {{a}, {a,b}} *)
definition Pair :: "[i, i] => i"
- where "Pair(a,b) == {{a,a}, {a,b}}"
+ where "Pair(a,b) \<equiv> {{a,a}, {a,b}}"
definition fst :: "i \<Rightarrow> i"
- where "fst(p) == THE a. \<exists>b. p = Pair(a, b)"
+ where "fst(p) \<equiv> THE a. \<exists>b. p = Pair(a, b)"
definition snd :: "i \<Rightarrow> i"
- where "snd(p) == THE b. \<exists>a. p = Pair(a, b)"
+ where "snd(p) \<equiv> THE b. \<exists>a. p = Pair(a, b)"
definition split :: "[[i, i] \<Rightarrow> 'a, i] \<Rightarrow> 'a::{}" \<comment> \<open>for pattern-matching\<close>
- where "split(c) == \<lambda>p. c(fst(p), snd(p))"
+ where "split(c) \<equiv> \<lambda>p. c(fst(p), snd(p))"
(* Patterns -- extends pre-defined type "pttrn" used in abstractions *)
nonterminal patterns
@@ -193,7 +193,7 @@
"\<lambda>\<langle>x,y\<rangle>.b" == "CONST split(\<lambda>x y. b)"
definition Sigma :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
- where "Sigma(A,B) == \<Union>x\<in>A. \<Union>y\<in>B(x). {\<langle>x,y\<rangle>}"
+ where "Sigma(A,B) \<equiv> \<Union>x\<in>A. \<Union>y\<in>B(x). {\<langle>x,y\<rangle>}"
abbreviation cart_prod :: "[i, i] => i" (infixr \<open>\<times>\<close> 80) \<comment> \<open>Cartesian product\<close>
where "A \<times> B \<equiv> Sigma(A, \<lambda>_. B)"
@@ -203,44 +203,44 @@
(*converse of relation r, inverse of function*)
definition converse :: "i \<Rightarrow> i"
- where "converse(r) == {z. w\<in>r, \<exists>x y. w=\<langle>x,y\<rangle> \<and> z=\<langle>y,x\<rangle>}"
+ where "converse(r) \<equiv> {z. w\<in>r, \<exists>x y. w=\<langle>x,y\<rangle> \<and> z=\<langle>y,x\<rangle>}"
definition domain :: "i \<Rightarrow> i"
- where "domain(r) == {x. w\<in>r, \<exists>y. w=\<langle>x,y\<rangle>}"
+ where "domain(r) \<equiv> {x. w\<in>r, \<exists>y. w=\<langle>x,y\<rangle>}"
definition range :: "i \<Rightarrow> i"
- where "range(r) == domain(converse(r))"
+ where "range(r) \<equiv> domain(converse(r))"
definition field :: "i \<Rightarrow> i"
- where "field(r) == domain(r) \<union> range(r)"
+ where "field(r) \<equiv> domain(r) \<union> range(r)"
definition relation :: "i \<Rightarrow> o" \<comment> \<open>recognizes sets of pairs\<close>
- where "relation(r) == \<forall>z\<in>r. \<exists>x y. z = \<langle>x,y\<rangle>"
+ where "relation(r) \<equiv> \<forall>z\<in>r. \<exists>x y. z = \<langle>x,y\<rangle>"
definition "function" :: "i \<Rightarrow> o" \<comment> \<open>recognizes functions; can have non-pairs\<close>
- where "function(r) == \<forall>x y. \<langle>x,y\<rangle> \<in> r \<longrightarrow> (\<forall>y'. \<langle>x,y'\<rangle> \<in> r \<longrightarrow> y = y')"
+ where "function(r) \<equiv> \<forall>x y. \<langle>x,y\<rangle> \<in> r \<longrightarrow> (\<forall>y'. \<langle>x,y'\<rangle> \<in> r \<longrightarrow> y = y')"
definition Image :: "[i, i] \<Rightarrow> i" (infixl \<open>``\<close> 90) \<comment> \<open>image\<close>
- where image_def: "r `` A == {y \<in> range(r). \<exists>x\<in>A. \<langle>x,y\<rangle> \<in> r}"
+ where image_def: "r `` A \<equiv> {y \<in> range(r). \<exists>x\<in>A. \<langle>x,y\<rangle> \<in> r}"
definition vimage :: "[i, i] \<Rightarrow> i" (infixl \<open>-``\<close> 90) \<comment> \<open>inverse image\<close>
- where vimage_def: "r -`` A == converse(r)``A"
+ where vimage_def: "r -`` A \<equiv> converse(r)``A"
(* Restrict the relation r to the domain A *)
definition restrict :: "[i, i] \<Rightarrow> i"
- where "restrict(r,A) == {z \<in> r. \<exists>x\<in>A. \<exists>y. z = \<langle>x,y\<rangle>}"
+ where "restrict(r,A) \<equiv> {z \<in> r. \<exists>x\<in>A. \<exists>y. z = \<langle>x,y\<rangle>}"
(* Abstraction, application and Cartesian product of a family of sets *)
definition Lambda :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
- where lam_def: "Lambda(A,b) == {\<langle>x,b(x)\<rangle>. x\<in>A}"
+ where lam_def: "Lambda(A,b) \<equiv> {\<langle>x,b(x)\<rangle>. x\<in>A}"
definition "apply" :: "[i, i] \<Rightarrow> i" (infixl \<open>`\<close> 90) \<comment> \<open>function application\<close>
- where "f`a == \<Union>(f``{a})"
+ where "f`a \<equiv> \<Union>(f``{a})"
definition Pi :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
- where "Pi(A,B) == {f\<in>Pow(Sigma(A,B)). A\<subseteq>domain(f) & function(f)}"
+ where "Pi(A,B) \<equiv> {f\<in>Pow(Sigma(A,B)). A\<subseteq>domain(f) & function(f)}"
abbreviation function_space :: "[i, i] \<Rightarrow> i" (infixr \<open>\<rightarrow>\<close> 60) \<comment> \<open>function space\<close>
where "A \<rightarrow> B \<equiv> Pi(A, \<lambda>_. B)"
@@ -267,7 +267,7 @@
function_space (infixr \<open>->\<close> 60) and
Subset (infixl \<open><=\<close> 50) and
mem (infixl \<open>:\<close> 50) and
- not_mem (infixl \<open>~:\<close> 50)
+ not_mem (infixl \<open>\<not>:\<close> 50)
syntax (ASCII)
"_Ball" :: "[pttrn, i, o] => o" (\<open>(3ALL _:_./ _)\<close> 10)
@@ -287,30 +287,30 @@
subsection \<open>Substitution\<close>
(*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *)
-lemma subst_elem: "[| b\<in>A; a=b |] ==> a\<in>A"
+lemma subst_elem: "\<lbrakk>b\<in>A; a=b\<rbrakk> \<Longrightarrow> a\<in>A"
by (erule ssubst, assumption)
subsection\<open>Bounded universal quantifier\<close>
-lemma ballI [intro!]: "[| !!x. x\<in>A ==> P(x) |] ==> \<forall>x\<in>A. P(x)"
+lemma ballI [intro!]: "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> P(x)\<rbrakk> \<Longrightarrow> \<forall>x\<in>A. P(x)"
by (simp add: Ball_def)
lemmas strip = impI allI ballI
-lemma bspec [dest?]: "[| \<forall>x\<in>A. P(x); x: A |] ==> P(x)"
+lemma bspec [dest?]: "\<lbrakk>\<forall>x\<in>A. P(x); x: A\<rbrakk> \<Longrightarrow> P(x)"
by (simp add: Ball_def)
(*Instantiates x first: better for automatic theorem proving?*)
lemma rev_ballE [elim]:
- "[| \<forall>x\<in>A. P(x); x\<notin>A ==> Q; P(x) ==> Q |] ==> Q"
+ "\<lbrakk>\<forall>x\<in>A. P(x); x\<notin>A \<Longrightarrow> Q; P(x) \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: Ball_def, blast)
-lemma ballE: "[| \<forall>x\<in>A. P(x); P(x) ==> Q; x\<notin>A ==> Q |] ==> Q"
+lemma ballE: "\<lbrakk>\<forall>x\<in>A. P(x); P(x) \<Longrightarrow> Q; x\<notin>A \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by blast
(*Used in the datatype package*)
-lemma rev_bspec: "[| x: A; \<forall>x\<in>A. P(x) |] ==> P(x)"
+lemma rev_bspec: "\<lbrakk>x: A; \<forall>x\<in>A. P(x)\<rbrakk> \<Longrightarrow> P(x)"
by (simp add: Ball_def)
(*Trival rewrite rule; @{term"(\<forall>x\<in>A.P)<->P"} holds only if A is nonempty!*)
@@ -319,11 +319,11 @@
(*Congruence rule for rewriting*)
lemma ball_cong [cong]:
- "[| A=A'; !!x. x\<in>A' ==> P(x) <-> P'(x) |] ==> (\<forall>x\<in>A. P(x)) <-> (\<forall>x\<in>A'. P'(x))"
+ "\<lbrakk>A=A'; \<And>x. x\<in>A' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk> \<Longrightarrow> (\<forall>x\<in>A. P(x)) <-> (\<forall>x\<in>A'. P'(x))"
by (simp add: Ball_def)
lemma atomize_ball:
- "(!!x. x \<in> A ==> P(x)) == Trueprop (\<forall>x\<in>A. P(x))"
+ "(\<And>x. x \<in> A \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x\<in>A. P(x))"
by (simp only: Ball_def atomize_all atomize_imp)
lemmas [symmetric, rulify] = atomize_ball
@@ -332,27 +332,27 @@
subsection\<open>Bounded existential quantifier\<close>
-lemma bexI [intro]: "[| P(x); x: A |] ==> \<exists>x\<in>A. P(x)"
+lemma bexI [intro]: "\<lbrakk>P(x); x: A\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. P(x)"
by (simp add: Bex_def, blast)
(*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)"
+lemma rev_bexI: "\<lbrakk>x\<in>A; P(x)\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. P(x)"
by blast
(*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)"
+lemma bexCI: "\<lbrakk>\<forall>x\<in>A. \<not>P(x) \<Longrightarrow> P(a); a: A\<rbrakk> \<Longrightarrow> \<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"
+lemma bexE [elim!]: "\<lbrakk>\<exists>x\<in>A. P(x); \<And>x. \<lbrakk>x\<in>A; P(x)\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
by (simp add: Bex_def, blast)
-(*We do not even have @{term"(\<exists>x\<in>A. True) <-> True"} unless @{term"A" is nonempty!!*)
+(*We do not even have @{term"(\<exists>x\<in>A. True) <-> True"} unless @{term"A" is nonempty\<And>*)
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) |]
- ==> (\<exists>x\<in>A. P(x)) <-> (\<exists>x\<in>A'. P'(x))"
+ "\<lbrakk>A=A'; \<And>x. x\<in>A' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk>
+ \<Longrightarrow> (\<exists>x\<in>A. P(x)) <-> (\<exists>x\<in>A'. P'(x))"
by (simp add: Bex_def cong: conj_cong)
@@ -360,34 +360,34 @@
subsection\<open>Rules for subsets\<close>
lemma subsetI [intro!]:
- "(!!x. x\<in>A ==> x\<in>B) ==> A \<subseteq> B"
+ "(\<And>x. x\<in>A \<Longrightarrow> x\<in>B) \<Longrightarrow> A \<subseteq> B"
by (simp add: subset_def)
(*Rule in Modus Ponens style [was called subsetE] *)
-lemma subsetD [elim]: "[| A \<subseteq> B; c\<in>A |] ==> c\<in>B"
+lemma subsetD [elim]: "\<lbrakk>A \<subseteq> B; c\<in>A\<rbrakk> \<Longrightarrow> c\<in>B"
apply (unfold subset_def)
apply (erule bspec, assumption)
done
(*Classical elimination rule*)
lemma subsetCE [elim]:
- "[| A \<subseteq> B; c\<notin>A ==> P; c\<in>B ==> P |] ==> P"
+ "\<lbrakk>A \<subseteq> B; c\<notin>A \<Longrightarrow> P; c\<in>B \<Longrightarrow> P\<rbrakk> \<Longrightarrow> 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"
+lemma rev_subsetD: "\<lbrakk>c\<in>A; A<=B\<rbrakk> \<Longrightarrow> c\<in>B"
by blast
-lemma contra_subsetD: "[| A \<subseteq> B; c \<notin> B |] ==> c \<notin> A"
+lemma contra_subsetD: "\<lbrakk>A \<subseteq> B; c \<notin> B\<rbrakk> \<Longrightarrow> c \<notin> A"
by blast
-lemma rev_contra_subsetD: "[| c \<notin> B; A \<subseteq> B |] ==> c \<notin> A"
+lemma rev_contra_subsetD: "\<lbrakk>c \<notin> B; A \<subseteq> B\<rbrakk> \<Longrightarrow> c \<notin> A"
by blast
lemma subset_refl [simp]: "A \<subseteq> A"
by blast
-lemma subset_trans: "[| A<=B; B<=C |] ==> A<=C"
+lemma subset_trans: "\<lbrakk>A<=B; B<=C\<rbrakk> \<Longrightarrow> A<=C"
by blast
(*Useful for proving A<=B by rewriting in some cases*)
@@ -404,25 +404,25 @@
subsection\<open>Rules for equality\<close>
(*Anti-symmetry of the subset relation*)
-lemma equalityI [intro]: "[| A \<subseteq> B; B \<subseteq> A |] ==> A = B"
+lemma equalityI [intro]: "\<lbrakk>A \<subseteq> B; B \<subseteq> A\<rbrakk> \<Longrightarrow> A = B"
by (rule extension [THEN iffD2], rule conjI)
-lemma equality_iffI: "(!!x. x\<in>A <-> x\<in>B) ==> A = B"
+lemma equality_iffI: "(\<And>x. x\<in>A <-> x\<in>B) \<Longrightarrow> A = B"
by (rule equalityI, blast+)
lemmas equalityD1 = extension [THEN iffD1, THEN conjunct1]
lemmas equalityD2 = extension [THEN iffD1, THEN conjunct2]
-lemma equalityE: "[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"
+lemma equalityE: "\<lbrakk>A = B; \<lbrakk>A<=B; B<=A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (blast dest: equalityD1 equalityD2)
lemma equalityCE:
- "[| A = B; [| c\<in>A; c\<in>B |] ==> P; [| c\<notin>A; c\<notin>B |] ==> P |] ==> P"
+ "\<lbrakk>A = B; \<lbrakk>c\<in>A; c\<in>B\<rbrakk> \<Longrightarrow> P; \<lbrakk>c\<notin>A; c\<notin>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (erule equalityE, blast)
lemma equality_iffD:
- "A = B ==> (!!x. x \<in> A <-> x \<in> B)"
+ "A = B \<Longrightarrow> (\<And>x. x \<in> A <-> x \<in> B)"
by auto
@@ -436,26 +436,26 @@
(*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 |] ==>
+ "\<lbrakk>P(x,b); x: A; \<And>y. P(x,y) \<Longrightarrow> y=b\<rbrakk> \<Longrightarrow>
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 \<in> {y. x\<in>A, P(x,y)};
- !!x. [| x: A; P(x,b); \<forall>y. P(x,y)\<longrightarrow>y=b |] ==> R
- |] ==> R"
+ "\<lbrakk>b \<in> {y. x\<in>A, P(x,y)};
+ \<And>x. \<lbrakk>x: A; P(x,b); \<forall>y. P(x,y)\<longrightarrow>y=b\<rbrakk> \<Longrightarrow> R
+\<rbrakk> \<Longrightarrow> R"
by (rule Replace_iff [THEN iffD1, THEN bexE], simp+)
(*As above but without the (generally useless) 3rd assumption*)
lemma ReplaceE2 [elim!]:
- "[| b \<in> {y. x\<in>A, P(x,y)};
- !!x. [| x: A; P(x,b) |] ==> R
- |] ==> R"
+ "\<lbrakk>b \<in> {y. x\<in>A, P(x,y)};
+ \<And>x. \<lbrakk>x: A; P(x,b)\<rbrakk> \<Longrightarrow> R
+\<rbrakk> \<Longrightarrow> R"
by (erule ReplaceE, blast)
lemma Replace_cong [cong]:
- "[| A=B; !!x y. x\<in>B ==> P(x,y) <-> Q(x,y) |] ==>
+ "\<lbrakk>A=B; \<And>x y. x\<in>B \<Longrightarrow> P(x,y) <-> Q(x,y)\<rbrakk> \<Longrightarrow>
Replace(A,P) = Replace(B,Q)"
apply (rule equality_iffI)
apply (simp add: Replace_iff)
@@ -464,23 +464,23 @@
subsection\<open>Rules for RepFun\<close>
-lemma RepFunI: "a \<in> A ==> f(a) \<in> {f(x). x\<in>A}"
+lemma RepFunI: "a \<in> A \<Longrightarrow> 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 \<in> {f(x). x\<in>A}"
+lemma RepFun_eqI [intro]: "\<lbrakk>b=f(a); a \<in> A\<rbrakk> \<Longrightarrow> b \<in> {f(x). x\<in>A}"
apply (erule ssubst)
apply (erule RepFunI)
done
lemma RepFunE [elim!]:
- "[| b \<in> {f(x). x\<in>A};
- !!x.[| x\<in>A; b=f(x) |] ==> P |] ==>
+ "\<lbrakk>b \<in> {f(x). x\<in>A};
+ \<And>x.\<lbrakk>x\<in>A; b=f(x)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow>
P"
by (simp add: RepFun_def Replace_iff, blast)
lemma RepFun_cong [cong]:
- "[| A=B; !!x. x\<in>B ==> f(x)=g(x) |] ==> RepFun(A,f) = RepFun(B,g)"
+ "\<lbrakk>A=B; \<And>x. x\<in>B \<Longrightarrow> f(x)=g(x)\<rbrakk> \<Longrightarrow> RepFun(A,f) = RepFun(B,g)"
by (simp add: RepFun_def)
lemma RepFun_iff [simp]: "b \<in> {f(x). x\<in>A} <-> (\<exists>x\<in>A. b=f(x))"
@@ -496,21 +496,21 @@
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 \<in> {x\<in>A. P(x)}"
+lemma CollectI [intro!]: "\<lbrakk>a\<in>A; P(a)\<rbrakk> \<Longrightarrow> a \<in> {x\<in>A. P(x)}"
by simp
-lemma CollectE [elim!]: "[| a \<in> {x\<in>A. P(x)}; [| a\<in>A; P(a) |] ==> R |] ==> R"
+lemma CollectE [elim!]: "\<lbrakk>a \<in> {x\<in>A. P(x)}; \<lbrakk>a\<in>A; P(a)\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by simp
-lemma CollectD1: "a \<in> {x\<in>A. P(x)} ==> a\<in>A"
+lemma CollectD1: "a \<in> {x\<in>A. P(x)} \<Longrightarrow> a\<in>A"
by (erule CollectE, assumption)
-lemma CollectD2: "a \<in> {x\<in>A. P(x)} ==> P(a)"
+lemma CollectD2: "a \<in> {x\<in>A. P(x)} \<Longrightarrow> P(a)"
by (erule CollectE, assumption)
lemma Collect_cong [cong]:
- "[| A=B; !!x. x\<in>B ==> P(x) <-> Q(x) |]
- ==> Collect(A, %x. P(x)) = Collect(B, %x. Q(x))"
+ "\<lbrakk>A=B; \<And>x. x\<in>B \<Longrightarrow> P(x) <-> Q(x)\<rbrakk>
+ \<Longrightarrow> Collect(A, %x. P(x)) = Collect(B, %x. Q(x))"
by (simp add: Collect_def)
@@ -519,10 +519,10 @@
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]: "\<lbrakk>B: C; A: B\<rbrakk> \<Longrightarrow> A: \<Union>(C)"
by (simp, blast)
-lemma UnionE [elim!]: "[| A \<in> \<Union>(C); !!B.[| A: B; B: C |] ==> R |] ==> R"
+lemma UnionE [elim!]: "\<lbrakk>A \<in> \<Union>(C); \<And>B.\<lbrakk>A: B; B: C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by (simp, blast)
@@ -533,16 +533,16 @@
by (simp add: Bex_def, blast)
(*The order of the premises presupposes that A is rigid; b may be flexible*)
-lemma UN_I: "[| a: A; b: B(a) |] ==> b: (\<Union>x\<in>A. B(x))"
+lemma UN_I: "\<lbrakk>a: A; b: B(a)\<rbrakk> \<Longrightarrow> b: (\<Union>x\<in>A. B(x))"
by (simp, blast)
lemma UN_E [elim!]:
- "[| b \<in> (\<Union>x\<in>A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R"
+ "\<lbrakk>b \<in> (\<Union>x\<in>A. B(x)); \<And>x.\<lbrakk>x: A; b: B(x)\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by blast
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))"
+ "\<lbrakk>A=B; \<And>x. x\<in>B \<Longrightarrow> C(x)=D(x)\<rbrakk> \<Longrightarrow> (\<Union>x\<in>A. C(x)) = (\<Union>x\<in>B. D(x))"
by simp
@@ -568,18 +568,18 @@
lemma empty_subsetI [simp]: "0 \<subseteq> A"
by blast
-lemma equals0I: "[| !!y. y\<in>A ==> False |] ==> A=0"
+lemma equals0I: "\<lbrakk>\<And>y. y\<in>A \<Longrightarrow> False\<rbrakk> \<Longrightarrow> A=0"
by blast
-lemma equals0D [dest]: "A=0 ==> a \<notin> A"
+lemma equals0D [dest]: "A=0 \<Longrightarrow> a \<notin> A"
by blast
declare sym [THEN equals0D, dest]
-lemma not_emptyI: "a\<in>A ==> A \<noteq> 0"
+lemma not_emptyI: "a\<in>A \<Longrightarrow> A \<noteq> 0"
by blast
-lemma not_emptyE: "[| A \<noteq> 0; !!x. x\<in>A ==> R |] ==> R"
+lemma not_emptyE: "\<lbrakk>A \<noteq> 0; \<And>x. x\<in>A \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by blast
@@ -591,17 +591,17 @@
(* 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)"
+ "\<lbrakk>\<And>x. x: C \<Longrightarrow> A: x; C\<noteq>0\<rbrakk> \<Longrightarrow> 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]: "\<lbrakk>A \<in> \<Inter>(C); B \<in> C\<rbrakk> \<Longrightarrow> A \<in> B"
by (unfold 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"
+ "\<lbrakk>A \<in> \<Inter>(C); B\<notin>C \<Longrightarrow> R; A\<in>B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by (simp add: Inter_def, blast)
@@ -612,14 +612,14 @@
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))"
+lemma INT_I: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b: B(x); A\<noteq>0\<rbrakk> \<Longrightarrow> b: (\<Inter>x\<in>A. B(x))"
by blast
-lemma INT_E: "[| b \<in> (\<Inter>x\<in>A. B(x)); a: A |] ==> b \<in> B(a)"
+lemma INT_E: "\<lbrakk>b \<in> (\<Inter>x\<in>A. B(x)); a: A\<rbrakk> \<Longrightarrow> 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))"
+ "\<lbrakk>A=B; \<And>x. x\<in>B \<Longrightarrow> C(x)=D(x)\<rbrakk> \<Longrightarrow> (\<Inter>x\<in>A. C(x)) = (\<Inter>x\<in>B. D(x))"
by simp
(*No "Addcongs [INT_cong]" because @{term\<Inter>} is a combination of constants*)
@@ -627,10 +627,10 @@
subsection\<open>Rules for Powersets\<close>
-lemma PowI: "A \<subseteq> B ==> A \<in> Pow(B)"
+lemma PowI: "A \<subseteq> B \<Longrightarrow> A \<in> Pow(B)"
by (erule Pow_iff [THEN iffD2])
-lemma PowD: "A \<in> Pow(B) ==> A<=B"
+lemma PowD: "A \<in> Pow(B) \<Longrightarrow> A<=B"
by (erule Pow_iff [THEN iffD1])
declare Pow_iff [iff]
--- a/src/ZF/Zorn.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/Zorn.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,27 +12,27 @@
definition
Subset_rel :: "i=>i" where
- "Subset_rel(A) == {z \<in> A*A . \<exists>x y. z=<x,y> & x<=y & x\<noteq>y}"
+ "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=<x,y> & x<=y & x\<noteq>y}"
definition
chain :: "i=>i" where
- "chain(A) == {F \<in> Pow(A). \<forall>X\<in>F. \<forall>Y\<in>F. X<=Y | Y<=X}"
+ "chain(A) \<equiv> {F \<in> Pow(A). \<forall>X\<in>F. \<forall>Y\<in>F. X<=Y | Y<=X}"
definition
super :: "[i,i]=>i" where
- "super(A,c) == {d \<in> chain(A). c<=d & c\<noteq>d}"
+ "super(A,c) \<equiv> {d \<in> chain(A). c<=d & c\<noteq>d}"
definition
maxchain :: "i=>i" where
- "maxchain(A) == {c \<in> chain(A). super(A,c)=0}"
+ "maxchain(A) \<equiv> {c \<in> chain(A). super(A,c)=0}"
definition
increasing :: "i=>i" where
- "increasing(A) == {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A \<longrightarrow> x<=f`x}"
+ "increasing(A) \<equiv> {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A \<longrightarrow> x<=f`x}"
text\<open>Lemma for the inductive definition below\<close>
-lemma Union_in_Pow: "Y \<in> Pow(Pow(A)) ==> \<Union>(Y) \<in> Pow(A)"
+lemma Union_in_Pow: "Y \<in> Pow(Pow(A)) \<Longrightarrow> \<Union>(Y) \<in> Pow(A)"
by blast
@@ -47,10 +47,10 @@
inductive
domains "TFin(S,next)" \<subseteq> "Pow(S)"
intros
- nextI: "[| x \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> next`x \<in> TFin(S,next)"
+ nextI: "\<lbrakk>x \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk>
+ \<Longrightarrow> 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)) \<Longrightarrow> \<Union>(Y) \<in> TFin(S,next)"
monos Pow_mono
con_defs increasing_def
@@ -59,22 +59,22 @@
subsection\<open>Mathematical Preamble\<close>
-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) \<Longrightarrow> \<Union>(C)<=A | B<=\<Union>(C)"
by blast
lemma Inter_lemma0:
- "[| c \<in> C; \<forall>x\<in>C. A<=x | x<=B |] ==> A \<subseteq> \<Inter>(C) | \<Inter>(C) \<subseteq> B"
+ "\<lbrakk>c \<in> C; \<forall>x\<in>C. A<=x | x<=B\<rbrakk> \<Longrightarrow> A \<subseteq> \<Inter>(C) | \<Inter>(C) \<subseteq> B"
by blast
subsection\<open>The Transfinite Construction\<close>
-lemma increasingD1: "f \<in> increasing(A) ==> f \<in> Pow(A)->Pow(A)"
+lemma increasingD1: "f \<in> increasing(A) \<Longrightarrow> f \<in> Pow(A)->Pow(A)"
apply (unfold increasing_def)
apply (erule CollectD1)
done
-lemma increasingD2: "[| f \<in> increasing(A); x<=A |] ==> x \<subseteq> f`x"
+lemma increasingD2: "\<lbrakk>f \<in> increasing(A); x<=A\<rbrakk> \<Longrightarrow> x \<subseteq> f`x"
by (unfold increasing_def, blast)
lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI]
@@ -84,10 +84,10 @@
text\<open>Structural induction on \<^term>\<open>TFin(S,next)\<close>\<close>
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 \<subseteq> TFin(S,next); \<forall>y\<in>Y. P(y) |] ==> P(\<Union>(Y))
- |] ==> P(n)"
+ "\<lbrakk>n \<in> TFin(S,next);
+ \<And>x. \<lbrakk>x \<in> TFin(S,next); P(x); next \<in> increasing(S)\<rbrakk> \<Longrightarrow> P(next`x);
+ \<And>Y. \<lbrakk>Y \<subseteq> TFin(S,next); \<forall>y\<in>Y. P(y)\<rbrakk> \<Longrightarrow> P(\<Union>(Y))
+\<rbrakk> \<Longrightarrow> P(n)"
by (erule TFin.induct, blast+)
@@ -98,9 +98,9 @@
text\<open>Lemma 1 of section 3.1\<close>
lemma TFin_linear_lemma1:
- "[| n \<in> TFin(S,next); m \<in> TFin(S,next);
- \<forall>x \<in> TFin(S,next) . x<=m \<longrightarrow> x=m | next`x<=m |]
- ==> n<=m | next`m<=n"
+ "\<lbrakk>n \<in> TFin(S,next); m \<in> TFin(S,next);
+ \<forall>x \<in> TFin(S,next) . x<=m \<longrightarrow> x=m | next`x<=m\<rbrakk>
+ \<Longrightarrow> n<=m | next`m<=n"
apply (erule TFin_induct)
apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*)
(*downgrade subsetI from intro! to intro*)
@@ -110,8 +110,8 @@
text\<open>Lemma 2 of section 3.2. Interesting in its own right!
Requires \<^term>\<open>next \<in> increasing(S)\<close> in the second induction step.\<close>
lemma TFin_linear_lemma2:
- "[| m \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> \<forall>n \<in> TFin(S,next). n<=m \<longrightarrow> n=m | next`n \<subseteq> m"
+ "\<lbrakk>m \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk>
+ \<Longrightarrow> \<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\<open>case split using \<open>TFin_linear_lemma1\<close>\<close>
@@ -135,14 +135,14 @@
text\<open>a more convenient form for Lemma 2\<close>
lemma TFin_subsetD:
- "[| n<=m; m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> n=m | next`n \<subseteq> m"
+ "\<lbrakk>n<=m; m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk>
+ \<Longrightarrow> n=m | next`n \<subseteq> m"
by (blast dest: TFin_linear_lemma2 [rule_format])
text\<open>Consequences from section 3.3 -- Property 3.2, the ordering is total\<close>
lemma TFin_subset_linear:
- "[| m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> n \<subseteq> m | m<=n"
+ "\<lbrakk>m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk>
+ \<Longrightarrow> 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\<open>Lemma 3 of section 3.3\<close>
lemma equal_next_upper:
- "[| n \<in> TFin(S,next); m \<in> TFin(S,next); m = next`m |] ==> n \<subseteq> m"
+ "\<lbrakk>n \<in> TFin(S,next); m \<in> TFin(S,next); m = next`m\<rbrakk> \<Longrightarrow> n \<subseteq> m"
apply (erule TFin_induct)
apply (drule TFin_subsetD)
apply (assumption+, force, blast)
@@ -161,8 +161,8 @@
text\<open>Property 3.3 of section 3.3\<close>
lemma equal_next_Union:
- "[| m \<in> TFin(S,next); next \<in> increasing(S) |]
- ==> m = next`m <-> m = \<Union>(TFin(S,next))"
+ "\<lbrakk>m \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk>
+ \<Longrightarrow> 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])
@@ -197,15 +197,15 @@
done
lemma choice_super:
- "[| ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S); X \<notin> maxchain(S) |]
- ==> ch ` super(S,X) \<in> super(S,X)"
+ "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S); X \<notin> maxchain(S)\<rbrakk>
+ \<Longrightarrow> ch ` super(S,X) \<in> super(S,X)"
apply (erule apply_type)
apply (unfold super_def maxchain_def, blast)
done
lemma choice_not_equals:
- "[| ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S); X \<notin> maxchain(S) |]
- ==> ch ` super(S,X) \<noteq> X"
+ "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S); X \<notin> maxchain(S)\<rbrakk>
+ \<Longrightarrow> ch ` super(S,X) \<noteq> X"
apply (rule notI)
apply (drule choice_super, assumption, assumption)
apply (simp add: super_def)
@@ -213,7 +213,7 @@
text\<open>This justifies Definition 4.4\<close>
lemma Hausdorff_next_exists:
- "ch \<in> (\<Prod>X \<in> Pow(chain(S))-{0}. X) ==>
+ "ch \<in> (\<Prod>X \<in> Pow(chain(S))-{0}. X) \<Longrightarrow>
\<exists>next \<in> increasing(S). \<forall>X \<in> Pow(S).
next`X = if(X \<in> chain(S)-maxchain(S), ch`super(S,X), X)"
apply (rule_tac x="\<lambda>X\<in>Pow(S).
@@ -236,12 +236,12 @@
text\<open>Lemma 4\<close>
lemma TFin_chain_lemma4:
- "[| c \<in> TFin(S,next);
+ "\<lbrakk>c \<in> TFin(S,next);
ch \<in> (\<Prod>X \<in> Pow(chain(S))-{0}. X);
next \<in> increasing(S);
\<forall>X \<in> Pow(S). next`X =
- if(X \<in> chain(S)-maxchain(S), ch`super(S,X), X) |]
- ==> c \<in> chain(S)"
+ if(X \<in> chain(S)-maxchain(S), ch`super(S,X), X)\<rbrakk>
+ \<Longrightarrow> c \<in> chain(S)"
apply (erule TFin_induct)
apply (simp (no_asm_simp) add: chain_subset_Pow [THEN subsetD, THEN PowD]
choice_super [THEN super_subset_chain [THEN subsetD]])
@@ -277,10 +277,10 @@
text\<open>Used in the proof of Zorn's Lemma\<close>
lemma chain_extend:
- "[| c \<in> chain(A); z \<in> A; \<forall>x \<in> c. x<=z |] ==> cons(z,c) \<in> chain(A)"
+ "\<lbrakk>c \<in> chain(A); z \<in> A; \<forall>x \<in> c. x<=z\<rbrakk> \<Longrightarrow> 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 \<longrightarrow> y=z"
+lemma Zorn: "\<forall>c \<in> chain(S). \<Union>(c) \<in> S \<Longrightarrow> \<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)
@@ -299,7 +299,7 @@
text \<open>Alternative version of Zorn's Lemma\<close>
theorem Zorn2:
- "\<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"
+ "\<forall>c \<in> chain(S). \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y \<Longrightarrow> \<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,8 +321,8 @@
text\<open>Lemma 5\<close>
lemma TFin_well_lemma5:
- "[| n \<in> TFin(S,next); Z \<subseteq> TFin(S,next); z:Z; ~ \<Inter>(Z) \<in> Z |]
- ==> \<forall>m \<in> Z. n \<subseteq> m"
+ "\<lbrakk>n \<in> TFin(S,next); Z \<subseteq> TFin(S,next); z:Z; \<not> \<Inter>(Z) \<in> Z\<rbrakk>
+ \<Longrightarrow> \<forall>m \<in> Z. n \<subseteq> m"
apply (erule TFin_induct)
prefer 2 apply blast txt\<open>second induction step is easy\<close>
apply (rule ballI)
@@ -332,7 +332,7 @@
done
text\<open>Well-ordering of \<^term>\<open>TFin(S,next)\<close>\<close>
-lemma well_ord_TFin_lemma: "[| Z \<subseteq> TFin(S,next); z \<in> Z |] ==> \<Inter>(Z) \<in> Z"
+lemma well_ord_TFin_lemma: "\<lbrakk>Z \<subseteq> TFin(S,next); z \<in> Z\<rbrakk> \<Longrightarrow> \<Inter>(Z) \<in> Z"
apply (rule classical)
apply (subgoal_tac "Z = {\<Union>(TFin (S,next))}")
apply (simp (no_asm_simp) add: Inter_singleton)
@@ -344,7 +344,7 @@
text\<open>This theorem just packages the previous result\<close>
lemma well_ord_TFin:
"next \<in> increasing(S)
- ==> well_ord(TFin(S,next), Subset_rel(TFin(S,next)))"
+ \<Longrightarrow> well_ord(TFin(S,next), Subset_rel(TFin(S,next)))"
apply (rule well_ordI)
apply (unfold Subset_rel_def linear_def)
txt\<open>Prove the well-foundedness goal\<close>
@@ -364,14 +364,14 @@
text\<open>* Defining the "next" operation for Zermelo's Theorem *\<close>
lemma choice_Diff:
- "[| ch \<in> (\<Prod>X \<in> Pow(S) - {0}. X); X \<subseteq> S; X\<noteq>S |] ==> ch ` (S-X) \<in> S-X"
+ "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(S) - {0}. X); X \<subseteq> S; X\<noteq>S\<rbrakk> \<Longrightarrow> ch ` (S-X) \<in> S-X"
apply (erule apply_type)
apply (blast elim!: equalityE)
done
text\<open>This justifies Definition 6.1\<close>
lemma Zermelo_next_exists:
- "ch \<in> (\<Prod>X \<in> Pow(S)-{0}. X) ==>
+ "ch \<in> (\<Prod>X \<in> Pow(S)-{0}. X) \<Longrightarrow>
\<exists>next \<in> increasing(S). \<forall>X \<in> Pow(S).
next`X = (if X=S then S else cons(ch`(S-X), X))"
apply (rule_tac x="\<lambda>X\<in>Pow(S). if X=S then S else cons(ch`(S-X), X)"
@@ -391,10 +391,10 @@
text\<open>The construction of the injection\<close>
lemma choice_imp_injection:
- "[| ch \<in> (\<Prod>X \<in> Pow(S)-{0}. X);
+ "\<lbrakk>ch \<in> (\<Prod>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}))
+ \<forall>X \<in> Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X))\<rbrakk>
+ \<Longrightarrow> (\<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)
@@ -406,7 +406,7 @@
prefer 2 apply (blast elim: equalityE)
txt\<open>For proving \<open>x \<in> next`\<Union>(...)\<close>.
Abrial and Laffitte's justification appears to be faulty.\<close>
-apply (subgoal_tac "~ next ` Union({y \<in> TFin (S,next) . x \<notin> y})
+apply (subgoal_tac "\<not> 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
@@ -439,7 +439,7 @@
"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)"
+ "r \<subseteq> s \<Longrightarrow> Chain(r) \<subseteq> Chain(s)"
unfolding Chain_def
by blast
--- a/src/ZF/equalities.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/equalities.thy Tue Sep 27 16:51:35 2022 +0100
@@ -10,7 +10,7 @@
text\<open>These cover union, intersection, converse, domain, range, etc. Philippe
de Groote proved many of the inclusions.\<close>
-lemma in_mono: "A\<subseteq>B ==> x\<in>A \<longrightarrow> x\<in>B"
+lemma in_mono: "A\<subseteq>B \<Longrightarrow> x\<in>A \<longrightarrow> x\<in>B"
by blast
lemma the_eq_0 [simp]: "(THE x. False) = 0"
@@ -39,22 +39,22 @@
lemma converse_iff [simp]: "<a,b>\<in> converse(r) \<longleftrightarrow> <b,a>\<in>r"
by (unfold converse_def, blast)
-lemma converseI [intro!]: "<a,b>\<in>r ==> <b,a>\<in>converse(r)"
+lemma converseI [intro!]: "<a,b>\<in>r \<Longrightarrow> <b,a>\<in>converse(r)"
by (unfold converse_def, blast)
-lemma converseD: "<a,b> \<in> converse(r) ==> <b,a> \<in> r"
+lemma converseD: "<a,b> \<in> converse(r) \<Longrightarrow> <b,a> \<in> r"
by (unfold converse_def, blast)
lemma converseE [elim!]:
- "[| yx \<in> converse(r);
- !!x y. [| yx=<y,x>; <x,y>\<in>r |] ==> P |]
- ==> P"
+ "\<lbrakk>yx \<in> converse(r);
+ \<And>x y. \<lbrakk>yx=<y,x>; <x,y>\<in>r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+ \<Longrightarrow> P"
by (unfold converse_def, blast)
-lemma converse_converse: "r\<subseteq>Sigma(A,B) ==> converse(converse(r)) = r"
+lemma converse_converse: "r\<subseteq>Sigma(A,B) \<Longrightarrow> converse(converse(r)) = r"
by blast
-lemma converse_type: "r\<subseteq>A*B ==> converse(r)\<subseteq>B*A"
+lemma converse_type: "r\<subseteq>A*B \<Longrightarrow> converse(r)\<subseteq>B*A"
by blast
lemma converse_prod [simp]: "converse(A*B) = B*A"
@@ -64,13 +64,13 @@
by blast
lemma converse_subset_iff:
- "A \<subseteq> Sigma(X,Y) ==> converse(A) \<subseteq> converse(B) \<longleftrightarrow> A \<subseteq> B"
+ "A \<subseteq> Sigma(X,Y) \<Longrightarrow> converse(A) \<subseteq> converse(B) \<longleftrightarrow> A \<subseteq> B"
by blast
subsection\<open>Finite Set Constructions Using \<^term>\<open>cons\<close>\<close>
-lemma cons_subsetI: "[| a\<in>C; B\<subseteq>C |] ==> cons(a,B) \<subseteq> C"
+lemma cons_subsetI: "\<lbrakk>a\<in>C; B\<subseteq>C\<rbrakk> \<Longrightarrow> cons(a,B) \<subseteq> C"
by blast
lemma subset_consI: "B \<subseteq> cons(a,B)"
@@ -80,7 +80,7 @@
by blast
(*A safe special case of subset elimination, adding no new variables
- [| cons(a,B) \<subseteq> C; [| a \<in> C; B \<subseteq> C |] ==> R |] ==> R *)
+ \<lbrakk>cons(a,B) \<subseteq> C; \<lbrakk>a \<in> C; B \<subseteq> C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R *)
lemmas cons_subsetE = cons_subset_iff [THEN iffD1, THEN conjE]
lemma subset_empty_iff: "A\<subseteq>0 \<longleftrightarrow> A=0"
@@ -96,16 +96,16 @@
lemma cons_commute: "cons(a, cons(b, C)) = cons(b, cons(a, C))"
by blast
-lemma cons_absorb: "a: B ==> cons(a,B) = B"
+lemma cons_absorb: "a: B \<Longrightarrow> cons(a,B) = B"
by blast
-lemma cons_Diff: "a: B ==> cons(a, B-{a}) = B"
+lemma cons_Diff: "a: B \<Longrightarrow> 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))"
by auto
-lemma equal_singleton: "[| a: C; \<And>y. y \<in>C \<Longrightarrow> y=b |] ==> C = {b}"
+lemma equal_singleton: "\<lbrakk>a: C; \<And>y. y \<in>C \<Longrightarrow> y=b\<rbrakk> \<Longrightarrow> C = {b}"
by blast
lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)"
@@ -113,10 +113,10 @@
(** singletons **)
-lemma singleton_subsetI: "a\<in>C ==> {a} \<subseteq> C"
+lemma singleton_subsetI: "a\<in>C \<Longrightarrow> {a} \<subseteq> C"
by blast
-lemma singleton_subsetD: "{a} \<subseteq> C ==> a\<in>C"
+lemma singleton_subsetD: "{a} \<subseteq> C \<Longrightarrow> a\<in>C"
by blast
@@ -127,11 +127,11 @@
(*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"
+lemma succ_subsetI: "\<lbrakk>i\<in>j; i\<subseteq>j\<rbrakk> \<Longrightarrow> succ(i)\<subseteq>j"
by (unfold succ_def, blast)
lemma succ_subsetE:
- "[| succ(i) \<subseteq> j; [| i\<in>j; i\<subseteq>j |] ==> P |] ==> P"
+ "\<lbrakk>succ(i) \<subseteq> j; \<lbrakk>i\<in>j; i\<subseteq>j\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold succ_def, blast)
lemma succ_subset_iff: "succ(a) \<subseteq> B \<longleftrightarrow> (a \<subseteq> B & a \<in> B)"
@@ -151,7 +151,7 @@
lemma Int_lower2: "A \<inter> B \<subseteq> B"
by blast
-lemma Int_greatest: "[| C\<subseteq>A; C\<subseteq>B |] ==> C \<subseteq> A \<inter> B"
+lemma Int_greatest: "\<lbrakk>C\<subseteq>A; C\<subseteq>B\<rbrakk> \<Longrightarrow> C \<subseteq> A \<inter> B"
by blast
lemma Int_cons: "cons(a,B) \<inter> C \<subseteq> cons(a, B \<inter> C)"
@@ -175,10 +175,10 @@
(*Intersection is an AC-operator*)
lemmas Int_ac= Int_assoc Int_left_absorb Int_commute Int_left_commute
-lemma Int_absorb1: "B \<subseteq> A ==> A \<inter> B = B"
+lemma Int_absorb1: "B \<subseteq> A \<Longrightarrow> A \<inter> B = B"
by blast
-lemma Int_absorb2: "A \<subseteq> B ==> A \<inter> B = A"
+lemma Int_absorb2: "A \<subseteq> B \<Longrightarrow> A \<inter> B = A"
by blast
lemma Int_Un_distrib: "A \<inter> (B \<union> C) = (A \<inter> B) \<union> (A \<inter> C)"
@@ -193,7 +193,7 @@
lemma subset_Int_iff2: "A\<subseteq>B \<longleftrightarrow> B \<inter> A = A"
by (blast elim!: equalityE)
-lemma Int_Diff_eq: "C\<subseteq>A ==> (A-B) \<inter> C = C-B"
+lemma Int_Diff_eq: "C\<subseteq>A \<Longrightarrow> (A-B) \<inter> C = C-B"
by blast
lemma Int_cons_left:
@@ -220,7 +220,7 @@
lemma Un_upper2: "B \<subseteq> A \<union> B"
by blast
-lemma Un_least: "[| A\<subseteq>C; B\<subseteq>C |] ==> A \<union> B \<subseteq> C"
+lemma Un_least: "\<lbrakk>A\<subseteq>C; B\<subseteq>C\<rbrakk> \<Longrightarrow> A \<union> B \<subseteq> C"
by blast
lemma Un_cons: "cons(a,B) \<union> C = cons(a, B \<union> C)"
@@ -244,10 +244,10 @@
(*Union is an AC-operator*)
lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute
-lemma Un_absorb1: "A \<subseteq> B ==> A \<union> B = B"
+lemma Un_absorb1: "A \<subseteq> B \<Longrightarrow> A \<union> B = B"
by blast
-lemma Un_absorb2: "B \<subseteq> A ==> A \<union> B = A"
+lemma Un_absorb2: "B \<subseteq> A \<Longrightarrow> A \<union> B = A"
by blast
lemma Un_Int_distrib: "(A \<inter> B) \<union> C = (A \<union> C) \<inter> (B \<union> C)"
@@ -270,7 +270,7 @@
lemma Diff_subset: "A-B \<subseteq> A"
by blast
-lemma Diff_contains: "[| C\<subseteq>A; C \<inter> B = 0 |] ==> C \<subseteq> A-B"
+lemma Diff_contains: "\<lbrakk>C\<subseteq>A; C \<inter> B = 0\<rbrakk> \<Longrightarrow> C \<subseteq> A-B"
by blast
lemma subset_Diff_cons_iff: "B \<subseteq> A - cons(c,C) \<longleftrightarrow> B\<subseteq>A-C & c \<notin> B"
@@ -279,7 +279,7 @@
lemma Diff_cancel: "A - A = 0"
by blast
-lemma Diff_triv: "A \<inter> B = 0 ==> A - B = A"
+lemma Diff_triv: "A \<inter> B = 0 \<Longrightarrow> A - B = A"
by blast
lemma empty_Diff [simp]: "0 - A = 0"
@@ -291,24 +291,24 @@
lemma Diff_eq_0_iff: "A - B = 0 \<longleftrightarrow> A \<subseteq> B"
by (blast elim: equalityE)
-(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
+(*NOT SUITABLE FOR REWRITING since {a} \<equiv> cons(a,0)*)
lemma Diff_cons: "A - cons(a,B) = A - B - {a}"
by blast
-(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
+(*NOT SUITABLE FOR REWRITING since {a} \<equiv> cons(a,0)*)
lemma Diff_cons2: "A - cons(a,B) = A - {a} - B"
by blast
lemma Diff_disjoint: "A \<inter> (B-A) = 0"
by blast
-lemma Diff_partition: "A\<subseteq>B ==> A \<union> (B-A) = B"
+lemma Diff_partition: "A\<subseteq>B \<Longrightarrow> A \<union> (B-A) = B"
by blast
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"
+lemma double_complement: "\<lbrakk>A\<subseteq>B; B\<subseteq>C\<rbrakk> \<Longrightarrow> B-(C-A) = A"
by blast
lemma double_complement_Un: "(A \<union> B) - (B-A) = A"
@@ -349,10 +349,10 @@
lemma Union_subset_iff: "\<Union>(A) \<subseteq> C \<longleftrightarrow> (\<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 \<Longrightarrow> B \<subseteq> \<Union>(A)"
by blast
-lemma Union_least: "[| !!x. x\<in>A ==> x\<subseteq>C |] ==> \<Union>(A) \<subseteq> C"
+lemma Union_least: "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> x\<subseteq>C\<rbrakk> \<Longrightarrow> \<Union>(A) \<subseteq> C"
by blast
lemma Union_cons [simp]: "\<Union>(cons(a,B)) = a \<union> \<Union>(B)"
@@ -375,33 +375,33 @@
(** Big Intersection is the greatest lower bound of a nonempty set **)
-lemma Inter_subset_iff: "A\<noteq>0 ==> C \<subseteq> \<Inter>(A) \<longleftrightarrow> (\<forall>x\<in>A. C \<subseteq> x)"
+lemma Inter_subset_iff: "A\<noteq>0 \<Longrightarrow> C \<subseteq> \<Inter>(A) \<longleftrightarrow> (\<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 \<Longrightarrow> \<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: "\<lbrakk>A\<noteq>0; \<And>x. x\<in>A \<Longrightarrow> C\<subseteq>x\<rbrakk> \<Longrightarrow> C \<subseteq> \<Inter>(A)"
by blast
(** Intersection of a family of sets **)
-lemma INT_lower: "x\<in>A ==> (\<Inter>x\<in>A. B(x)) \<subseteq> B(x)"
+lemma INT_lower: "x\<in>A \<Longrightarrow> (\<Inter>x\<in>A. B(x)) \<subseteq> B(x)"
by blast
-lemma INT_greatest: "[| A\<noteq>0; !!x. x\<in>A ==> C\<subseteq>B(x) |] ==> C \<subseteq> (\<Inter>x\<in>A. B(x))"
+lemma INT_greatest: "\<lbrakk>A\<noteq>0; \<And>x. x\<in>A \<Longrightarrow> C\<subseteq>B(x)\<rbrakk> \<Longrightarrow> C \<subseteq> (\<Inter>x\<in>A. B(x))"
by force
lemma Inter_0 [simp]: "\<Inter>(0) = 0"
by (unfold Inter_def, blast)
lemma Inter_Un_subset:
- "[| z\<in>A; z\<in>B |] ==> \<Inter>(A) \<union> \<Inter>(B) \<subseteq> \<Inter>(A \<inter> B)"
+ "\<lbrakk>z\<in>A; z\<in>B\<rbrakk> \<Longrightarrow> \<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 \<union> B) = \<Inter>(A) \<inter> \<Inter>(B)"
+ "\<lbrakk>A\<noteq>0; B\<noteq>0\<rbrakk> \<Longrightarrow> \<Inter>(A \<union> B) = \<Inter>(A) \<inter> \<Inter>(B)"
by blast
lemma Union_singleton: "\<Union>({b}) = b"
@@ -422,10 +422,10 @@
lemma UN_subset_iff: "(\<Union>x\<in>A. B(x)) \<subseteq> C \<longleftrightarrow> (\<forall>x\<in>A. B(x) \<subseteq> C)"
by blast
-lemma UN_upper: "x\<in>A ==> B(x) \<subseteq> (\<Union>x\<in>A. B(x))"
+lemma UN_upper: "x\<in>A \<Longrightarrow> B(x) \<subseteq> (\<Union>x\<in>A. B(x))"
by (erule RepFunI [THEN Union_upper])
-lemma UN_least: "[| !!x. x\<in>A ==> B(x)\<subseteq>C |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> C"
+lemma UN_least: "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> B(x)\<subseteq>C\<rbrakk> \<Longrightarrow> (\<Union>x\<in>A. B(x)) \<subseteq> C"
by blast
lemma Union_eq_UN: "\<Union>(A) = (\<Union>x\<in>A. x)"
@@ -456,14 +456,14 @@
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 \<union> (\<Inter>i\<in>I. A(i)) = (\<Inter>i\<in>I. B \<union> A(i))"
+lemma Un_INT_distrib: "I\<noteq>0 \<Longrightarrow> 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)) \<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 |] ==>
+lemma Un_INT_distrib2: "\<lbrakk>I\<noteq>0; J\<noteq>0\<rbrakk> \<Longrightarrow>
(\<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
@@ -480,7 +480,7 @@
by (auto simp add: Inter_def)
lemma INT_Union_eq:
- "0 \<notin> A ==> (\<Inter>x\<in> \<Union>(A). B(x)) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B(x))"
+ "0 \<notin> A \<Longrightarrow> (\<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)
@@ -488,7 +488,7 @@
lemma INT_UN_eq:
"(\<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))"
+ \<Longrightarrow> (\<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
@@ -502,7 +502,7 @@
by blast
lemma INT_Int_distrib:
- "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))"
+ "I\<noteq>0 \<Longrightarrow> (\<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:
@@ -511,10 +511,10 @@
(** Devlin, page 12, exercise 5: Complements **)
-lemma Diff_UN: "I\<noteq>0 ==> B - (\<Union>i\<in>I. A(i)) = (\<Inter>i\<in>I. B - A(i))"
+lemma Diff_UN: "I\<noteq>0 \<Longrightarrow> B - (\<Union>i\<in>I. A(i)) = (\<Inter>i\<in>I. B - A(i))"
by (blast elim!: not_emptyE)
-lemma Diff_INT: "I\<noteq>0 ==> B - (\<Inter>i\<in>I. A(i)) = (\<Union>i\<in>I. B - A(i))"
+lemma Diff_INT: "I\<noteq>0 \<Longrightarrow> B - (\<Inter>i\<in>I. A(i)) = (\<Union>i\<in>I. B - A(i))"
by (blast elim!: not_emptyE)
@@ -583,17 +583,17 @@
lemma domain_iff: "a: domain(r) \<longleftrightarrow> (\<exists>y. <a,y>\<in> r)"
by (unfold domain_def, blast)
-lemma domainI [intro]: "<a,b>\<in> r ==> a: domain(r)"
+lemma domainI [intro]: "<a,b>\<in> r \<Longrightarrow> a: domain(r)"
by (unfold domain_def, blast)
lemma domainE [elim!]:
- "[| a \<in> domain(r); !!y. <a,y>\<in> r ==> P |] ==> P"
+ "\<lbrakk>a \<in> domain(r); \<And>y. <a,y>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold domain_def, blast)
lemma domain_subset: "domain(Sigma(A,B)) \<subseteq> A"
by blast
-lemma domain_of_prod: "b\<in>B ==> domain(A*B) = A"
+lemma domain_of_prod: "b\<in>B \<Longrightarrow> domain(A*B) = A"
by blast
lemma domain_0 [simp]: "domain(0) = 0"
@@ -620,12 +620,12 @@
(** Range **)
-lemma rangeI [intro]: "<a,b>\<in> r ==> b \<in> range(r)"
+lemma rangeI [intro]: "<a,b>\<in> r \<Longrightarrow> b \<in> range(r)"
apply (unfold range_def)
apply (erule converseI [THEN domainI])
done
-lemma rangeE [elim!]: "[| b \<in> range(r); !!x. <x,b>\<in> r ==> P |] ==> P"
+lemma rangeE [elim!]: "\<lbrakk>b \<in> range(r); \<And>x. <x,b>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold range_def, blast)
lemma range_subset: "range(A*B) \<subseteq> B"
@@ -634,7 +634,7 @@
apply (rule domain_subset)
done
-lemma range_of_prod: "a\<in>A ==> range(A*B) = B"
+lemma range_of_prod: "a\<in>A \<Longrightarrow> range(A*B) = B"
by blast
lemma range_0 [simp]: "range(0) = 0"
@@ -661,21 +661,21 @@
(** Field **)
-lemma fieldI1: "<a,b>\<in> r ==> a \<in> field(r)"
+lemma fieldI1: "<a,b>\<in> r \<Longrightarrow> a \<in> field(r)"
by (unfold field_def, blast)
-lemma fieldI2: "<a,b>\<in> r ==> b \<in> field(r)"
+lemma fieldI2: "<a,b>\<in> r \<Longrightarrow> b \<in> field(r)"
by (unfold field_def, blast)
lemma fieldCI [intro]:
- "(~ <c,a>\<in>r ==> <a,b>\<in> r) ==> a \<in> field(r)"
+ "(\<not> <c,a>\<in>r \<Longrightarrow> <a,b>\<in> r) \<Longrightarrow> a \<in> field(r)"
apply (unfold field_def, blast)
done
lemma fieldE [elim!]:
- "[| a \<in> field(r);
- !!x. <a,x>\<in> r ==> P;
- !!x. <x,a>\<in> r ==> P |] ==> P"
+ "\<lbrakk>a \<in> field(r);
+ \<And>x. <a,x>\<in> r \<Longrightarrow> P;
+ \<And>x. <x,a>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold field_def, blast)
lemma field_subset: "field(A*B) \<subseteq> A \<union> B"
@@ -691,13 +691,13 @@
apply (rule Un_upper2)
done
-lemma domain_times_range: "r \<subseteq> Sigma(A,B) ==> r \<subseteq> domain(r)*range(r)"
+lemma domain_times_range: "r \<subseteq> Sigma(A,B) \<Longrightarrow> r \<subseteq> domain(r)*range(r)"
by blast
-lemma field_times_field: "r \<subseteq> Sigma(A,B) ==> r \<subseteq> field(r)*field(r)"
+lemma field_times_field: "r \<subseteq> Sigma(A,B) \<Longrightarrow> r \<subseteq> field(r)*field(r)"
by blast
-lemma relation_field_times_field: "relation(r) ==> r \<subseteq> field(r)*field(r)"
+lemma relation_field_times_field: "relation(r) \<Longrightarrow> r \<subseteq> field(r)*field(r)"
by (simp add: relation_def, blast)
lemma field_of_prod: "field(A*A) = A"
@@ -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. \<exists>A B. x \<subseteq> A*B) ==>
+lemma rel_Union: "(\<forall>x\<in>S. \<exists>A B. x \<subseteq> A*B) \<Longrightarrow>
\<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 \<union> s) \<subseteq> (A \<union> C) * (B \<union> D)"
+lemma rel_Un: "\<lbrakk>r \<subseteq> A*B; s \<subseteq> C*D\<rbrakk> \<Longrightarrow> (r \<union> s) \<subseteq> (A \<union> C) * (B \<union> D)"
by blast
-lemma domain_Diff_eq: "[| <a,c> \<in> r; c\<noteq>b |] ==> domain(r-{<a,b>}) = domain(r)"
+lemma domain_Diff_eq: "\<lbrakk><a,c> \<in> r; c\<noteq>b\<rbrakk> \<Longrightarrow> domain(r-{<a,b>}) = domain(r)"
by blast
-lemma range_Diff_eq: "[| <c,b> \<in> r; c\<noteq>a |] ==> range(r-{<a,b>}) = range(r)"
+lemma range_Diff_eq: "\<lbrakk><c,b> \<in> r; c\<noteq>a\<rbrakk> \<Longrightarrow> range(r-{<a,b>}) = range(r)"
by blast
@@ -745,14 +745,14 @@
lemma image_singleton_iff: "b \<in> r``{a} \<longleftrightarrow> <a,b>\<in>r"
by (rule image_iff [THEN iff_trans], blast)
-lemma imageI [intro]: "[| <a,b>\<in> r; a\<in>A |] ==> b \<in> r``A"
+lemma imageI [intro]: "\<lbrakk><a,b>\<in> r; a\<in>A\<rbrakk> \<Longrightarrow> b \<in> r``A"
by (unfold image_def, blast)
lemma imageE [elim!]:
- "[| b: r``A; !!x.[| <x,b>\<in> r; x\<in>A |] ==> P |] ==> P"
+ "\<lbrakk>b: r``A; \<And>x.\<lbrakk><x,b>\<in> r; x\<in>A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold image_def, blast)
-lemma image_subset: "r \<subseteq> A*B ==> r``C \<subseteq> B"
+lemma image_subset: "r \<subseteq> A*B \<Longrightarrow> r``C \<subseteq> B"
by blast
lemma image_0 [simp]: "r``0 = 0"
@@ -774,7 +774,7 @@
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 \<inter> A*A)``B = (r``B) \<inter> A"
+lemma image_Int_square: "B\<subseteq>A \<Longrightarrow> (r \<inter> A*A)``B = (r``B) \<inter> A"
by blast
@@ -798,15 +798,15 @@
lemma vimage_singleton_iff: "a \<in> r-``{b} \<longleftrightarrow> <a,b>\<in>r"
by (rule vimage_iff [THEN iff_trans], blast)
-lemma vimageI [intro]: "[| <a,b>\<in> r; b\<in>B |] ==> a \<in> r-``B"
+lemma vimageI [intro]: "\<lbrakk><a,b>\<in> r; b\<in>B\<rbrakk> \<Longrightarrow> a \<in> r-``B"
by (unfold vimage_def, blast)
lemma vimageE [elim!]:
- "[| a: r-``B; !!x.[| <a,x>\<in> r; x\<in>B |] ==> P |] ==> P"
+ "\<lbrakk>a: r-``B; \<And>x.\<lbrakk><a,x>\<in> r; x\<in>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
apply (unfold vimage_def, blast)
done
-lemma vimage_subset: "r \<subseteq> A*B ==> r-``C \<subseteq> A"
+lemma vimage_subset: "r \<subseteq> A*B \<Longrightarrow> r-``C \<subseteq> A"
apply (unfold vimage_def)
apply (erule converse_type [THEN image_subset])
done
@@ -825,19 +825,19 @@
by blast
lemma function_vimage_Int:
- "function(f) ==> f-``(A \<inter> B) = (f-``A) \<inter> (f-``B)"
+ "function(f) \<Longrightarrow> 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)"
+lemma function_vimage_Diff: "function(f) \<Longrightarrow> f-``(A-B) = (f-``A) - (f-``B)"
by (unfold function_def, blast)
-lemma function_image_vimage: "function(f) ==> f `` (f-`` A) \<subseteq> A"
+lemma function_image_vimage: "function(f) \<Longrightarrow> f `` (f-`` A) \<subseteq> A"
by (unfold function_def, blast)
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 \<inter> A*A)-``B = (r-``B) \<inter> A"
+lemma vimage_Int_square: "B\<subseteq>A \<Longrightarrow> (r \<inter> A*A)-``B = (r-``B) \<inter> A"
by blast
@@ -903,13 +903,13 @@
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)))"
+lemma Pow_INT_eq: "A\<noteq>0 \<Longrightarrow> Pow(\<Inter>x\<in>A. B(x)) = (\<Inter>x\<in>A. Pow(B(x)))"
by (blast elim!: not_emptyE)
subsection\<open>RepFun\<close>
-lemma RepFun_subset: "[| !!x. x\<in>A ==> f(x) \<in> B |] ==> {f(x). x\<in>A} \<subseteq> B"
+lemma RepFun_subset: "\<lbrakk>\<And>x. x\<in>A \<Longrightarrow> f(x) \<in> B\<rbrakk> \<Longrightarrow> {f(x). x\<in>A} \<subseteq> B"
by blast
lemma RepFun_eq_0_iff [simp]: "{f(x).x\<in>A}=0 \<longleftrightarrow> A=0"
--- a/src/ZF/ex/BinEx.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/BinEx.thy Tue Sep 27 16:51:35 2022 +0100
@@ -23,7 +23,7 @@
lemma "$- #65745 = #-65745"
by simp (*80 msec*)
-(* negation of ~54321 *)
+(* negation of \<not>54321 *)
lemma "$- #-54321 = #54321"
by simp (*90 msec*)
@@ -62,7 +62,7 @@
by simp
-(*** Quotient and remainder!! [they could be faster] ***)
+(*** Quotient and remainder\<And>[they could be faster] ***)
lemma "#23 zdiv #3 = #7"
by simp
--- a/src/ZF/ex/CoUnit.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/CoUnit.thy Tue Sep 27 16:51:35 2022 +0100
@@ -25,7 +25,7 @@
"counit" = Con ("x \<in> counit")
inductive_cases ConE: "Con(x) \<in> counit"
- \<comment> \<open>USELESS because folding on \<^term>\<open>Con(xa) == xa\<close> fails.\<close>
+ \<comment> \<open>USELESS because folding on \<^term>\<open>Con(xa) \<equiv> xa\<close> fails.\<close>
lemma Con_iff: "Con(x) = Con(y) \<longleftrightarrow> x = y"
\<comment> \<open>Proving freeness results.\<close>
@@ -73,7 +73,7 @@
done
lemma counit2_Int_Vset_subset [rule_format]:
- "Ord(i) ==> \<forall>x y. x \<in> counit2 \<longrightarrow> y \<in> counit2 \<longrightarrow> x \<inter> Vset(i) \<subseteq> y"
+ "Ord(i) \<Longrightarrow> \<forall>x y. x \<in> counit2 \<longrightarrow> y \<in> counit2 \<longrightarrow> x \<inter> Vset(i) \<subseteq> y"
\<comment> \<open>Lemma for proving finality.\<close>
apply (erule trans_induct)
apply (tactic "safe_tac (put_claset subset_cs \<^context>)")
@@ -85,7 +85,7 @@
addSEs [@{thm Ord_in_Ord}, @{thm Pair_inject}]) 1\<close>)
done
-lemma counit2_implies_equal: "[| x \<in> counit2; y \<in> counit2 |] ==> x = y"
+lemma counit2_implies_equal: "\<lbrakk>x \<in> counit2; y \<in> counit2\<rbrakk> \<Longrightarrow> x = y"
apply (rule equalityI)
apply (assumption | rule conjI counit2_Int_Vset_subset [THEN Int_Vset_subset])+
done
--- a/src/ZF/ex/Commutation.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Commutation.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,40 +9,40 @@
definition
square :: "[i, i, i, i] => o" where
- "square(r,s,t,u) ==
+ "square(r,s,t,u) \<equiv>
(\<forall>a b. <a,b> \<in> r \<longrightarrow> (\<forall>c. <a, c> \<in> s \<longrightarrow> (\<exists>x. <b,x> \<in> t & <c,x> \<in> u)))"
definition
commute :: "[i, i] => o" where
- "commute(r,s) == square(r,s,s,r)"
+ "commute(r,s) \<equiv> square(r,s,s,r)"
definition
diamond :: "i=>o" where
- "diamond(r) == commute(r, r)"
+ "diamond(r) \<equiv> commute(r, r)"
definition
strip :: "i=>o" where
- "strip(r) == commute(r^*, r)"
+ "strip(r) \<equiv> commute(r^*, r)"
definition
Church_Rosser :: "i => o" where
- "Church_Rosser(r) == (\<forall>x y. <x,y> \<in> (r \<union> converse(r))^* \<longrightarrow>
+ "Church_Rosser(r) \<equiv> (\<forall>x y. <x,y> \<in> (r \<union> converse(r))^* \<longrightarrow>
(\<exists>z. <x,z> \<in> r^* & <y,z> \<in> r^*))"
definition
confluent :: "i=>o" where
- "confluent(r) == diamond(r^*)"
+ "confluent(r) \<equiv> diamond(r^*)"
-lemma square_sym: "square(r,s,t,u) ==> square(s,r,u,t)"
+lemma square_sym: "square(r,s,t,u) \<Longrightarrow> square(s,r,u,t)"
unfolding square_def by blast
-lemma square_subset: "[| square(r,s,t,u); t \<subseteq> t' |] ==> square(r,s,t',u)"
+lemma square_subset: "\<lbrakk>square(r,s,t,u); t \<subseteq> t'\<rbrakk> \<Longrightarrow> square(r,s,t',u)"
unfolding square_def by blast
lemma square_rtrancl:
- "square(r,s,s,t) ==> field(s)<=field(t) ==> square(r^*,s,s,t^*)"
+ "square(r,s,s,t) \<Longrightarrow> field(s)<=field(t) \<Longrightarrow> square(r^*,s,s,t^*)"
apply (unfold square_def, clarify)
apply (erule rtrancl_induct)
apply (blast intro: rtrancl_refl)
@@ -51,18 +51,18 @@
(* A special case of square_rtrancl_on *)
lemma diamond_strip:
- "diamond(r) ==> strip(r)"
+ "diamond(r) \<Longrightarrow> strip(r)"
apply (unfold diamond_def commute_def strip_def)
apply (rule square_rtrancl, simp_all)
done
(*** commute ***)
-lemma commute_sym: "commute(r,s) ==> commute(s,r)"
+lemma commute_sym: "commute(r,s) \<Longrightarrow> commute(s,r)"
unfolding commute_def by (blast intro: square_sym)
lemma commute_rtrancl:
- "commute(r,s) ==> field(r)=field(s) ==> commute(r^*,s^*)"
+ "commute(r,s) \<Longrightarrow> field(r)=field(s) \<Longrightarrow> commute(r^*,s^*)"
apply (unfold commute_def)
apply (rule square_rtrancl)
apply (rule square_sym [THEN square_rtrancl, THEN square_sym])
@@ -70,31 +70,31 @@
done
-lemma confluentD: "confluent(r) ==> diamond(r^*)"
+lemma confluentD: "confluent(r) \<Longrightarrow> diamond(r^*)"
by (simp add: confluent_def)
-lemma strip_confluent: "strip(r) ==> confluent(r)"
+lemma strip_confluent: "strip(r) \<Longrightarrow> confluent(r)"
apply (unfold strip_def confluent_def diamond_def)
apply (drule commute_rtrancl)
apply (simp_all add: rtrancl_field)
done
-lemma commute_Un: "[| commute(r,t); commute(s,t) |] ==> commute(r \<union> s, t)"
+lemma commute_Un: "\<lbrakk>commute(r,t); commute(s,t)\<rbrakk> \<Longrightarrow> commute(r \<union> s, t)"
unfolding commute_def square_def by blast
lemma diamond_Un:
- "[| diamond(r); diamond(s); commute(r, s) |] ==> diamond(r \<union> s)"
+ "\<lbrakk>diamond(r); diamond(s); commute(r, s)\<rbrakk> \<Longrightarrow> diamond(r \<union> s)"
unfolding diamond_def by (blast intro: commute_Un commute_sym)
lemma diamond_confluent:
- "diamond(r) ==> confluent(r)"
+ "diamond(r) \<Longrightarrow> confluent(r)"
apply (unfold diamond_def confluent_def)
apply (erule commute_rtrancl, simp)
done
lemma confluent_Un:
- "[| confluent(r); confluent(s); commute(r^*, s^*);
- relation(r); relation(s) |] ==> confluent(r \<union> s)"
+ "\<lbrakk>confluent(r); confluent(s); commute(r^*, s^*);
+ relation(r); relation(s)\<rbrakk> \<Longrightarrow> confluent(r \<union> s)"
apply (unfold confluent_def)
apply (rule rtrancl_Un_rtrancl [THEN subst], auto)
apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD])
@@ -102,7 +102,7 @@
lemma diamond_to_confluence:
- "[| diamond(r); s \<subseteq> r; r<= s^* |] ==> confluent(s)"
+ "\<lbrakk>diamond(r); s \<subseteq> r; r<= s^*\<rbrakk> \<Longrightarrow> confluent(s)"
apply (drule rtrancl_subset [symmetric], assumption)
apply (simp_all add: confluent_def)
apply (blast intro: diamond_confluent [THEN confluentD])
@@ -112,7 +112,7 @@
(*** Church_Rosser ***)
lemma Church_Rosser1:
- "Church_Rosser(r) ==> confluent(r)"
+ "Church_Rosser(r) \<Longrightarrow> confluent(r)"
apply (unfold confluent_def Church_Rosser_def square_def
commute_def diamond_def, auto)
apply (drule converseI)
@@ -125,7 +125,7 @@
lemma Church_Rosser2:
- "confluent(r) ==> Church_Rosser(r)"
+ "confluent(r) \<Longrightarrow> Church_Rosser(r)"
apply (unfold confluent_def Church_Rosser_def square_def
commute_def diamond_def, auto)
apply (frule fieldI1)
--- a/src/ZF/ex/Group.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Group.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,23 +19,23 @@
definition
carrier :: "i => i" where
- "carrier(M) == fst(M)"
+ "carrier(M) \<equiv> fst(M)"
definition
mmult :: "[i, i, i] => i" (infixl \<open>\<cdot>\<index>\<close> 70) where
- "mmult(M,x,y) == fst(snd(M)) ` <x,y>"
+ "mmult(M,x,y) \<equiv> fst(snd(M)) ` <x,y>"
definition
one :: "i => i" (\<open>\<one>\<index>\<close>) where
- "one(M) == fst(snd(snd(M)))"
+ "one(M) \<equiv> fst(snd(snd(M)))"
definition
update_carrier :: "[i,i] => i" where
- "update_carrier(M,A) == <A,snd(M)>"
+ "update_carrier(M,A) \<equiv> <A,snd(M)>"
definition
m_inv :: "i => i => i" (\<open>inv\<index> _\<close> [81] 80) where
- "inv\<^bsub>G\<^esub> x == (THE y. y \<in> carrier(G) & y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> & x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
+ "inv\<^bsub>G\<^esub> x \<equiv> (THE y. y \<in> carrier(G) & y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> & x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
locale monoid = fixes G (structure)
assumes m_closed [intro, simp]:
@@ -133,7 +133,7 @@
with x xG show "x \<cdot> \<one> = x" by simp
qed
have inv_ex:
- "!!x. x \<in> carrier(G) ==> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> & x \<cdot> y = \<one>"
+ "\<And>x. x \<in> carrier(G) \<Longrightarrow> \<exists>y \<in> carrier(G). y \<cdot> x = \<one> & x \<cdot> y = \<one>"
proof -
fix x
assume x: "x \<in> carrier(G)"
@@ -295,7 +295,7 @@
monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp]
lemma subgroup_nonempty:
- "~ subgroup(0,G)"
+ "\<not> subgroup(0,G)"
by (blast dest: subgroup.one_closed)
@@ -303,7 +303,7 @@
definition
DirProdGroup :: "[i,i] => i" (infixr \<open>\<Otimes>\<close> 80) where
- "G \<Otimes> H == <carrier(G) \<times> carrier(H),
+ "G \<Otimes> H \<equiv> <carrier(G) \<times> carrier(H),
(\<lambda><<g,h>, <g', h'>>
\<in> (carrier(G) \<times> carrier(H)) \<times> (carrier(G) \<times> carrier(H)).
<g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>),
@@ -328,8 +328,8 @@
by (simp add: DirProdGroup_def)
lemma mult_DirProdGroup [simp]:
- "[|g \<in> carrier(G); h \<in> carrier(H); g' \<in> carrier(G); h' \<in> carrier(H)|]
- ==> <g, h> \<cdot>\<^bsub>G \<Otimes> H\<^esub> <g', h'> = <g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>"
+ "\<lbrakk>g \<in> carrier(G); h \<in> carrier(H); g' \<in> carrier(G); h' \<in> carrier(H)\<rbrakk>
+ \<Longrightarrow> <g, h> \<cdot>\<^bsub>G \<Otimes> H\<^esub> <g', h'> = <g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>"
by (simp add: DirProdGroup_def)
lemma inv_DirProdGroup [simp]:
@@ -345,7 +345,7 @@
definition
hom :: "[i,i] => i" where
- "hom(G,H) ==
+ "hom(G,H) \<equiv>
{h \<in> carrier(G) -> carrier(H).
(\<forall>x \<in> carrier(G). \<forall>y \<in> carrier(G). h ` (x \<cdot>\<^bsub>G\<^esub> y) = (h ` x) \<cdot>\<^bsub>H\<^esub> (h ` y))}"
@@ -371,7 +371,7 @@
definition
iso :: "[i,i] => i" (infixr \<open>\<cong>\<close> 60) where
- "G \<cong> H == hom(G,H) \<inter> bij(carrier(G), carrier(H))"
+ "G \<cong> H \<equiv> hom(G,H) \<inter> bij(carrier(G), carrier(H))"
lemma (in group) iso_refl: "id(carrier(G)) \<in> G \<cong> G"
by (simp add: iso_def hom_def id_type id_bij)
@@ -490,8 +490,8 @@
lemma (in group) subgroupI:
assumes subset: "H \<subseteq> carrier(G)" and non_empty: "H \<noteq> 0"
- and "!!a. a \<in> H ==> inv a \<in> H"
- and "!!a b. [|a \<in> H; b \<in> H|] ==> a \<cdot> b \<in> H"
+ and "\<And>a. a \<in> H \<Longrightarrow> inv a \<in> H"
+ and "\<And>a b. \<lbrakk>a \<in> H; b \<in> H\<rbrakk> \<Longrightarrow> a \<cdot> b \<in> H"
shows "subgroup(H,G)"
proof (simp add: subgroup_def assms)
show "\<one> \<in> H"
@@ -503,7 +503,7 @@
definition
BijGroup :: "i=>i" where
- "BijGroup(S) ==
+ "BijGroup(S) \<equiv>
<bij(S,S),
\<lambda><g,f> \<in> bij(S,S) \<times> bij(S,S). g O f,
id(S), 0>"
@@ -532,17 +532,17 @@
left_comp_inverse [OF bij_is_inj])
done
-lemma iso_is_bij: "h \<in> G \<cong> H ==> h \<in> bij(carrier(G), carrier(H))"
+lemma iso_is_bij: "h \<in> G \<cong> H \<Longrightarrow> h \<in> bij(carrier(G), carrier(H))"
by (simp add: iso_def)
definition
auto :: "i=>i" where
- "auto(G) == iso(G,G)"
+ "auto(G) \<equiv> iso(G,G)"
definition
AutoGroup :: "i=>i" where
- "AutoGroup(G) == update_carrier(BijGroup(carrier(G)), auto(G))"
+ "AutoGroup(G) \<equiv> update_carrier(BijGroup(carrier(G)), auto(G))"
lemma (in group) id_in_auto: "id(carrier(G)) \<in> auto(G)"
@@ -577,23 +577,23 @@
definition
r_coset :: "[i,i,i] => i" (infixl \<open>#>\<index>\<close> 60) where
- "H #>\<^bsub>G\<^esub> a == \<Union>h\<in>H. {h \<cdot>\<^bsub>G\<^esub> a}"
+ "H #>\<^bsub>G\<^esub> a \<equiv> \<Union>h\<in>H. {h \<cdot>\<^bsub>G\<^esub> a}"
definition
l_coset :: "[i,i,i] => i" (infixl \<open><#\<index>\<close> 60) where
- "a <#\<^bsub>G\<^esub> H == \<Union>h\<in>H. {a \<cdot>\<^bsub>G\<^esub> h}"
+ "a <#\<^bsub>G\<^esub> H \<equiv> \<Union>h\<in>H. {a \<cdot>\<^bsub>G\<^esub> h}"
definition
RCOSETS :: "[i,i] => i" (\<open>rcosets\<index> _\<close> [81] 80) where
- "rcosets\<^bsub>G\<^esub> H == \<Union>a\<in>carrier(G). {H #>\<^bsub>G\<^esub> a}"
+ "rcosets\<^bsub>G\<^esub> H \<equiv> \<Union>a\<in>carrier(G). {H #>\<^bsub>G\<^esub> a}"
definition
set_mult :: "[i,i,i] => i" (infixl \<open><#>\<index>\<close> 60) where
- "H <#>\<^bsub>G\<^esub> K == \<Union>h\<in>H. \<Union>k\<in>K. {h \<cdot>\<^bsub>G\<^esub> k}"
+ "H <#>\<^bsub>G\<^esub> K \<equiv> \<Union>h\<in>H. \<Union>k\<in>K. {h \<cdot>\<^bsub>G\<^esub> k}"
definition
SET_INV :: "[i,i] => i" (\<open>set'_inv\<index> _\<close> [81] 80) where
- "set_inv\<^bsub>G\<^esub> H == \<Union>h\<in>H. {inv\<^bsub>G\<^esub> h}"
+ "set_inv\<^bsub>G\<^esub> H \<equiv> \<Union>h\<in>H. {inv\<^bsub>G\<^esub> h}"
locale normal = subgroup + group +
@@ -654,7 +654,7 @@
subsection \<open>Normal subgroups\<close>
-lemma normal_imp_subgroup: "H \<lhd> G ==> subgroup(H,G)"
+lemma normal_imp_subgroup: "H \<lhd> G \<Longrightarrow> subgroup(H,G)"
by (simp add: normal_def subgroup_def)
lemma (in group) normalI:
@@ -861,7 +861,7 @@
definition
r_congruent :: "[i,i] => i" (\<open>rcong\<index> _\<close> [60] 60) where
- "rcong\<^bsub>G\<^esub> H == {<x,y> \<in> carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \<cdot>\<^bsub>G\<^esub> y \<in> H}"
+ "rcong\<^bsub>G\<^esub> H \<equiv> {<x,y> \<in> carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \<cdot>\<^bsub>G\<^esub> y \<in> H}"
lemma (in subgroup) equiv_rcong:
@@ -941,7 +941,7 @@
definition
order :: "i => i" where
- "order(S) == |carrier(S)|"
+ "order(S) \<equiv> |carrier(S)|"
lemma (in group) rcos_self:
assumes "subgroup(H, G)"
@@ -1022,7 +1022,7 @@
definition
FactGroup :: "[i,i] => i" (infixl \<open>Mod\<close> 65) where
\<comment> \<open>Actually defined for groups rather than monoids\<close>
- "G Mod H ==
+ "G Mod H \<equiv>
<rcosets\<^bsub>G\<^esub> H, \<lambda><K1,K2> \<in> (rcosets\<^bsub>G\<^esub> H) \<times> (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>"
lemma (in normal) setmult_closed:
@@ -1086,7 +1086,7 @@
definition
kernel :: "[i,i,i] => i" where
\<comment> \<open>the kernel of a homomorphism\<close>
- "kernel(G,H,h) == {x \<in> carrier(G). h ` x = \<one>\<^bsub>H\<^esub>}"
+ "kernel(G,H,h) \<equiv> {x \<in> carrier(G). h ` x = \<one>\<^bsub>H\<^esub>}"
lemma (in group_hom) subgroup_kernel: "subgroup (kernel(G,H,h), G)"
apply (rule subgroup.intro)
@@ -1127,13 +1127,13 @@
qed
lemma mult_FactGroup:
- "[|X \<in> carrier(G Mod H); X' \<in> carrier(G Mod H)|]
- ==> X \<cdot>\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'"
+ "\<lbrakk>X \<in> carrier(G Mod H); X' \<in> carrier(G Mod H)\<rbrakk>
+ \<Longrightarrow> X \<cdot>\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'"
by (simp add: FactGroup_def)
lemma (in normal) FactGroup_m_closed:
- "[|X \<in> carrier(G Mod H); X' \<in> carrier(G Mod H)|]
- ==> X <#>\<^bsub>G\<^esub> X' \<in> carrier(G Mod H)"
+ "\<lbrakk>X \<in> carrier(G Mod H); X' \<in> carrier(G Mod H)\<rbrakk>
+ \<Longrightarrow> X <#>\<^bsub>G\<^esub> X' \<in> carrier(G Mod H)"
by (simp add: FactGroup_def setmult_closed)
lemma (in group_hom) FactGroup_hom:
--- a/src/ZF/ex/LList.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/LList.thy Tue Sep 27 16:51:35 2022 +0100
@@ -33,21 +33,21 @@
domains "lleq(A)" \<subseteq> "llist(A) * llist(A)"
intros
LNil: "<LNil, LNil> \<in> lleq(A)"
- LCons: "[| a \<in> A; <l,l'> \<in> lleq(A) |]
- ==> <LCons(a,l), LCons(a,l')> \<in> lleq(A)"
+ LCons: "\<lbrakk>a \<in> A; <l,l'> \<in> lleq(A)\<rbrakk>
+ \<Longrightarrow> <LCons(a,l), LCons(a,l')> \<in> lleq(A)"
type_intros llist.intros
(*Lazy list functions; flip is not definitional!*)
definition
lconst :: "i => i" where
- "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
+ "lconst(a) \<equiv> lfp(univ(a), %l. LCons(a,l))"
axiomatization flip :: "i => i"
where
flip_LNil: "flip(LNil) = LNil" and
- flip_LCons: "[| x \<in> bool; l \<in> llist(bool) |]
- ==> flip(LCons(x,l)) = LCons(not(x), flip(l))"
+ flip_LCons: "\<lbrakk>x \<in> bool; l \<in> llist(bool)\<rbrakk>
+ \<Longrightarrow> flip(LCons(x,l)) = LCons(not(x), flip(l))"
(*These commands cause classical reasoning to regard the subset relation
@@ -77,7 +77,7 @@
lemma LCons_iff: "LCons(a,l)=LCons(a',l') \<longleftrightarrow> a=a' & l=l'"
by auto
-lemma LNil_LCons_iff: "~ LNil=LCons(a,l)"
+lemma LNil_LCons_iff: "\<not> LNil=LCons(a,l)"
by auto
(*
@@ -90,7 +90,7 @@
(*** Lemmas to justify using "llist" in other recursive type definitions ***)
-lemma llist_mono: "A \<subseteq> B ==> llist(A) \<subseteq> llist(B)"
+lemma llist_mono: "A \<subseteq> B \<Longrightarrow> llist(A) \<subseteq> llist(B)"
apply (unfold llist.defs )
apply (rule gfp_mono)
apply (rule llist.bnd_mono)
@@ -107,7 +107,7 @@
declare Ord_in_Ord [elim!]
lemma llist_quniv_lemma:
- "Ord(i) ==> l \<in> llist(quniv(A)) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> univ(eclose(A))"
+ "Ord(i) \<Longrightarrow> l \<in> llist(quniv(A)) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> univ(eclose(A))"
proof (induct i arbitrary: l rule: trans_induct)
case (step i l)
show ?case using \<open>l \<in> llist(quniv(A))\<close>
@@ -142,7 +142,7 @@
(*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*)
lemma lleq_Int_Vset_subset:
- "Ord(i) ==> <l,l'> \<in> lleq(A) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> l'"
+ "Ord(i) \<Longrightarrow> <l,l'> \<in> lleq(A) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> l'"
proof (induct i arbitrary: l l' rule: trans_induct)
case (step i l l')
show ?case using \<open>\<langle>l, l'\<rangle> \<in> lleq(A)\<close>
@@ -156,20 +156,20 @@
qed
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
-lemma lleq_symmetric: "<l,l'> \<in> lleq(A) ==> <l',l> \<in> lleq(A)"
+lemma lleq_symmetric: "<l,l'> \<in> lleq(A) \<Longrightarrow> <l',l> \<in> lleq(A)"
apply (erule lleq.coinduct [OF converseI])
apply (rule lleq.dom_subset [THEN converse_type], safe)
apply (erule lleq.cases, blast+)
done
-lemma lleq_implies_equal: "<l,l'> \<in> lleq(A) ==> l=l'"
+lemma lleq_implies_equal: "<l,l'> \<in> lleq(A) \<Longrightarrow> l=l'"
apply (rule equalityI)
apply (assumption | rule lleq_Int_Vset_subset [THEN Int_Vset_subset] |
erule lleq_symmetric)+
done
lemma equal_llist_implies_leq:
- "[| l=l'; l \<in> llist(A) |] ==> <l,l'> \<in> lleq(A)"
+ "\<lbrakk>l=l'; l \<in> llist(A)\<rbrakk> \<Longrightarrow> <l,l'> \<in> lleq(A)"
apply (rule_tac X = "{<l,l>. l \<in> llist (A) }" in lleq.coinduct)
apply blast
apply safe
@@ -195,12 +195,12 @@
lemmas lconst_subset = lconst_def [THEN def_lfp_subset]
lemmas member_subset_Union_eclose = arg_into_eclose [THEN Union_upper]
-lemma lconst_in_quniv: "a \<in> A ==> lconst(a) \<in> quniv(A)"
+lemma lconst_in_quniv: "a \<in> A \<Longrightarrow> lconst(a) \<in> quniv(A)"
apply (rule lconst_subset [THEN subset_trans, THEN qunivI])
apply (erule arg_into_eclose [THEN eclose_subset, THEN univ_mono])
done
-lemma lconst_type: "a \<in> A ==> lconst(a): llist(A)"
+lemma lconst_type: "a \<in> A \<Longrightarrow> lconst(a): llist(A)"
apply (rule singletonI [THEN llist.coinduct])
apply (erule lconst_in_quniv [THEN singleton_subsetI])
apply (fast intro!: lconst)
@@ -212,7 +212,7 @@
flip_LCons [simp]
not_type [simp]
-lemma bool_Int_subset_univ: "b \<in> bool ==> b \<inter> X \<subseteq> univ(eclose(A))"
+lemma bool_Int_subset_univ: "b \<in> bool \<Longrightarrow> b \<inter> X \<subseteq> univ(eclose(A))"
by (fast intro: Int_lower1 [THEN subset_trans] elim!: boolE)
declare not_type [intro!]
@@ -221,7 +221,7 @@
(*Reasoning borrowed from lleq.ML; a similar proof works for all
"productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
lemma flip_llist_quniv_lemma:
- "Ord(i) ==> l \<in> llist(bool) \<Longrightarrow> flip(l) \<inter> Vset(i) \<subseteq> univ(eclose(bool))"
+ "Ord(i) \<Longrightarrow> l \<in> llist(bool) \<Longrightarrow> flip(l) \<inter> Vset(i) \<subseteq> univ(eclose(bool))"
proof (induct i arbitrary: l rule: trans_induct)
case (step i l)
show ?case using \<open>l \<in> llist(bool)\<close>
@@ -237,10 +237,10 @@
qed
qed
-lemma flip_in_quniv: "l \<in> llist(bool) ==> flip(l) \<in> quniv(bool)"
+lemma flip_in_quniv: "l \<in> llist(bool) \<Longrightarrow> flip(l) \<in> quniv(bool)"
by (rule flip_llist_quniv_lemma [THEN Int_Vset_subset, THEN qunivI], assumption+)
-lemma flip_type: "l \<in> llist(bool) ==> flip(l): llist(bool)"
+lemma flip_type: "l \<in> llist(bool) \<Longrightarrow> flip(l): llist(bool)"
apply (rule_tac X = "{flip (l) . l \<in> llist (bool) }" in llist.coinduct)
apply blast
apply (fast intro!: flip_in_quniv)
@@ -248,7 +248,7 @@
apply (erule_tac a=la in llist.cases, auto)
done
-lemma flip_flip: "l \<in> llist(bool) ==> flip(flip(l)) = l"
+lemma flip_flip: "l \<in> llist(bool) \<Longrightarrow> flip(flip(l)) = l"
apply (rule_tac X1 = "{<flip (flip (l)),l> . l \<in> llist (bool) }" in
lleq.coinduct [THEN lleq_implies_equal])
apply blast
--- a/src/ZF/ex/Limit.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Limit.thy Tue Sep 27 16:51:35 2022 +0100
@@ -21,15 +21,15 @@
definition
rel :: "[i,i,i]=>o" where
- "rel(D,x,y) == <x,y>:snd(D)"
+ "rel(D,x,y) \<equiv> <x,y>:snd(D)"
definition
set :: "i=>i" where
- "set(D) == fst(D)"
+ "set(D) \<equiv> fst(D)"
definition
po :: "i=>o" where
- "po(D) ==
+ "po(D) \<equiv>
(\<forall>x \<in> set(D). rel(D,x,x)) &
(\<forall>x \<in> set(D). \<forall>y \<in> set(D). \<forall>z \<in> set(D).
rel(D,x,y) \<longrightarrow> rel(D,y,z) \<longrightarrow> rel(D,x,z)) &
@@ -38,65 +38,65 @@
definition
chain :: "[i,i]=>o" where
(* Chains are object level functions nat->set(D) *)
- "chain(D,X) == X \<in> nat->set(D) & (\<forall>n \<in> nat. rel(D,X`n,X`(succ(n))))"
+ "chain(D,X) \<equiv> X \<in> nat->set(D) & (\<forall>n \<in> nat. rel(D,X`n,X`(succ(n))))"
definition
isub :: "[i,i,i]=>o" where
- "isub(D,X,x) == x \<in> set(D) & (\<forall>n \<in> nat. rel(D,X`n,x))"
+ "isub(D,X,x) \<equiv> x \<in> set(D) & (\<forall>n \<in> nat. rel(D,X`n,x))"
definition
islub :: "[i,i,i]=>o" where
- "islub(D,X,x) == isub(D,X,x) & (\<forall>y. isub(D,X,y) \<longrightarrow> rel(D,x,y))"
+ "islub(D,X,x) \<equiv> isub(D,X,x) & (\<forall>y. isub(D,X,y) \<longrightarrow> rel(D,x,y))"
definition
lub :: "[i,i]=>i" where
- "lub(D,X) == THE x. islub(D,X,x)"
+ "lub(D,X) \<equiv> THE x. islub(D,X,x)"
definition
cpo :: "i=>o" where
- "cpo(D) == po(D) & (\<forall>X. chain(D,X) \<longrightarrow> (\<exists>x. islub(D,X,x)))"
+ "cpo(D) \<equiv> po(D) & (\<forall>X. chain(D,X) \<longrightarrow> (\<exists>x. islub(D,X,x)))"
definition
pcpo :: "i=>o" where
- "pcpo(D) == cpo(D) & (\<exists>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y))"
+ "pcpo(D) \<equiv> cpo(D) & (\<exists>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y))"
definition
bot :: "i=>i" where
- "bot(D) == THE x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
+ "bot(D) \<equiv> THE x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
definition
mono :: "[i,i]=>i" where
- "mono(D,E) ==
+ "mono(D,E) \<equiv>
{f \<in> set(D)->set(E).
\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longrightarrow> rel(E,f`x,f`y)}"
definition
cont :: "[i,i]=>i" where
- "cont(D,E) ==
+ "cont(D,E) \<equiv>
{f \<in> mono(D,E).
\<forall>X. chain(D,X) \<longrightarrow> f`(lub(D,X)) = lub(E,\<lambda>n \<in> nat. f`(X`n))}"
definition
cf :: "[i,i]=>i" where
- "cf(D,E) ==
+ "cf(D,E) \<equiv>
<cont(D,E),
{y \<in> cont(D,E)*cont(D,E). \<forall>x \<in> set(D). rel(E,(fst(y))`x,(snd(y))`x)}>"
definition
suffix :: "[i,i]=>i" where
- "suffix(X,n) == \<lambda>m \<in> nat. X`(n #+ m)"
+ "suffix(X,n) \<equiv> \<lambda>m \<in> nat. X`(n #+ m)"
definition
subchain :: "[i,i]=>o" where
- "subchain(X,Y) == \<forall>m \<in> nat. \<exists>n \<in> nat. X`m = Y`(m #+ n)"
+ "subchain(X,Y) \<equiv> \<forall>m \<in> nat. \<exists>n \<in> nat. X`m = Y`(m #+ n)"
definition
dominate :: "[i,i,i]=>o" where
- "dominate(D,X,Y) == \<forall>m \<in> nat. \<exists>n \<in> nat. rel(D,X`m,Y`n)"
+ "dominate(D,X,Y) \<equiv> \<forall>m \<in> nat. \<exists>n \<in> nat. rel(D,X`m,Y`n)"
definition
matrix :: "[i,i]=>o" where
- "matrix(D,M) ==
+ "matrix(D,M) \<equiv>
M \<in> nat -> (nat -> set(D)) &
(\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`m)) &
(\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`n`succ(m))) &
@@ -104,22 +104,22 @@
definition
projpair :: "[i,i,i,i]=>o" where
- "projpair(D,E,e,p) ==
+ "projpair(D,E,e,p) \<equiv>
e \<in> cont(D,E) & p \<in> cont(E,D) &
p O e = id(set(D)) & rel(cf(E,E),e O p,id(set(E)))"
definition
emb :: "[i,i,i]=>o" where
- "emb(D,E,e) == \<exists>p. projpair(D,E,e,p)"
+ "emb(D,E,e) \<equiv> \<exists>p. projpair(D,E,e,p)"
definition
Rp :: "[i,i,i]=>i" where
- "Rp(D,E,e) == THE p. projpair(D,E,e,p)"
+ "Rp(D,E,e) \<equiv> THE p. projpair(D,E,e,p)"
definition
(* Twice, constructions on cpos are more difficult. *)
iprod :: "i=>i" where
- "iprod(DD) ==
+ "iprod(DD) \<equiv>
<(\<Prod>n \<in> nat. set(DD`n)),
{x:(\<Prod>n \<in> nat. set(DD`n))*(\<Prod>n \<in> nat. set(DD`n)).
\<forall>n \<in> nat. rel(DD`n,fst(x)`n,snd(x)`n)}>"
@@ -127,66 +127,66 @@
definition
mkcpo :: "[i,i=>o]=>i" where
(* Cannot use rel(D), is meta fun, need two more args *)
- "mkcpo(D,P) ==
+ "mkcpo(D,P) \<equiv>
<{x \<in> set(D). P(x)},{x \<in> set(D)*set(D). rel(D,fst(x),snd(x))}>"
definition
subcpo :: "[i,i]=>o" where
- "subcpo(D,E) ==
+ "subcpo(D,E) \<equiv>
set(D) \<subseteq> set(E) &
(\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longleftrightarrow> rel(E,x,y)) &
(\<forall>X. chain(D,X) \<longrightarrow> lub(E,X):set(D))"
definition
subpcpo :: "[i,i]=>o" where
- "subpcpo(D,E) == subcpo(D,E) & bot(E):set(D)"
+ "subpcpo(D,E) \<equiv> subcpo(D,E) & bot(E):set(D)"
definition
emb_chain :: "[i,i]=>o" where
- "emb_chain(DD,ee) ==
+ "emb_chain(DD,ee) \<equiv>
(\<forall>n \<in> nat. cpo(DD`n)) & (\<forall>n \<in> nat. emb(DD`n,DD`succ(n),ee`n))"
definition
Dinf :: "[i,i]=>i" where
- "Dinf(DD,ee) ==
+ "Dinf(DD,ee) \<equiv>
mkcpo(iprod(DD))
(%x. \<forall>n \<in> nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)"
definition
e_less :: "[i,i,i,i]=>i" where
(* Valid for m \<le> n only. *)
- "e_less(DD,ee,m,n) == rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)"
+ "e_less(DD,ee,m,n) \<equiv> rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)"
definition
e_gr :: "[i,i,i,i]=>i" where
(* Valid for n \<le> m only. *)
- "e_gr(DD,ee,m,n) ==
+ "e_gr(DD,ee,m,n) \<equiv>
rec(m#-n,id(set(DD`n)),
%x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))"
definition
eps :: "[i,i,i,i]=>i" where
- "eps(DD,ee,m,n) == if(m \<le> n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))"
+ "eps(DD,ee,m,n) \<equiv> if(m \<le> n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))"
definition
rho_emb :: "[i,i,i]=>i" where
- "rho_emb(DD,ee,n) == \<lambda>x \<in> set(DD`n). \<lambda>m \<in> nat. eps(DD,ee,n,m)`x"
+ "rho_emb(DD,ee,n) \<equiv> \<lambda>x \<in> set(DD`n). \<lambda>m \<in> nat. eps(DD,ee,n,m)`x"
definition
rho_proj :: "[i,i,i]=>i" where
- "rho_proj(DD,ee,n) == \<lambda>x \<in> set(Dinf(DD,ee)). x`n"
+ "rho_proj(DD,ee,n) \<equiv> \<lambda>x \<in> set(Dinf(DD,ee)). x`n"
definition
commute :: "[i,i,i,i=>i]=>o" where
- "commute(DD,ee,E,r) ==
+ "commute(DD,ee,E,r) \<equiv>
(\<forall>n \<in> nat. emb(DD`n,E,r(n))) &
(\<forall>m \<in> nat. \<forall>n \<in> nat. m \<le> n \<longrightarrow> r(n) O eps(DD,ee,m,n) = r(m))"
definition
mediating :: "[i,i,i=>i,i=>i,i]=>o" where
- "mediating(E,G,r,f,t) == emb(E,G,t) & (\<forall>n \<in> nat. f(n) = t O r(n))"
+ "mediating(E,G,r,f,t) \<equiv> emb(E,G,t) & (\<forall>n \<in> nat. f(n) = t O r(n))"
lemmas nat_linear_le = Ord_linear_le [OF nat_into_Ord nat_into_Ord]
@@ -196,102 +196,102 @@
(* Basic results. *)
(*----------------------------------------------------------------------*)
-lemma set_I: "x \<in> fst(D) ==> x \<in> set(D)"
+lemma set_I: "x \<in> fst(D) \<Longrightarrow> x \<in> set(D)"
by (simp add: set_def)
-lemma rel_I: "<x,y>:snd(D) ==> rel(D,x,y)"
+lemma rel_I: "<x,y>:snd(D) \<Longrightarrow> rel(D,x,y)"
by (simp add: rel_def)
-lemma rel_E: "rel(D,x,y) ==> <x,y>:snd(D)"
+lemma rel_E: "rel(D,x,y) \<Longrightarrow> <x,y>:snd(D)"
by (simp add: rel_def)
(*----------------------------------------------------------------------*)
(* I/E/D rules for po and cpo. *)
(*----------------------------------------------------------------------*)
-lemma po_refl: "[|po(D); x \<in> set(D)|] ==> rel(D,x,x)"
+lemma po_refl: "\<lbrakk>po(D); x \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,x)"
by (unfold po_def, blast)
-lemma po_trans: "[|po(D); rel(D,x,y); rel(D,y,z); x \<in> set(D);
- y \<in> set(D); z \<in> set(D)|] ==> rel(D,x,z)"
+lemma po_trans: "\<lbrakk>po(D); rel(D,x,y); rel(D,y,z); x \<in> set(D);
+ y \<in> set(D); z \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,z)"
by (unfold po_def, blast)
lemma po_antisym:
- "[|po(D); rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)|] ==> x = y"
+ "\<lbrakk>po(D); rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> x = y"
by (unfold po_def, blast)
lemma poI:
- "[| !!x. x \<in> set(D) ==> rel(D,x,x);
- !!x y z. [| rel(D,x,y); rel(D,y,z); x \<in> set(D); y \<in> set(D); z \<in> set(D)|]
- ==> rel(D,x,z);
- !!x y. [| rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)|] ==> x=y |]
- ==> po(D)"
+ "\<lbrakk>\<And>x. x \<in> set(D) \<Longrightarrow> rel(D,x,x);
+ \<And>x y z. \<lbrakk>rel(D,x,y); rel(D,y,z); x \<in> set(D); y \<in> set(D); z \<in> set(D)\<rbrakk>
+ \<Longrightarrow> rel(D,x,z);
+ \<And>x y. \<lbrakk>rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> x=y\<rbrakk>
+ \<Longrightarrow> po(D)"
by (unfold po_def, blast)
-lemma cpoI: "[| po(D); !!X. chain(D,X) ==> islub(D,X,x(D,X))|] ==> cpo(D)"
+lemma cpoI: "\<lbrakk>po(D); \<And>X. chain(D,X) \<Longrightarrow> islub(D,X,x(D,X))\<rbrakk> \<Longrightarrow> cpo(D)"
by (simp add: cpo_def, blast)
-lemma cpo_po: "cpo(D) ==> po(D)"
+lemma cpo_po: "cpo(D) \<Longrightarrow> po(D)"
by (simp add: cpo_def)
-lemma cpo_refl [simp,intro!,TC]: "[|cpo(D); x \<in> set(D)|] ==> rel(D,x,x)"
+lemma cpo_refl [simp,intro!,TC]: "\<lbrakk>cpo(D); x \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,x)"
by (blast intro: po_refl cpo_po)
lemma cpo_trans:
- "[|cpo(D); rel(D,x,y); rel(D,y,z); x \<in> set(D);
- y \<in> set(D); z \<in> set(D)|] ==> rel(D,x,z)"
+ "\<lbrakk>cpo(D); rel(D,x,y); rel(D,y,z); x \<in> set(D);
+ y \<in> set(D); z \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,z)"
by (blast intro: cpo_po po_trans)
lemma cpo_antisym:
- "[|cpo(D); rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)|] ==> x = y"
+ "\<lbrakk>cpo(D); rel(D,x,y); rel(D,y,x); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> x = y"
by (blast intro: cpo_po po_antisym)
-lemma cpo_islub: "[|cpo(D); chain(D,X); !!x. islub(D,X,x) ==> R|] ==> R"
+lemma cpo_islub: "\<lbrakk>cpo(D); chain(D,X); \<And>x. islub(D,X,x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by (simp add: cpo_def, blast)
(*----------------------------------------------------------------------*)
(* Theorems about isub and islub. *)
(*----------------------------------------------------------------------*)
-lemma islub_isub: "islub(D,X,x) ==> isub(D,X,x)"
+lemma islub_isub: "islub(D,X,x) \<Longrightarrow> isub(D,X,x)"
by (simp add: islub_def)
-lemma islub_in: "islub(D,X,x) ==> x \<in> set(D)"
+lemma islub_in: "islub(D,X,x) \<Longrightarrow> x \<in> set(D)"
by (simp add: islub_def isub_def)
-lemma islub_ub: "[|islub(D,X,x); n \<in> nat|] ==> rel(D,X`n,x)"
+lemma islub_ub: "\<lbrakk>islub(D,X,x); n \<in> nat\<rbrakk> \<Longrightarrow> rel(D,X`n,x)"
by (simp add: islub_def isub_def)
-lemma islub_least: "[|islub(D,X,x); isub(D,X,y)|] ==> rel(D,x,y)"
+lemma islub_least: "\<lbrakk>islub(D,X,x); isub(D,X,y)\<rbrakk> \<Longrightarrow> rel(D,x,y)"
by (simp add: islub_def)
lemma islubI:
- "[|isub(D,X,x); !!y. isub(D,X,y) ==> rel(D,x,y)|] ==> islub(D,X,x)"
+ "\<lbrakk>isub(D,X,x); \<And>y. isub(D,X,y) \<Longrightarrow> rel(D,x,y)\<rbrakk> \<Longrightarrow> islub(D,X,x)"
by (simp add: islub_def)
lemma isubI:
- "[|x \<in> set(D); !!n. n \<in> nat ==> rel(D,X`n,x)|] ==> isub(D,X,x)"
+ "\<lbrakk>x \<in> set(D); \<And>n. n \<in> nat \<Longrightarrow> rel(D,X`n,x)\<rbrakk> \<Longrightarrow> isub(D,X,x)"
by (simp add: isub_def)
lemma isubE:
- "[|isub(D,X,x); [|x \<in> set(D); !!n. n \<in> nat==>rel(D,X`n,x)|] ==> P
- |] ==> P"
+ "\<lbrakk>isub(D,X,x); \<lbrakk>x \<in> set(D); \<And>n. n \<in> nat\<Longrightarrow>rel(D,X`n,x)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: isub_def)
-lemma isubD1: "isub(D,X,x) ==> x \<in> set(D)"
+lemma isubD1: "isub(D,X,x) \<Longrightarrow> x \<in> set(D)"
by (simp add: isub_def)
-lemma isubD2: "[|isub(D,X,x); n \<in> nat|]==>rel(D,X`n,x)"
+lemma isubD2: "\<lbrakk>isub(D,X,x); n \<in> nat\<rbrakk>\<Longrightarrow>rel(D,X`n,x)"
by (simp add: isub_def)
-lemma islub_unique: "[|islub(D,X,x); islub(D,X,y); cpo(D)|] ==> x = y"
+lemma islub_unique: "\<lbrakk>islub(D,X,x); islub(D,X,y); cpo(D)\<rbrakk> \<Longrightarrow> x = y"
by (blast intro: cpo_antisym islub_least islub_isub islub_in)
(*----------------------------------------------------------------------*)
(* lub gives the least upper bound of chains. *)
(*----------------------------------------------------------------------*)
-lemma cpo_lub: "[|chain(D,X); cpo(D)|] ==> islub(D,X,lub(D,X))"
+lemma cpo_lub: "\<lbrakk>chain(D,X); cpo(D)\<rbrakk> \<Longrightarrow> islub(D,X,lub(D,X))"
apply (simp add: lub_def)
apply (best elim: cpo_islub intro: theI islub_unique)
done
@@ -301,29 +301,29 @@
(*----------------------------------------------------------------------*)
lemma chainI:
- "[|X \<in> nat->set(D); !!n. n \<in> nat ==> rel(D,X`n,X`succ(n))|] ==> chain(D,X)"
+ "\<lbrakk>X \<in> nat->set(D); \<And>n. n \<in> nat \<Longrightarrow> rel(D,X`n,X`succ(n))\<rbrakk> \<Longrightarrow> chain(D,X)"
by (simp add: chain_def)
-lemma chain_fun: "chain(D,X) ==> X \<in> nat -> set(D)"
+lemma chain_fun: "chain(D,X) \<Longrightarrow> X \<in> nat -> set(D)"
by (simp add: chain_def)
-lemma chain_in [simp,TC]: "[|chain(D,X); n \<in> nat|] ==> X`n \<in> set(D)"
+lemma chain_in [simp,TC]: "\<lbrakk>chain(D,X); n \<in> nat\<rbrakk> \<Longrightarrow> X`n \<in> set(D)"
apply (simp add: chain_def)
apply (blast dest: apply_type)
done
lemma chain_rel [simp,TC]:
- "[|chain(D,X); n \<in> nat|] ==> rel(D, X ` n, X ` succ(n))"
+ "\<lbrakk>chain(D,X); n \<in> nat\<rbrakk> \<Longrightarrow> rel(D, X ` n, X ` succ(n))"
by (simp add: chain_def)
lemma chain_rel_gen_add:
- "[|chain(D,X); cpo(D); n \<in> nat; m \<in> nat|] ==> rel(D,X`n,(X`(m #+ n)))"
+ "\<lbrakk>chain(D,X); cpo(D); n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> rel(D,X`n,(X`(m #+ n)))"
apply (induct_tac m)
apply (auto intro: cpo_trans)
done
lemma chain_rel_gen:
- "[|n \<le> m; chain(D,X); cpo(D); m \<in> nat|] ==> rel(D,X`n,X`m)"
+ "\<lbrakk>n \<le> m; chain(D,X); cpo(D); m \<in> nat\<rbrakk> \<Longrightarrow> rel(D,X`n,X`m)"
apply (frule lt_nat_in_nat, erule nat_succI)
apply (erule rev_mp) (*prepare the induction*)
apply (induct_tac m)
@@ -335,46 +335,46 @@
(*----------------------------------------------------------------------*)
lemma pcpoI:
- "[|!!y. y \<in> set(D)==>rel(D,x,y); x \<in> set(D); cpo(D)|]==>pcpo(D)"
+ "\<lbrakk>\<And>y. y \<in> set(D)\<Longrightarrow>rel(D,x,y); x \<in> set(D); cpo(D)\<rbrakk>\<Longrightarrow>pcpo(D)"
by (simp add: pcpo_def, auto)
-lemma pcpo_cpo [TC]: "pcpo(D) ==> cpo(D)"
+lemma pcpo_cpo [TC]: "pcpo(D) \<Longrightarrow> cpo(D)"
by (simp add: pcpo_def)
lemma pcpo_bot_ex1:
- "pcpo(D) ==> \<exists>! x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
+ "pcpo(D) \<Longrightarrow> \<exists>! x. x \<in> set(D) & (\<forall>y \<in> set(D). rel(D,x,y))"
apply (simp add: pcpo_def)
apply (blast intro: cpo_antisym)
done
lemma bot_least [TC]:
- "[| pcpo(D); y \<in> set(D)|] ==> rel(D,bot(D),y)"
+ "\<lbrakk>pcpo(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,bot(D),y)"
apply (simp add: bot_def)
apply (best intro: pcpo_bot_ex1 [THEN theI2])
done
lemma bot_in [TC]:
- "pcpo(D) ==> bot(D):set(D)"
+ "pcpo(D) \<Longrightarrow> bot(D):set(D)"
apply (simp add: bot_def)
apply (best intro: pcpo_bot_ex1 [THEN theI2])
done
lemma bot_unique:
- "[| pcpo(D); x \<in> set(D); !!y. y \<in> set(D) ==> rel(D,x,y)|] ==> x = bot(D)"
+ "\<lbrakk>pcpo(D); x \<in> set(D); \<And>y. y \<in> set(D) \<Longrightarrow> rel(D,x,y)\<rbrakk> \<Longrightarrow> x = bot(D)"
by (blast intro: cpo_antisym pcpo_cpo bot_in bot_least)
(*----------------------------------------------------------------------*)
(* Constant chains and lubs and cpos. *)
(*----------------------------------------------------------------------*)
-lemma chain_const: "[|x \<in> set(D); cpo(D)|] ==> chain(D,(\<lambda>n \<in> nat. x))"
+lemma chain_const: "\<lbrakk>x \<in> set(D); cpo(D)\<rbrakk> \<Longrightarrow> chain(D,(\<lambda>n \<in> nat. x))"
by (simp add: chain_def)
lemma islub_const:
- "[|x \<in> set(D); cpo(D)|] ==> islub(D,(\<lambda>n \<in> nat. x),x)"
+ "\<lbrakk>x \<in> set(D); cpo(D)\<rbrakk> \<Longrightarrow> islub(D,(\<lambda>n \<in> nat. x),x)"
by (simp add: islub_def isub_def, blast)
-lemma lub_const: "[|x \<in> set(D); cpo(D)|] ==> lub(D,\<lambda>n \<in> nat. x) = x"
+lemma lub_const: "\<lbrakk>x \<in> set(D); cpo(D)\<rbrakk> \<Longrightarrow> lub(D,\<lambda>n \<in> nat. x) = x"
by (blast intro: islub_unique cpo_lub chain_const islub_const)
(*----------------------------------------------------------------------*)
@@ -382,18 +382,18 @@
(*----------------------------------------------------------------------*)
lemma isub_suffix:
- "[| chain(D,X); cpo(D) |] ==> isub(D,suffix(X,n),x) \<longleftrightarrow> isub(D,X,x)"
+ "\<lbrakk>chain(D,X); cpo(D)\<rbrakk> \<Longrightarrow> isub(D,suffix(X,n),x) \<longleftrightarrow> isub(D,X,x)"
apply (simp add: isub_def suffix_def, safe)
apply (drule_tac x = na in bspec)
apply (auto intro: cpo_trans chain_rel_gen_add)
done
lemma islub_suffix:
- "[|chain(D,X); cpo(D)|] ==> islub(D,suffix(X,n),x) \<longleftrightarrow> islub(D,X,x)"
+ "\<lbrakk>chain(D,X); cpo(D)\<rbrakk> \<Longrightarrow> islub(D,suffix(X,n),x) \<longleftrightarrow> islub(D,X,x)"
by (simp add: islub_def isub_suffix)
lemma lub_suffix:
- "[|chain(D,X); cpo(D)|] ==> lub(D,suffix(X,n)) = lub(D,X)"
+ "\<lbrakk>chain(D,X); cpo(D)\<rbrakk> \<Longrightarrow> lub(D,suffix(X,n)) = lub(D,X)"
by (simp add: lub_def islub_suffix)
(*----------------------------------------------------------------------*)
@@ -401,31 +401,31 @@
(*----------------------------------------------------------------------*)
lemma dominateI:
- "[| !!m. m \<in> nat ==> n(m):nat; !!m. m \<in> nat ==> rel(D,X`m,Y`n(m))|]
- ==> dominate(D,X,Y)"
+ "\<lbrakk>\<And>m. m \<in> nat \<Longrightarrow> n(m):nat; \<And>m. m \<in> nat \<Longrightarrow> rel(D,X`m,Y`n(m))\<rbrakk>
+ \<Longrightarrow> dominate(D,X,Y)"
by (simp add: dominate_def, blast)
lemma dominate_isub:
- "[|dominate(D,X,Y); isub(D,Y,x); cpo(D);
- X \<in> nat->set(D); Y \<in> nat->set(D)|] ==> isub(D,X,x)"
+ "\<lbrakk>dominate(D,X,Y); isub(D,Y,x); cpo(D);
+ X \<in> nat->set(D); Y \<in> nat->set(D)\<rbrakk> \<Longrightarrow> isub(D,X,x)"
apply (simp add: isub_def dominate_def)
apply (blast intro: cpo_trans intro!: apply_funtype)
done
lemma dominate_islub:
- "[|dominate(D,X,Y); islub(D,X,x); islub(D,Y,y); cpo(D);
- X \<in> nat->set(D); Y \<in> nat->set(D)|] ==> rel(D,x,y)"
+ "\<lbrakk>dominate(D,X,Y); islub(D,X,x); islub(D,Y,y); cpo(D);
+ X \<in> nat->set(D); Y \<in> nat->set(D)\<rbrakk> \<Longrightarrow> rel(D,x,y)"
apply (simp add: islub_def)
apply (blast intro: dominate_isub)
done
lemma subchain_isub:
- "[|subchain(Y,X); isub(D,X,x)|] ==> isub(D,Y,x)"
+ "\<lbrakk>subchain(Y,X); isub(D,X,x)\<rbrakk> \<Longrightarrow> isub(D,Y,x)"
by (simp add: isub_def subchain_def, force)
lemma dominate_islub_eq:
- "[|dominate(D,X,Y); subchain(Y,X); islub(D,X,x); islub(D,Y,y); cpo(D);
- X \<in> nat->set(D); Y \<in> nat->set(D)|] ==> x = y"
+ "\<lbrakk>dominate(D,X,Y); subchain(Y,X); islub(D,X,x); islub(D,Y,y); cpo(D);
+ X \<in> nat->set(D); Y \<in> nat->set(D)\<rbrakk> \<Longrightarrow> x = y"
by (blast intro: cpo_antisym dominate_islub islub_least subchain_isub
islub_isub islub_in)
@@ -434,55 +434,55 @@
(* Matrix. *)
(*----------------------------------------------------------------------*)
-lemma matrix_fun: "matrix(D,M) ==> M \<in> nat -> (nat -> set(D))"
+lemma matrix_fun: "matrix(D,M) \<Longrightarrow> M \<in> nat -> (nat -> set(D))"
by (simp add: matrix_def)
-lemma matrix_in_fun: "[|matrix(D,M); n \<in> nat|] ==> M`n \<in> nat -> set(D)"
+lemma matrix_in_fun: "\<lbrakk>matrix(D,M); n \<in> nat\<rbrakk> \<Longrightarrow> M`n \<in> nat -> set(D)"
by (blast intro: apply_funtype matrix_fun)
-lemma matrix_in: "[|matrix(D,M); n \<in> nat; m \<in> nat|] ==> M`n`m \<in> set(D)"
+lemma matrix_in: "\<lbrakk>matrix(D,M); n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> M`n`m \<in> set(D)"
by (blast intro: apply_funtype matrix_in_fun)
lemma matrix_rel_1_0:
- "[|matrix(D,M); n \<in> nat; m \<in> nat|] ==> rel(D,M`n`m,M`succ(n)`m)"
+ "\<lbrakk>matrix(D,M); n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> rel(D,M`n`m,M`succ(n)`m)"
by (simp add: matrix_def)
lemma matrix_rel_0_1:
- "[|matrix(D,M); n \<in> nat; m \<in> nat|] ==> rel(D,M`n`m,M`n`succ(m))"
+ "\<lbrakk>matrix(D,M); n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> rel(D,M`n`m,M`n`succ(m))"
by (simp add: matrix_def)
lemma matrix_rel_1_1:
- "[|matrix(D,M); n \<in> nat; m \<in> nat|] ==> rel(D,M`n`m,M`succ(n)`succ(m))"
+ "\<lbrakk>matrix(D,M); n \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> rel(D,M`n`m,M`succ(n)`succ(m))"
by (simp add: matrix_def)
-lemma fun_swap: "f \<in> X->Y->Z ==> (\<lambda>y \<in> Y. \<lambda>x \<in> X. f`x`y):Y->X->Z"
+lemma fun_swap: "f \<in> X->Y->Z \<Longrightarrow> (\<lambda>y \<in> Y. \<lambda>x \<in> X. f`x`y):Y->X->Z"
by (blast intro: lam_type apply_funtype)
lemma matrix_sym_axis:
- "matrix(D,M) ==> matrix(D,\<lambda>m \<in> nat. \<lambda>n \<in> nat. M`n`m)"
+ "matrix(D,M) \<Longrightarrow> matrix(D,\<lambda>m \<in> nat. \<lambda>n \<in> nat. M`n`m)"
by (simp add: matrix_def fun_swap)
lemma matrix_chain_diag:
- "matrix(D,M) ==> chain(D,\<lambda>n \<in> nat. M`n`n)"
+ "matrix(D,M) \<Longrightarrow> chain(D,\<lambda>n \<in> nat. M`n`n)"
apply (simp add: chain_def)
apply (auto intro: lam_type matrix_in matrix_rel_1_1)
done
lemma matrix_chain_left:
- "[|matrix(D,M); n \<in> nat|] ==> chain(D,M`n)"
+ "\<lbrakk>matrix(D,M); n \<in> nat\<rbrakk> \<Longrightarrow> chain(D,M`n)"
apply (unfold chain_def)
apply (auto intro: matrix_fun [THEN apply_type] matrix_in matrix_rel_0_1)
done
lemma matrix_chain_right:
- "[|matrix(D,M); m \<in> nat|] ==> chain(D,\<lambda>n \<in> nat. M`n`m)"
+ "\<lbrakk>matrix(D,M); m \<in> nat\<rbrakk> \<Longrightarrow> chain(D,\<lambda>n \<in> nat. M`n`m)"
apply (simp add: chain_def)
apply (auto intro: lam_type matrix_in matrix_rel_1_0)
done
lemma matrix_chainI:
- assumes xprem: "!!x. x \<in> nat==>chain(D,M`x)"
- and yprem: "!!y. y \<in> nat==>chain(D,\<lambda>x \<in> nat. M`x`y)"
+ assumes xprem: "\<And>x. x \<in> nat\<Longrightarrow>chain(D,M`x)"
+ and yprem: "\<And>y. y \<in> nat\<Longrightarrow>chain(D,\<lambda>x \<in> nat. M`x`y)"
and Mfun: "M \<in> nat->nat->set(D)"
and cpoD: "cpo(D)"
shows "matrix(D,M)"
@@ -506,13 +506,13 @@
qed
lemma lemma2:
- "[|x \<in> nat; m \<in> nat; rel(D,(\<lambda>n \<in> nat. M`n`m1)`x,(\<lambda>n \<in> nat. M`n`m1)`m)|]
- ==> rel(D,M`x`m1,M`m`m1)"
+ "\<lbrakk>x \<in> nat; m \<in> nat; rel(D,(\<lambda>n \<in> nat. M`n`m1)`x,(\<lambda>n \<in> nat. M`n`m1)`m)\<rbrakk>
+ \<Longrightarrow> rel(D,M`x`m1,M`m`m1)"
by simp
lemma isub_lemma:
- "[|isub(D, \<lambda>n \<in> nat. M`n`n, y); matrix(D,M); cpo(D)|]
- ==> isub(D, \<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m), y)"
+ "\<lbrakk>isub(D, \<lambda>n \<in> nat. M`n`n, y); matrix(D,M); cpo(D)\<rbrakk>
+ \<Longrightarrow> isub(D, \<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m), y)"
proof (simp add: isub_def, safe)
fix n
assume DM: "matrix(D, M)" and D: "cpo(D)" and n: "n \<in> nat" and y: "y \<in> set(D)"
@@ -546,7 +546,7 @@
qed
lemma matrix_chain_lub:
- "[|matrix(D,M); cpo(D)|] ==> chain(D,\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))"
+ "\<lbrakk>matrix(D,M); cpo(D)\<rbrakk> \<Longrightarrow> chain(D,\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))"
proof (simp add: chain_def, intro conjI ballI)
assume "matrix(D, M)" "cpo(D)"
thus "(\<lambda>x\<in>nat. lub(D, Lambda(nat, (`)(M ` x)))) \<in> nat \<rightarrow> set(D)"
@@ -591,16 +591,16 @@
by (simp add: lub_def)
lemma lub_matrix_diag:
- "[|matrix(D,M); cpo(D)|]
- ==> lub(D,(\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))) =
+ "\<lbrakk>matrix(D,M); cpo(D)\<rbrakk>
+ \<Longrightarrow> lub(D,(\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))) =
lub(D,(\<lambda>n \<in> nat. M`n`n))"
apply (simp (no_asm) add: lub_matrix_diag_aux1 lub_matrix_diag_aux2)
apply (simp add: islub_def isub_eq)
done
lemma lub_matrix_diag_sym:
- "[|matrix(D,M); cpo(D)|]
- ==> lub(D,(\<lambda>m \<in> nat. lub(D,\<lambda>n \<in> nat. M`n`m))) =
+ "\<lbrakk>matrix(D,M); cpo(D)\<rbrakk>
+ \<Longrightarrow> lub(D,(\<lambda>m \<in> nat. lub(D,\<lambda>n \<in> nat. M`n`m))) =
lub(D,(\<lambda>n \<in> nat. M`n`n))"
by (drule matrix_sym_axis [THEN lub_matrix_diag], auto)
@@ -609,49 +609,49 @@
(*----------------------------------------------------------------------*)
lemma monoI:
- "[|f \<in> set(D)->set(E);
- !!x y. [|rel(D,x,y); x \<in> set(D); y \<in> set(D)|] ==> rel(E,f`x,f`y)|]
- ==> f \<in> mono(D,E)"
+ "\<lbrakk>f \<in> set(D)->set(E);
+ \<And>x y. \<lbrakk>rel(D,x,y); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(E,f`x,f`y)\<rbrakk>
+ \<Longrightarrow> f \<in> mono(D,E)"
by (simp add: mono_def)
-lemma mono_fun: "f \<in> mono(D,E) ==> f \<in> set(D)->set(E)"
+lemma mono_fun: "f \<in> mono(D,E) \<Longrightarrow> f \<in> set(D)->set(E)"
by (simp add: mono_def)
-lemma mono_map: "[|f \<in> mono(D,E); x \<in> set(D)|] ==> f`x \<in> set(E)"
+lemma mono_map: "\<lbrakk>f \<in> mono(D,E); x \<in> set(D)\<rbrakk> \<Longrightarrow> f`x \<in> set(E)"
by (blast intro!: mono_fun [THEN apply_type])
lemma mono_mono:
- "[|f \<in> mono(D,E); rel(D,x,y); x \<in> set(D); y \<in> set(D)|] ==> rel(E,f`x,f`y)"
+ "\<lbrakk>f \<in> mono(D,E); rel(D,x,y); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(E,f`x,f`y)"
by (simp add: mono_def)
lemma contI:
- "[|f \<in> set(D)->set(E);
- !!x y. [|rel(D,x,y); x \<in> set(D); y \<in> set(D)|] ==> rel(E,f`x,f`y);
- !!X. chain(D,X) ==> f`lub(D,X) = lub(E,\<lambda>n \<in> nat. f`(X`n))|]
- ==> f \<in> cont(D,E)"
+ "\<lbrakk>f \<in> set(D)->set(E);
+ \<And>x y. \<lbrakk>rel(D,x,y); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(E,f`x,f`y);
+ \<And>X. chain(D,X) \<Longrightarrow> f`lub(D,X) = lub(E,\<lambda>n \<in> nat. f`(X`n))\<rbrakk>
+ \<Longrightarrow> f \<in> cont(D,E)"
by (simp add: cont_def mono_def)
-lemma cont2mono: "f \<in> cont(D,E) ==> f \<in> mono(D,E)"
+lemma cont2mono: "f \<in> cont(D,E) \<Longrightarrow> f \<in> mono(D,E)"
by (simp add: cont_def)
-lemma cont_fun [TC]: "f \<in> cont(D,E) ==> f \<in> set(D)->set(E)"
+lemma cont_fun [TC]: "f \<in> cont(D,E) \<Longrightarrow> f \<in> set(D)->set(E)"
apply (simp add: cont_def)
apply (rule mono_fun, blast)
done
-lemma cont_map [TC]: "[|f \<in> cont(D,E); x \<in> set(D)|] ==> f`x \<in> set(E)"
+lemma cont_map [TC]: "\<lbrakk>f \<in> cont(D,E); x \<in> set(D)\<rbrakk> \<Longrightarrow> f`x \<in> set(E)"
by (blast intro!: cont_fun [THEN apply_type])
declare comp_fun [TC]
lemma cont_mono:
- "[|f \<in> cont(D,E); rel(D,x,y); x \<in> set(D); y \<in> set(D)|] ==> rel(E,f`x,f`y)"
+ "\<lbrakk>f \<in> cont(D,E); rel(D,x,y); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(E,f`x,f`y)"
apply (simp add: cont_def)
apply (blast intro!: mono_mono)
done
lemma cont_lub:
- "[|f \<in> cont(D,E); chain(D,X)|] ==> f`(lub(D,X)) = lub(E,\<lambda>n \<in> nat. f`(X`n))"
+ "\<lbrakk>f \<in> cont(D,E); chain(D,X)\<rbrakk> \<Longrightarrow> f`(lub(D,X)) = lub(E,\<lambda>n \<in> nat. f`(X`n))"
by (simp add: cont_def)
(*----------------------------------------------------------------------*)
@@ -659,13 +659,13 @@
(*----------------------------------------------------------------------*)
lemma mono_chain:
- "[|f \<in> mono(D,E); chain(D,X)|] ==> chain(E,\<lambda>n \<in> nat. f`(X`n))"
+ "\<lbrakk>f \<in> mono(D,E); chain(D,X)\<rbrakk> \<Longrightarrow> chain(E,\<lambda>n \<in> nat. f`(X`n))"
apply (simp (no_asm) add: chain_def)
apply (blast intro: lam_type mono_map chain_in mono_mono chain_rel)
done
lemma cont_chain:
- "[|f \<in> cont(D,E); chain(D,X)|] ==> chain(E,\<lambda>n \<in> nat. f`(X`n))"
+ "\<lbrakk>f \<in> cont(D,E); chain(D,X)\<rbrakk> \<Longrightarrow> chain(E,\<lambda>n \<in> nat. f`(X`n))"
by (blast intro: mono_chain cont2mono)
(*----------------------------------------------------------------------*)
@@ -674,21 +674,21 @@
(* The following development more difficult with cpo-as-relation approach. *)
-lemma cf_cont: "f \<in> set(cf(D,E)) ==> f \<in> cont(D,E)"
+lemma cf_cont: "f \<in> set(cf(D,E)) \<Longrightarrow> f \<in> cont(D,E)"
by (simp add: set_def cf_def)
lemma cont_cf: (* Non-trivial with relation *)
- "f \<in> cont(D,E) ==> f \<in> set(cf(D,E))"
+ "f \<in> cont(D,E) \<Longrightarrow> f \<in> set(cf(D,E))"
by (simp add: set_def cf_def)
(* rel_cf originally an equality. Now stated as two rules. Seemed easiest. *)
lemma rel_cfI:
- "[|!!x. x \<in> set(D) ==> rel(E,f`x,g`x); f \<in> cont(D,E); g \<in> cont(D,E)|]
- ==> rel(cf(D,E),f,g)"
+ "\<lbrakk>\<And>x. x \<in> set(D) \<Longrightarrow> rel(E,f`x,g`x); f \<in> cont(D,E); g \<in> cont(D,E)\<rbrakk>
+ \<Longrightarrow> rel(cf(D,E),f,g)"
by (simp add: rel_I cf_def)
-lemma rel_cf: "[|rel(cf(D,E),f,g); x \<in> set(D)|] ==> rel(E,f`x,g`x)"
+lemma rel_cf: "\<lbrakk>rel(cf(D,E),f,g); x \<in> set(D)\<rbrakk> \<Longrightarrow> rel(E,f`x,g`x)"
by (simp add: rel_def cf_def)
(*----------------------------------------------------------------------*)
@@ -696,15 +696,15 @@
(*----------------------------------------------------------------------*)
lemma chain_cf:
- "[| chain(cf(D,E),X); x \<in> set(D)|] ==> chain(E,\<lambda>n \<in> nat. X`n`x)"
+ "\<lbrakk>chain(cf(D,E),X); x \<in> set(D)\<rbrakk> \<Longrightarrow> chain(E,\<lambda>n \<in> nat. X`n`x)"
apply (rule chainI)
apply (blast intro: lam_type apply_funtype cont_fun cf_cont chain_in, simp)
apply (blast intro: rel_cf chain_rel)
done
lemma matrix_lemma:
- "[|chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E) |]
- ==> matrix(E,\<lambda>x \<in> nat. \<lambda>xa \<in> nat. X`x`(Xa`xa))"
+ "\<lbrakk>chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E)\<rbrakk>
+ \<Longrightarrow> matrix(E,\<lambda>x \<in> nat. \<lambda>xa \<in> nat. X`x`(Xa`xa))"
apply (rule matrix_chainI, auto)
apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont cont_mono)
apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont rel_cf)
@@ -745,8 +745,8 @@
qed
lemma islub_cf:
- "[| chain(cf(D,E),X); cpo(D); cpo(E)|]
- ==> islub(cf(D,E), X, \<lambda>x \<in> set(D). lub(E,\<lambda>n \<in> nat. X`n`x))"
+ "\<lbrakk>chain(cf(D,E),X); cpo(D); cpo(E)\<rbrakk>
+ \<Longrightarrow> islub(cf(D,E), X, \<lambda>x \<in> set(D). lub(E,\<lambda>n \<in> nat. X`n`x))"
apply (rule islubI)
apply (rule isubI)
apply (rule chain_cf_lub_cont [THEN cont_cf], assumption+)
@@ -762,7 +762,7 @@
done
lemma cpo_cf [TC]:
- "[| cpo(D); cpo(E)|] ==> cpo(cf(D,E))"
+ "\<lbrakk>cpo(D); cpo(E)\<rbrakk> \<Longrightarrow> cpo(cf(D,E))"
apply (rule poI [THEN cpoI])
apply (rule rel_cfI)
apply (assumption | rule cpo_refl cf_cont [THEN cont_fun, THEN apply_type]
@@ -780,13 +780,13 @@
done
lemma lub_cf:
- "[| chain(cf(D,E),X); cpo(D); cpo(E)|]
- ==> lub(cf(D,E), X) = (\<lambda>x \<in> set(D). lub(E,\<lambda>n \<in> nat. X`n`x))"
+ "\<lbrakk>chain(cf(D,E),X); cpo(D); cpo(E)\<rbrakk>
+ \<Longrightarrow> lub(cf(D,E), X) = (\<lambda>x \<in> set(D). lub(E,\<lambda>n \<in> nat. X`n`x))"
by (blast intro: islub_unique cpo_lub islub_cf cpo_cf)
lemma const_cont [TC]:
- "[|y \<in> set(E); cpo(D); cpo(E)|] ==> (\<lambda>x \<in> set(D).y) \<in> cont(D,E)"
+ "\<lbrakk>y \<in> set(E); cpo(D); cpo(E)\<rbrakk> \<Longrightarrow> (\<lambda>x \<in> set(D).y) \<in> cont(D,E)"
apply (rule contI)
prefer 2 apply simp
apply (blast intro: lam_type)
@@ -794,19 +794,19 @@
done
lemma cf_least:
- "[|cpo(D); pcpo(E); y \<in> cont(D,E)|]==>rel(cf(D,E),(\<lambda>x \<in> set(D).bot(E)),y)"
+ "\<lbrakk>cpo(D); pcpo(E); y \<in> cont(D,E)\<rbrakk>\<Longrightarrow>rel(cf(D,E),(\<lambda>x \<in> set(D).bot(E)),y)"
apply (rule rel_cfI, simp, typecheck)
done
lemma pcpo_cf:
- "[|cpo(D); pcpo(E)|] ==> pcpo(cf(D,E))"
+ "\<lbrakk>cpo(D); pcpo(E)\<rbrakk> \<Longrightarrow> pcpo(cf(D,E))"
apply (rule pcpoI)
apply (assumption |
rule cf_least bot_in const_cont [THEN cont_cf] cf_cont cpo_cf pcpo_cpo)+
done
lemma bot_cf:
- "[|cpo(D); pcpo(E)|] ==> bot(cf(D,E)) = (\<lambda>x \<in> set(D).bot(E))"
+ "\<lbrakk>cpo(D); pcpo(E)\<rbrakk> \<Longrightarrow> bot(cf(D,E)) = (\<lambda>x \<in> set(D).bot(E))"
by (blast intro: bot_unique [symmetric] pcpo_cf cf_least
bot_in [THEN const_cont, THEN cont_cf] cf_cont pcpo_cpo)
@@ -815,13 +815,13 @@
(*----------------------------------------------------------------------*)
lemma id_cont [TC,intro!]:
- "cpo(D) ==> id(set(D)) \<in> cont(D,D)"
+ "cpo(D) \<Longrightarrow> id(set(D)) \<in> cont(D,D)"
by (simp add: id_type contI cpo_lub [THEN islub_in] chain_fun [THEN eta])
lemmas comp_cont_apply = cont_fun [THEN comp_fun_apply]
lemma comp_pres_cont [TC]:
- "[| f \<in> cont(D',E); g \<in> cont(D,D'); cpo(D)|] ==> f O g \<in> cont(D,E)"
+ "\<lbrakk>f \<in> cont(D',E); g \<in> cont(D,D'); cpo(D)\<rbrakk> \<Longrightarrow> f O g \<in> cont(D,E)"
apply (rule contI)
apply (rule_tac [2] comp_cont_apply [THEN ssubst])
apply (rule_tac [4] comp_cont_apply [THEN ssubst])
@@ -837,9 +837,9 @@
lemma comp_mono:
- "[| f \<in> cont(D',E); g \<in> cont(D,D'); f':cont(D',E); g':cont(D,D');
- rel(cf(D',E),f,f'); rel(cf(D,D'),g,g'); cpo(D); cpo(E) |]
- ==> rel(cf(D,E),f O g,f' O g')"
+ "\<lbrakk>f \<in> cont(D',E); g \<in> cont(D,D'); f':cont(D',E); g':cont(D,D');
+ rel(cf(D',E),f,f'); rel(cf(D,D'),g,g'); cpo(D); cpo(E)\<rbrakk>
+ \<Longrightarrow> rel(cf(D,E),f O g,f' O g')"
apply (rule rel_cfI)
apply (subst comp_cont_apply)
apply (rule_tac [3] comp_cont_apply [THEN ssubst])
@@ -848,8 +848,8 @@
done
lemma chain_cf_comp:
- "[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)|]
- ==> chain(cf(D,E),\<lambda>n \<in> nat. X`n O Y`n)"
+ "\<lbrakk>chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)\<rbrakk>
+ \<Longrightarrow> chain(cf(D,E),\<lambda>n \<in> nat. X`n O Y`n)"
apply (rule chainI)
defer 1
apply simp
@@ -864,8 +864,8 @@
done
lemma comp_lubs:
- "[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(D'); cpo(E)|]
- ==> lub(cf(D',E),X) O lub(cf(D,D'),Y) = lub(cf(D,E),\<lambda>n \<in> nat. X`n O Y`n)"
+ "\<lbrakk>chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(D'); cpo(E)\<rbrakk>
+ \<Longrightarrow> lub(cf(D',E),X) O lub(cf(D,D'),Y) = lub(cf(D,E),\<lambda>n \<in> nat. X`n O Y`n)"
apply (rule fun_extension)
apply (rule_tac [3] lub_cf [THEN ssubst])
apply (assumption |
@@ -891,23 +891,23 @@
(*----------------------------------------------------------------------*)
lemma projpairI:
- "[| e \<in> cont(D,E); p \<in> cont(E,D); p O e = id(set(D));
- rel(cf(E,E))(e O p)(id(set(E)))|] ==> projpair(D,E,e,p)"
+ "\<lbrakk>e \<in> cont(D,E); p \<in> cont(E,D); p O e = id(set(D));
+ rel(cf(E,E))(e O p)(id(set(E)))\<rbrakk> \<Longrightarrow> projpair(D,E,e,p)"
by (simp add: projpair_def)
-lemma projpair_e_cont: "projpair(D,E,e,p) ==> e \<in> cont(D,E)"
+lemma projpair_e_cont: "projpair(D,E,e,p) \<Longrightarrow> e \<in> cont(D,E)"
by (simp add: projpair_def)
-lemma projpair_p_cont: "projpair(D,E,e,p) ==> p \<in> cont(E,D)"
+lemma projpair_p_cont: "projpair(D,E,e,p) \<Longrightarrow> p \<in> cont(E,D)"
by (simp add: projpair_def)
-lemma projpair_ep_cont: "projpair(D,E,e,p) ==> e \<in> cont(D,E) & p \<in> cont(E,D)"
+lemma projpair_ep_cont: "projpair(D,E,e,p) \<Longrightarrow> e \<in> cont(D,E) & p \<in> cont(E,D)"
by (simp add: projpair_def)
-lemma projpair_eq: "projpair(D,E,e,p) ==> p O e = id(set(D))"
+lemma projpair_eq: "projpair(D,E,e,p) \<Longrightarrow> p O e = id(set(D))"
by (simp add: projpair_def)
-lemma projpair_rel: "projpair(D,E,e,p) ==> rel(cf(E,E))(e O p)(id(set(E)))"
+lemma projpair_rel: "projpair(D,E,e,p) \<Longrightarrow> rel(cf(E,E))(e O p)(id(set(E)))"
by (simp add: projpair_def)
@@ -924,8 +924,8 @@
and comp_id = fun_is_rel [THEN right_comp_id]
lemma projpair_unique_aux1:
- "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p');
- rel(cf(D,E),e,e')|] ==> rel(cf(E,D),p',p)"
+ "\<lbrakk>cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p');
+ rel(cf(D,E),e,e')\<rbrakk> \<Longrightarrow> rel(cf(E,D),p',p)"
apply (rule_tac b=p' in
projpair_p_cont [THEN cont_fun, THEN id_comp, THEN subst], assumption)
apply (rule projpair_eq [THEN subst], assumption)
@@ -947,8 +947,8 @@
text\<open>Proof's very like the previous one. Is there a pattern that
could be exploited?\<close>
lemma projpair_unique_aux2:
- "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p');
- rel(cf(E,D),p',p)|] ==> rel(cf(D,E),e,e')"
+ "\<lbrakk>cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p');
+ rel(cf(E,D),p',p)\<rbrakk> \<Longrightarrow> rel(cf(D,E),e,e')"
apply (rule_tac b=e
in projpair_e_cont [THEN cont_fun, THEN comp_id, THEN subst],
assumption)
@@ -968,26 +968,26 @@
lemma projpair_unique:
- "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p')|]
- ==> (e=e')\<longleftrightarrow>(p=p')"
+ "\<lbrakk>cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p')\<rbrakk>
+ \<Longrightarrow> (e=e')\<longleftrightarrow>(p=p')"
by (blast intro: cpo_antisym projpair_unique_aux1 projpair_unique_aux2 cpo_cf cont_cf
dest: projpair_ep_cont)
(* Slightly different, more asms, since THE chooses the unique element. *)
lemma embRp:
- "[|emb(D,E,e); cpo(D); cpo(E)|] ==> projpair(D,E,e,Rp(D,E,e))"
+ "\<lbrakk>emb(D,E,e); cpo(D); cpo(E)\<rbrakk> \<Longrightarrow> projpair(D,E,e,Rp(D,E,e))"
apply (simp add: emb_def Rp_def)
apply (blast intro: theI2 projpair_unique [THEN iffD1])
done
-lemma embI: "projpair(D,E,e,p) ==> emb(D,E,e)"
+lemma embI: "projpair(D,E,e,p) \<Longrightarrow> emb(D,E,e)"
by (simp add: emb_def, auto)
-lemma Rp_unique: "[|projpair(D,E,e,p); cpo(D); cpo(E)|] ==> Rp(D,E,e) = p"
+lemma Rp_unique: "\<lbrakk>projpair(D,E,e,p); cpo(D); cpo(E)\<rbrakk> \<Longrightarrow> Rp(D,E,e) = p"
by (blast intro: embRp embI projpair_unique [THEN iffD1])
-lemma emb_cont [TC]: "emb(D,E,e) ==> e \<in> cont(D,E)"
+lemma emb_cont [TC]: "emb(D,E,e) \<Longrightarrow> e \<in> cont(D,E)"
apply (simp add: emb_def)
apply (blast intro: projpair_e_cont)
done
@@ -1000,7 +1000,7 @@
lemma embRp_eq_thm:
- "[|emb(D,E,e); x \<in> set(D); cpo(D); cpo(E)|] ==> Rp(D,E,e)`(e`x) = x"
+ "\<lbrakk>emb(D,E,e); x \<in> set(D); cpo(D); cpo(E)\<rbrakk> \<Longrightarrow> Rp(D,E,e)`(e`x) = x"
apply (rule comp_fun_apply [THEN subst])
apply (assumption | rule Rp_cont emb_cont cont_fun)+
apply (subst embRp_eq)
@@ -1013,17 +1013,17 @@
(*----------------------------------------------------------------------*)
lemma projpair_id:
- "cpo(D) ==> projpair(D,D,id(set(D)),id(set(D)))"
+ "cpo(D) \<Longrightarrow> projpair(D,D,id(set(D)),id(set(D)))"
apply (simp add: projpair_def)
apply (blast intro: cpo_cf cont_cf)
done
lemma emb_id:
- "cpo(D) ==> emb(D,D,id(set(D)))"
+ "cpo(D) \<Longrightarrow> emb(D,D,id(set(D)))"
by (auto intro: embI projpair_id)
lemma Rp_id:
- "cpo(D) ==> Rp(D,D,id(set(D))) = id(set(D))"
+ "cpo(D) \<Longrightarrow> Rp(D,D,id(set(D))) = id(set(D))"
by (auto intro: Rp_unique projpair_id)
@@ -1036,8 +1036,8 @@
(* Proof in Isa/ZF: 23 lines (compared to 56: 60% reduction). *)
lemma comp_lemma:
- "[|emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)|]
- ==> projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))"
+ "\<lbrakk>emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)\<rbrakk>
+ \<Longrightarrow> projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))"
apply (simp add: projpair_def, safe)
apply (assumption | rule comp_pres_cont Rp_cont emb_cont)+
apply (rule comp_assoc [THEN subst])
@@ -1071,27 +1071,27 @@
(*----------------------------------------------------------------------*)
lemma iprodI:
- "x:(\<Prod>n \<in> nat. set(DD`n)) ==> x \<in> set(iprod(DD))"
+ "x:(\<Prod>n \<in> nat. set(DD`n)) \<Longrightarrow> x \<in> set(iprod(DD))"
by (simp add: set_def iprod_def)
lemma iprodE:
- "x \<in> set(iprod(DD)) ==> x:(\<Prod>n \<in> nat. set(DD`n))"
+ "x \<in> set(iprod(DD)) \<Longrightarrow> x:(\<Prod>n \<in> nat. set(DD`n))"
by (simp add: set_def iprod_def)
(* Contains typing conditions in contrast to HOL-ST *)
lemma rel_iprodI:
- "[|!!n. n \<in> nat ==> rel(DD`n,f`n,g`n); f:(\<Prod>n \<in> nat. set(DD`n));
- g:(\<Prod>n \<in> nat. set(DD`n))|] ==> rel(iprod(DD),f,g)"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> rel(DD`n,f`n,g`n); f:(\<Prod>n \<in> nat. set(DD`n));
+ g:(\<Prod>n \<in> nat. set(DD`n))\<rbrakk> \<Longrightarrow> rel(iprod(DD),f,g)"
by (simp add: iprod_def rel_I)
lemma rel_iprodE:
- "[|rel(iprod(DD),f,g); n \<in> nat|] ==> rel(DD`n,f`n,g`n)"
+ "\<lbrakk>rel(iprod(DD),f,g); n \<in> nat\<rbrakk> \<Longrightarrow> rel(DD`n,f`n,g`n)"
by (simp add: iprod_def rel_def)
lemma chain_iprod:
- "[|chain(iprod(DD),X); !!n. n \<in> nat ==> cpo(DD`n); n \<in> nat|]
- ==> chain(DD`n,\<lambda>m \<in> nat. X`m`n)"
+ "\<lbrakk>chain(iprod(DD),X); \<And>n. n \<in> nat \<Longrightarrow> cpo(DD`n); n \<in> nat\<rbrakk>
+ \<Longrightarrow> chain(DD`n,\<lambda>m \<in> nat. X`m`n)"
apply (unfold chain_def, safe)
apply (rule lam_type)
apply (rule apply_type)
@@ -1101,8 +1101,8 @@
done
lemma islub_iprod:
- "[|chain(iprod(DD),X); !!n. n \<in> nat ==> cpo(DD`n)|]
- ==> islub(iprod(DD),X,\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
+ "\<lbrakk>chain(iprod(DD),X); \<And>n. n \<in> nat \<Longrightarrow> cpo(DD`n)\<rbrakk>
+ \<Longrightarrow> islub(iprod(DD),X,\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
apply (simp add: islub_def isub_def, safe)
apply (rule iprodI)
apply (blast intro: lam_type chain_iprod [THEN cpo_lub, THEN islub_in])
@@ -1123,7 +1123,7 @@
done
lemma cpo_iprod [TC]:
- "(!!n. n \<in> nat ==> cpo(DD`n)) ==> cpo(iprod(DD))"
+ "(\<And>n. n \<in> nat \<Longrightarrow> cpo(DD`n)) \<Longrightarrow> cpo(iprod(DD))"
apply (assumption | rule cpoI poI)+
apply (rule rel_iprodI) (*not repeated: want to solve 1, leave 2 unchanged *)
apply (simp | rule cpo_refl iprodE [THEN apply_type] iprodE)+
@@ -1140,8 +1140,8 @@
lemma lub_iprod:
- "[|chain(iprod(DD),X); !!n. n \<in> nat ==> cpo(DD`n)|]
- ==> lub(iprod(DD),X) = (\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
+ "\<lbrakk>chain(iprod(DD),X); \<And>n. n \<in> nat \<Longrightarrow> cpo(DD`n)\<rbrakk>
+ \<Longrightarrow> lub(iprod(DD),X) = (\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
by (blast intro: cpo_lub [THEN islub_unique] islub_iprod cpo_iprod)
@@ -1150,39 +1150,39 @@
(*----------------------------------------------------------------------*)
lemma subcpoI:
- "[|set(D)<=set(E);
- !!x y. [|x \<in> set(D); y \<in> set(D)|] ==> rel(D,x,y)\<longleftrightarrow>rel(E,x,y);
- !!X. chain(D,X) ==> lub(E,X) \<in> set(D)|] ==> subcpo(D,E)"
+ "\<lbrakk>set(D)<=set(E);
+ \<And>x y. \<lbrakk>x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,y)\<longleftrightarrow>rel(E,x,y);
+ \<And>X. chain(D,X) \<Longrightarrow> lub(E,X) \<in> set(D)\<rbrakk> \<Longrightarrow> subcpo(D,E)"
by (simp add: subcpo_def)
-lemma subcpo_subset: "subcpo(D,E) ==> set(D)<=set(E)"
+lemma subcpo_subset: "subcpo(D,E) \<Longrightarrow> set(D)<=set(E)"
by (simp add: subcpo_def)
lemma subcpo_rel_eq:
- "[|subcpo(D,E); x \<in> set(D); y \<in> set(D)|] ==> rel(D,x,y)\<longleftrightarrow>rel(E,x,y)"
+ "\<lbrakk>subcpo(D,E); x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(D,x,y)\<longleftrightarrow>rel(E,x,y)"
by (simp add: subcpo_def)
lemmas subcpo_relD1 = subcpo_rel_eq [THEN iffD1]
lemmas subcpo_relD2 = subcpo_rel_eq [THEN iffD2]
-lemma subcpo_lub: "[|subcpo(D,E); chain(D,X)|] ==> lub(E,X) \<in> set(D)"
+lemma subcpo_lub: "\<lbrakk>subcpo(D,E); chain(D,X)\<rbrakk> \<Longrightarrow> lub(E,X) \<in> set(D)"
by (simp add: subcpo_def)
-lemma chain_subcpo: "[|subcpo(D,E); chain(D,X)|] ==> chain(E,X)"
+lemma chain_subcpo: "\<lbrakk>subcpo(D,E); chain(D,X)\<rbrakk> \<Longrightarrow> chain(E,X)"
by (blast intro: Pi_type [THEN chainI] chain_fun subcpo_relD1
subcpo_subset [THEN subsetD]
chain_in chain_rel)
-lemma ub_subcpo: "[|subcpo(D,E); chain(D,X); isub(D,X,x)|] ==> isub(E,X,x)"
+lemma ub_subcpo: "\<lbrakk>subcpo(D,E); chain(D,X); isub(D,X,x)\<rbrakk> \<Longrightarrow> isub(E,X,x)"
by (blast intro: isubI subcpo_relD1 subcpo_relD1 chain_in isubD1 isubD2
subcpo_subset [THEN subsetD] chain_in chain_rel)
lemma islub_subcpo:
- "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> islub(D,X,lub(E,X))"
+ "\<lbrakk>subcpo(D,E); cpo(E); chain(D,X)\<rbrakk> \<Longrightarrow> islub(D,X,lub(E,X))"
by (blast intro: islubI isubI subcpo_lub subcpo_relD2 chain_in islub_ub
islub_least cpo_lub chain_subcpo isubD1 ub_subcpo)
-lemma subcpo_cpo: "[|subcpo(D,E); cpo(E)|] ==> cpo(D)"
+lemma subcpo_cpo: "\<lbrakk>subcpo(D,E); cpo(E)\<rbrakk> \<Longrightarrow> cpo(D)"
apply (assumption | rule cpoI poI)+
apply (simp add: subcpo_rel_eq)
apply (assumption | rule cpo_refl subcpo_subset [THEN subsetD])+
@@ -1193,39 +1193,39 @@
apply (fast intro: islub_subcpo)
done
-lemma lub_subcpo: "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> lub(D,X) = lub(E,X)"
+lemma lub_subcpo: "\<lbrakk>subcpo(D,E); cpo(E); chain(D,X)\<rbrakk> \<Longrightarrow> lub(D,X) = lub(E,X)"
by (blast intro: cpo_lub [THEN islub_unique] islub_subcpo subcpo_cpo)
(*----------------------------------------------------------------------*)
(* Making subcpos using mkcpo. *)
(*----------------------------------------------------------------------*)
-lemma mkcpoI: "[|x \<in> set(D); P(x)|] ==> x \<in> set(mkcpo(D,P))"
+lemma mkcpoI: "\<lbrakk>x \<in> set(D); P(x)\<rbrakk> \<Longrightarrow> x \<in> set(mkcpo(D,P))"
by (simp add: set_def mkcpo_def)
-lemma mkcpoD1: "x \<in> set(mkcpo(D,P))==> x \<in> set(D)"
+lemma mkcpoD1: "x \<in> set(mkcpo(D,P))\<Longrightarrow> x \<in> set(D)"
by (simp add: set_def mkcpo_def)
-lemma mkcpoD2: "x \<in> set(mkcpo(D,P))==> P(x)"
+lemma mkcpoD2: "x \<in> set(mkcpo(D,P))\<Longrightarrow> P(x)"
by (simp add: set_def mkcpo_def)
-lemma rel_mkcpoE: "rel(mkcpo(D,P),x,y) ==> rel(D,x,y)"
+lemma rel_mkcpoE: "rel(mkcpo(D,P),x,y) \<Longrightarrow> rel(D,x,y)"
by (simp add: rel_def mkcpo_def)
lemma rel_mkcpo:
- "[|x \<in> set(D); y \<in> set(D)|] ==> rel(mkcpo(D,P),x,y) \<longleftrightarrow> rel(D,x,y)"
+ "\<lbrakk>x \<in> set(D); y \<in> set(D)\<rbrakk> \<Longrightarrow> rel(mkcpo(D,P),x,y) \<longleftrightarrow> rel(D,x,y)"
by (simp add: mkcpo_def rel_def set_def)
lemma chain_mkcpo:
- "chain(mkcpo(D,P),X) ==> chain(D,X)"
+ "chain(mkcpo(D,P),X) \<Longrightarrow> chain(D,X)"
apply (rule chainI)
apply (blast intro: Pi_type chain_fun chain_in [THEN mkcpoD1])
apply (blast intro: rel_mkcpo [THEN iffD1] chain_rel mkcpoD1 chain_in)
done
lemma subcpo_mkcpo:
- "[|!!X. chain(mkcpo(D,P),X) ==> P(lub(D,X)); cpo(D)|]
- ==> subcpo(mkcpo(D,P),D)"
+ "\<lbrakk>\<And>X. chain(mkcpo(D,P),X) \<Longrightarrow> P(lub(D,X)); cpo(D)\<rbrakk>
+ \<Longrightarrow> subcpo(mkcpo(D,P),D)"
apply (intro subcpoI subsetI rel_mkcpo)
apply (erule mkcpoD1)+
apply (blast intro: mkcpoI cpo_lub [THEN islub_in] chain_mkcpo)
@@ -1236,15 +1236,15 @@
(*----------------------------------------------------------------------*)
lemma emb_chainI:
- "[|!!n. n \<in> nat ==> cpo(DD`n);
- !!n. n \<in> nat ==> emb(DD`n,DD`succ(n),ee`n)|] ==> emb_chain(DD,ee)"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> cpo(DD`n);
+ \<And>n. n \<in> nat \<Longrightarrow> emb(DD`n,DD`succ(n),ee`n)\<rbrakk> \<Longrightarrow> emb_chain(DD,ee)"
by (simp add: emb_chain_def)
-lemma emb_chain_cpo [TC]: "[|emb_chain(DD,ee); n \<in> nat|] ==> cpo(DD`n)"
+lemma emb_chain_cpo [TC]: "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk> \<Longrightarrow> cpo(DD`n)"
by (simp add: emb_chain_def)
lemma emb_chain_emb:
- "[|emb_chain(DD,ee); n \<in> nat|] ==> emb(DD`n,DD`succ(n),ee`n)"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk> \<Longrightarrow> emb(DD`n,DD`succ(n),ee`n)"
by (simp add: emb_chain_def)
(*----------------------------------------------------------------------*)
@@ -1252,44 +1252,44 @@
(*----------------------------------------------------------------------*)
lemma DinfI:
- "[|x:(\<Prod>n \<in> nat. set(DD`n));
- !!n. n \<in> nat ==> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n|]
- ==> x \<in> set(Dinf(DD,ee))"
+ "\<lbrakk>x:(\<Prod>n \<in> nat. set(DD`n));
+ \<And>n. n \<in> nat \<Longrightarrow> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n\<rbrakk>
+ \<Longrightarrow> x \<in> set(Dinf(DD,ee))"
apply (simp add: Dinf_def)
apply (blast intro: mkcpoI iprodI)
done
-lemma Dinf_prod: "x \<in> set(Dinf(DD,ee)) ==> x:(\<Prod>n \<in> nat. set(DD`n))"
+lemma Dinf_prod: "x \<in> set(Dinf(DD,ee)) \<Longrightarrow> x:(\<Prod>n \<in> nat. set(DD`n))"
apply (simp add: Dinf_def)
apply (erule mkcpoD1 [THEN iprodE])
done
lemma Dinf_eq:
- "[|x \<in> set(Dinf(DD,ee)); n \<in> nat|]
- ==> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n"
+ "\<lbrakk>x \<in> set(Dinf(DD,ee)); n \<in> nat\<rbrakk>
+ \<Longrightarrow> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n"
apply (simp add: Dinf_def)
apply (blast dest: mkcpoD2)
done
lemma rel_DinfI:
- "[|!!n. n \<in> nat ==> rel(DD`n,x`n,y`n);
- x:(\<Prod>n \<in> nat. set(DD`n)); y:(\<Prod>n \<in> nat. set(DD`n))|]
- ==> rel(Dinf(DD,ee),x,y)"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> rel(DD`n,x`n,y`n);
+ x:(\<Prod>n \<in> nat. set(DD`n)); y:(\<Prod>n \<in> nat. set(DD`n))\<rbrakk>
+ \<Longrightarrow> rel(Dinf(DD,ee),x,y)"
apply (simp add: Dinf_def)
apply (blast intro: rel_mkcpo [THEN iffD2] rel_iprodI iprodI)
done
-lemma rel_Dinf: "[|rel(Dinf(DD,ee),x,y); n \<in> nat|] ==> rel(DD`n,x`n,y`n)"
+lemma rel_Dinf: "\<lbrakk>rel(Dinf(DD,ee),x,y); n \<in> nat\<rbrakk> \<Longrightarrow> rel(DD`n,x`n,y`n)"
apply (simp add: Dinf_def)
apply (erule rel_mkcpoE [THEN rel_iprodE], assumption)
done
-lemma chain_Dinf: "chain(Dinf(DD,ee),X) ==> chain(iprod(DD),X)"
+lemma chain_Dinf: "chain(Dinf(DD,ee),X) \<Longrightarrow> chain(iprod(DD),X)"
apply (simp add: Dinf_def)
apply (erule chain_mkcpo)
done
-lemma subcpo_Dinf: "emb_chain(DD,ee) ==> subcpo(Dinf(DD,ee),iprod(DD))"
+lemma subcpo_Dinf: "emb_chain(DD,ee) \<Longrightarrow> subcpo(Dinf(DD,ee),iprod(DD))"
apply (simp add: Dinf_def)
apply (rule subcpo_mkcpo)
apply (simp add: Dinf_def [symmetric])
@@ -1308,7 +1308,7 @@
(* Simple example of existential reasoning in Isabelle versus HOL. *)
-lemma cpo_Dinf: "emb_chain(DD,ee) ==> cpo(Dinf(DD,ee))"
+lemma cpo_Dinf: "emb_chain(DD,ee) \<Longrightarrow> cpo(Dinf(DD,ee))"
apply (rule subcpo_cpo)
apply (erule subcpo_Dinf)
apply (auto intro: cpo_iprod emb_chain_cpo)
@@ -1318,8 +1318,8 @@
the proof steps are essentially the same (I think). *)
lemma lub_Dinf:
- "[|chain(Dinf(DD,ee),X); emb_chain(DD,ee)|]
- ==> lub(Dinf(DD,ee),X) = (\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
+ "\<lbrakk>chain(Dinf(DD,ee),X); emb_chain(DD,ee)\<rbrakk>
+ \<Longrightarrow> lub(Dinf(DD,ee),X) = (\<lambda>n \<in> nat. lub(DD`n,\<lambda>m \<in> nat. X`m`n))"
apply (subst subcpo_Dinf [THEN lub_subcpo])
apply (auto intro: cpo_iprod emb_chain_cpo lub_iprod chain_Dinf)
done
@@ -1330,7 +1330,7 @@
(*----------------------------------------------------------------------*)
lemma e_less_eq:
- "m \<in> nat ==> e_less(DD,ee,m,m) = id(set(DD`m))"
+ "m \<in> nat \<Longrightarrow> e_less(DD,ee,m,m) = id(set(DD`m))"
by (simp add: e_less_def)
lemma lemma_succ_sub: "succ(m#+n)#-m = succ(natify(n))"
@@ -1341,12 +1341,12 @@
by (simp add: e_less_def)
lemma le_exists:
- "[| m \<le> n; !!x. [|n=m#+x; x \<in> nat|] ==> Q; n \<in> nat |] ==> Q"
+ "\<lbrakk>m \<le> n; \<And>x. \<lbrakk>n=m#+x; x \<in> nat\<rbrakk> \<Longrightarrow> Q; n \<in> nat\<rbrakk> \<Longrightarrow> Q"
apply (drule less_imp_succ_add, auto)
done
-lemma e_less_le: "[| m \<le> n; n \<in> nat |]
- ==> e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)"
+lemma e_less_le: "\<lbrakk>m \<le> n; n \<in> nat\<rbrakk>
+ \<Longrightarrow> e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)"
apply (rule le_exists, assumption)
apply (simp add: e_less_add, assumption)
done
@@ -1354,12 +1354,12 @@
(* All theorems assume variables m and n are natural numbers. *)
lemma e_less_succ:
- "m \<in> nat ==> e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))"
+ "m \<in> nat \<Longrightarrow> e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))"
by (simp add: e_less_le e_less_eq)
lemma e_less_succ_emb:
- "[|!!n. n \<in> nat ==> emb(DD`n,DD`succ(n),ee`n); m \<in> nat|]
- ==> e_less(DD,ee,m,succ(m)) = ee`m"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> emb(DD`n,DD`succ(n),ee`n); m \<in> nat\<rbrakk>
+ \<Longrightarrow> e_less(DD,ee,m,succ(m)) = ee`m"
apply (simp add: e_less_succ)
apply (blast intro: emb_cont cont_fun comp_id)
done
@@ -1368,8 +1368,8 @@
(* In any case the one below was very easy to write. *)
lemma emb_e_less_add:
- "[| emb_chain(DD,ee); m \<in> nat |]
- ==> emb(DD`m, DD`(m#+k), e_less(DD,ee,m,m#+k))"
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat\<rbrakk>
+ \<Longrightarrow> emb(DD`m, DD`(m#+k), e_less(DD,ee,m,m#+k))"
apply (subgoal_tac "emb (DD`m, DD` (m#+natify (k)),
e_less (DD,ee,m,m#+natify (k))) ")
apply (rule_tac [2] n = "natify (k) " in nat_induct)
@@ -1380,8 +1380,8 @@
done
lemma emb_e_less:
- "[| m \<le> n; emb_chain(DD,ee); n \<in> nat |]
- ==> emb(DD`m, DD`n, e_less(DD,ee,m,n))"
+ "\<lbrakk>m \<le> n; emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> emb(DD`m, DD`n, e_less(DD,ee,m,n))"
apply (frule lt_nat_in_nat)
apply (erule nat_succI)
(* same proof as e_less_le *)
@@ -1389,7 +1389,7 @@
apply (simp add: emb_e_less_add, assumption)
done
-lemma comp_mono_eq: "[|f=f'; g=g'|] ==> f O g = f' O g'"
+lemma comp_mono_eq: "\<lbrakk>f=f'; g=g'\<rbrakk> \<Longrightarrow> f O g = f' O g'"
by simp
(* Note the object-level implication for induction on k. This
@@ -1397,8 +1397,8 @@
Therefore this theorem is only a lemma. *)
lemma e_less_split_add_lemma [rule_format]:
- "[| emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> n \<le> k \<longrightarrow>
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> n \<le> k \<longrightarrow>
e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)"
apply (induct_tac k)
apply (simp add: e_less_eq id_type [THEN id_comp])
@@ -1419,41 +1419,41 @@
done
lemma e_less_split_add:
- "[| n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)"
+ "\<lbrakk>n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)"
by (blast intro: e_less_split_add_lemma)
lemma e_gr_eq:
- "m \<in> nat ==> e_gr(DD,ee,m,m) = id(set(DD`m))"
+ "m \<in> nat \<Longrightarrow> e_gr(DD,ee,m,m) = id(set(DD`m))"
by (simp add: e_gr_def)
lemma e_gr_add:
- "[|n \<in> nat; k \<in> nat|]
- ==> e_gr(DD,ee,succ(n#+k),n) =
+ "\<lbrakk>n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,succ(n#+k),n) =
e_gr(DD,ee,n#+k,n) O Rp(DD`(n#+k),DD`succ(n#+k),ee`(n#+k))"
by (simp add: e_gr_def)
lemma e_gr_le:
- "[|n \<le> m; m \<in> nat; n \<in> nat|]
- ==> e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)"
+ "\<lbrakk>n \<le> m; m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)"
apply (erule le_exists)
apply (simp add: e_gr_add, assumption+)
done
lemma e_gr_succ:
- "m \<in> nat ==> e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)"
+ "m \<in> nat \<Longrightarrow> e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)"
by (simp add: e_gr_le e_gr_eq)
(* Cpo asm's due to THE uniqueness. *)
-lemma e_gr_succ_emb: "[|emb_chain(DD,ee); m \<in> nat|]
- ==> e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"
+lemma e_gr_succ_emb: "\<lbrakk>emb_chain(DD,ee); m \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"
apply (simp add: e_gr_succ)
apply (blast intro: id_comp Rp_cont cont_fun emb_chain_cpo emb_chain_emb)
done
lemma e_gr_fun_add:
- "[|emb_chain(DD,ee); n \<in> nat; k \<in> nat|]
- ==> e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)"
apply (induct_tac k)
apply (simp add: e_gr_eq id_type)
apply (simp add: e_gr_add)
@@ -1461,15 +1461,15 @@
done
lemma e_gr_fun:
- "[|n \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat|]
- ==> e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)"
+ "\<lbrakk>n \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)"
apply (rule le_exists, assumption)
apply (simp add: e_gr_fun_add, assumption+)
done
lemma e_gr_split_add_lemma:
- "[| emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> m \<le> k \<longrightarrow>
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> m \<le> k \<longrightarrow>
e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)"
apply (induct_tac k)
apply (rule impI)
@@ -1491,20 +1491,20 @@
done
lemma e_gr_split_add:
- "[| m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)"
+ "\<lbrakk>m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)"
apply (blast intro: e_gr_split_add_lemma [THEN mp])
done
lemma e_less_cont:
- "[|m \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat|]
- ==> e_less(DD,ee,m,n):cont(DD`m,DD`n)"
+ "\<lbrakk>m \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> e_less(DD,ee,m,n):cont(DD`m,DD`n)"
apply (blast intro: emb_cont emb_e_less)
done
lemma e_gr_cont:
- "[|n \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat|]
- ==> e_gr(DD,ee,m,n):cont(DD`m,DD`n)"
+ "\<lbrakk>n \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,m,n):cont(DD`m,DD`n)"
apply (erule rev_mp)
apply (induct_tac m)
apply (simp add: le0_iff e_gr_eq nat_0I)
@@ -1520,8 +1520,8 @@
(* Considerably shorter.... 57 against 26 *)
lemma e_less_e_gr_split_add:
- "[|n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> e_less(DD,ee,m,m#+n) = e_gr(DD,ee,m#+k,m#+n) O e_less(DD,ee,m,m#+k)"
+ "\<lbrakk>n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_less(DD,ee,m,m#+n) = e_gr(DD,ee,m#+k,m#+n) O e_less(DD,ee,m,m#+k)"
(* Use mp to prepare for induction. *)
apply (erule rev_mp)
apply (induct_tac k)
@@ -1547,8 +1547,8 @@
(* Again considerably shorter, and easy to obtain from the previous thm. *)
lemma e_gr_e_less_split_add:
- "[|m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> e_gr(DD,ee,n#+m,n) = e_gr(DD,ee,n#+k,n) O e_less(DD,ee,n#+m,n#+k)"
+ "\<lbrakk>m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> e_gr(DD,ee,n#+m,n) = e_gr(DD,ee,n#+k,n) O e_less(DD,ee,n#+m,n#+k)"
(* Use mp to prepare for induction. *)
apply (erule rev_mp)
apply (induct_tac k)
@@ -1572,54 +1572,54 @@
lemma emb_eps:
- "[|m \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat|]
- ==> emb(DD`m,DD`n,eps(DD,ee,m,n))"
+ "\<lbrakk>m \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> emb(DD`m,DD`n,eps(DD,ee,m,n))"
apply (simp add: eps_def)
apply (blast intro: emb_e_less)
done
lemma eps_fun:
- "[|emb_chain(DD,ee); m \<in> nat; n \<in> nat|]
- ==> eps(DD,ee,m,n): set(DD`m)->set(DD`n)"
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n): set(DD`m)->set(DD`n)"
apply (simp add: eps_def)
apply (auto intro: e_less_cont [THEN cont_fun]
not_le_iff_lt [THEN iffD1, THEN leI]
e_gr_fun nat_into_Ord)
done
-lemma eps_id: "n \<in> nat ==> eps(DD,ee,n,n) = id(set(DD`n))"
+lemma eps_id: "n \<in> nat \<Longrightarrow> eps(DD,ee,n,n) = id(set(DD`n))"
by (simp add: eps_def e_less_eq)
lemma eps_e_less_add:
- "[|m \<in> nat; n \<in> nat|] ==> eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)"
by (simp add: eps_def add_le_self)
lemma eps_e_less:
- "[|m \<le> n; m \<in> nat; n \<in> nat|] ==> eps(DD,ee,m,n) = e_less(DD,ee,m,n)"
+ "\<lbrakk>m \<le> n; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> eps(DD,ee,m,n) = e_less(DD,ee,m,n)"
by (simp add: eps_def)
lemma eps_e_gr_add:
- "[|n \<in> nat; k \<in> nat|] ==> eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)"
+ "\<lbrakk>n \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)"
by (simp add: eps_def e_less_eq e_gr_eq)
lemma eps_e_gr:
- "[|n \<le> m; m \<in> nat; n \<in> nat|] ==> eps(DD,ee,m,n) = e_gr(DD,ee,m,n)"
+ "\<lbrakk>n \<le> m; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> eps(DD,ee,m,n) = e_gr(DD,ee,m,n)"
apply (erule le_exists)
apply (simp_all add: eps_e_gr_add)
done
lemma eps_succ_ee:
- "[|!!n. n \<in> nat ==> emb(DD`n,DD`succ(n),ee`n); m \<in> nat|]
- ==> eps(DD,ee,m,succ(m)) = ee`m"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> emb(DD`n,DD`succ(n),ee`n); m \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,succ(m)) = ee`m"
by (simp add: eps_e_less le_succ_iff e_less_succ_emb)
lemma eps_succ_Rp:
- "[|emb_chain(DD,ee); m \<in> nat|]
- ==> eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"
by (simp add: eps_e_gr le_succ_iff e_gr_succ_emb)
lemma eps_cont:
- "[|emb_chain(DD,ee); m \<in> nat; n \<in> nat|] ==> eps(DD,ee,m,n): cont(DD`m,DD`n)"
+ "\<lbrakk>emb_chain(DD,ee); m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> eps(DD,ee,m,n): cont(DD`m,DD`n)"
apply (rule_tac i = m and j = n in nat_linear_le)
apply (simp_all add: eps_e_less e_less_cont eps_e_gr e_gr_cont)
done
@@ -1627,29 +1627,29 @@
(* Theorems about splitting. *)
lemma eps_split_add_left:
- "[|n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)"
+ "\<lbrakk>n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)"
apply (simp add: eps_e_less add_le_self add_le_mono)
apply (auto intro: e_less_split_add)
done
lemma eps_split_add_left_rev:
- "[|n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)"
+ "\<lbrakk>n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)"
apply (simp add: eps_e_less_add eps_e_gr add_le_self add_le_mono)
apply (auto intro: e_less_e_gr_split_add)
done
lemma eps_split_add_right:
- "[|m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)"
+ "\<lbrakk>m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)"
apply (simp add: eps_e_gr add_le_self add_le_mono)
apply (auto intro: e_gr_split_add)
done
lemma eps_split_add_right_rev:
- "[|m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)"
+ "\<lbrakk>m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)"
apply (simp add: eps_e_gr_add eps_e_less add_le_self add_le_mono)
apply (auto intro: e_gr_e_less_split_add)
done
@@ -1657,38 +1657,38 @@
(* Arithmetic *)
lemma le_exists_lemma:
- "[| n \<le> k; k \<le> m;
- !!p q. [|p \<le> q; k=n#+p; m=n#+q; p \<in> nat; q \<in> nat|] ==> R;
- m \<in> nat |]==>R"
+ "\<lbrakk>n \<le> k; k \<le> m;
+ \<And>p q. \<lbrakk>p \<le> q; k=n#+p; m=n#+q; p \<in> nat; q \<in> nat\<rbrakk> \<Longrightarrow> R;
+ m \<in> nat\<rbrakk>\<Longrightarrow>R"
apply (rule le_exists, assumption)
prefer 2 apply (simp add: lt_nat_in_nat)
apply (rule le_trans [THEN le_exists], assumption+, force+)
done
lemma eps_split_left_le:
- "[|m \<le> k; k \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>m \<le> k; k \<le> n; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule le_exists_lemma, assumption+)
apply (auto intro: eps_split_add_left)
done
lemma eps_split_left_le_rev:
- "[|m \<le> n; n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>m \<le> n; n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule le_exists_lemma, assumption+)
apply (auto intro: eps_split_add_left_rev)
done
lemma eps_split_right_le:
- "[|n \<le> k; k \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>n \<le> k; k \<le> m; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule le_exists_lemma, assumption+)
apply (auto intro: eps_split_add_right)
done
lemma eps_split_right_le_rev:
- "[|n \<le> m; m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>n \<le> m; m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule le_exists_lemma, assumption+)
apply (auto intro: eps_split_add_right_rev)
done
@@ -1696,8 +1696,8 @@
(* The desired two theorems about `splitting'. *)
lemma eps_split_left:
- "[|m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>m \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule nat_linear_le)
apply (rule_tac [4] eps_split_right_le_rev)
prefer 4 apply assumption
@@ -1708,8 +1708,8 @@
done
lemma eps_split_right:
- "[|n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat|]
- ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
+ "\<lbrakk>n \<le> k; emb_chain(DD,ee); m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk>
+ \<Longrightarrow> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"
apply (rule nat_linear_le)
apply (rule_tac [3] eps_split_left_le_rev)
prefer 3 apply assumption
@@ -1726,8 +1726,8 @@
(* Considerably shorter. *)
lemma rho_emb_fun:
- "[|emb_chain(DD,ee); n \<in> nat|]
- ==> rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))"
apply (simp add: rho_emb_def)
apply (assumption |
rule lam_type DinfI eps_cont [THEN cont_fun, THEN apply_type])+
@@ -1760,21 +1760,21 @@
done
lemma rho_emb_apply1:
- "x \<in> set(DD`n) ==> rho_emb(DD,ee,n)`x = (\<lambda>m \<in> nat. eps(DD,ee,n,m)`x)"
+ "x \<in> set(DD`n) \<Longrightarrow> rho_emb(DD,ee,n)`x = (\<lambda>m \<in> nat. eps(DD,ee,n,m)`x)"
by (simp add: rho_emb_def)
lemma rho_emb_apply2:
- "[|x \<in> set(DD`n); m \<in> nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x"
+ "\<lbrakk>x \<in> set(DD`n); m \<in> nat\<rbrakk> \<Longrightarrow> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x"
by (simp add: rho_emb_def)
-lemma rho_emb_id: "[| x \<in> set(DD`n); n \<in> nat|] ==> rho_emb(DD,ee,n)`x`n = x"
+lemma rho_emb_id: "\<lbrakk>x \<in> set(DD`n); n \<in> nat\<rbrakk> \<Longrightarrow> rho_emb(DD,ee,n)`x`n = x"
by (simp add: rho_emb_apply2 eps_id)
(* Shorter proof, 23 against 62. *)
lemma rho_emb_cont:
- "[|emb_chain(DD,ee); n \<in> nat|]
- ==> rho_emb(DD,ee,n): cont(DD`n,Dinf(DD,ee))"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> rho_emb(DD,ee,n): cont(DD`n,Dinf(DD,ee))"
apply (rule contI)
apply (assumption | rule rho_emb_fun)+
apply (rule rel_DinfI)
@@ -1807,8 +1807,8 @@
(* 32 vs 61, using safe_tac with imp in asm would be unfortunate (5steps) *)
lemma eps1_aux1:
- "[|m \<le> n; emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat|]
- ==> rel(DD`n,e_less(DD,ee,m,n)`(x`m),x`n)"
+ "\<lbrakk>m \<le> n; emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> rel(DD`n,e_less(DD,ee,m,n)`(x`m),x`n)"
apply (erule rev_mp) (* For induction proof *)
apply (induct_tac n)
apply (rule impI)
@@ -1851,8 +1851,8 @@
(* 18 vs 40 *)
lemma eps1_aux2:
- "[|n \<le> m; emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat|]
- ==> rel(DD`n,e_gr(DD,ee,m,n)`(x`m),x`n)"
+ "\<lbrakk>n \<le> m; emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> rel(DD`n,e_gr(DD,ee,m,n)`(x`m),x`n)"
apply (erule rev_mp) (* For induction proof *)
apply (induct_tac m)
apply (rule impI)
@@ -1875,8 +1875,8 @@
done
lemma eps1:
- "[|emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat|]
- ==> rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)"
+ "\<lbrakk>emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)"
apply (simp add: eps_def)
apply (blast intro: eps1_aux1 not_le_iff_lt [THEN iffD1, THEN leI, THEN eps1_aux2]
nat_into_Ord)
@@ -1887,8 +1887,8 @@
Look for occurrences of rel_cfI, rel_DinfI, etc to evaluate the problem. *)
lemma lam_Dinf_cont:
- "[| emb_chain(DD,ee); n \<in> nat |]
- ==> (\<lambda>x \<in> set(Dinf(DD,ee)). x`n) \<in> cont(Dinf(DD,ee),DD`n)"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> (\<lambda>x \<in> set(Dinf(DD,ee)). x`n) \<in> cont(Dinf(DD,ee),DD`n)"
apply (rule contI)
apply (assumption | rule lam_type apply_type Dinf_prod)+
apply simp
@@ -1899,8 +1899,8 @@
done
lemma rho_projpair:
- "[| emb_chain(DD,ee); n \<in> nat |]
- ==> projpair(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n),rho_proj(DD,ee,n))"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> projpair(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n),rho_proj(DD,ee,n))"
apply (simp add: rho_proj_def)
apply (rule projpairI)
apply (assumption | rule rho_emb_cont)+
@@ -1934,11 +1934,11 @@
done
lemma emb_rho_emb:
- "[| emb_chain(DD,ee); n \<in> nat |] ==> emb(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))"
+ "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk> \<Longrightarrow> emb(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))"
by (auto simp add: emb_def intro: exI rho_projpair)
-lemma rho_proj_cont: "[| emb_chain(DD,ee); n \<in> nat |]
- ==> rho_proj(DD,ee,n) \<in> cont(Dinf(DD,ee),DD`n)"
+lemma rho_proj_cont: "\<lbrakk>emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> rho_proj(DD,ee,n) \<in> cont(Dinf(DD,ee),DD`n)"
by (auto intro: rho_projpair projpair_p_cont)
(*----------------------------------------------------------------------*)
@@ -1946,24 +1946,24 @@
(*----------------------------------------------------------------------*)
lemma commuteI:
- "[| !!n. n \<in> nat ==> emb(DD`n,E,r(n));
- !!m n. [|m \<le> n; m \<in> nat; n \<in> nat|] ==> r(n) O eps(DD,ee,m,n) = r(m) |]
- ==> commute(DD,ee,E,r)"
+ "\<lbrakk>\<And>n. n \<in> nat \<Longrightarrow> emb(DD`n,E,r(n));
+ \<And>m n. \<lbrakk>m \<le> n; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> r(n) O eps(DD,ee,m,n) = r(m)\<rbrakk>
+ \<Longrightarrow> commute(DD,ee,E,r)"
by (simp add: commute_def)
lemma commute_emb [TC]:
- "[| commute(DD,ee,E,r); n \<in> nat |] ==> emb(DD`n,E,r(n))"
+ "\<lbrakk>commute(DD,ee,E,r); n \<in> nat\<rbrakk> \<Longrightarrow> emb(DD`n,E,r(n))"
by (simp add: commute_def)
lemma commute_eq:
- "[| commute(DD,ee,E,r); m \<le> n; m \<in> nat; n \<in> nat |]
- ==> r(n) O eps(DD,ee,m,n) = r(m) "
+ "\<lbrakk>commute(DD,ee,E,r); m \<le> n; m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> r(n) O eps(DD,ee,m,n) = r(m) "
by (simp add: commute_def)
(* Shorter proof: 11 vs 46 lines. *)
lemma rho_emb_commute:
- "emb_chain(DD,ee) ==> commute(DD,ee,Dinf(DD,ee),rho_emb(DD,ee))"
+ "emb_chain(DD,ee) \<Longrightarrow> commute(DD,ee,Dinf(DD,ee),rho_emb(DD,ee))"
apply (rule commuteI)
apply (assumption | rule emb_rho_emb)+
apply (rule fun_extension) (* Manual instantiation in HOL. *)
@@ -1976,14 +1976,14 @@
apply (auto intro: eps_fun)
done
-lemma le_succ: "n \<in> nat ==> n \<le> succ(n)"
+lemma le_succ: "n \<in> nat \<Longrightarrow> n \<le> succ(n)"
by (simp add: le_succ_iff)
(* Shorter proof: 21 vs 83 (106 - 23, due to OAssoc complication) *)
lemma commute_chain:
- "[| commute(DD,ee,E,r); emb_chain(DD,ee); cpo(E) |]
- ==> chain(cf(E,E),\<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n)))"
+ "\<lbrakk>commute(DD,ee,E,r); emb_chain(DD,ee); cpo(E)\<rbrakk>
+ \<Longrightarrow> chain(cf(E,E),\<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n)))"
apply (rule chainI)
apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont
emb_cont emb_chain_cpo,
@@ -2007,25 +2007,25 @@
done
lemma rho_emb_chain:
- "emb_chain(DD,ee) ==>
+ "emb_chain(DD,ee) \<Longrightarrow>
chain(cf(Dinf(DD,ee),Dinf(DD,ee)),
\<lambda>n \<in> nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))"
by (auto intro: commute_chain rho_emb_commute cpo_Dinf)
lemma rho_emb_chain_apply1:
- "[| emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)) |]
- ==> chain(Dinf(DD,ee),
+ "\<lbrakk>emb_chain(DD,ee); x \<in> set(Dinf(DD,ee))\<rbrakk>
+ \<Longrightarrow> chain(Dinf(DD,ee),
\<lambda>n \<in> nat.
(rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))`x)"
by (drule rho_emb_chain [THEN chain_cf], assumption, simp)
lemma chain_iprod_emb_chain:
- "[| chain(iprod(DD),X); emb_chain(DD,ee); n \<in> nat |]
- ==> chain(DD`n,\<lambda>m \<in> nat. X `m `n)"
+ "\<lbrakk>chain(iprod(DD),X); emb_chain(DD,ee); n \<in> nat\<rbrakk>
+ \<Longrightarrow> chain(DD`n,\<lambda>m \<in> nat. X `m `n)"
by (auto intro: chain_iprod emb_chain_cpo)
lemma rho_emb_chain_apply2:
- "[| emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); n \<in> nat |] ==>
+ "\<lbrakk>emb_chain(DD,ee); x \<in> set(Dinf(DD,ee)); n \<in> nat\<rbrakk> \<Longrightarrow>
chain
(DD`n,
\<lambda>xa \<in> nat.
@@ -2037,7 +2037,7 @@
(* Shorter proof: 32 vs 72 (roughly), Isabelle proof has lemmas. *)
lemma rho_emb_lub:
- "emb_chain(DD,ee) ==>
+ "emb_chain(DD,ee) \<Longrightarrow>
lub(cf(Dinf(DD,ee),Dinf(DD,ee)),
\<lambda>n \<in> nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))) =
id(set(Dinf(DD,ee)))"
@@ -2078,9 +2078,9 @@
done
lemma theta_chain: (* almost same proof as commute_chain *)
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> chain(cf(E,G),\<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n)))"
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> chain(cf(E,G),\<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n)))"
apply (rule chainI)
apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont
emb_cont emb_chain_cpo,
@@ -2105,9 +2105,9 @@
done
lemma theta_proj_chain: (* similar proof to theta_chain *)
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> chain(cf(G,E),\<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n)))"
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> chain(cf(G,E),\<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n)))"
apply (rule chainI)
apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont
emb_cont emb_chain_cpo, simp)
@@ -2138,9 +2138,9 @@
(* Controlled simplification inside lambda: introduce lemmas *)
lemma commute_O_lemma:
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G); x \<in> nat |]
- ==> r(x) O Rp(DD ` x, G, f(x)) O f(x) O Rp(DD ` x, E, r(x)) =
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G); x \<in> nat\<rbrakk>
+ \<Longrightarrow> r(x) O Rp(DD ` x, G, f(x)) O f(x) O Rp(DD ` x, E, r(x)) =
r(x) O Rp(DD ` x, E, r(x))"
apply (rule_tac s1 = "f (x) " in comp_assoc [THEN subst])
apply (subst embRp_eq)
@@ -2152,10 +2152,10 @@
(* Shorter proof (but lemmas): 19 vs 79 (103 - 24, due to OAssoc) *)
lemma theta_projpair:
- "[| lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
+ "\<lbrakk>lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> projpair
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> projpair
(E,G,
lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))),
lub(cf(G,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,G,f(n))))"
@@ -2182,17 +2182,17 @@
done
lemma emb_theta:
- "[| lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
+ "\<lbrakk>lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> emb(E,G,lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))))"
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> emb(E,G,lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))))"
apply (simp add: emb_def)
apply (blast intro: theta_projpair)
done
lemma mono_lemma:
- "[| g \<in> cont(D,D'); cpo(D); cpo(D'); cpo(E) |]
- ==> (\<lambda>f \<in> cont(D',E). f O g) \<in> mono(cf(D',E),cf(D,E))"
+ "\<lbrakk>g \<in> cont(D,D'); cpo(D); cpo(D'); cpo(E)\<rbrakk>
+ \<Longrightarrow> (\<lambda>f \<in> cont(D',E). f O g) \<in> mono(cf(D',E),cf(D,E))"
apply (rule monoI)
apply (simp add: set_def cf_def)
apply (drule cf_cont)+
@@ -2201,9 +2201,9 @@
done
lemma commute_lam_lemma:
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G); n \<in> nat |]
- ==> (\<lambda>na \<in> nat. (\<lambda>f \<in> cont(E, G). f O r(n)) `
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G); n \<in> nat\<rbrakk>
+ \<Longrightarrow> (\<lambda>na \<in> nat. (\<lambda>f \<in> cont(E, G). f O r(n)) `
((\<lambda>n \<in> nat. f(n) O Rp(DD ` n, E, r(n))) ` na)) =
(\<lambda>na \<in> nat. (f(na) O Rp(DD ` na, E, r(na))) O r(n))"
apply (rule fun_extension)
@@ -2212,18 +2212,18 @@
done
lemma chain_lemma:
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G); n \<in> nat |]
- ==> chain(cf(DD`n,G),\<lambda>x \<in> nat. (f(x) O Rp(DD ` x, E, r(x))) O r(n))"
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G); n \<in> nat\<rbrakk>
+ \<Longrightarrow> chain(cf(DD`n,G),\<lambda>x \<in> nat. (f(x) O Rp(DD ` x, E, r(x))) O r(n))"
apply (rule commute_lam_lemma [THEN subst])
apply (blast intro: theta_chain emb_chain_cpo
commute_emb [THEN emb_cont, THEN mono_lemma, THEN mono_chain])+
done
lemma suffix_lemma:
- "[| commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x \<in> nat |]
- ==> suffix(\<lambda>n \<in> nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) =
+ "\<lbrakk>commute(DD,ee,E,r); commute(DD,ee,G,f);
+ emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x \<in> nat\<rbrakk>
+ \<Longrightarrow> suffix(\<lambda>n \<in> nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) =
(\<lambda>n \<in> nat. f(x))"
apply (simp add: suffix_def)
apply (rule lam_type [THEN fun_extension])
@@ -2241,20 +2241,20 @@
lemma mediatingI:
- "[|emb(E,G,t); !!n. n \<in> nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)"
+ "\<lbrakk>emb(E,G,t); \<And>n. n \<in> nat \<Longrightarrow> f(n) = t O r(n)\<rbrakk>\<Longrightarrow>mediating(E,G,r,f,t)"
by (simp add: mediating_def)
-lemma mediating_emb: "mediating(E,G,r,f,t) ==> emb(E,G,t)"
+lemma mediating_emb: "mediating(E,G,r,f,t) \<Longrightarrow> emb(E,G,t)"
by (simp add: mediating_def)
-lemma mediating_eq: "[| mediating(E,G,r,f,t); n \<in> nat |] ==> f(n) = t O r(n)"
+lemma mediating_eq: "\<lbrakk>mediating(E,G,r,f,t); n \<in> nat\<rbrakk> \<Longrightarrow> f(n) = t O r(n)"
by (simp add: mediating_def)
lemma lub_universal_mediating:
- "[| lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
+ "\<lbrakk>lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> mediating(E,G,r,f,lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))))"
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> mediating(E,G,r,f,lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n))))"
apply (assumption | rule mediatingI emb_theta)+
apply (rule_tac b = "r (n) " in lub_const [THEN subst])
apply (rule_tac [3] comp_lubs [THEN ssubst])
@@ -2268,11 +2268,11 @@
done
lemma lub_universal_unique:
- "[| mediating(E,G,r,f,t);
+ "\<lbrakk>mediating(E,G,r,f,t);
lub(cf(E,E), \<lambda>n \<in> nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E));
commute(DD,ee,E,r); commute(DD,ee,G,f);
- emb_chain(DD,ee); cpo(E); cpo(G) |]
- ==> t = lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n)))"
+ emb_chain(DD,ee); cpo(E); cpo(G)\<rbrakk>
+ \<Longrightarrow> t = lub(cf(E,G), \<lambda>n \<in> nat. f(n) O Rp(DD`n,E,r(n)))"
apply (rule_tac b = t in comp_id [THEN subst])
apply (erule_tac [2] subst)
apply (rule_tac [2] b = t in lub_const [THEN subst])
@@ -2289,7 +2289,7 @@
(*---------------------------------------------------------------------*)
theorem Dinf_universal:
- "[| commute(DD,ee,G,f); emb_chain(DD,ee); cpo(G) |] ==>
+ "\<lbrakk>commute(DD,ee,G,f); emb_chain(DD,ee); cpo(G)\<rbrakk> \<Longrightarrow>
mediating
(Dinf(DD,ee),G,rho_emb(DD,ee),f,
lub(cf(Dinf(DD,ee),G),
--- a/src/ZF/ex/NatSum.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/NatSum.thy Tue Sep 27 16:51:35 2022 +0100
@@ -12,7 +12,7 @@
thanks to new simprocs.
Thanks to Sloane's On-Line Encyclopedia of Integer Sequences,
- http://www.research.att.com/~njas/sequences/
+ http://www.research.att.com/\<not>njas/sequences/
*)
@@ -27,41 +27,41 @@
declare zdiff_zmult_distrib [simp] zdiff_zmult_distrib2 [simp]
(*The sum of the first n odd numbers equals n squared.*)
-lemma sum_of_odds: "n \<in> nat ==> sum (%i. i $+ i $+ #1, n) = $#n $* $#n"
+lemma sum_of_odds: "n \<in> nat \<Longrightarrow> sum (%i. i $+ i $+ #1, n) = $#n $* $#n"
by (induct_tac "n", auto)
(*The sum of the first n odd squares*)
lemma sum_of_odd_squares:
- "n \<in> nat ==> #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) =
+ "n \<in> nat \<Longrightarrow> #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) =
$#n $* (#4 $* $#n $* $#n $- #1)"
by (induct_tac "n", auto)
(*The sum of the first n odd cubes*)
lemma sum_of_odd_cubes:
"n \<in> nat
- ==> sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) =
+ \<Longrightarrow> sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) =
$#n $* $#n $* (#2 $* $#n $* $#n $- #1)"
by (induct_tac "n", auto)
(*The sum of the first n positive integers equals n(n+1)/2.*)
lemma sum_of_naturals:
- "n \<in> nat ==> #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)"
+ "n \<in> nat \<Longrightarrow> #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)"
by (induct_tac "n", auto)
lemma sum_of_squares:
- "n \<in> nat ==> #6 $* sum (%i. i$*i, succ(n)) =
+ "n \<in> nat \<Longrightarrow> #6 $* sum (%i. i$*i, succ(n)) =
$#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1)"
by (induct_tac "n", auto)
lemma sum_of_cubes:
- "n \<in> nat ==> #4 $* sum (%i. i$*i$*i, succ(n)) =
+ "n \<in> nat \<Longrightarrow> #4 $* sum (%i. i$*i$*i, succ(n)) =
$#n $* $#n $* ($#n $+ #1) $* ($#n $+ #1)"
by (induct_tac "n", auto)
(** Sum of fourth powers **)
lemma sum_of_fourth_powers:
- "n \<in> nat ==> #30 $* sum (%i. i$*i$*i$*i, succ(n)) =
+ "n \<in> nat \<Longrightarrow> #30 $* sum (%i. i$*i$*i$*i, succ(n)) =
$#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1) $*
(#3 $* $#n $* $#n $+ #3 $* $#n $- #1)"
by (induct_tac "n", auto)
--- a/src/ZF/ex/Primes.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Primes.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,67 +9,67 @@
definition
divides :: "[i,i]=>o" (infixl \<open>dvd\<close> 50) where
- "m dvd n == m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
+ "m dvd n \<equiv> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
definition
is_gcd :: "[i,i,i]=>o" \<comment> \<open>definition of great common divisor\<close> where
- "is_gcd(p,m,n) == ((p dvd m) & (p dvd n)) &
+ "is_gcd(p,m,n) \<equiv> ((p dvd m) & (p dvd n)) &
(\<forall>d\<in>nat. (d dvd m) & (d dvd n) \<longrightarrow> d dvd p)"
definition
gcd :: "[i,i]=>i" \<comment> \<open>Euclid's algorithm for the gcd\<close> where
- "gcd(m,n) == transrec(natify(n),
+ "gcd(m,n) \<equiv> transrec(natify(n),
%n f. \<lambda>m \<in> nat.
if n=0 then m else f`(m mod n)`n) ` natify(m)"
definition
coprime :: "[i,i]=>o" \<comment> \<open>the coprime relation\<close> where
- "coprime(m,n) == gcd(m,n) = 1"
+ "coprime(m,n) \<equiv> gcd(m,n) = 1"
definition
prime :: i \<comment> \<open>the set of prime numbers\<close> where
- "prime == {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p \<longrightarrow> m=1 | m=p)}"
+ "prime \<equiv> {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p \<longrightarrow> m=1 | m=p)}"
subsection\<open>The Divides Relation\<close>
-lemma dvdD: "m dvd n ==> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
+lemma dvdD: "m dvd n \<Longrightarrow> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
by (unfold divides_def, assumption)
lemma dvdE:
- "[|m dvd n; !!k. [|m \<in> nat; n \<in> nat; k \<in> nat; n = m#*k|] ==> P|] ==> P"
+ "\<lbrakk>m dvd n; \<And>k. \<lbrakk>m \<in> nat; n \<in> nat; k \<in> nat; n = m#*k\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (blast dest!: dvdD)
lemmas dvd_imp_nat1 = dvdD [THEN conjunct1]
lemmas dvd_imp_nat2 = dvdD [THEN conjunct2, THEN conjunct1]
-lemma dvd_0_right [simp]: "m \<in> nat ==> m dvd 0"
+lemma dvd_0_right [simp]: "m \<in> nat \<Longrightarrow> m dvd 0"
apply (simp add: divides_def)
apply (fast intro: nat_0I mult_0_right [symmetric])
done
-lemma dvd_0_left: "0 dvd m ==> m = 0"
+lemma dvd_0_left: "0 dvd m \<Longrightarrow> m = 0"
by (simp add: divides_def)
-lemma dvd_refl [simp]: "m \<in> nat ==> m dvd m"
+lemma dvd_refl [simp]: "m \<in> nat \<Longrightarrow> m dvd m"
apply (simp add: divides_def)
apply (fast intro: nat_1I mult_1_right [symmetric])
done
-lemma dvd_trans: "[| m dvd n; n dvd p |] ==> m dvd p"
+lemma dvd_trans: "\<lbrakk>m dvd n; n dvd p\<rbrakk> \<Longrightarrow> m dvd p"
by (auto simp add: divides_def intro: mult_assoc mult_type)
-lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m=n"
+lemma dvd_anti_sym: "\<lbrakk>m dvd n; n dvd m\<rbrakk> \<Longrightarrow> m=n"
apply (simp add: divides_def)
apply (force dest: mult_eq_self_implies_10
simp add: mult_assoc mult_eq_1_iff)
done
-lemma dvd_mult_left: "[|(i#*j) dvd k; i \<in> nat|] ==> i dvd k"
+lemma dvd_mult_left: "\<lbrakk>(i#*j) dvd k; i \<in> nat\<rbrakk> \<Longrightarrow> i dvd k"
by (auto simp add: divides_def mult_assoc)
-lemma dvd_mult_right: "[|(i#*j) dvd k; j \<in> nat|] ==> j dvd k"
+lemma dvd_mult_right: "\<lbrakk>(i#*j) dvd k; j \<in> nat\<rbrakk> \<Longrightarrow> j dvd k"
apply (simp add: divides_def, clarify)
apply (rule_tac x = "i#*ka" in bexI)
apply (simp add: mult_ac)
@@ -91,14 +91,14 @@
by (simp add: gcd_def)
lemma gcd_non_0_raw:
- "[| 0<n; n \<in> nat |] ==> gcd(m,n) = gcd(n, m mod n)"
+ "\<lbrakk>0<n; n \<in> nat\<rbrakk> \<Longrightarrow> gcd(m,n) = gcd(n, m mod n)"
apply (simp add: gcd_def)
apply (rule_tac P = "%z. left (z) = right" for left right in transrec [THEN ssubst])
apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym]
mod_less_divisor [THEN ltD])
done
-lemma gcd_non_0: "0 < natify(n) ==> gcd(m,n) = gcd(n, m mod n)"
+lemma gcd_non_0: "0 < natify(n) \<Longrightarrow> gcd(m,n) = gcd(n, m mod n)"
apply (cut_tac m = m and n = "natify (n) " in gcd_non_0_raw)
apply auto
done
@@ -106,17 +106,17 @@
lemma gcd_1 [simp]: "gcd(m,1) = 1"
by (simp (no_asm_simp) add: gcd_non_0)
-lemma dvd_add: "[| k dvd a; k dvd b |] ==> k dvd (a #+ b)"
+lemma dvd_add: "\<lbrakk>k dvd a; k dvd b\<rbrakk> \<Longrightarrow> k dvd (a #+ b)"
apply (simp add: divides_def)
apply (fast intro: add_mult_distrib_left [symmetric] add_type)
done
-lemma dvd_mult: "k dvd n ==> k dvd (m #* n)"
+lemma dvd_mult: "k dvd n \<Longrightarrow> k dvd (m #* n)"
apply (simp add: divides_def)
apply (fast intro: mult_left_commute mult_type)
done
-lemma dvd_mult2: "k dvd m ==> k dvd (m #* n)"
+lemma dvd_mult2: "k dvd m \<Longrightarrow> k dvd (m #* n)"
apply (subst mult_commute)
apply (blast intro: dvd_mult)
done
@@ -126,7 +126,7 @@
lemmas dvdI2 [simp] = dvd_refl [THEN dvd_mult2]
lemma dvd_mod_imp_dvd_raw:
- "[| a \<in> nat; b \<in> nat; k dvd b; k dvd (a mod b) |] ==> k dvd a"
+ "\<lbrakk>a \<in> nat; b \<in> nat; k dvd b; k dvd (a mod b)\<rbrakk> \<Longrightarrow> k dvd a"
apply (case_tac "b=0")
apply (simp add: DIVISION_BY_ZERO_MOD)
apply (blast intro: mod_div_equality [THEN subst]
@@ -134,17 +134,17 @@
intro!: dvd_add dvd_mult mult_type mod_type div_type)
done
-lemma dvd_mod_imp_dvd: "[| k dvd (a mod b); k dvd b; a \<in> nat |] ==> k dvd a"
+lemma dvd_mod_imp_dvd: "\<lbrakk>k dvd (a mod b); k dvd b; a \<in> nat\<rbrakk> \<Longrightarrow> k dvd a"
apply (cut_tac b = "natify (b)" in dvd_mod_imp_dvd_raw)
apply auto
apply (simp add: divides_def)
done
(*Imitating TFL*)
-lemma gcd_induct_lemma [rule_format (no_asm)]: "[| n \<in> nat;
+lemma gcd_induct_lemma [rule_format (no_asm)]: "\<lbrakk>n \<in> nat;
\<forall>m \<in> nat. P(m,0);
- \<forall>m \<in> nat. \<forall>n \<in> nat. 0<n \<longrightarrow> P(n, m mod n) \<longrightarrow> P(m,n) |]
- ==> \<forall>m \<in> nat. P (m,n)"
+ \<forall>m \<in> nat. \<forall>n \<in> nat. 0<n \<longrightarrow> P(n, m mod n) \<longrightarrow> P(m,n)\<rbrakk>
+ \<Longrightarrow> \<forall>m \<in> nat. P (m,n)"
apply (erule_tac i = n in complete_induct)
apply (case_tac "x=0")
apply (simp (no_asm_simp))
@@ -154,10 +154,10 @@
apply (blast intro: mod_less_divisor [THEN ltD])
done
-lemma gcd_induct: "!!P. [| m \<in> nat; n \<in> nat;
- !!m. m \<in> nat ==> P(m,0);
- !!m n. [|m \<in> nat; n \<in> nat; 0<n; P(n, m mod n)|] ==> P(m,n) |]
- ==> P (m,n)"
+lemma gcd_induct: "\<And>P. \<lbrakk>m \<in> nat; n \<in> nat;
+ \<And>m. m \<in> nat \<Longrightarrow> P(m,0);
+ \<And>m n. \<lbrakk>m \<in> nat; n \<in> nat; 0<n; P(n, m mod n)\<rbrakk> \<Longrightarrow> P(m,n)\<rbrakk>
+ \<Longrightarrow> P (m,n)"
by (blast intro: gcd_induct_lemma)
@@ -176,25 +176,25 @@
text\<open>Property 1: gcd(a,b) divides a and b\<close>
lemma gcd_dvd_both:
- "[| m \<in> nat; n \<in> nat |] ==> gcd (m, n) dvd m & gcd (m, n) dvd n"
+ "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> gcd (m, n) dvd m & gcd (m, n) dvd n"
apply (rule_tac m = m and n = n in gcd_induct)
apply (simp_all add: gcd_non_0)
apply (blast intro: dvd_mod_imp_dvd_raw nat_into_Ord [THEN Ord_0_lt])
done
-lemma gcd_dvd1 [simp]: "m \<in> nat ==> gcd(m,n) dvd m"
+lemma gcd_dvd1 [simp]: "m \<in> nat \<Longrightarrow> gcd(m,n) dvd m"
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both)
apply auto
done
-lemma gcd_dvd2 [simp]: "n \<in> nat ==> gcd(m,n) dvd n"
+lemma gcd_dvd2 [simp]: "n \<in> nat \<Longrightarrow> gcd(m,n) dvd n"
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both)
apply auto
done
text\<open>if f divides a and b then f divides gcd(a,b)\<close>
-lemma dvd_mod: "[| f dvd a; f dvd b |] ==> f dvd (a mod b)"
+lemma dvd_mod: "\<lbrakk>f dvd a; f dvd b\<rbrakk> \<Longrightarrow> f dvd (a mod b)"
apply (simp add: divides_def)
apply (case_tac "b=0")
apply (simp add: DIVISION_BY_ZERO_MOD, auto)
@@ -205,19 +205,19 @@
if f divides a and f divides b then f divides gcd(a,b)\<close>
lemma gcd_greatest_raw [rule_format]:
- "[| m \<in> nat; n \<in> nat; f \<in> nat |]
- ==> (f dvd m) \<longrightarrow> (f dvd n) \<longrightarrow> f dvd gcd(m,n)"
+ "\<lbrakk>m \<in> nat; n \<in> nat; f \<in> nat\<rbrakk>
+ \<Longrightarrow> (f dvd m) \<longrightarrow> (f dvd n) \<longrightarrow> f dvd gcd(m,n)"
apply (rule_tac m = m and n = n in gcd_induct)
apply (simp_all add: gcd_non_0 dvd_mod)
done
-lemma gcd_greatest: "[| f dvd m; f dvd n; f \<in> nat |] ==> f dvd gcd(m,n)"
+lemma gcd_greatest: "\<lbrakk>f dvd m; f dvd n; f \<in> nat\<rbrakk> \<Longrightarrow> f dvd gcd(m,n)"
apply (rule gcd_greatest_raw)
apply (auto simp add: divides_def)
done
-lemma gcd_greatest_iff [simp]: "[| k \<in> nat; m \<in> nat; n \<in> nat |]
- ==> (k dvd gcd (m, n)) \<longleftrightarrow> (k dvd m & k dvd n)"
+lemma gcd_greatest_iff [simp]: "\<lbrakk>k \<in> nat; m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> (k dvd gcd (m, n)) \<longleftrightarrow> (k dvd m & k dvd n)"
by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans)
@@ -225,12 +225,12 @@
text\<open>The GCD exists and function gcd computes it.\<close>
-lemma is_gcd: "[| m \<in> nat; n \<in> nat |] ==> is_gcd(gcd(m,n), m, n)"
+lemma is_gcd: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> is_gcd(gcd(m,n), m, n)"
by (simp add: is_gcd_def)
text\<open>The GCD is unique\<close>
-lemma is_gcd_unique: "[|is_gcd(m,a,b); is_gcd(n,a,b); m\<in>nat; n\<in>nat|] ==> m=n"
+lemma is_gcd_unique: "\<lbrakk>is_gcd(m,a,b); is_gcd(n,a,b); m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> m=n"
apply (simp add: is_gcd_def)
apply (blast intro: dvd_anti_sym)
done
@@ -238,7 +238,7 @@
lemma is_gcd_commute: "is_gcd(k,m,n) \<longleftrightarrow> is_gcd(k,n,m)"
by (simp add: is_gcd_def, blast)
-lemma gcd_commute_raw: "[| m \<in> nat; n \<in> nat |] ==> gcd(m,n) = gcd(n,m)"
+lemma gcd_commute_raw: "\<lbrakk>m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> gcd(m,n) = gcd(n,m)"
apply (rule is_gcd_unique)
apply (rule is_gcd)
apply (rule_tac [3] is_gcd_commute [THEN iffD1])
@@ -250,8 +250,8 @@
apply auto
done
-lemma gcd_assoc_raw: "[| k \<in> nat; m \<in> nat; n \<in> nat |]
- ==> gcd (gcd (k, m), n) = gcd (k, gcd (m, n))"
+lemma gcd_assoc_raw: "\<lbrakk>k \<in> nat; m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> gcd (gcd (k, m), n) = gcd (k, gcd (m, n))"
apply (rule is_gcd_unique)
apply (rule is_gcd)
apply (simp_all add: is_gcd_def)
@@ -289,7 +289,7 @@
lemma gcd_add2' [simp]: "gcd (m, n #+ m) = gcd (m, n)"
by (subst add_commute, rule gcd_add2)
-lemma gcd_add_mult_raw: "k \<in> nat ==> gcd (m, k #* m #+ n) = gcd (m, n)"
+lemma gcd_add_mult_raw: "k \<in> nat \<Longrightarrow> gcd (m, k #* m #+ n) = gcd (m, n)"
apply (erule nat_induct)
apply (auto simp add: gcd_add2 add_assoc)
done
@@ -303,8 +303,8 @@
subsection\<open>Multiplication Laws\<close>
lemma gcd_mult_distrib2_raw:
- "[| k \<in> nat; m \<in> nat; n \<in> nat |]
- ==> k #* gcd (m, n) = gcd (k #* m, k #* n)"
+ "\<lbrakk>k \<in> nat; m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> k #* gcd (m, n) = gcd (k #* m, k #* n)"
apply (erule_tac m = m and n = n in gcd_induct, assumption)
apply simp
apply (case_tac "k = 0", simp)
@@ -324,25 +324,25 @@
by (cut_tac k = k and n = 1 in gcd_mult, auto)
lemma relprime_dvd_mult:
- "[| gcd (k,n) = 1; k dvd (m #* n); m \<in> nat |] ==> k dvd m"
+ "\<lbrakk>gcd (k,n) = 1; k dvd (m #* n); m \<in> nat\<rbrakk> \<Longrightarrow> k dvd m"
apply (cut_tac k = m and m = k and n = n in gcd_mult_distrib2, auto)
apply (erule_tac b = m in ssubst)
apply (simp add: dvd_imp_nat1)
done
lemma relprime_dvd_mult_iff:
- "[| gcd (k,n) = 1; m \<in> nat |] ==> k dvd (m #* n) \<longleftrightarrow> k dvd m"
+ "\<lbrakk>gcd (k,n) = 1; m \<in> nat\<rbrakk> \<Longrightarrow> k dvd (m #* n) \<longleftrightarrow> k dvd m"
by (blast intro: dvdI2 relprime_dvd_mult dvd_trans)
lemma prime_imp_relprime:
- "[| p \<in> prime; ~ (p dvd n); n \<in> nat |] ==> gcd (p, n) = 1"
+ "\<lbrakk>p \<in> prime; \<not> (p dvd n); n \<in> nat\<rbrakk> \<Longrightarrow> gcd (p, n) = 1"
apply (simp add: prime_def, clarify)
apply (drule_tac x = "gcd (p,n)" in bspec)
apply auto
apply (cut_tac m = p and n = n in gcd_dvd2, auto)
done
-lemma prime_into_nat: "p \<in> prime ==> p \<in> nat"
+lemma prime_into_nat: "p \<in> prime \<Longrightarrow> p \<in> nat"
by (simp add: prime_def)
lemma prime_nonzero: "p \<in> prime \<Longrightarrow> p\<noteq>0"
@@ -354,12 +354,12 @@
one of those primes.\<close>
lemma prime_dvd_mult:
- "[|p dvd m #* n; p \<in> prime; m \<in> nat; n \<in> nat |] ==> p dvd m \<or> p dvd n"
+ "\<lbrakk>p dvd m #* n; p \<in> prime; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> p dvd m \<or> p dvd n"
by (blast intro: relprime_dvd_mult prime_imp_relprime prime_into_nat)
lemma gcd_mult_cancel_raw:
- "[|gcd (k,n) = 1; m \<in> nat; n \<in> nat|] ==> gcd (k #* m, n) = gcd (m, n)"
+ "\<lbrakk>gcd (k,n) = 1; m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> gcd (k #* m, n) = gcd (m, n)"
apply (rule dvd_anti_sym)
apply (rule gcd_greatest)
apply (rule relprime_dvd_mult [of _ k])
@@ -369,7 +369,7 @@
apply (blast intro: dvdI1 gcd_dvd1 dvd_trans)
done
-lemma gcd_mult_cancel: "gcd (k,n) = 1 ==> gcd (k #* m, n) = gcd (m, n)"
+lemma gcd_mult_cancel: "gcd (k,n) = 1 \<Longrightarrow> gcd (k #* m, n) = gcd (m, n)"
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_mult_cancel_raw)
apply auto
done
--- a/src/ZF/ex/Ramsey.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Ramsey.thy Tue Sep 27 16:51:35 2022 +0100
@@ -28,23 +28,23 @@
definition
Symmetric :: "i=>o" where
- "Symmetric(E) == (\<forall>x y. <x,y>:E \<longrightarrow> <y,x>:E)"
+ "Symmetric(E) \<equiv> (\<forall>x y. <x,y>:E \<longrightarrow> <y,x>:E)"
definition
Atleast :: "[i,i]=>o" where \<comment> \<open>not really necessary: ZF defines cardinality\<close>
- "Atleast(n,S) == (\<exists>f. f \<in> inj(n,S))"
+ "Atleast(n,S) \<equiv> (\<exists>f. f \<in> inj(n,S))"
definition
Clique :: "[i,i,i]=>o" where
- "Clique(C,V,E) == (C \<subseteq> V) & (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y \<longrightarrow> <x,y> \<in> E)"
+ "Clique(C,V,E) \<equiv> (C \<subseteq> V) & (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y \<longrightarrow> <x,y> \<in> E)"
definition
Indept :: "[i,i,i]=>o" where
- "Indept(I,V,E) == (I \<subseteq> V) & (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y \<longrightarrow> <x,y> \<notin> E)"
+ "Indept(I,V,E) \<equiv> (I \<subseteq> V) & (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y \<longrightarrow> <x,y> \<notin> E)"
definition
Ramsey :: "[i,i,i]=>o" where
- "Ramsey(n,i,j) == \<forall>V E. Symmetric(E) & Atleast(n,V) \<longrightarrow>
+ "Ramsey(n,i,j) \<equiv> \<forall>V E. Symmetric(E) & Atleast(n,V) \<longrightarrow>
(\<exists>C. Clique(C,V,E) & Atleast(i,C)) |
(\<exists>I. Indept(I,V,E) & Atleast(j,I))"
@@ -53,13 +53,13 @@
lemma Clique0 [intro]: "Clique(0,V,E)"
by (unfold Clique_def, blast)
-lemma Clique_superset: "[| Clique(C,V',E); V'<=V |] ==> Clique(C,V,E)"
+lemma Clique_superset: "\<lbrakk>Clique(C,V',E); V'<=V\<rbrakk> \<Longrightarrow> Clique(C,V,E)"
by (unfold Clique_def, blast)
lemma Indept0 [intro]: "Indept(0,V,E)"
by (unfold Indept_def, blast)
-lemma Indept_superset: "[| Indept(I,V',E); V'<=V |] ==> Indept(I,V,E)"
+lemma Indept_superset: "\<lbrakk>Indept(I,V',E); V'<=V\<rbrakk> \<Longrightarrow> Indept(I,V,E)"
by (unfold Indept_def, blast)
(*** Atleast ***)
@@ -68,31 +68,31 @@
by (unfold Atleast_def inj_def Pi_def function_def, blast)
lemma Atleast_succD:
- "Atleast(succ(m),A) ==> \<exists>x \<in> A. Atleast(m, A-{x})"
+ "Atleast(succ(m),A) \<Longrightarrow> \<exists>x \<in> A. Atleast(m, A-{x})"
apply (unfold Atleast_def)
apply (blast dest: inj_is_fun [THEN apply_type] inj_succ_restrict)
done
lemma Atleast_superset:
- "[| Atleast(n,A); A \<subseteq> B |] ==> Atleast(n,B)"
+ "\<lbrakk>Atleast(n,A); A \<subseteq> B\<rbrakk> \<Longrightarrow> Atleast(n,B)"
by (unfold Atleast_def, blast intro: inj_weaken_type)
lemma Atleast_succI:
- "[| Atleast(m,B); b\<notin> B |] ==> Atleast(succ(m), cons(b,B))"
+ "\<lbrakk>Atleast(m,B); b\<notin> B\<rbrakk> \<Longrightarrow> Atleast(succ(m), cons(b,B))"
apply (unfold Atleast_def succ_def)
apply (blast intro: inj_extend elim: mem_irrefl)
done
lemma Atleast_Diff_succI:
- "[| Atleast(m, B-{x}); x \<in> B |] ==> Atleast(succ(m), B)"
+ "\<lbrakk>Atleast(m, B-{x}); x \<in> B\<rbrakk> \<Longrightarrow> Atleast(succ(m), B)"
by (blast intro: Atleast_succI [THEN Atleast_superset])
(*** Main Cardinality Lemma ***)
(*The #-succ(0) strengthens the original theorem statement, but precisely
- the same proof could be used!!*)
+ the same proof could be used\<And>*)
lemma pigeon2 [rule_format]:
- "m \<in> nat ==>
+ "m \<in> nat \<Longrightarrow>
\<forall>n \<in> nat. \<forall>A B. Atleast((m#+n) #- succ(0), A \<union> B) \<longrightarrow>
Atleast(m,A) | Atleast(n,B)"
apply (induct_tac "m")
@@ -137,16 +137,16 @@
(*The use of succ(m) here, rather than #-succ(0), simplifies the proof of
Ramsey_step_lemma.*)
-lemma Atleast_partition: "[| Atleast(m #+ n, A); m \<in> nat; n \<in> nat |]
- ==> Atleast(succ(m), {x \<in> A. ~P(x)}) | Atleast(n, {x \<in> A. P(x)})"
+lemma Atleast_partition: "\<lbrakk>Atleast(m #+ n, A); m \<in> nat; n \<in> nat\<rbrakk>
+ \<Longrightarrow> Atleast(succ(m), {x \<in> A. \<not>P(x)}) | Atleast(n, {x \<in> A. P(x)})"
apply (rule nat_succI [THEN pigeon2], assumption+)
apply (rule Atleast_superset, auto)
done
-(*For the Atleast part, proves ~(a \<in> I) from the second premise!*)
+(*For the Atleast part, proves \<not>(a \<in> I) from the second premise!*)
lemma Indept_succ:
- "[| Indept(I, {z \<in> V-{a}. <a,z> \<notin> E}, E); Symmetric(E); a \<in> V;
- Atleast(j,I) |] ==>
+ "\<lbrakk>Indept(I, {z \<in> V-{a}. <a,z> \<notin> E}, E); Symmetric(E); a \<in> V;
+ Atleast(j,I)\<rbrakk> \<Longrightarrow>
Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))"
apply (unfold Symmetric_def Indept_def)
apply (blast intro!: Atleast_succI)
@@ -154,8 +154,8 @@
lemma Clique_succ:
- "[| Clique(C, {z \<in> V-{a}. <a,z>:E}, E); Symmetric(E); a \<in> V;
- Atleast(j,C) |] ==>
+ "\<lbrakk>Clique(C, {z \<in> V-{a}. <a,z>:E}, E); Symmetric(E); a \<in> V;
+ Atleast(j,C)\<rbrakk> \<Longrightarrow>
Clique(cons(a,C), V, E) & Atleast(succ(j), cons(a,C))"
apply (unfold Symmetric_def Clique_def)
apply (blast intro!: Atleast_succI)
@@ -165,8 +165,8 @@
(*Published proofs gloss over the need for Ramsey numbers to be POSITIVE.*)
lemma Ramsey_step_lemma:
- "[| Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j));
- m \<in> nat; n \<in> nat |] ==> Ramsey(succ(m#+n), succ(i), succ(j))"
+ "\<lbrakk>Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j));
+ m \<in> nat; n \<in> nat\<rbrakk> \<Longrightarrow> Ramsey(succ(m#+n), succ(i), succ(j))"
apply (unfold Ramsey_def, clarify)
apply (erule Atleast_succD [THEN bexE])
apply (erule_tac P1 = "%z.<x,z>:E" in Atleast_partition [THEN disjE],
@@ -181,7 +181,7 @@
(** The actual proof **)
(*Again, the induction requires Ramsey numbers to be positive.*)
-lemma ramsey_lemma: "i \<in> nat ==> \<forall>j \<in> nat. \<exists>n \<in> nat. Ramsey(succ(n), i, j)"
+lemma ramsey_lemma: "i \<in> nat \<Longrightarrow> \<forall>j \<in> nat. \<exists>n \<in> nat. Ramsey(succ(n), i, j)"
apply (induct_tac "i")
apply (blast intro!: Ramsey0j)
apply (rule ballI)
@@ -191,7 +191,7 @@
done
(*Final statement in a tidy form, without succ(...) *)
-lemma ramsey: "[| i \<in> nat; j \<in> nat |] ==> \<exists>n \<in> nat. Ramsey(n,i,j)"
+lemma ramsey: "\<lbrakk>i \<in> nat; j \<in> nat\<rbrakk> \<Longrightarrow> \<exists>n \<in> nat. Ramsey(n,i,j)"
by (blast dest: ramsey_lemma)
end
--- a/src/ZF/ex/Ring.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/Ring.thy Tue Sep 27 16:51:35 2022 +0100
@@ -43,7 +43,7 @@
definition
a_inv :: "[i,i] => i" (\<open>\<ominus>\<index> _\<close> [81] 80) where
- "a_inv(R) == m_inv (<carrier(R), add_field(R), zero(R), 0>)"
+ "a_inv(R) \<equiv> m_inv (<carrier(R), add_field(R), zero(R), 0>)"
definition
minus :: "[i,i,i] => i" (\<open>(_ \<ominus>\<index> _)\<close> [65,66] 65) where
@@ -105,7 +105,7 @@
by (simp add: a_inv_def group.inv_closed [OF a_group, simplified])
lemma (in abelian_monoid) a_closed [intro, simp]:
- "[| x \<in> carrier(G); y \<in> carrier(G) |] ==> x \<oplus> y \<in> carrier(G)"
+ "\<lbrakk>x \<in> carrier(G); y \<in> carrier(G)\<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier(G)"
by (rule monoid.m_closed [OF a_monoid,
simplified, simplified ring_add_def [symmetric]])
@@ -173,7 +173,7 @@
lemmas (in abelian_monoid) a_ac = a_assoc a_comm a_lcomm
text \<open>
- The following proofs are from Jacobson, Basic Algebra I, pp.~88--89
+ The following proofs are from Jacobson, Basic Algebra I, pp.\<not>88--89
\<close>
context ring
@@ -232,7 +232,7 @@
definition
ring_hom :: "[i,i] => i" where
- "ring_hom(R,S) ==
+ "ring_hom(R,S) \<equiv>
{h \<in> carrier(R) -> carrier(S).
(\<forall>x y. x \<in> carrier(R) & y \<in> carrier(R) \<longrightarrow>
h ` (x \<cdot>\<^bsub>R\<^esub> y) = (h ` x) \<cdot>\<^bsub>S\<^esub> (h ` y) &
--- a/src/ZF/ex/misc.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/ex/misc.thy Tue Sep 27 16:51:35 2022 +0100
@@ -31,7 +31,7 @@
text\<open>A weird property of ordered pairs.\<close>
-lemma "b\<noteq>c ==> <a,b> \<inter> <a,c> = <a,a>"
+lemma "b\<noteq>c \<Longrightarrow> <a,b> \<inter> <a,c> = <a,a>"
by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast)
text\<open>These two are cited in Benzmueller and Kohlhase's system description of
@@ -44,15 +44,15 @@
by (blast intro!: equalityI)
text\<open>trivial example of term synthesis: apparently hard for some provers!\<close>
-schematic_goal "a \<noteq> b ==> a:?X & b \<notin> ?X"
+schematic_goal "a \<noteq> b \<Longrightarrow> a:?X & b \<notin> ?X"
by blast
text\<open>Nice blast benchmark. Proved in 0.3s; old tactics can't manage it!\<close>
-lemma "\<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y ==> \<exists>z. S \<subseteq> {z}"
+lemma "\<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
by blast
text\<open>variant of the benchmark above\<close>
-lemma "\<forall>x \<in> S. \<Union>(S) \<subseteq> x ==> \<exists>z. S \<subseteq> {z}"
+lemma "\<forall>x \<in> S. \<Union>(S) \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
by blast
(*Example 12 (credited to Peter Andrews) from
@@ -87,10 +87,10 @@
by force
text\<open>Another version, with meta-level rewriting\<close>
-lemma "(!! A f B g. hom(A,f,B,g) ==
+lemma "(\<And>A f B g. hom(A,f,B,g) \<equiv>
{H \<in> A->B. f \<in> A*A->A & g \<in> B*B->B &
(\<forall>x \<in> A. \<forall>y \<in> A. H`(f`<x,y>) = g`<H`x,H`y>)})
- ==> J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<longrightarrow> (K O J) \<in> hom(A,f,C,h)"
+ \<Longrightarrow> J \<in> hom(A,f,B,g) & K \<in> hom(B,g,C,h) \<longrightarrow> (K O J) \<in> hom(A,f,C,h)"
by force
@@ -105,38 +105,38 @@
comp_mem_injD2 comp_mem_surjD2
lemma pastre1:
- "[| (h O g O f) \<in> inj(A,A);
+ "\<lbrakk>(h O g O f) \<in> inj(A,A);
(f O h O g) \<in> surj(B,B);
(g O f O h) \<in> surj(C,C);
- f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)"
+ f \<in> A->B; g \<in> B->C; h \<in> C->A\<rbrakk> \<Longrightarrow> h \<in> bij(C,A)"
by (unfold bij_def, blast)
lemma pastre3:
- "[| (h O g O f) \<in> surj(A,A);
+ "\<lbrakk>(h O g O f) \<in> surj(A,A);
(f O h O g) \<in> surj(B,B);
(g O f O h) \<in> inj(C,C);
- f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)"
+ f \<in> A->B; g \<in> B->C; h \<in> C->A\<rbrakk> \<Longrightarrow> h \<in> bij(C,A)"
by (unfold bij_def, blast)
lemma pastre4:
- "[| (h O g O f) \<in> surj(A,A);
+ "\<lbrakk>(h O g O f) \<in> surj(A,A);
(f O h O g) \<in> inj(B,B);
(g O f O h) \<in> inj(C,C);
- f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)"
+ f \<in> A->B; g \<in> B->C; h \<in> C->A\<rbrakk> \<Longrightarrow> h \<in> bij(C,A)"
by (unfold bij_def, blast)
lemma pastre5:
- "[| (h O g O f) \<in> inj(A,A);
+ "\<lbrakk>(h O g O f) \<in> inj(A,A);
(f O h O g) \<in> surj(B,B);
(g O f O h) \<in> inj(C,C);
- f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)"
+ f \<in> A->B; g \<in> B->C; h \<in> C->A\<rbrakk> \<Longrightarrow> h \<in> bij(C,A)"
by (unfold bij_def, blast)
lemma pastre6:
- "[| (h O g O f) \<in> inj(A,A);
+ "\<lbrakk>(h O g O f) \<in> inj(A,A);
(f O h O g) \<in> inj(B,B);
(g O f O h) \<in> surj(C,C);
- f \<in> A->B; g \<in> B->C; h \<in> C->A |] ==> h \<in> bij(C,A)"
+ f \<in> A->B; g \<in> B->C; h \<in> C->A\<rbrakk> \<Longrightarrow> h \<in> bij(C,A)"
by (unfold bij_def, blast)
--- a/src/ZF/func.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/func.thy Tue Sep 27 16:51:35 2022 +0100
@@ -9,11 +9,11 @@
subsection\<open>The Pi Operator: Dependent Function Space\<close>
-lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) ==> relation(r)"
+lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) \<Longrightarrow> relation(r)"
by (simp add: relation_def, blast)
lemma relation_converse_converse [simp]:
- "relation(r) ==> converse(converse(r)) = r"
+ "relation(r) \<Longrightarrow> converse(converse(r)) = r"
by (simp add: relation_def, blast)
lemma relation_restrict [simp]: "relation(restrict(r,A))"
@@ -28,23 +28,23 @@
"f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) & (\<forall>x\<in>A. \<exists>!y. <x,y>: f)"
by (unfold Pi_def function_def, blast)
-lemma fun_is_function: "f \<in> Pi(A,B) ==> function(f)"
+lemma fun_is_function: "f \<in> Pi(A,B) \<Longrightarrow> function(f)"
by (simp only: Pi_iff)
lemma function_imp_Pi:
- "[|function(f); relation(f)|] ==> f \<in> domain(f) -> range(f)"
+ "\<lbrakk>function(f); relation(f)\<rbrakk> \<Longrightarrow> f \<in> domain(f) -> range(f)"
by (simp add: Pi_iff relation_def, blast)
lemma functionI:
- "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
+ "\<lbrakk>\<And>x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
by (simp add: function_def, blast)
(*Functions are relations*)
-lemma fun_is_rel: "f \<in> Pi(A,B) ==> f \<subseteq> Sigma(A,B)"
+lemma fun_is_rel: "f \<in> Pi(A,B) \<Longrightarrow> f \<subseteq> Sigma(A,B)"
by (unfold Pi_def, blast)
lemma Pi_cong:
- "[| A=A'; !!x. x \<in> A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')"
+ "\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow> Pi(A,B) = Pi(A',B')"
by (simp add: Pi_def cong add: Sigma_cong)
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
@@ -52,57 +52,57 @@
Sigmas and Pis are abbreviated as * or -> *)
(*Weakening one function type to another; see also Pi_type*)
-lemma fun_weaken_type: "[| f \<in> A->B; B<=D |] ==> f \<in> A->D"
+lemma fun_weaken_type: "\<lbrakk>f \<in> A->B; B<=D\<rbrakk> \<Longrightarrow> f \<in> A->D"
by (unfold Pi_def, best)
subsection\<open>Function Application\<close>
-lemma apply_equality2: "[| <a,b>: f; <a,c>: f; f \<in> Pi(A,B) |] ==> b=c"
+lemma apply_equality2: "\<lbrakk><a,b>: f; <a,c>: f; f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b=c"
by (unfold Pi_def function_def, blast)
-lemma function_apply_equality: "[| <a,b>: f; function(f) |] ==> f`a = b"
+lemma function_apply_equality: "\<lbrakk><a,b>: f; function(f)\<rbrakk> \<Longrightarrow> f`a = b"
by (unfold apply_def function_def, blast)
-lemma apply_equality: "[| <a,b>: f; f \<in> Pi(A,B) |] ==> f`a = b"
+lemma apply_equality: "\<lbrakk><a,b>: f; f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> f`a = b"
apply (unfold Pi_def)
apply (blast intro: function_apply_equality)
done
(*Applying a function outside its domain yields 0*)
-lemma apply_0: "a \<notin> domain(f) ==> f`a = 0"
+lemma apply_0: "a \<notin> domain(f) \<Longrightarrow> f`a = 0"
by (unfold apply_def, blast)
-lemma Pi_memberD: "[| f \<in> Pi(A,B); c \<in> f |] ==> \<exists>x\<in>A. c = <x,f`x>"
+lemma Pi_memberD: "\<lbrakk>f \<in> Pi(A,B); c \<in> f\<rbrakk> \<Longrightarrow> \<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 \<in> domain(f)|] ==> <a,f`a>: f"
+lemma function_apply_Pair: "\<lbrakk>function(f); a \<in> domain(f)\<rbrakk> \<Longrightarrow> <a,f`a>: f"
apply (simp add: function_def, clarify)
apply (subgoal_tac "f`a = y", blast)
apply (simp add: apply_def, blast)
done
-lemma apply_Pair: "[| f \<in> Pi(A,B); a \<in> A |] ==> <a,f`a>: f"
+lemma apply_Pair: "\<lbrakk>f \<in> Pi(A,B); a \<in> A\<rbrakk> \<Longrightarrow> <a,f`a>: f"
apply (simp add: Pi_iff)
apply (blast intro: function_apply_Pair)
done
(*Conclusion is flexible -- use rule_tac or else apply_funtype below!*)
-lemma apply_type [TC]: "[| f \<in> Pi(A,B); a \<in> A |] ==> f`a \<in> B(a)"
+lemma apply_type [TC]: "\<lbrakk>f \<in> Pi(A,B); a \<in> A\<rbrakk> \<Longrightarrow> 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 \<in> A->B; a \<in> A |] ==> f`a \<in> B"
+lemma apply_funtype: "\<lbrakk>f \<in> A->B; a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> B"
by (blast dest: apply_type)
-lemma apply_iff: "f \<in> Pi(A,B) ==> <a,b>: f \<longleftrightarrow> a \<in> A & f`a = b"
+lemma apply_iff: "f \<in> Pi(A,B) \<Longrightarrow> <a,b>: f \<longleftrightarrow> a \<in> A & f`a = b"
apply (frule fun_is_rel)
apply (blast intro!: apply_Pair apply_equality)
done
(*Refining one Pi type to another*)
-lemma Pi_type: "[| f \<in> Pi(A,C); !!x. x \<in> A ==> f`x \<in> B(x) |] ==> f \<in> Pi(A,B)"
+lemma Pi_type: "\<lbrakk>f \<in> Pi(A,C); \<And>x. x \<in> A \<Longrightarrow> f`x \<in> B(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,B)"
apply (simp only: Pi_iff)
apply (blast dest: function_apply_equality)
done
@@ -114,38 +114,38 @@
by (blast intro: Pi_type dest: apply_type)
lemma Pi_weaken_type:
- "[| f \<in> Pi(A,B); !!x. x \<in> A ==> B(x)<=C(x) |] ==> f \<in> Pi(A,C)"
+ "\<lbrakk>f \<in> Pi(A,B); \<And>x. x \<in> A \<Longrightarrow> B(x)<=C(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,C)"
by (blast intro: Pi_type dest: apply_type)
(** Elimination of membership in a function **)
-lemma domain_type: "[| <a,b> \<in> f; f \<in> Pi(A,B) |] ==> a \<in> A"
+lemma domain_type: "\<lbrakk><a,b> \<in> f; f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A"
by (blast dest: fun_is_rel)
-lemma range_type: "[| <a,b> \<in> f; f \<in> Pi(A,B) |] ==> b \<in> B(a)"
+lemma range_type: "\<lbrakk><a,b> \<in> f; f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
by (blast dest: fun_is_rel)
-lemma Pair_mem_PiD: "[| <a,b>: f; f \<in> Pi(A,B) |] ==> a \<in> A & b \<in> B(a) & f`a = b"
+lemma Pair_mem_PiD: "\<lbrakk><a,b>: f; f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A & b \<in> B(a) & f`a = b"
by (blast intro: domain_type range_type apply_equality)
subsection\<open>Lambda Abstraction\<close>
-lemma lamI: "a \<in> A ==> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
+lemma lamI: "a \<in> A \<Longrightarrow> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
apply (unfold lam_def)
apply (erule RepFunI)
done
lemma lamE:
- "[| p: (\<lambda>x\<in>A. b(x)); !!x.[| x \<in> A; p=<x,b(x)> |] ==> P
- |] ==> P"
+ "\<lbrakk>p: (\<lambda>x\<in>A. b(x)); \<And>x.\<lbrakk>x \<in> A; p=<x,b(x)>\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (simp add: lam_def, blast)
-lemma lamD: "[| <a,c>: (\<lambda>x\<in>A. b(x)) |] ==> c = b(a)"
+lemma lamD: "\<lbrakk><a,c>: (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> c = b(a)"
by (simp add: lam_def)
lemma lam_type [TC]:
- "[| !!x. x \<in> A ==> b(x): B(x) |] ==> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
+ "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> b(x): B(x)\<rbrakk> \<Longrightarrow> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
by (simp add: lam_def Pi_def function_def, blast)
lemma lam_funtype: "(\<lambda>x\<in>A. b(x)) \<in> A -> {b(x). x \<in> A}"
@@ -160,7 +160,7 @@
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 \<in> A ==> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
+lemma beta: "a \<in> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
by (simp add: apply_def lam_def, blast)
lemma lam_empty [simp]: "(\<lambda>x\<in>0. b(x)) = 0"
@@ -171,17 +171,17 @@
(*congruence rule for lambda abstraction*)
lemma lam_cong [cong]:
- "[| A=A'; !!x. x \<in> A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')"
+ "\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> b(x)=b'(x)\<rbrakk> \<Longrightarrow> Lambda(A,b) = Lambda(A',b')"
by (simp only: lam_def cong add: RepFun_cong)
lemma lam_theI:
- "(!!x. x \<in> A ==> \<exists>!y. Q(x,y)) ==> \<exists>f. \<forall>x\<in>A. Q(x, f`x)"
+ "(\<And>x. x \<in> A \<Longrightarrow> \<exists>!y. Q(x,y)) \<Longrightarrow> \<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: "[| (\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x)); a \<in> A |] ==> f(a)=g(a)"
+lemma lam_eqE: "\<lbrakk>(\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x)); a \<in> A\<rbrakk> \<Longrightarrow> f(a)=g(a)"
by (fast intro!: lamI elim: equalityE lamE)
@@ -207,26 +207,26 @@
(*Semi-extensionality!*)
lemma fun_subset:
- "[| f \<in> Pi(A,B); g \<in> Pi(C,D); A<=C;
- !!x. x \<in> A ==> f`x = g`x |] ==> f<=g"
+ "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(C,D); A<=C;
+ \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f<=g"
by (force dest: Pi_memberD intro: apply_Pair)
lemma fun_extension:
- "[| f \<in> Pi(A,B); g \<in> Pi(A,D);
- !!x. x \<in> A ==> f`x = g`x |] ==> f=g"
+ "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,D);
+ \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f=g"
by (blast del: subsetI intro: subset_refl sym fun_subset)
-lemma eta [simp]: "f \<in> Pi(A,B) ==> (\<lambda>x\<in>A. f`x) = f"
+lemma eta [simp]: "f \<in> Pi(A,B) \<Longrightarrow> (\<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 \<in> Pi(A,B); g \<in> Pi(A,C) |] ==> (\<forall>a\<in>A. f`a = g`a) \<longleftrightarrow> f=g"
+ "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> (\<forall>a\<in>A. f`a = g`a) \<longleftrightarrow> f=g"
by (blast intro: fun_extension)
(*thm by Mark Staples, proof by lcp*)
-lemma fun_subset_eq: "[| f \<in> Pi(A,B); g \<in> Pi(A,C) |] ==> f \<subseteq> g \<longleftrightarrow> (f = g)"
+lemma fun_subset_eq: "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> f \<subseteq> g \<longleftrightarrow> (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 \<in> Pi(A,B)"
- and minor: "!!b. [| \<forall>x\<in>A. b(x):B(x); f = (\<lambda>x\<in>A. b(x)) |] ==> P"
+ and minor: "\<And>b. \<lbrakk>\<forall>x\<in>A. b(x):B(x); f = (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> P"
shows "P"
apply (rule minor)
apply (rule_tac [2] eta [symmetric])
@@ -244,12 +244,12 @@
subsection\<open>Images of Functions\<close>
-lemma image_lam: "C \<subseteq> A ==> (\<lambda>x\<in>A. b(x)) `` C = {b(x). x \<in> C}"
+lemma image_lam: "C \<subseteq> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) `` C = {b(x). x \<in> C}"
by (unfold lam_def, blast)
lemma Repfun_function_if:
"function(f)
- ==> {f`x. x \<in> C} = (if C \<subseteq> domain(f) then f``C else cons(0,f``C))"
+ \<Longrightarrow> {f`x. x \<in> 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 \<subseteq> domain(f) |] ==> f``C = {f`x. x \<in> C}"
+ "\<lbrakk>function(f); C \<subseteq> domain(f)\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
by (simp add: Repfun_function_if)
-lemma image_fun: "[| f \<in> Pi(A,B); C \<subseteq> A |] ==> f``C = {f`x. x \<in> C}"
+lemma image_fun: "\<lbrakk>f \<in> Pi(A,B); C \<subseteq> A\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
apply (simp add: Pi_iff)
apply (blast intro: image_function)
done
@@ -274,7 +274,7 @@
by (auto simp add: image_fun [OF f])
lemma Pi_image_cons:
- "[| f \<in> Pi(A,B); x \<in> A |] ==> f `` cons(x,y) = cons(f`x, f``y)"
+ "\<lbrakk>f \<in> Pi(A,B); x \<in> A\<rbrakk> \<Longrightarrow> f `` cons(x,y) = cons(f`x, f``y)"
by (blast dest: apply_equality apply_Pair)
@@ -284,10 +284,10 @@
by (unfold restrict_def, blast)
lemma function_restrictI:
- "function(f) ==> function(restrict(f,A))"
+ "function(f) \<Longrightarrow> function(restrict(f,A))"
by (unfold restrict_def function_def, blast)
-lemma restrict_type2: "[| f \<in> Pi(C,B); A<=C |] ==> restrict(f,A) \<in> Pi(A,B)"
+lemma restrict_type2: "\<lbrakk>f \<in> Pi(C,B); A<=C\<rbrakk> \<Longrightarrow> 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 \<in> A then f`a else 0)"
@@ -308,13 +308,13 @@
apply (auto simp add: domain_def)
done
-lemma restrict_idem: "f \<subseteq> Sigma(A,B) ==> restrict(f,A) = f"
+lemma restrict_idem: "f \<subseteq> Sigma(A,B) \<Longrightarrow> restrict(f,A) = f"
by (simp add: restrict_def, blast)
(*converse probably holds too*)
lemma domain_restrict_idem:
- "[| domain(r) \<subseteq> A; relation(r) |] ==> restrict(r,A) = r"
+ "\<lbrakk>domain(r) \<subseteq> A; relation(r)\<rbrakk> \<Longrightarrow> restrict(r,A) = r"
by (simp add: restrict_def relation_def, blast)
lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \<inter> C"
@@ -327,11 +327,11 @@
by (simp add: restrict apply_0)
lemma restrict_lam_eq:
- "A<=C ==> restrict(\<lambda>x\<in>C. b(x), A) = (\<lambda>x\<in>A. b(x))"
+ "A<=C \<Longrightarrow> 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 \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"
+ "f \<in> cons(a, b) -> B \<Longrightarrow> 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,14 +343,14 @@
(** The Union of a set of COMPATIBLE functions is a function **)
lemma function_Union:
- "[| \<forall>x\<in>S. function(x);
- \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x |]
- ==> function(\<Union>(S))"
+ "\<lbrakk>\<forall>x\<in>S. function(x);
+ \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x\<rbrakk>
+ \<Longrightarrow> function(\<Union>(S))"
by (unfold function_def, blast)
lemma fun_Union:
- "[| \<forall>f\<in>S. \<exists>C D. f \<in> C->D;
- \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f |] ==>
+ "\<lbrakk>\<forall>f\<in>S. \<exists>C D. f \<in> C->D;
+ \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f\<rbrakk> \<Longrightarrow>
\<Union>(S) \<in> domain(\<Union>(S)) -> range(\<Union>(S))"
apply (unfold Pi_def)
apply (blast intro!: rel_Union function_Union)
@@ -368,48 +368,48 @@
subset_trans [OF _ Un_upper2]
lemma fun_disjoint_Un:
- "[| f \<in> A->B; g \<in> C->D; A \<inter> C = 0 |]
- ==> (f \<union> g) \<in> (A \<union> C) -> (B \<union> D)"
+ "\<lbrakk>f \<in> A->B; g \<in> C->D; A \<inter> C = 0\<rbrakk>
+ \<Longrightarrow> (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 \<union> g)`a = f`a"
+lemma fun_disjoint_apply1: "a \<notin> domain(g) \<Longrightarrow> (f \<union> g)`a = f`a"
by (simp add: apply_def, blast)
-lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f \<union> g)`c = g`c"
+lemma fun_disjoint_apply2: "c \<notin> domain(f) \<Longrightarrow> (f \<union> g)`c = g`c"
by (simp add: apply_def, blast)
subsection\<open>Domain and Range of a Function or Relation\<close>
-lemma domain_of_fun: "f \<in> Pi(A,B) ==> domain(f)=A"
+lemma domain_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> domain(f)=A"
by (unfold Pi_def, blast)
-lemma apply_rangeI: "[| f \<in> Pi(A,B); a \<in> A |] ==> f`a \<in> range(f)"
+lemma apply_rangeI: "\<lbrakk>f \<in> Pi(A,B); a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> range(f)"
by (erule apply_Pair [THEN rangeI], assumption)
-lemma range_of_fun: "f \<in> Pi(A,B) ==> f \<in> A->range(f)"
+lemma range_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> f \<in> A->range(f)"
by (blast intro: Pi_type apply_rangeI)
subsection\<open>Extensions of Functions\<close>
lemma fun_extend:
- "[| f \<in> A->B; c\<notin>A |] ==> cons(<c,b>,f) \<in> cons(c,A) -> cons(b,B)"
+ "\<lbrakk>f \<in> A->B; c\<notin>A\<rbrakk> \<Longrightarrow> 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 \<in> A->B; c\<notin>A; b \<in> B |] ==> cons(<c,b>,f) \<in> cons(c,A) -> B"
+ "\<lbrakk>f \<in> A->B; c\<notin>A; b \<in> B\<rbrakk> \<Longrightarrow> cons(<c,b>,f) \<in> cons(c,A) -> B"
by (blast intro: fun_extend [THEN fun_weaken_type])
lemma extend_apply:
- "c \<notin> domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+ "c \<notin> domain(f) \<Longrightarrow> 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 \<in> A->B; c\<notin>A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+ "\<lbrakk>f \<in> A->B; c\<notin>A\<rbrakk> \<Longrightarrow> 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,7 +418,7 @@
(*For Finite.ML. Inclusion of right into left is easy*)
lemma cons_fun_eq:
- "c \<notin> A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
+ "c \<notin> A \<Longrightarrow> 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*)
@@ -443,7 +443,7 @@
definition
update :: "[i,i,i] => i" where
- "update(f,a,b) == \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
+ "update(f,a,b) \<equiv> \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
nonterminal updbinds and updbind
@@ -467,7 +467,7 @@
apply (simp_all add: apply_0)
done
-lemma update_idem: "[| f`x = y; f \<in> Pi(A,B); x \<in> A |] ==> f(x:=y) = f"
+lemma update_idem: "\<lbrakk>f`x = y; f \<in> Pi(A,B); x \<in> A\<rbrakk> \<Longrightarrow> f(x:=y) = f"
apply (unfold update_def)
apply (simp add: domain_of_fun cons_absorb)
apply (rule fun_extension)
@@ -475,13 +475,13 @@
done
-(* [| f \<in> Pi(A, B); x \<in> A |] ==> f(x := f`x) = f *)
+(* \<lbrakk>f \<in> Pi(A, B); x \<in> A\<rbrakk> \<Longrightarrow> f(x := f`x) = f *)
declare refl [THEN update_idem, simp]
lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))"
by (unfold update_def, simp)
-lemma update_type: "[| f \<in> Pi(A,B); x \<in> A; y \<in> B(x) |] ==> f(x:=y) \<in> Pi(A, B)"
+lemma update_type: "\<lbrakk>f \<in> Pi(A,B); x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> 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,96 +493,96 @@
(*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) \<subseteq> Replace(B,P)"
+lemma Replace_mono: "A<=B \<Longrightarrow> Replace(A,P) \<subseteq> Replace(B,P)"
by (blast elim!: ReplaceE)
-lemma RepFun_mono: "A<=B ==> {f(x). x \<in> A} \<subseteq> {f(x). x \<in> B}"
+lemma RepFun_mono: "A<=B \<Longrightarrow> {f(x). x \<in> A} \<subseteq> {f(x). x \<in> B}"
by blast
-lemma Pow_mono: "A<=B ==> Pow(A) \<subseteq> Pow(B)"
+lemma Pow_mono: "A<=B \<Longrightarrow> Pow(A) \<subseteq> Pow(B)"
by blast
-lemma Union_mono: "A<=B ==> \<Union>(A) \<subseteq> \<Union>(B)"
+lemma Union_mono: "A<=B \<Longrightarrow> \<Union>(A) \<subseteq> \<Union>(B)"
by blast
lemma UN_mono:
- "[| A<=C; !!x. x \<in> A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> (\<Union>x\<in>C. D(x))"
+ "\<lbrakk>A<=C; \<And>x. x \<in> A \<Longrightarrow> B(x)<=D(x)\<rbrakk> \<Longrightarrow> (\<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) \<subseteq> \<Inter>(A)"
+lemma Inter_anti_mono: "\<lbrakk>A<=B; A\<noteq>0\<rbrakk> \<Longrightarrow> \<Inter>(B) \<subseteq> \<Inter>(A)"
by blast
-lemma cons_mono: "C<=D ==> cons(a,C) \<subseteq> cons(a,D)"
+lemma cons_mono: "C<=D \<Longrightarrow> cons(a,C) \<subseteq> cons(a,D)"
by blast
-lemma Un_mono: "[| A<=C; B<=D |] ==> A \<union> B \<subseteq> C \<union> D"
+lemma Un_mono: "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> A \<union> B \<subseteq> C \<union> D"
by blast
-lemma Int_mono: "[| A<=C; B<=D |] ==> A \<inter> B \<subseteq> C \<inter> D"
+lemma Int_mono: "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> A \<inter> B \<subseteq> C \<inter> D"
by blast
-lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B \<subseteq> C-D"
+lemma Diff_mono: "\<lbrakk>A<=C; D<=B\<rbrakk> \<Longrightarrow> A-B \<subseteq> C-D"
by blast
subsubsection\<open>Standard Products, Sums and Function Spaces\<close>
lemma Sigma_mono [rule_format]:
- "[| A<=C; !!x. x \<in> A \<longrightarrow> B(x) \<subseteq> D(x) |] ==> Sigma(A,B) \<subseteq> Sigma(C,D)"
+ "\<lbrakk>A<=C; \<And>x. x \<in> A \<longrightarrow> B(x) \<subseteq> D(x)\<rbrakk> \<Longrightarrow> Sigma(A,B) \<subseteq> Sigma(C,D)"
by blast
-lemma sum_mono: "[| A<=C; B<=D |] ==> A+B \<subseteq> C+D"
+lemma sum_mono: "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> 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 \<subseteq> A->C"
+lemma Pi_mono: "B<=C \<Longrightarrow> A->B \<subseteq> A->C"
by (blast intro: lam_type elim: Pi_lamE)
-lemma lam_mono: "A<=B ==> Lambda(A,c) \<subseteq> Lambda(B,c)"
+lemma lam_mono: "A<=B \<Longrightarrow> Lambda(A,c) \<subseteq> Lambda(B,c)"
apply (unfold lam_def)
apply (erule RepFun_mono)
done
subsubsection\<open>Converse, Domain, Range, Field\<close>
-lemma converse_mono: "r<=s ==> converse(r) \<subseteq> converse(s)"
+lemma converse_mono: "r<=s \<Longrightarrow> converse(r) \<subseteq> converse(s)"
by blast
-lemma domain_mono: "r<=s ==> domain(r)<=domain(s)"
+lemma domain_mono: "r<=s \<Longrightarrow> domain(r)<=domain(s)"
by blast
lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset]
-lemma range_mono: "r<=s ==> range(r)<=range(s)"
+lemma range_mono: "r<=s \<Longrightarrow> range(r)<=range(s)"
by blast
lemmas range_rel_subset = subset_trans [OF range_mono range_subset]
-lemma field_mono: "r<=s ==> field(r)<=field(s)"
+lemma field_mono: "r<=s \<Longrightarrow> field(r)<=field(s)"
by blast
-lemma field_rel_subset: "r \<subseteq> A*A ==> field(r) \<subseteq> A"
+lemma field_rel_subset: "r \<subseteq> A*A \<Longrightarrow> field(r) \<subseteq> A"
by (erule field_mono [THEN subset_trans], blast)
subsubsection\<open>Images\<close>
lemma image_pair_mono:
- "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A \<subseteq> s``B"
+ "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s; A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
by blast
lemma vimage_pair_mono:
- "[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A \<subseteq> s-``B"
+ "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s; A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
by blast
-lemma image_mono: "[| r<=s; A<=B |] ==> r``A \<subseteq> s``B"
+lemma image_mono: "\<lbrakk>r<=s; A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
by blast
-lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A \<subseteq> s-``B"
+lemma vimage_mono: "\<lbrakk>r<=s; A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
by blast
lemma Collect_mono:
- "[| A<=B; !!x. x \<in> A ==> P(x) \<longrightarrow> Q(x) |] ==> Collect(A,P) \<subseteq> Collect(B,Q)"
+ "\<lbrakk>A<=B; \<And>x. x \<in> A \<Longrightarrow> P(x) \<longrightarrow> Q(x)\<rbrakk> \<Longrightarrow> 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 \<in> Pi(X, Y); A \<subseteq> X |] ==> (\<exists>x\<in>f``A. P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(f`x))"
+ "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk> \<Longrightarrow> (\<exists>x\<in>f``A. P(x)) \<longleftrightarrow> (\<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 \<in> Pi(X, Y); A \<subseteq> X |] ==> (\<forall>x\<in>f``A. P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(f`x))"
+ "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk> \<Longrightarrow> (\<forall>x\<in>f``A. P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(f`x))"
apply safe
apply (blast intro: apply_Pair)
apply (drule bspec) apply assumption
--- a/src/ZF/pair.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/pair.thy Tue Sep 27 16:51:35 2022 +0100
@@ -54,7 +54,7 @@
declare sym [THEN Pair_neq_0, elim!]
-lemma Pair_neq_fst: "<a,b>=a ==> P"
+lemma Pair_neq_fst: "<a,b>=a \<Longrightarrow> P"
proof (unfold Pair_def)
assume eq: "{{a, a}, {a, b}} = a"
have "{a, a} \<in> {{a, a}, {a, b}}" by (rule consI1)
@@ -63,7 +63,7 @@
ultimately show "P" by (rule mem_asym)
qed
-lemma Pair_neq_snd: "<a,b>=b ==> P"
+lemma Pair_neq_snd: "<a,b>=b \<Longrightarrow> P"
proof (unfold Pair_def)
assume eq: "{{a, a}, {a, b}} = b"
have "{a, b} \<in> {{a, a}, {a, b}}" by blast
@@ -80,7 +80,7 @@
lemma Sigma_iff [simp]: "<a,b>: Sigma(A,B) \<longleftrightarrow> a \<in> A & b \<in> B(a)"
by (simp add: Sigma_def)
-lemma SigmaI [TC,intro!]: "[| a \<in> A; b \<in> B(a) |] ==> <a,b> \<in> Sigma(A,B)"
+lemma SigmaI [TC,intro!]: "\<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> <a,b> \<in> Sigma(A,B)"
by simp
lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1]
@@ -88,19 +88,19 @@
(*The general elimination rule*)
lemma SigmaE [elim!]:
- "[| c \<in> Sigma(A,B);
- !!x y.[| x \<in> A; y \<in> B(x); c=<x,y> |] ==> P
- |] ==> P"
+ "\<lbrakk>c \<in> Sigma(A,B);
+ \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x); c=<x,y>\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (unfold Sigma_def, blast)
lemma SigmaE2 [elim!]:
- "[| <a,b> \<in> Sigma(A,B);
- [| a \<in> A; b \<in> B(a) |] ==> P
- |] ==> P"
+ "\<lbrakk><a,b> \<in> Sigma(A,B);
+ \<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (unfold Sigma_def, blast)
lemma Sigma_cong:
- "[| A=A'; !!x. x \<in> A' ==> B(x)=B'(x) |] ==>
+ "\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow>
Sigma(A,B) = Sigma(A',B')"
by (simp add: Sigma_def)
@@ -126,46 +126,46 @@
lemma snd_conv [simp]: "snd(<a,b>) = b"
by (simp add: snd_def)
-lemma fst_type [TC]: "p \<in> Sigma(A,B) ==> fst(p) \<in> A"
+lemma fst_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> fst(p) \<in> A"
by auto
-lemma snd_type [TC]: "p \<in> Sigma(A,B) ==> snd(p) \<in> B(fst(p))"
+lemma snd_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> snd(p) \<in> B(fst(p))"
by auto
-lemma Pair_fst_snd_eq: "a \<in> Sigma(A,B) ==> <fst(a),snd(a)> = a"
+lemma Pair_fst_snd_eq: "a \<in> Sigma(A,B) \<Longrightarrow> <fst(a),snd(a)> = a"
by auto
subsection\<open>The Eliminator, \<^term>\<open>split\<close>\<close>
(*A META-equality, so that it applies to higher types as well...*)
-lemma split [simp]: "split(%x y. c(x,y), <a,b>) == c(a,b)"
+lemma split [simp]: "split(%x y. c(x,y), <a,b>) \<equiv> c(a,b)"
by (simp add: split_def)
lemma split_type [TC]:
- "[| p \<in> Sigma(A,B);
- !!x y.[| x \<in> A; y \<in> B(x) |] ==> c(x,y):C(<x,y>)
- |] ==> split(%x y. c(x,y), p) \<in> C(p)"
+ "\<lbrakk>p \<in> Sigma(A,B);
+ \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(<x,y>)
+\<rbrakk> \<Longrightarrow> split(%x y. c(x,y), p) \<in> C(p)"
by (erule SigmaE, auto)
lemma expand_split:
- "u \<in> A*B ==>
+ "u \<in> A*B \<Longrightarrow>
R(split(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x,y> \<longrightarrow> R(c(x,y)))"
by (auto simp add: split_def)
subsection\<open>A version of \<^term>\<open>split\<close> for Formulae: Result Type \<^typ>\<open>o\<close>\<close>
-lemma splitI: "R(a,b) ==> split(R, <a,b>)"
+lemma splitI: "R(a,b) \<Longrightarrow> split(R, <a,b>)"
by (simp add: split_def)
lemma splitE:
- "[| split(R,z); z \<in> Sigma(A,B);
- !!x y. [| z = <x,y>; R(x,y) |] ==> P
- |] ==> P"
+ "\<lbrakk>split(R,z); z \<in> Sigma(A,B);
+ \<And>x y. \<lbrakk>z = <x,y>; R(x,y)\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow> P"
by (auto simp add: split_def)
-lemma splitD: "split(R,<a,b>) ==> R(a,b)"
+lemma splitD: "split(R,<a,b>) \<Longrightarrow> R(a,b)"
by (simp add: split_def)
text \<open>
--- a/src/ZF/upair.thy Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/upair.thy Tue Sep 27 16:51:35 2022 +0100
@@ -19,7 +19,7 @@
ML_file \<open>Tools/typechk.ML\<close>
lemma atomize_ball [symmetric, rulify]:
- "(!!x. x \<in> A ==> P(x)) == Trueprop (\<forall>x\<in>A. P(x))"
+ "(\<And>x. x \<in> A \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x\<in>A. P(x))"
by (simp add: Ball_def atomize_all atomize_imp)
@@ -34,7 +34,7 @@
lemma UpairI2: "b \<in> Upair(a,b)"
by simp
-lemma UpairE: "[| a \<in> Upair(b,c); a=b ==> P; a=c ==> P |] ==> P"
+lemma UpairE: "\<lbrakk>a \<in> Upair(b,c); a=b \<Longrightarrow> P; a=c \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp, blast)
subsection\<open>Rules for Binary Union, Defined via \<^term>\<open>Upair\<close>\<close>
@@ -44,23 +44,23 @@
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
done
-lemma UnI1: "c \<in> A ==> c \<in> A \<union> B"
+lemma UnI1: "c \<in> A \<Longrightarrow> c \<in> A \<union> B"
by simp
-lemma UnI2: "c \<in> B ==> c \<in> A \<union> B"
+lemma UnI2: "c \<in> B \<Longrightarrow> c \<in> A \<union> B"
by simp
declare UnI1 [elim?] UnI2 [elim?]
-lemma UnE [elim!]: "[| c \<in> A \<union> B; c \<in> A ==> P; c \<in> B ==> P |] ==> P"
+lemma UnE [elim!]: "\<lbrakk>c \<in> A \<union> B; c \<in> A \<Longrightarrow> P; c \<in> B \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp, blast)
(*Stronger version of the rule above*)
-lemma UnE': "[| c \<in> A \<union> B; c \<in> A ==> P; [| c \<in> B; c\<notin>A |] ==> P |] ==> P"
+lemma UnE': "\<lbrakk>c \<in> A \<union> B; c \<in> A \<Longrightarrow> P; \<lbrakk>c \<in> B; c\<notin>A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp, blast)
(*Classical introduction rule: no commitment to A vs B*)
-lemma UnCI [intro!]: "(c \<notin> B ==> c \<in> A) ==> c \<in> A \<union> B"
+lemma UnCI [intro!]: "(c \<notin> B \<Longrightarrow> c \<in> A) \<Longrightarrow> c \<in> A \<union> B"
by (simp, blast)
subsection\<open>Rules for Binary Intersection, Defined via \<^term>\<open>Upair\<close>\<close>
@@ -70,16 +70,16 @@
apply (blast intro: UpairI1 UpairI2 elim: UpairE)
done
-lemma IntI [intro!]: "[| c \<in> A; c \<in> B |] ==> c \<in> A \<inter> B"
+lemma IntI [intro!]: "\<lbrakk>c \<in> A; c \<in> B\<rbrakk> \<Longrightarrow> c \<in> A \<inter> B"
by simp
-lemma IntD1: "c \<in> A \<inter> B ==> c \<in> A"
+lemma IntD1: "c \<in> A \<inter> B \<Longrightarrow> c \<in> A"
by simp
-lemma IntD2: "c \<in> A \<inter> B ==> c \<in> B"
+lemma IntD2: "c \<in> A \<inter> B \<Longrightarrow> c \<in> B"
by simp
-lemma IntE [elim!]: "[| c \<in> A \<inter> B; [| c \<in> A; c \<in> B |] ==> P |] ==> P"
+lemma IntE [elim!]: "\<lbrakk>c \<in> A \<inter> B; \<lbrakk>c \<in> A; c \<in> B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by simp
@@ -88,16 +88,16 @@
lemma Diff_iff [simp]: "c \<in> A-B \<longleftrightarrow> (c \<in> A & c\<notin>B)"
by (unfold Diff_def, blast)
-lemma DiffI [intro!]: "[| c \<in> A; c \<notin> B |] ==> c \<in> A - B"
+lemma DiffI [intro!]: "\<lbrakk>c \<in> A; c \<notin> B\<rbrakk> \<Longrightarrow> c \<in> A - B"
by simp
-lemma DiffD1: "c \<in> A - B ==> c \<in> A"
+lemma DiffD1: "c \<in> A - B \<Longrightarrow> c \<in> A"
by simp
-lemma DiffD2: "c \<in> A - B ==> c \<notin> B"
+lemma DiffD2: "c \<in> A - B \<Longrightarrow> c \<notin> B"
by simp
-lemma DiffE [elim!]: "[| c \<in> A - B; [| c \<in> A; c\<notin>B |] ==> P |] ==> P"
+lemma DiffE [elim!]: "\<lbrakk>c \<in> A - B; \<lbrakk>c \<in> A; c\<notin>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by simp
@@ -114,19 +114,19 @@
by simp
-lemma consI2: "a \<in> B ==> a \<in> cons(b,B)"
+lemma consI2: "a \<in> B \<Longrightarrow> a \<in> cons(b,B)"
by simp
-lemma consE [elim!]: "[| a \<in> cons(b,A); a=b ==> P; a \<in> A ==> P |] ==> P"
+lemma consE [elim!]: "\<lbrakk>a \<in> cons(b,A); a=b \<Longrightarrow> P; a \<in> A \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp, blast)
(*Stronger version of the rule above*)
lemma consE':
- "[| a \<in> cons(b,A); a=b ==> P; [| a \<in> A; a\<noteq>b |] ==> P |] ==> P"
+ "\<lbrakk>a \<in> cons(b,A); a=b \<Longrightarrow> P; \<lbrakk>a \<in> A; a\<noteq>b\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (simp, blast)
(*Classical introduction rule*)
-lemma consCI [intro!]: "(a\<notin>B ==> a=b) ==> a \<in> cons(b,B)"
+lemma consCI [intro!]: "(a\<notin>B \<Longrightarrow> a=b) \<Longrightarrow> a \<in> cons(b,B)"
by (simp, blast)
lemma cons_not_0 [simp]: "cons(a,B) \<noteq> 0"
@@ -151,16 +151,16 @@
subsection\<open>Descriptions\<close>
lemma the_equality [intro]:
- "[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
+ "\<lbrakk>P(a); \<And>x. P(x) \<Longrightarrow> x=a\<rbrakk> \<Longrightarrow> (THE x. P(x)) = a"
apply (unfold the_def)
apply (fast dest: subst)
done
(* Only use this if you already know \<exists>!x. P(x) *)
-lemma the_equality2: "[| \<exists>!x. P(x); P(a) |] ==> (THE x. P(x)) = a"
+lemma the_equality2: "\<lbrakk>\<exists>!x. P(x); P(a)\<rbrakk> \<Longrightarrow> (THE x. P(x)) = a"
by blast
-lemma theI: "\<exists>!x. P(x) ==> P(THE x. P(x))"
+lemma theI: "\<exists>!x. P(x) \<Longrightarrow> P(THE x. P(x))"
apply (erule ex1E)
apply (subst the_equality)
apply (blast+)
@@ -170,15 +170,15 @@
@{term "THE x.P(x)"} rewrites to @{term "THE x.Q(x)"} *)
(*If it's "undefined", it's zero!*)
-lemma the_0: "~ (\<exists>!x. P(x)) ==> (THE x. P(x))=0"
+lemma the_0: "\<not> (\<exists>!x. P(x)) \<Longrightarrow> (THE x. P(x))=0"
apply (unfold the_def)
apply (blast elim!: ReplaceE)
done
(*Easier to apply than theI: conclusion has only one occurrence of P*)
lemma theI2:
- assumes p1: "~ Q(0) ==> \<exists>!x. P(x)"
- and p2: "!!x. P(x) ==> Q(x)"
+ assumes p1: "\<not> Q(0) \<Longrightarrow> \<exists>!x. P(x)"
+ and p2: "\<And>x. P(x) \<Longrightarrow> Q(x)"
shows "Q(THE x. P(x))"
apply (rule classical)
apply (rule p2)
@@ -205,25 +205,25 @@
(*Never use with case splitting, or if P is known to be true or false*)
lemma if_cong:
- "[| P\<longleftrightarrow>Q; Q ==> a=c; ~Q ==> b=d |]
- ==> (if P then a else b) = (if Q then c else d)"
+ "\<lbrakk>P\<longleftrightarrow>Q; Q \<Longrightarrow> a=c; \<not>Q \<Longrightarrow> b=d\<rbrakk>
+ \<Longrightarrow> (if P then a else b) = (if Q then c else d)"
by (simp add: if_def cong add: conj_cong)
(*Prevents simplification of x and y \<in> faster and allows the execution
of functional programs. NOW THE DEFAULT.*)
-lemma if_weak_cong: "P\<longleftrightarrow>Q ==> (if P then x else y) = (if Q then x else y)"
+lemma if_weak_cong: "P\<longleftrightarrow>Q \<Longrightarrow> (if P then x else y) = (if Q then x else y)"
by simp
(*Not needed for rewriting, since P would rewrite to True anyway*)
-lemma if_P: "P ==> (if P then a else b) = a"
+lemma if_P: "P \<Longrightarrow> (if P then a else b) = a"
by (unfold if_def, blast)
(*Not needed for rewriting, since P would rewrite to False anyway*)
-lemma if_not_P: "~P ==> (if P then a else b) = b"
+lemma if_not_P: "\<not>P \<Longrightarrow> (if P then a else b) = b"
by (unfold if_def, blast)
lemma split_if [split]:
- "P(if Q then x else y) \<longleftrightarrow> ((Q \<longrightarrow> P(x)) & (~Q \<longrightarrow> P(y)))"
+ "P(if Q then x else y) \<longleftrightarrow> ((Q \<longrightarrow> P(x)) & (\<not>Q \<longrightarrow> P(y)))"
by (case_tac Q, simp_all)
(** Rewrite rules for boolean case-splitting: faster than split_if [split]
@@ -238,16 +238,16 @@
lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2
(*Logically equivalent to split_if_mem2*)
-lemma if_iff: "a: (if P then x else y) \<longleftrightarrow> P & a \<in> x | ~P & a \<in> y"
+lemma if_iff: "a: (if P then x else y) \<longleftrightarrow> P & a \<in> x | \<not>P & a \<in> y"
by simp
lemma if_type [TC]:
- "[| P ==> a \<in> A; ~P ==> b \<in> A |] ==> (if P then a else b): A"
+ "\<lbrakk>P \<Longrightarrow> a \<in> A; \<not>P \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> (if P then a else b): A"
by simp
(** Splitting IFs in the assumptions **)
-lemma split_if_asm: "P(if Q then x else y) \<longleftrightarrow> (~((Q & ~P(x)) | (~Q & ~P(y))))"
+lemma split_if_asm: "P(if Q then x else y) \<longleftrightarrow> (\<not>((Q & \<not>P(x)) | (\<not>Q & \<not>P(y))))"
by simp
lemmas if_splits = split_if split_if_asm
@@ -256,14 +256,14 @@
subsection\<open>Consequences of Foundation\<close>
(*was called mem_anti_sym*)
-lemma mem_asym: "[| a \<in> b; ~P ==> b \<in> a |] ==> P"
+lemma mem_asym: "\<lbrakk>a \<in> b; \<not>P \<Longrightarrow> b \<in> a\<rbrakk> \<Longrightarrow> P"
apply (rule classical)
apply (rule_tac A1 = "{a,b}" in foundation [THEN disjE])
apply (blast elim!: equalityE)+
done
(*was called mem_anti_refl*)
-lemma mem_irrefl: "a \<in> a ==> P"
+lemma mem_irrefl: "a \<in> a \<Longrightarrow> P"
by (blast intro: mem_asym)
(*mem_irrefl should NOT be added to default databases:
@@ -275,10 +275,10 @@
done
(*Good for proving inequalities by rewriting*)
-lemma mem_imp_not_eq: "a \<in> A ==> a \<noteq> A"
+lemma mem_imp_not_eq: "a \<in> A \<Longrightarrow> a \<noteq> A"
by (blast elim!: mem_irrefl)
-lemma eq_imp_not_mem: "a=A ==> a \<notin> A"
+lemma eq_imp_not_mem: "a=A \<Longrightarrow> a \<notin> A"
by (blast intro: elim: mem_irrefl)
subsection\<open>Rules for Successor\<close>
@@ -289,16 +289,16 @@
lemma succI1 [simp]: "i \<in> succ(i)"
by (simp add: succ_iff)
-lemma succI2: "i \<in> j ==> i \<in> succ(j)"
+lemma succI2: "i \<in> j \<Longrightarrow> i \<in> succ(j)"
by (simp add: succ_iff)
lemma succE [elim!]:
- "[| i \<in> succ(j); i=j ==> P; i \<in> j ==> P |] ==> P"
+ "\<lbrakk>i \<in> succ(j); i=j \<Longrightarrow> P; i \<in> j \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
apply (simp add: succ_iff, blast)
done
(*Classical introduction rule*)
-lemma succCI [intro!]: "(i\<notin>j ==> i=j) ==> i \<in> succ(j)"
+lemma succCI [intro!]: "(i\<notin>j \<Longrightarrow> i=j) \<Longrightarrow> i \<in> succ(j)"
by (simp add: succ_iff, blast)
lemma succ_not_0 [simp]: "succ(n) \<noteq> 0"
@@ -309,7 +309,7 @@
declare succ_not_0 [THEN not_sym, simp]
declare sym [THEN succ_neq_0, elim!]
-(* @{term"succ(c) \<subseteq> B ==> c \<in> B"} *)
+(* @{term"succ(c) \<subseteq> B \<Longrightarrow> c \<in> B"} *)
lemmas succ_subsetD = succI1 [THEN [2] subsetD]
(* @{term"succ(b) \<noteq> b"} *)
@@ -327,7 +327,7 @@
"(\<forall>x\<in>A. P(x) & Q) \<longleftrightarrow> (\<forall>x\<in>A. P(x)) & (A=0 | Q)"
"(\<forall>x\<in>A. P(x) | Q) \<longleftrightarrow> ((\<forall>x\<in>A. P(x)) | Q)"
"(\<forall>x\<in>A. P(x) \<longrightarrow> Q) \<longleftrightarrow> ((\<exists>x\<in>A. P(x)) \<longrightarrow> Q)"
- "(~(\<forall>x\<in>A. P(x))) \<longleftrightarrow> (\<exists>x\<in>A. ~P(x))"
+ "(\<not>(\<forall>x\<in>A. P(x))) \<longleftrightarrow> (\<exists>x\<in>A. \<not>P(x))"
"(\<forall>x\<in>0.P(x)) \<longleftrightarrow> True"
"(\<forall>x\<in>succ(i).P(x)) \<longleftrightarrow> P(i) & (\<forall>x\<in>i. P(x))"
"(\<forall>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) & (\<forall>x\<in>B. P(x))"
@@ -363,7 +363,7 @@
"(\<exists>x\<in>cons(a,B).P(x)) \<longleftrightarrow> P(a) | (\<exists>x\<in>B. P(x))"
"(\<exists>x\<in>RepFun(A,f). P(x)) \<longleftrightarrow> (\<exists>y\<in>A. P(f(y)))"
"(\<exists>x\<in>\<Union>(A).P(x)) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y. P(x))"
- "(~(\<exists>x\<in>A. P(x))) \<longleftrightarrow> (\<forall>x\<in>A. ~P(x))"
+ "(\<not>(\<exists>x\<in>A. P(x))) \<longleftrightarrow> (\<forall>x\<in>A. \<not>P(x))"
by blast+
lemma bex_simps2: