# HG changeset patch # User wenzelm # Date 976230274 -3600 # Node ID b115460e5c505528d35814ffe860fa33726ab7f4 # Parent b4c3af8ebada194ebd8c6f0621411cdb270fdab2 unsymbolize; diff -r b4c3af8ebada -r b115460e5c50 src/ZF/Integ/Bin.ML --- a/src/ZF/Integ/Bin.ML Thu Dec 07 22:35:25 2000 +0100 +++ b/src/ZF/Integ/Bin.ML Fri Dec 08 00:04:34 2000 +0100 @@ -532,7 +532,7 @@ by Auto_tac; qed "nat_le_int0"; -Goal "$# n = #0 \\ natify(n) = 0"; +Goal "$# n = #0 ==> natify(n) = 0"; by (rtac not_znegative_imp_zero 1); by Auto_tac; qed "int_of_eq_0_imp_natify_eq_0"; @@ -545,7 +545,7 @@ by (blast_tac (claset() addIs [int_of_eq_0_imp_natify_eq_0, sym]) 1); qed "nat_of_zminus_int_of"; -Goal "#0 $<= z \\ $# nat_of(z) = intify(z)"; +Goal "#0 $<= z ==> $# nat_of(z) = intify(z)"; by (rtac not_zneg_nat_of_intify 1); by (asm_simp_tac (simpset() addsimps [znegative_iff_zless_0, not_zless_iff_zle]) 1); diff -r b4c3af8ebada -r b115460e5c50 src/ZF/Integ/IntDiv.ML --- a/src/ZF/Integ/IntDiv.ML Thu Dec 07 22:35:25 2000 +0100 +++ b/src/ZF/Integ/IntDiv.ML Fri Dec 08 00:04:34 2000 +0100 @@ -383,8 +383,8 @@ Addsimps [adjust_eq]; -Goal "\\#0 $< b; \\ a $< b\\ \ -\ \\ nat_of(a $- #2 $\\ b $+ #1) < nat_of(a $- b $+ #1)"; +Goal "[| #0 $< b; \\ a $< b |] \ +\ ==> nat_of(a $- #2 $\\ b $+ #1) < nat_of(a $- b $+ #1)"; by (simp_tac (simpset() addsimps [zless_nat_conj]) 1); by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle, zless_add1_iff_zle]@zcompare_rls) 1); @@ -404,7 +404,7 @@ Goal "[| !!a b. [| a: int; b: int; \ \ ~ (a $< b | b $<= #0) --> P() |] \ \ ==> P() |] \ -\ ==> : int*int \\ P()"; +\ ==> : int*int --> P()"; by (res_inst_tac [("a","")] wf_induct 1); by (res_inst_tac [("A","int*int"), ("f","%.nat_of (a $- b $+ #1)")] wf_measure 1);