--- a/src/LCF/IsaMakefile Thu Jun 01 21:14:54 2006 +0200
+++ b/src/LCF/IsaMakefile Thu Jun 01 23:07:51 2006 +0200
@@ -26,7 +26,7 @@
FOL:
@cd $(SRC)/FOL; $(ISATOOL) make FOL
-$(OUT)/LCF: $(OUT)/FOL LCF_lemmas.ML LCF.thy ROOT.ML fix.ML pair.ML
+$(OUT)/LCF: $(OUT)/FOL LCF.thy ROOT.ML
@$(ISATOOL) usedir -b -r $(OUT)/FOL LCF
--- a/src/LCF/LCF.thy Thu Jun 01 21:14:54 2006 +0200
+++ b/src/LCF/LCF.thy Thu Jun 01 23:07:51 2006 +0200
@@ -1,4 +1,4 @@
-(* Title: LCF/lcf.thy
+(* Title: LCF/LCF.thy
ID: $Id$
Author: Tobias Nipkow
Copyright 1992 University of Cambridge
@@ -8,7 +8,6 @@
theory LCF
imports FOL
-uses ("LCF_lemmas.ML") ("pair.ML") ("fix.ML")
begin
text {* This theory is based on Lawrence Paulson's book Logic and Computation. *}
@@ -118,18 +117,277 @@
adm_imp: "[| adm(%x.~P(x)); adm(Q) |] ==> adm(%x. P(x)-->Q(x))"
adm_all: "(!!y. adm(P(y))) ==> adm(%x. ALL y. P(y,x))"
-ML {* use_legacy_bindings (the_context ()) *}
+
+lemma eq_imp_less1: "x = y ==> x << y"
+ by (simp add: eq_def)
+
+lemma eq_imp_less2: "x = y ==> y << x"
+ by (simp add: eq_def)
+
+lemma less_refl [simp]: "x << x"
+ apply (rule eq_imp_less1)
+ apply (rule refl)
+ done
+
+lemma less_anti_sym: "[| x << y; y << x |] ==> x=y"
+ by (simp add: eq_def)
+
+lemma ext: "(!!x::'a::cpo. f(x)=(g(x)::'b::cpo)) ==> (%x. f(x))=(%x. g(x))"
+ apply (rule less_anti_sym)
+ apply (rule less_ext)
+ apply simp
+ apply simp
+ done
+
+lemma cong: "[| f=g; x=y |] ==> f(x)=g(y)"
+ by simp
+
+lemma less_ap_term: "x << y ==> f(x) << f(y)"
+ by (rule less_refl [THEN mono])
+
+lemma less_ap_thm: "f << g ==> f(x) << g(x)"
+ by (rule less_refl [THEN [2] mono])
+
+lemma ap_term: "(x::'a::cpo) = y ==> (f(x)::'b::cpo) = f(y)"
+ apply (rule cong [OF refl])
+ apply simp
+ done
+
+lemma ap_thm: "f = g ==> f(x) = g(x)"
+ apply (erule cong)
+ apply (rule refl)
+ done
+
+
+lemma UU_abs: "(%x::'a::cpo. UU) = UU"
+ apply (rule less_anti_sym)
+ prefer 2
+ apply (rule minimal)
+ apply (rule less_ext)
+ apply (rule allI)
+ apply (rule minimal)
+ done
+
+lemma UU_app: "UU(x) = UU"
+ by (rule UU_abs [symmetric, THEN ap_thm])
+
+lemma less_UU: "x << UU ==> x=UU"
+ apply (rule less_anti_sym)
+ apply assumption
+ apply (rule minimal)
+ done
-use "LCF_lemmas.ML"
+lemma tr_induct: "[| P(UU); P(TT); P(FF) |] ==> ALL b. P(b)"
+ apply (rule allI)
+ apply (rule mp)
+ apply (rule_tac [2] p = b in tr_cases)
+ apply blast
+ done
+
+lemma Contrapos: "~ B ==> (A ==> B) ==> ~A"
+ by blast
+
+lemma not_less_imp_not_eq1: "~ x << y \<Longrightarrow> x \<noteq> y"
+ apply (erule Contrapos)
+ apply simp
+ done
+
+lemma not_less_imp_not_eq2: "~ y << x \<Longrightarrow> x \<noteq> y"
+ apply (erule Contrapos)
+ apply simp
+ done
+
+lemma not_UU_eq_TT: "UU \<noteq> TT"
+ by (rule not_less_imp_not_eq2) (rule not_TT_less_UU)
+lemma not_UU_eq_FF: "UU \<noteq> FF"
+ by (rule not_less_imp_not_eq2) (rule not_FF_less_UU)
+lemma not_TT_eq_UU: "TT \<noteq> UU"
+ by (rule not_less_imp_not_eq1) (rule not_TT_less_UU)
+lemma not_TT_eq_FF: "TT \<noteq> FF"
+ by (rule not_less_imp_not_eq1) (rule not_TT_less_FF)
+lemma not_FF_eq_UU: "FF \<noteq> UU"
+ by (rule not_less_imp_not_eq1) (rule not_FF_less_UU)
+lemma not_FF_eq_TT: "FF \<noteq> TT"
+ by (rule not_less_imp_not_eq1) (rule not_FF_less_TT)
+
+
+lemma COND_cases_iff [rule_format]:
+ "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))"
+ apply (insert not_UU_eq_TT not_UU_eq_FF not_TT_eq_UU
+ not_TT_eq_FF not_FF_eq_UU not_FF_eq_TT)
+ apply (rule tr_induct)
+ apply (simplesubst COND_UU)
+ apply blast
+ apply (simplesubst COND_TT)
+ apply blast
+ apply (simplesubst COND_FF)
+ apply blast
+ done
+
+lemma COND_cases:
+ "[| x = UU --> P(UU); x = TT --> P(xa); x = FF --> P(y) |] ==> P(x => xa | y)"
+ apply (rule COND_cases_iff [THEN iffD2])
+ apply blast
+ done
+
+lemmas [simp] =
+ minimal
+ UU_app
+ UU_app [THEN ap_thm]
+ UU_app [THEN ap_thm, THEN ap_thm]
+ not_TT_less_FF not_FF_less_TT not_TT_less_UU not_FF_less_UU not_UU_eq_TT
+ not_UU_eq_FF not_TT_eq_UU not_TT_eq_FF not_FF_eq_UU not_FF_eq_TT
+ COND_UU COND_TT COND_FF
+ surj_pairing FST SND
subsection {* Ordered pairs and products *}
-use "pair.ML"
+lemma expand_all_PROD: "(ALL p. P(p)) <-> (ALL x y. P(<x,y>))"
+ apply (rule iffI)
+ apply blast
+ apply (rule allI)
+ apply (rule surj_pairing [THEN subst])
+ apply blast
+ done
+
+lemma PROD_less: "(p::'a*'b) << q <-> FST(p) << FST(q) & SND(p) << SND(q)"
+ apply (rule iffI)
+ apply (rule conjI)
+ apply (erule less_ap_term)
+ apply (erule less_ap_term)
+ apply (erule conjE)
+ apply (rule surj_pairing [of p, THEN subst])
+ apply (rule surj_pairing [of q, THEN subst])
+ apply (rule mono, erule less_ap_term, assumption)
+ done
+
+lemma PROD_eq: "p=q <-> FST(p)=FST(q) & SND(p)=SND(q)"
+ apply (rule iffI)
+ apply simp
+ apply (unfold eq_def)
+ apply (simp add: PROD_less)
+ done
+
+lemma PAIR_less [simp]: "<a,b> << <c,d> <-> a<<c & b<<d"
+ by (simp add: PROD_less)
+
+lemma PAIR_eq [simp]: "<a,b> = <c,d> <-> a=c & b=d"
+ by (simp add: PROD_eq)
+
+lemma UU_is_UU_UU [simp]: "<UU,UU> = UU"
+ by (rule less_UU) (simp add: PROD_less)
+
+lemma FST_STRICT [simp]: "FST(UU) = UU"
+ apply (rule subst [OF UU_is_UU_UU])
+ apply (simp del: UU_is_UU_UU)
+ done
+
+lemma SND_STRICT [simp]: "SND(UU) = UU"
+ apply (rule subst [OF UU_is_UU_UU])
+ apply (simp del: UU_is_UU_UU)
+ done
subsection {* Fixedpoint theory *}
-use "fix.ML"
+lemma adm_eq: "adm(%x. t(x)=(u(x)::'a::cpo))"
+ apply (unfold eq_def)
+ apply (rule adm_conj adm_less)+
+ done
+
+lemma adm_not_not: "adm(P) ==> adm(%x.~~P(x))"
+ by simp
+
+lemma not_eq_TT: "ALL p. ~p=TT <-> (p=FF | p=UU)"
+ and not_eq_FF: "ALL p. ~p=FF <-> (p=TT | p=UU)"
+ and not_eq_UU: "ALL p. ~p=UU <-> (p=TT | p=FF)"
+ by (rule tr_induct, simp_all)+
+
+lemma adm_not_eq_tr: "ALL p::tr. adm(%x. ~t(x)=p)"
+ apply (rule tr_induct)
+ apply (simp_all add: not_eq_TT not_eq_FF not_eq_UU)
+ apply (rule adm_disj adm_eq)+
+ done
+
+lemmas adm_lemmas =
+ adm_not_free adm_eq adm_less adm_not_less
+ adm_not_eq_tr adm_conj adm_disj adm_imp adm_all
+
+
+ML {*
+local
+ val adm_lemmas = thms "adm_lemmas"
+ val induct = thm "induct"
+in
+ fun induct_tac v i =
+ res_inst_tac[("f",v)] induct i THEN REPEAT (resolve_tac adm_lemmas i)
+end
+*}
+
+lemma least_FIX: "f(p) = p ==> FIX(f) << p"
+ apply (tactic {* induct_tac "f" 1 *})
+ apply (rule minimal)
+ apply (intro strip)
+ apply (erule subst)
+ apply (erule less_ap_term)
+ done
+
+lemma lfp_is_FIX:
+ assumes 1: "f(p) = p"
+ and 2: "ALL q. f(q)=q --> p << q"
+ shows "p = FIX(f)"
+ apply (rule less_anti_sym)
+ apply (rule 2 [THEN spec, THEN mp])
+ apply (rule FIX_eq)
+ apply (rule least_FIX)
+ apply (rule 1)
+ done
+
+
+lemma FIX_pair: "<FIX(f),FIX(g)> = FIX(%p.<f(FST(p)),g(SND(p))>)"
+ apply (rule lfp_is_FIX)
+ apply (simp add: FIX_eq [of f] FIX_eq [of g])
+ apply (intro strip)
+ apply (simp add: PROD_less)
+ apply (rule conjI)
+ apply (rule least_FIX)
+ apply (erule subst, rule FST [symmetric])
+ apply (rule least_FIX)
+ apply (erule subst, rule SND [symmetric])
+ done
+
+lemma FIX1: "FIX(f) = FST(FIX(%p. <f(FST(p)),g(SND(p))>))"
+ by (rule FIX_pair [unfolded PROD_eq FST SND, THEN conjunct1])
+
+lemma FIX2: "FIX(g) = SND(FIX(%p. <f(FST(p)),g(SND(p))>))"
+ by (rule FIX_pair [unfolded PROD_eq FST SND, THEN conjunct2])
+
+lemma induct2:
+ assumes 1: "adm(%p. P(FST(p),SND(p)))"
+ and 2: "P(UU::'a,UU::'b)"
+ and 3: "ALL x y. P(x,y) --> P(f(x),g(y))"
+ shows "P(FIX(f),FIX(g))"
+ apply (rule FIX1 [THEN ssubst, of _ f g])
+ apply (rule FIX2 [THEN ssubst, of _ f g])
+ apply (rule induct [OF 1, where ?f = "%x. <f(FST(x)),g(SND(x))>"])
+ apply simp
+ apply (rule 2)
+ apply (simp add: expand_all_PROD)
+ apply (rule 3)
+ done
+
+ML {*
+local
+ val induct2 = thm "induct2"
+ val adm_lemmas = thms "adm_lemmas"
+in
+
+fun induct2_tac (f,g) i =
+ res_inst_tac[("f",f),("g",g)] induct2 i THEN
+ REPEAT(resolve_tac adm_lemmas i)
end
+*}
+
+end
--- a/src/LCF/LCF_lemmas.ML Thu Jun 01 21:14:54 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(* Title: LCF/lcf.ML
- ID: $Id$
- Author: Tobias Nipkow
- Copyright 1992 University of Cambridge
-*)
-
-(* Standard abbreviations *)
-
-val rstac = resolve_tac;
-fun stac th = rtac(th RS sym RS subst);
-fun sstac ths = EVERY' (map stac ths);
-
-fun strip_tac i = REPEAT(rstac [impI,allI] i);
-
-bind_thm ("eq_imp_less1", prove_goal (the_context ()) "x=y ==> x << y"
- (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]));
-
-bind_thm ("eq_imp_less2", prove_goal (the_context ()) "x=y ==> y << x"
- (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]));
-
-bind_thm ("less_refl", refl RS eq_imp_less1);
-
-bind_thm ("less_anti_sym", prove_goal (the_context ()) "[| x << y; y << x |] ==> x=y"
- (fn prems => [rewtac eq_def,
- REPEAT(rstac(conjI::prems)1)]));
-
-bind_thm ("ext", prove_goal (the_context ())
- "(!!x::'a::cpo. f(x)=(g(x)::'b::cpo)) ==> (%x. f(x))=(%x. g(x))"
- (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI,
- prem RS eq_imp_less1,
- prem RS eq_imp_less2]1)]));
-
-bind_thm ("cong", prove_goal (the_context ()) "[| f=g; x=y |] ==> f(x)=g(y)"
- (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1,
- rtac refl 1]));
-
-bind_thm ("less_ap_term", less_refl RS mono);
-bind_thm ("less_ap_thm", less_refl RSN (2,mono));
-bind_thm ("ap_term", refl RS cong);
-bind_thm ("ap_thm", refl RSN (2,cong));
-
-bind_thm ("UU_abs", prove_goal (the_context ()) "(%x::'a::cpo. UU) = UU"
- (fn _ => [rtac less_anti_sym 1, rtac minimal 2,
- rtac less_ext 1, rtac allI 1, rtac minimal 1]));
-
-bind_thm ("UU_app", UU_abs RS sym RS ap_thm);
-
-bind_thm ("less_UU", prove_goal (the_context ()) "x << UU ==> x=UU"
- (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]));
-
-
-bind_thm ("tr_induct", prove_goal (the_context ()) "[| P(UU); P(TT); P(FF) |] ==> ALL b. P(b)"
- (fn prems => [rtac allI 1, rtac mp 1,
- res_inst_tac[("p","b")]tr_cases 2,
- fast_tac (FOL_cs addIs prems) 1]));
-
-
-bind_thm ("Contrapos", prove_goal (the_context ()) "(A ==> B) ==> (~B ==> ~A)"
- (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1,
- rstac prems 1, atac 1]));
-
-bind_thm ("not_less_imp_not_eq1", eq_imp_less1 COMP Contrapos);
-bind_thm ("not_less_imp_not_eq2", eq_imp_less2 COMP Contrapos);
-
-bind_thm ("not_UU_eq_TT", not_TT_less_UU RS not_less_imp_not_eq2);
-bind_thm ("not_UU_eq_FF", not_FF_less_UU RS not_less_imp_not_eq2);
-bind_thm ("not_TT_eq_UU", not_TT_less_UU RS not_less_imp_not_eq1);
-bind_thm ("not_TT_eq_FF", not_TT_less_FF RS not_less_imp_not_eq1);
-bind_thm ("not_FF_eq_UU", not_FF_less_UU RS not_less_imp_not_eq1);
-bind_thm ("not_FF_eq_TT", not_FF_less_TT RS not_less_imp_not_eq1);
-
-
-bind_thm ("COND_cases_iff", (prove_goal (the_context ())
- "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))"
- (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,
- not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1,
- rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2,
- stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)])) RS spec);
-
-val lemma = prove_goal (the_context ()) "A<->B ==> B ==> A"
- (fn prems => [cut_facts_tac prems 1, rewtac iff_def,
- fast_tac FOL_cs 1]);
-
-bind_thm ("COND_cases", conjI RSN (2,conjI RS (COND_cases_iff RS lemma)));
-
-
-val LCF_ss = FOL_ss addsimps
- [minimal,
- UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm,
- not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU,
- not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF,
- not_FF_eq_UU,not_FF_eq_TT,
- COND_UU,COND_TT,COND_FF,
- surj_pairing,FST,SND];
-
-change_simpset (fn _ => LCF_ss);
--- a/src/LCF/fix.ML Thu Jun 01 21:14:54 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,75 +0,0 @@
-(* Title: LCF/fix
- ID: $Id$
- Author: Tobias Nipkow
- Copyright 1992 University of Cambridge
-*)
-
-val adm_eq = prove_goal (the_context ()) "adm(%x. t(x)=(u(x)::'a::cpo))"
- (fn _ => [rewtac eq_def,
- REPEAT(rstac[adm_conj,adm_less]1)]);
-
-val adm_not_not = prove_goal (the_context ()) "adm(P) ==> adm(%x.~~P(x))"
- (fn prems => [simp_tac (LCF_ss addsimps prems) 1]);
-
-
-val tac = rtac tr_induct 1 THEN ALLGOALS (simp_tac LCF_ss);
-
-val not_eq_TT = prove_goal (the_context ()) "ALL p. ~p=TT <-> (p=FF | p=UU)"
- (fn _ => [tac]) RS spec;
-
-val not_eq_FF = prove_goal (the_context ()) "ALL p. ~p=FF <-> (p=TT | p=UU)"
- (fn _ => [tac]) RS spec;
-
-val not_eq_UU = prove_goal (the_context ()) "ALL p. ~p=UU <-> (p=TT | p=FF)"
- (fn _ => [tac]) RS spec;
-
-val adm_not_eq_tr = prove_goal (the_context ()) "ALL p::tr. adm(%x. ~t(x)=p)"
- (fn _ => [rtac tr_induct 1,
- REPEAT(simp_tac (LCF_ss addsimps [not_eq_TT,not_eq_FF,not_eq_UU]) 1 THEN
- REPEAT(rstac [adm_disj,adm_eq] 1))]) RS spec;
-
-val adm_lemmas = [adm_not_free,adm_eq,adm_less,adm_not_less,adm_not_eq_tr,
- adm_conj,adm_disj,adm_imp,adm_all];
-
-fun induct_tac v i = res_inst_tac[("f",v)] induct i THEN
- REPEAT(rstac adm_lemmas i);
-
-
-val least_FIX = prove_goal (the_context ()) "f(p) = p ==> FIX(f) << p"
- (fn [prem] => [induct_tac "f" 1, rtac minimal 1, strip_tac 1,
- stac (prem RS sym) 1, etac less_ap_term 1]);
-
-val lfp_is_FIX = prove_goal (the_context ())
- "[| f(p) = p; ALL q. f(q)=q --> p << q |] ==> p = FIX(f)"
- (fn [prem1,prem2] => [rtac less_anti_sym 1,
- rtac (prem2 RS spec RS mp) 1, rtac FIX_eq 1,
- rtac least_FIX 1, rtac prem1 1]);
-
-val ffix = read_instantiate [("f","f::?'a=>?'a")] FIX_eq;
-val gfix = read_instantiate [("f","g::?'a=>?'a")] FIX_eq;
-val ss = LCF_ss addsimps [ffix,gfix];
-
-val FIX_pair = prove_goal (the_context ())
- "<FIX(f),FIX(g)> = FIX(%p.<f(FST(p)),g(SND(p))>)"
- (fn _ => [rtac lfp_is_FIX 1, simp_tac ss 1,
- strip_tac 1, simp_tac (LCF_ss addsimps [PROD_less]) 1,
- rtac conjI 1, rtac least_FIX 1, etac subst 1, rtac (FST RS sym) 1,
- rtac least_FIX 1, etac subst 1, rtac (SND RS sym) 1]);
-
-val FIX_pair_conj = rewrite_rule (map mk_meta_eq [PROD_eq,FST,SND]) FIX_pair;
-
-val FIX1 = FIX_pair_conj RS conjunct1;
-val FIX2 = FIX_pair_conj RS conjunct2;
-
-val induct2 = prove_goal (the_context ())
- "[| adm(%p. P(FST(p),SND(p))); P(UU::'a,UU::'b);\
-\ ALL x y. P(x,y) --> P(f(x),g(y)) |] ==> P(FIX(f),FIX(g))"
- (fn prems => [EVERY1
- [res_inst_tac [("f","f"),("g","g")] (standard(FIX1 RS ssubst)),
- res_inst_tac [("f","f"),("g","g")] (standard(FIX2 RS ssubst)),
- res_inst_tac [("f","%x. <f(FST(x)),g(SND(x))>")] induct,
- rstac prems, simp_tac ss, rstac prems,
- simp_tac (LCF_ss addsimps [expand_all_PROD]), rstac prems]]);
-
-fun induct2_tac (f,g) i = res_inst_tac[("f",f),("g",g)] induct2 i THEN
- REPEAT(rstac adm_lemmas i);
--- a/src/LCF/pair.ML Thu Jun 01 21:14:54 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,39 +0,0 @@
-(* Title: LCF/pair
- ID: $Id$
- Author: Tobias Nipkow
- Copyright 1992 University of Cambridge
-*)
-
-val expand_all_PROD = prove_goal (the_context ())
- "(ALL p. P(p)) <-> (ALL x y. P(<x,y>))"
- (fn _ => [rtac iffI 1, fast_tac FOL_cs 1, rtac allI 1,
- rtac (surj_pairing RS subst) 1, fast_tac FOL_cs 1]);
-
-local
-val ppair = read_instantiate [("z","p::'a*'b")] surj_pairing;
-val qpair = read_instantiate [("z","q::'a*'b")] surj_pairing;
-in
-val PROD_less = prove_goal (the_context ())
- "(p::'a*'b) << q <-> FST(p) << FST(q) & SND(p) << SND(q)"
- (fn _ => [EVERY1[rtac iffI,
- rtac conjI, etac less_ap_term, etac less_ap_term,
- rtac (ppair RS subst), rtac (qpair RS subst),
- etac conjE, rtac mono, etac less_ap_term, atac]]);
-end;
-
-val PROD_eq = prove_goal (the_context ()) "p=q <-> FST(p)=FST(q) & SND(p)=SND(q)"
- (fn _ => [rtac iffI 1, asm_simp_tac LCF_ss 1,
- rewtac eq_def,
- asm_simp_tac (LCF_ss addsimps [PROD_less]) 1]);
-
-val PAIR_less = prove_goal (the_context ()) "<a,b> << <c,d> <-> a<<c & b<<d"
- (fn _ => [simp_tac (LCF_ss addsimps [PROD_less])1]);
-
-val PAIR_eq = prove_goal (the_context ()) "<a,b> = <c,d> <-> a=c & b=d"
- (fn _ => [simp_tac (LCF_ss addsimps [PROD_eq])1]);
-
-val UU_is_UU_UU = prove_goal (the_context ()) "<UU,UU> << UU"
- (fn _ => [simp_tac (LCF_ss addsimps [PROD_less]) 1])
- RS less_UU RS sym;
-
-val LCF_ss = LCF_ss addsimps [PAIR_less,PAIR_eq,UU_is_UU_UU];