# HG changeset patch # User paulson # Date 962882916 -7200 # Node ID 051592f4236aedfa11e783a03005bcf27becac0b # Parent 53e09e592278ef4cabaafb2b54f8d6dc0f422965 removal of batch style, and tidying diff -r 53e09e592278 -r 051592f4236a src/FOL/FOL_lemmas1.ML --- a/src/FOL/FOL_lemmas1.ML Thu Jul 06 13:11:32 2000 +0200 +++ b/src/FOL/FOL_lemmas1.ML Thu Jul 06 13:28:36 2000 +0200 @@ -12,39 +12,40 @@ (*** Classical introduction rules for | and EX ***) -qed_goal "disjCI" (the_context ()) - "(~Q ==> P) ==> P|Q" - (fn prems=> - [ (rtac classical 1), - (REPEAT (ares_tac (prems@[disjI1,notI]) 1)), - (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]); +val prems = Goal "(~Q ==> P) ==> P|Q"; +by (rtac classical 1); +by (REPEAT (ares_tac (prems@[disjI1,notI]) 1)); +by (REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ; +qed "disjCI"; (*introduction rule involving only EX*) -qed_goal "ex_classical" (the_context ()) - "( ~(EX x. P(x)) ==> P(a)) ==> EX x. P(x)" - (fn prems=> - [ (rtac classical 1), - (eresolve_tac (prems RL [exI]) 1) ]); +val prems = Goal "( ~(EX x. P(x)) ==> P(a)) ==> EX x. P(x)"; +by (rtac classical 1); +by (eresolve_tac (prems RL [exI]) 1) ; +qed "ex_classical"; (*version of above, simplifying ~EX to ALL~ *) -qed_goal "exCI" (the_context ()) - "(ALL x. ~P(x) ==> P(a)) ==> EX x. P(x)" - (fn [prem]=> - [ (rtac ex_classical 1), - (resolve_tac [notI RS allI RS prem] 1), - (etac notE 1), - (etac exI 1) ]); +val [prem]= Goal "(ALL x. ~P(x) ==> P(a)) ==> EX x. P(x)"; +by (rtac ex_classical 1); +by (resolve_tac [notI RS allI RS prem] 1); +by (etac notE 1); +by (etac exI 1) ; +qed "exCI"; -qed_goal "excluded_middle" (the_context ()) "~P | P" - (fn _=> [ rtac disjCI 1, assume_tac 1 ]); +Goal"~P | P"; +by (rtac disjCI 1); +by (assume_tac 1) ; +qed "excluded_middle"; (*For disjunctive case analysis*) fun excluded_middle_tac sP = res_inst_tac [("Q",sP)] (excluded_middle RS disjE); -qed_goal "case_split_thm" (the_context ()) "[| P ==> Q; ~P ==> Q |] ==> Q" - (fn [p1,p2] => [rtac (excluded_middle RS disjE) 1, - etac p2 1, etac p1 1]); +val [p1,p2] = Goal"[| P ==> Q; ~P ==> Q |] ==> Q"; +by (rtac (excluded_middle RS disjE) 1); +by (etac p2 1); +by (etac p1 1); +qed "case_split_thm"; (*HOL's more natural case analysis tactic*) fun case_tac a = res_inst_tac [("P",a)] case_split_thm; @@ -54,39 +55,41 @@ (*Classical implies (-->) elimination. *) -qed_goal "impCE" (the_context ()) - "[| P-->Q; ~P ==> R; Q ==> R |] ==> R" - (fn major::prems=> - [ (resolve_tac [excluded_middle RS disjE] 1), - (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); +val major::prems = Goal "[| P-->Q; ~P ==> R; Q ==> R |] ==> R"; +by (resolve_tac [excluded_middle RS disjE] 1); +by (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ; +qed "impCE"; (*This version of --> elimination works on Q before P. It works best for those cases in which P holds "almost everywhere". Can't install as default: would break old proofs.*) -qed_goal "impCE'" (the_context ()) - "[| P-->Q; Q ==> R; ~P ==> R |] ==> R" - (fn major::prems=> - [ (resolve_tac [excluded_middle RS disjE] 1), - (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ]); +val major::prems = Goal "[| P-->Q; Q ==> R; ~P ==> R |] ==> R"; +by (resolve_tac [excluded_middle RS disjE] 1); +by (DEPTH_SOLVE (ares_tac (prems@[major RS mp]) 1)) ; +qed "impCE'"; (*Double negation law*) -qed_goal "notnotD" (the_context ()) "~~P ==> P" - (fn [major]=> - [ (rtac classical 1), (eresolve_tac [major RS notE] 1) ]); +Goal"~~P ==> P"; +by (rtac classical 1); +by (etac notE 1); +by (assume_tac 1); +qed "notnotD"; -qed_goal "contrapos2" (the_context ()) "[| Q; ~ P ==> ~ Q |] ==> P" (fn [p1,p2] => [ - rtac classical 1, - dtac p2 1, - etac notE 1, - rtac p1 1]); +val [p1,p2] = Goal"[| Q; ~ P ==> ~ Q |] ==> P"; +by (rtac classical 1); +by (dtac p2 1); +by (etac notE 1); +by (rtac p1 1); +qed "contrapos2"; (*** Tactics for implication and contradiction ***) (*Classical <-> elimination. Proof substitutes P=Q in ~P ==> ~Q and P ==> Q *) -qed_goalw "iffCE" (the_context ()) [iff_def] - "[| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R" - (fn prems => - [ (rtac conjE 1), - (REPEAT (DEPTH_SOLVE_1 - (etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]); +val major::prems = +Goalw [iff_def] "[| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R"; +by (rtac (major RS conjE) 1); +by (REPEAT_FIRST (etac impCE)); +by (REPEAT (DEPTH_SOLVE_1 (mp_tac 1 ORELSE ares_tac prems 1))); +qed "iffCE"; + diff -r 53e09e592278 -r 051592f4236a src/FOL/FOL_lemmas2.ML --- a/src/FOL/FOL_lemmas2.ML Thu Jul 06 13:11:32 2000 +0200 +++ b/src/FOL/FOL_lemmas2.ML Thu Jul 06 13:28:36 2000 +0200 @@ -1,4 +1,4 @@ -Goal "!!a b c. [| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c"; +Goal "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c"; by (Blast_tac 1); qed "ex1_functional"; diff -r 53e09e592278 -r 051592f4236a src/FOL/IFOL_lemmas.ML --- a/src/FOL/IFOL_lemmas.ML Thu Jul 06 13:11:32 2000 +0200 +++ b/src/FOL/IFOL_lemmas.ML Thu Jul 06 13:28:36 2000 +0200 @@ -32,63 +32,78 @@ -qed_goalw "TrueI" (the_context ()) [True_def] "True" - (fn _ => [ (REPEAT (ares_tac [impI] 1)) ]); +Goalw [True_def] "True"; +by (REPEAT (ares_tac [impI] 1)) ; +qed "TrueI"; (*** Sequent-style elimination rules for & --> and ALL ***) -qed_goal "conjE" (the_context ()) - "[| P&Q; [| P; Q |] ==> R |] ==> R" - (fn prems=> - [ (REPEAT (resolve_tac prems 1 - ORELSE (resolve_tac [conjunct1, conjunct2] 1 THEN - resolve_tac prems 1))) ]); +val major::prems = Goal + "[| P&Q; [| P; Q |] ==> R |] ==> R"; +by (resolve_tac prems 1); +by (rtac (major RS conjunct1) 1); +by (rtac (major RS conjunct2) 1); +qed "conjE"; -qed_goal "impE" (the_context ()) - "[| P-->Q; P; Q ==> R |] ==> R" - (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); +val major::prems = Goal + "[| P-->Q; P; Q ==> R |] ==> R"; +by (resolve_tac prems 1); +by (rtac (major RS mp) 1); +by (resolve_tac prems 1); +qed "impE"; -qed_goal "allE" (the_context ()) - "[| ALL x. P(x); P(x) ==> R |] ==> R" - (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); +val major::prems = Goal + "[| ALL x. P(x); P(x) ==> R |] ==> R"; +by (resolve_tac prems 1); +by (rtac (major RS spec) 1); +qed "allE"; (*Duplicates the quantifier; for use with eresolve_tac*) -qed_goal "all_dupE" (the_context ()) +val major::prems = Goal "[| ALL x. P(x); [| P(x); ALL x. P(x) |] ==> R \ -\ |] ==> R" - (fn prems=> [ (REPEAT (resolve_tac (prems@[spec]) 1)) ]); +\ |] ==> R"; +by (resolve_tac prems 1); +by (rtac (major RS spec) 1); +by (rtac major 1); +qed "all_dupE"; (*** Negation rules, which translate between ~P and P-->False ***) -qed_goalw "notI" (the_context ()) [not_def] "(P ==> False) ==> ~P" - (fn prems=> [ (REPEAT (ares_tac (prems@[impI]) 1)) ]); +val prems = Goalw [not_def] "(P ==> False) ==> ~P"; +by (REPEAT (ares_tac (prems@[impI]) 1)) ; +qed "notI"; -qed_goalw "notE" (the_context ()) [not_def] "[| ~P; P |] ==> R" - (fn prems=> - [ (resolve_tac [mp RS FalseE] 1), - (REPEAT (resolve_tac prems 1)) ]); +Goalw [not_def] "[| ~P; P |] ==> R"; +by (etac (mp RS FalseE) 1); +by (assume_tac 1); +qed "notE"; -qed_goal "rev_notE" (the_context ()) "!!P R. [| P; ~P |] ==> R" - (fn _ => [REPEAT (ares_tac [notE] 1)]); +Goal "[| P; ~P |] ==> R"; +by (etac notE 1); +by (assume_tac 1); +qed "rev_notE"; (*This is useful with the special implication rules for each kind of P. *) -qed_goal "not_to_imp" (the_context ()) - "[| ~P; (P-->False) ==> Q |] ==> Q" - (fn prems=> [ (REPEAT (ares_tac (prems@[impI,notE]) 1)) ]); +val prems = Goal + "[| ~P; (P-->False) ==> Q |] ==> Q"; +by (REPEAT (ares_tac (prems@[impI,notE]) 1)) ; +qed "not_to_imp"; (* For substitution into an assumption P, reduce Q to P-->Q, substitute into this implication, then apply impI to move P back into the assumptions. To specify P use something like eres_inst_tac [ ("P","ALL y. ?S(x,y)") ] rev_mp 1 *) -qed_goal "rev_mp" (the_context ()) "[| P; P --> Q |] ==> Q" - (fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]); +Goal "[| P; P --> Q |] ==> Q"; +by (etac mp 1); +by (assume_tac 1); +qed "rev_mp"; (*Contrapositive of an inference rule*) -qed_goal "contrapos" (the_context ()) "[| ~Q; P==>Q |] ==> ~P" - (fn [major,minor]=> - [ (rtac (major RS notE RS notI) 1), - (etac minor 1) ]); +val [major,minor]= Goal "[| ~Q; P==>Q |] ==> ~P"; +by (rtac (major RS notE RS notI) 1); +by (etac minor 1) ; +qed "contrapos"; (*** Modus Ponens Tactics ***) @@ -102,44 +117,55 @@ (*** If-and-only-if ***) -qed_goalw "iffI" (the_context ()) [iff_def] - "[| P ==> Q; Q ==> P |] ==> P<->Q" - (fn prems=> [ (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ]); +val prems = Goalw [iff_def] + "[| P ==> Q; Q ==> P |] ==> P<->Q"; +by (REPEAT (ares_tac (prems@[conjI, impI]) 1)) ; +qed "iffI"; (*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) -qed_goalw "iffE" (the_context ()) [iff_def] - "[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R" - (fn prems => [ (rtac conjE 1), (REPEAT (ares_tac prems 1)) ]); +val prems = Goalw [iff_def] + "[| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R"; +by (rtac conjE 1); +by (REPEAT (ares_tac prems 1)) ; +qed "iffE"; (* Destruct rules for <-> similar to Modus Ponens *) -qed_goalw "iffD1" (the_context ()) [iff_def] "[| P <-> Q; P |] ==> Q" - (fn prems => [ (rtac (conjunct1 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); - -qed_goalw "iffD2" (the_context ()) [iff_def] "[| P <-> Q; Q |] ==> P" - (fn prems => [ (rtac (conjunct2 RS mp) 1), (REPEAT (ares_tac prems 1)) ]); +Goalw [iff_def] "[| P <-> Q; P |] ==> Q"; +by (etac (conjunct1 RS mp) 1); +by (assume_tac 1); +qed "iffD1"; -qed_goal "rev_iffD1" (the_context ()) "!!P. [| P; P <-> Q |] ==> Q" - (fn _ => [etac iffD1 1, assume_tac 1]); +val prems = Goalw [iff_def] "[| P <-> Q; Q |] ==> P"; +by (etac (conjunct2 RS mp) 1); +by (assume_tac 1); +qed "iffD2"; -qed_goal "rev_iffD2" (the_context ()) "!!P. [| Q; P <-> Q |] ==> P" - (fn _ => [etac iffD2 1, assume_tac 1]); +Goal "[| P; P <-> Q |] ==> Q"; +by (etac iffD1 1); +by (assume_tac 1); +qed "rev_iffD1"; -qed_goal "iff_refl" (the_context ()) "P <-> P" - (fn _ => [ (REPEAT (ares_tac [iffI] 1)) ]); +Goal "[| Q; P <-> Q |] ==> P"; +by (etac iffD2 1); +by (assume_tac 1); +qed "rev_iffD2"; + +Goal "P <-> P"; +by (REPEAT (ares_tac [iffI] 1)) ; +qed "iff_refl"; -qed_goal "iff_sym" (the_context ()) "Q <-> P ==> P <-> Q" - (fn [major] => - [ (rtac (major RS iffE) 1), - (rtac iffI 1), - (REPEAT (eresolve_tac [asm_rl,mp] 1)) ]); +Goal "Q <-> P ==> P <-> Q"; +by (etac iffE 1); +by (rtac iffI 1); +by (REPEAT (eresolve_tac [asm_rl,mp] 1)) ; +qed "iff_sym"; -qed_goal "iff_trans" (the_context ()) - "!!P Q R. [| P <-> Q; Q<-> R |] ==> P <-> R" - (fn _ => - [ (rtac iffI 1), - (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ]); +Goal "[| P <-> Q; Q<-> R |] ==> P <-> R"; +by (rtac iffI 1); +by (REPEAT (eresolve_tac [asm_rl,iffE] 1 ORELSE mp_tac 1)) ; +qed "iff_trans"; (*** Unique existence. NOTE THAT the following 2 quantifications @@ -148,21 +174,23 @@ do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. ***) -qed_goalw "ex1I" (the_context ()) [ex1_def] - "[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)" - (fn prems => [ (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ]); +val prems = Goalw [ex1_def] + "[| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)"; +by (REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)) ; +qed "ex1I"; (*Sometimes easier to use: the premises have no shared variables. Safe!*) -qed_goal "ex_ex1I" (the_context ()) - "[| EX x. P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)" - (fn [ex,eq] => [ (rtac (ex RS exE) 1), - (REPEAT (ares_tac [ex1I,eq] 1)) ]); +val [ex,eq] = Goal + "[| EX x. P(x); !!x y. [| P(x); P(y) |] ==> x=y |] ==> EX! x. P(x)"; +by (rtac (ex RS exE) 1); +by (REPEAT (ares_tac [ex1I,eq] 1)) ; +qed "ex_ex1I"; -qed_goalw "ex1E" (the_context ()) [ex1_def] - "[| EX! x. P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ]); +val prems = Goalw [ex1_def] + "[| EX! x. P(x); !!x. [| P(x); ALL y. P(y) --> y=x |] ==> R |] ==> R"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [exE,conjE] 1 ORELSE ares_tac prems 1)) ; +qed "ex1E"; (*** <-> congruence rules for simplification ***) @@ -172,83 +200,78 @@ resolve_tac (prems RL [iffE]) i THEN REPEAT1 (eresolve_tac [asm_rl,mp] i); -qed_goal "conj_cong" (the_context ()) - "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (ares_tac [iffI,conjI] 1 - ORELSE eresolve_tac [iffE,conjE,mp] 1 - ORELSE iff_tac prems 1)) ]); +val prems = Goal + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P&Q) <-> (P'&Q')"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 + ORELSE iff_tac prems 1)) ; +qed "conj_cong"; (*Reversed congruence rule! Used in ZF/Order*) -qed_goal "conj_cong2" (the_context ()) - "[| P <-> P'; P' ==> Q <-> Q' |] ==> (Q&P) <-> (Q'&P')" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (ares_tac [iffI,conjI] 1 - ORELSE eresolve_tac [iffE,conjE,mp] 1 - ORELSE iff_tac prems 1)) ]); +val prems = Goal + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (Q&P) <-> (Q'&P')"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [iffI,conjI] 1 + ORELSE eresolve_tac [iffE,conjE,mp] 1 ORELSE iff_tac prems 1)) ; +qed "conj_cong2"; -qed_goal "disj_cong" (the_context ()) - "[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 - ORELSE ares_tac [iffI] 1 - ORELSE mp_tac 1)) ]); +val prems = Goal + "[| P <-> P'; Q <-> Q' |] ==> (P|Q) <-> (P'|Q')"; +by (cut_facts_tac prems 1); +by (REPEAT (eresolve_tac [iffE,disjE,disjI1,disjI2] 1 + ORELSE ares_tac [iffI] 1 ORELSE mp_tac 1)) ; +qed "disj_cong"; -qed_goal "imp_cong" (the_context ()) - "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (ares_tac [iffI,impI] 1 - ORELSE etac iffE 1 - ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ]); +val prems = Goal + "[| P <-> P'; P' ==> Q <-> Q' |] ==> (P-->Q) <-> (P'-->Q')"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [iffI,impI] 1 + ORELSE etac iffE 1 ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ; +qed "imp_cong"; -qed_goal "iff_cong" (the_context ()) - "[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (etac iffE 1 - ORELSE ares_tac [iffI] 1 - ORELSE mp_tac 1)) ]); +val prems = Goal + "[| P <-> P'; Q <-> Q' |] ==> (P<->Q) <-> (P'<->Q')"; +by (cut_facts_tac prems 1); +by (REPEAT (etac iffE 1 ORELSE ares_tac [iffI] 1 ORELSE mp_tac 1)) ; +qed "iff_cong"; -qed_goal "not_cong" (the_context ()) - "P <-> P' ==> ~P <-> ~P'" - (fn prems => - [ (cut_facts_tac prems 1), - (REPEAT (ares_tac [iffI,notI] 1 - ORELSE mp_tac 1 - ORELSE eresolve_tac [iffE,notE] 1)) ]); +val prems = Goal + "P <-> P' ==> ~P <-> ~P'"; +by (cut_facts_tac prems 1); +by (REPEAT (ares_tac [iffI,notI] 1 + ORELSE mp_tac 1 ORELSE eresolve_tac [iffE,notE] 1)) ; +qed "not_cong"; -qed_goal "all_cong" (the_context ()) - "(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))" - (fn prems => - [ (REPEAT (ares_tac [iffI,allI] 1 - ORELSE mp_tac 1 - ORELSE etac allE 1 ORELSE iff_tac prems 1)) ]); +val prems = Goal + "(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))"; +by (REPEAT (ares_tac [iffI,allI] 1 + ORELSE mp_tac 1 ORELSE etac allE 1 ORELSE iff_tac prems 1)) ; +qed "all_cong"; -qed_goal "ex_cong" (the_context ()) - "(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))" - (fn prems => - [ (REPEAT (etac exE 1 ORELSE ares_tac [iffI,exI] 1 - ORELSE mp_tac 1 - ORELSE iff_tac prems 1)) ]); +val prems = Goal + "(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))"; +by (REPEAT (etac exE 1 ORELSE ares_tac [iffI,exI] 1 + ORELSE mp_tac 1 ORELSE iff_tac prems 1)) ; +qed "ex_cong"; -qed_goal "ex1_cong" (the_context ()) - "(!!x. P(x) <-> Q(x)) ==> (EX! x. P(x)) <-> (EX! x. Q(x))" - (fn prems => - [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 - ORELSE mp_tac 1 - ORELSE iff_tac prems 1)) ]); +val prems = Goal + "(!!x. P(x) <-> Q(x)) ==> (EX! x. P(x)) <-> (EX! x. Q(x))"; +by (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 + ORELSE ares_tac [iffI,ex1I] 1 ORELSE mp_tac 1 + ORELSE iff_tac prems 1)) ; +qed "ex1_cong"; (*** Equality rules ***) -qed_goal "sym" (the_context ()) "a=b ==> b=a" - (fn [major] => [ (rtac (major RS subst) 1), (rtac refl 1) ]); +Goal "a=b ==> b=a"; +by (etac subst 1); +by (rtac refl 1) ; +qed "sym"; -qed_goal "trans" (the_context ()) "[| a=b; b=c |] ==> a=c" - (fn [prem1,prem2] => [ (rtac (prem2 RS subst) 1), (rtac prem1 1) ]); +Goal "[| a=b; b=c |] ==> a=c"; +by (etac subst 1 THEN assume_tac 1) ; +qed "trans"; (** ~ b=a ==> ~ a=b **) bind_thm ("not_sym", hd (compose(sym,2,contrapos))); @@ -257,12 +280,12 @@ (* Two theorms for rewriting only one instance of a definition: the first for definitions of formulae and the second for terms *) -val prems = goal (the_context ()) "(A == B) ==> A <-> B"; +val prems = goal (the_context()) "(A == B) ==> A <-> B"; by (rewrite_goals_tac prems); by (rtac iff_refl 1); qed "def_imp_iff"; -val prems = goal (the_context ()) "(A == B) ==> A = B"; +val prems = goal (the_context()) "(A == B) ==> A = B"; by (rewrite_goals_tac prems); by (rtac refl 1); qed "meta_eq_to_obj_eq"; @@ -279,75 +302,67 @@ end; (*A special case of ex1E that would otherwise need quantifier expansion*) -qed_goal "ex1_equalsE" (the_context ()) - "[| EX! x. P(x); P(a); P(b) |] ==> a=b" - (fn prems => - [ (cut_facts_tac prems 1), - (etac ex1E 1), - (rtac trans 1), - (rtac sym 2), - (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ]); +val prems = Goal + "[| EX! x. P(x); P(a); P(b) |] ==> a=b"; +by (cut_facts_tac prems 1); +by (etac ex1E 1); +by (rtac trans 1); +by (rtac sym 2); +by (REPEAT (eresolve_tac [asm_rl, spec RS mp] 1)) ; +qed "ex1_equalsE"; (** Polymorphic congruence rules **) -qed_goal "subst_context" (the_context ()) - "[| a=b |] ==> t(a)=t(b)" - (fn prems=> - [ (resolve_tac (prems RL [ssubst]) 1), - (rtac refl 1) ]); +Goal "[| a=b |] ==> t(a)=t(b)"; +by (etac ssubst 1); +by (rtac refl 1) ; +qed "subst_context"; -qed_goal "subst_context2" (the_context ()) - "[| a=b; c=d |] ==> t(a,c)=t(b,d)" - (fn prems=> - [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); +Goal "[| a=b; c=d |] ==> t(a,c)=t(b,d)"; +by (REPEAT (etac ssubst 1)); +by (rtac refl 1) ; +qed "subst_context2"; -qed_goal "subst_context3" (the_context ()) - "[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)" - (fn prems=> - [ (EVERY1 (map rtac ((prems RL [ssubst]) @ [refl]))) ]); +Goal "[| a=b; c=d; e=f |] ==> t(a,c,e)=t(b,d,f)"; +by (REPEAT (etac ssubst 1)); +by (rtac refl 1) ; +qed "subst_context3"; (*Useful with eresolve_tac for proving equalties from known equalities. a = b | | c = d *) -qed_goal "box_equals" (the_context ()) - "[| a=b; a=c; b=d |] ==> c=d" - (fn prems=> - [ (rtac trans 1), - (rtac trans 1), - (rtac sym 1), - (REPEAT (resolve_tac prems 1)) ]); +Goal "[| a=b; a=c; b=d |] ==> c=d"; +by (rtac trans 1); +by (rtac trans 1); +by (rtac sym 1); +by (REPEAT (assume_tac 1)); +qed "box_equals"; (*Dual of box_equals: for proving equalities backwards*) -qed_goal "simp_equals" (the_context ()) - "[| a=c; b=d; c=d |] ==> a=b" - (fn prems=> - [ (rtac trans 1), - (rtac trans 1), - (REPEAT (resolve_tac (prems @ (prems RL [sym])) 1)) ]); +Goal "[| a=c; b=d; c=d |] ==> a=b"; +by (rtac trans 1); +by (rtac trans 1); +by (REPEAT (assume_tac 1)); +by (etac sym 1); +qed "simp_equals"; (** Congruence rules for predicate letters **) -qed_goal "pred1_cong" (the_context ()) - "a=a' ==> P(a) <-> P(a')" - (fn prems => - [ (cut_facts_tac prems 1), - (rtac iffI 1), - (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); +Goal "a=a' ==> P(a) <-> P(a')"; +by (rtac iffI 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; +qed "pred1_cong"; -qed_goal "pred2_cong" (the_context ()) - "[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')" - (fn prems => - [ (cut_facts_tac prems 1), - (rtac iffI 1), - (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); +Goal "[| a=a'; b=b' |] ==> P(a,b) <-> P(a',b')"; +by (rtac iffI 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; +qed "pred2_cong"; -qed_goal "pred3_cong" (the_context ()) - "[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')" - (fn prems => - [ (cut_facts_tac prems 1), - (rtac iffI 1), - (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ]); +Goal "[| a=a'; b=b'; c=c' |] ==> P(a,b,c) <-> P(a',b',c')"; +by (rtac iffI 1); +by (DEPTH_SOLVE (eresolve_tac [asm_rl, subst, ssubst] 1)) ; +qed "pred3_cong"; (*special cases for free variables P, Q, R, S -- up to 3 arguments*) @@ -368,52 +383,52 @@ R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic (preprint, University of St Andrews, 1991) ***) -qed_goal "conj_impE" (the_context ()) - "[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R"; +by (REPEAT (ares_tac ([conjI, impI, major RS mp]@prems) 1)) ; +qed "conj_impE"; -qed_goal "disj_impE" (the_context ()) - "[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R" - (fn major::prems=> - [ (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R"; +by (DEPTH_SOLVE (ares_tac ([disjI1, disjI2, impI, major RS mp]@prems) 1)) ; +qed "disj_impE"; (*Simplifies the implication. Classical version is stronger. Still UNSAFE since Q must be provable -- backtracking needed. *) -qed_goal "imp_impE" (the_context ()) - "[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R"; +by (REPEAT (ares_tac ([impI, major RS mp]@prems) 1)) ; +qed "imp_impE"; (*Simplifies the implication. Classical version is stronger. Still UNSAFE since ~P must be provable -- backtracking needed. *) -qed_goal "not_impE" (the_context ()) - "[| ~P --> S; P ==> False; S ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| ~P --> S; P ==> False; S ==> R |] ==> R"; +by (REPEAT (ares_tac ([notI, impI, major RS mp]@prems) 1)) ; +qed "not_impE"; (*Simplifies the implication. UNSAFE. *) -qed_goal "iff_impE" (the_context ()) +val major::prems= Goal "[| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P; \ -\ S ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ]); +\ S ==> R |] ==> R"; +by (REPEAT (ares_tac ([iffI, impI, major RS mp]@prems) 1)) ; +qed "iff_impE"; (*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) -qed_goal "all_impE" (the_context ()) - "[| (ALL x. P(x))-->S; !!x. P(x); S ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| (ALL x. P(x))-->S; !!x. P(x); S ==> R |] ==> R"; +by (REPEAT (ares_tac ([allI, impI, major RS mp]@prems) 1)) ; +qed "all_impE"; (*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) -qed_goal "ex_impE" (the_context ()) - "[| (EX x. P(x))-->S; P(x)-->S ==> R |] ==> R" - (fn major::prems=> - [ (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ]); +val major::prems= Goal + "[| (EX x. P(x))-->S; P(x)-->S ==> R |] ==> R"; +by (REPEAT (ares_tac ([exI, impI, major RS mp]@prems) 1)) ; +qed "ex_impE"; (*** Courtesy of Krzysztof Grabczewski ***) -val major::prems = goal (the_context ()) "[| P|Q; P==>R; Q==>S |] ==> R|S"; +val major::prems = Goal "[| P|Q; P==>R; Q==>S |] ==> R|S"; by (rtac (major RS disjE) 1); by (REPEAT (eresolve_tac (prems RL [disjI1, disjI2]) 1)); qed "disj_imp_disj"; diff -r 53e09e592278 -r 051592f4236a src/ZF/Resid/Substitution.ML --- a/src/ZF/Resid/Substitution.ML Thu Jul 06 13:11:32 2000 +0200 +++ b/src/ZF/Resid/Substitution.ML Thu Jul 06 13:28:36 2000 +0200 @@ -9,26 +9,23 @@ (* Arithmetic extensions *) (* ------------------------------------------------------------------------- *) -goal Arith.thy - "!!m.[| p < n; n:nat|] ==> n~=p"; +Goal "[| p < n; n:nat|] ==> n~=p"; by (Fast_tac 1); qed "gt_not_eq"; -val succ_pred = prove_goal Arith.thy - "!!i.[|j:nat; i:nat|] ==> i < j --> succ(j #- 1) = j" - (fn prems =>[(induct_tac "j" 1), - (Fast_tac 1), - (Asm_simp_tac 1)]); +Goal "[|j:nat; i:nat|] ==> i < j --> succ(j #- 1) = j"; +by (induct_tac "j" 1); +by (Fast_tac 1); +by (Asm_simp_tac 1); +qed "succ_pred"; -goal Arith.thy - "!!i.[|succ(x) x < n #- 1 "; +Goal "[|succ(x) x < n #- 1 "; by (rtac succ_leE 1); by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1); by (asm_simp_tac (simpset() addsimps [succ_pred]) 1); qed "lt_pred"; -goal Arith.thy - "!!i.[|n < succ(x); p n #- 1 < x "; +Goal "[|n < succ(x); p n #- 1 < x "; by (rtac succ_leE 1); by (asm_simp_tac (simpset() addsimps [succ_pred]) 1); qed "gt_pred"; diff -r 53e09e592278 -r 051592f4236a src/ZF/ex/Limit.ML --- a/src/ZF/ex/Limit.ML Thu Jul 06 13:11:32 2000 +0200 +++ b/src/ZF/ex/Limit.ML Thu Jul 06 13:28:36 2000 +0200 @@ -214,7 +214,7 @@ AddTCs [pcpo_cpo, bot_least, bot_in]; -val prems = goal Limit.thy (* bot_unique *) +val prems = Goal (* bot_unique *) "[| pcpo(D); x:set(D); !!y. y:set(D) ==> rel(D,x,y)|] ==> x = bot(D)"; by (blast_tac (claset() addIs ([cpo_antisym,pcpo_cpo,bot_in,bot_least]@ prems)) 1); @@ -364,14 +364,14 @@ xprem::yprem::prems)); qed "matrix_chainI"; -val lemma = prove_goal Limit.thy - "!!z.[|m : nat; rel(D, (lam n:nat. M`n`n)`m, y)|] ==> rel(D,M`m`m, y)" - (fn prems => [Asm_full_simp_tac 1]); +Goal "[|m : nat; rel(D, (lam n:nat. M`n`n)`m, y)|] ==> rel(D,M`m`m, y)"; +by (Asm_full_simp_tac 1); +qed "lemma"; -val lemma2 = prove_goal Limit.thy - "!!z.[|x:nat; m:nat; rel(D,(lam n:nat. M`n`m1)`x,(lam n:nat. M`n`m1)`m)|] ==> \ -\ rel(D,M`x`m1,M`m`m1)" - (fn prems => [Asm_full_simp_tac 1]); +Goal "[|x:nat; m:nat; rel(D,(lam n:nat. M`n`m1)`x,(lam n:nat. M`n`m1)`m)|] \ +\ ==> rel(D,M`x`m1,M`m`m1)"; +by (Asm_full_simp_tac 1); +qed "lemma2"; Goalw [isub_def] (* isub_lemma *) "[|isub(D, lam n:nat. M`n`n, y); matrix(D,M); cpo(D)|] ==> \ @@ -563,7 +563,7 @@ (* rel_cf originally an equality. Now stated as two rules. Seemed easiest. Besides, now complicated by typing assumptions. *) -val prems = goal Limit.thy +val prems = Goal "[|!!x. x:set(D) ==> rel(E,f`x,g`x); f:cont(D,E); g:cont(D,E)|] ==> \ \ rel(cf(D,E),f,g)"; by (asm_simp_tac (simpset() addsimps [rel_I, cf_def]@prems) 1); @@ -1080,7 +1080,7 @@ brr(lam_type::(chain_iprod RS cpo_lub RS islub_in)::iprodE::prems) 1; qed "islub_iprod"; -val prems = goal Limit.thy (* cpo_iprod *) +val prems = Goal (* cpo_iprod *) "(!!n. n:nat ==> cpo(DD`n)) ==> cpo(iprod(DD))"; brr[cpoI,poI] 1; by (rtac rel_iprodI 1); (* not repeated: want to solve 1 and leave 2 unchanged *) @@ -1230,7 +1230,7 @@ chain_in,nat_succI]) 1); qed "chain_mkcpo"; -val prems = goal Limit.thy (* subcpo_mkcpo *) +val prems = Goal (* subcpo_mkcpo *) "[|!!X. chain(mkcpo(D,P),X) ==> P(lub(D,X)); cpo(D)|] ==> \ \ subcpo(mkcpo(D,P),D)"; brr(subcpoI::subsetI::prems) 1; @@ -1250,15 +1250,16 @@ by (REPEAT (ares_tac prems 1)); qed "emb_chainI"; -val emb_chain_cpo = prove_goalw Limit.thy [emb_chain_def] - "!!x. [|emb_chain(DD,ee); n:nat|] ==> cpo(DD`n)" - (fn prems => [Fast_tac 1]); +Goalw [emb_chain_def] "[|emb_chain(DD,ee); n:nat|] ==> cpo(DD`n)"; +by (Fast_tac 1); +qed "emb_chain_cpo"; AddTCs [emb_chain_cpo]; -val emb_chain_emb = prove_goalw Limit.thy [emb_chain_def] - "!!x. [|emb_chain(DD,ee); n:nat|] ==> emb(DD`n,DD`succ(n),ee`n)" - (fn prems => [Fast_tac 1]); +Goalw [emb_chain_def] + "[|emb_chain(DD,ee); n:nat|] ==> emb(DD`n,DD`succ(n),ee`n)"; +by (Fast_tac 1); +qed "emb_chain_emb"; (*----------------------------------------------------------------------*) (* Dinf, the inverse Limit. *) @@ -1358,12 +1359,11 @@ (* Again, would like more theorems about arithmetic. *) -val add1 = prove_goal Limit.thy - "!!x. n:nat ==> succ(n) = n #+ 1" - (fn prems => [Asm_simp_tac 1]); +Goal "n:nat ==> succ(n) = n #+ 1"; +by (Asm_simp_tac 1); +qed "add1"; -Goal (* succ_sub1 *) - "x:nat ==> 0 < x --> succ(x #- 1)=x"; +Goal "x:nat ==> 0 < x --> succ(x #- 1)=x"; by (induct_tac "x" 1); by Auto_tac; qed "succ_sub1"; @@ -1416,7 +1416,7 @@ by (asm_simp_tac(simpset() addsimps[e_less_le, e_less_eq]) 1); qed "e_less_succ"; -val prems = goal Limit.thy (* e_less_succ_emb *) +val prems = Goal (* e_less_succ_emb *) "[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \ \ e_less(DD,ee,m,succ(m)) = ee`m"; by (asm_simp_tac(simpset() addsimps e_less_succ::prems) 1); @@ -1446,9 +1446,9 @@ by (REPEAT (assume_tac 1)); qed "emb_e_less"; -val comp_mono_eq = prove_goal Limit.thy - "!!z.[|f=f'; g=g'|] ==> f O g = f' O g'" - (fn prems => [Asm_simp_tac 1]); +Goal "[|f=f'; g=g'|] ==> f O g = f' O g'"; +by (Asm_simp_tac 1); +qed "comp_mono_eq"; (* Typing, typing, typing, three irritating assumptions. Extra theorems needed in proof, but no real difficulty. *) @@ -1695,7 +1695,7 @@ by (REPEAT (assume_tac 1)); qed "eps_e_gr"; -val prems = goal Limit.thy (* eps_succ_ee *) +val prems = Goal (* eps_succ_ee *) "[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \ \ eps(DD,ee,m,succ(m)) = ee`m"; by (asm_simp_tac(simpset() addsimps eps_e_less::le_succ_iff::e_less_succ_emb:: @@ -1752,7 +1752,7 @@ (* Arithmetic, little support in Isabelle/ZF. *) -val prems = goal Limit.thy (* le_exists_lemma *) +val prems = Goal (* le_exists_lemma *) "[|n le k; k le m; \ \ !!p q. [|p le q; k=n#+p; m=n#+q; p:nat; q:nat|] ==> R; \ \ m:nat; n:nat; k:nat|]==>R"; @@ -1871,17 +1871,19 @@ by (asm_full_simp_tac(simpset() addsimps[eps_succ_Rp, e_less_eq, id_conv, nat_succI]) 1); qed "rho_emb_fun"; -val rho_emb_apply1 = prove_goalw Limit.thy [rho_emb_def] - "!!z. x:set(DD`n) ==> rho_emb(DD,ee,n)`x = (lam m:nat. eps(DD,ee,n,m)`x)" - (fn prems => [Asm_simp_tac 1]); +Goalw [rho_emb_def] + "x:set(DD`n) ==> rho_emb(DD,ee,n)`x = (lam m:nat. eps(DD,ee,n,m)`x)"; +by (Asm_simp_tac 1); +qed "rho_emb_apply1"; -val rho_emb_apply2 = prove_goalw Limit.thy [rho_emb_def] - "!!z. [|x:set(DD`n); m:nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x" - (fn prems => [Asm_simp_tac 1]); +Goalw [rho_emb_def] + "[|x:set(DD`n); m:nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x"; +by (Asm_simp_tac 1); +qed "rho_emb_apply2"; -val rho_emb_id = prove_goal Limit.thy - "!!z. [| x:set(DD`n); n:nat|] ==> rho_emb(DD,ee,n)`x`n = x" - (fn prems => [asm_simp_tac(simpset() addsimps[rho_emb_apply2,eps_id]) 1]); +Goal "[| x:set(DD`n); n:nat|] ==> rho_emb(DD,ee,n)`x`n = x"; +by (asm_simp_tac(simpset() addsimps[rho_emb_apply2,eps_id]) 1); +qed "rho_emb_id"; (* Shorter proof, 23 against 62. *) @@ -2078,10 +2080,9 @@ by (auto_tac (claset() addIs [eps_fun], simpset())); qed "rho_emb_commute"; -val le_succ = prove_goal Arith.thy "n:nat ==> n le succ(n)" - (fn prems => - [REPEAT (ares_tac - ((disjI1 RS(le_succ_iff RS iffD2))::le_refl::nat_into_Ord::prems) 1)]); +val prems = goal Arith.thy "n:nat ==> n le succ(n)"; +by (REPEAT (ares_tac ((disjI1 RS(le_succ_iff RS iffD2))::le_refl::nat_into_Ord::prems) 1)); +qed "le_succ"; (* Shorter proof: 21 vs 83 (106 - 23, due to OAssoc complication) *) @@ -2348,17 +2349,19 @@ -val mediatingI = prove_goalw Limit.thy [mediating_def] - "[|emb(E,G,t); !!n. n:nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)" - (fn prems => [Safe_tac,REPEAT (ares_tac prems 1)]); +val prems = Goalw [mediating_def] + "[|emb(E,G,t); !!n. n:nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)"; +by (Safe_tac); +by (REPEAT (ares_tac prems 1)); +qed "mediatingI"; -val mediating_emb = prove_goalw Limit.thy [mediating_def] - "!!z. mediating(E,G,r,f,t) ==> emb(E,G,t)" - (fn prems => [Fast_tac 1]); +Goalw [mediating_def] "mediating(E,G,r,f,t) ==> emb(E,G,t)"; +by (Fast_tac 1); +qed "mediating_emb"; -val mediating_eq = prove_goalw Limit.thy [mediating_def] - "!!z. [| mediating(E,G,r,f,t); n:nat |] ==> f(n) = t O r(n)" - (fn prems => [Blast_tac 1]); +Goalw [mediating_def] "[| mediating(E,G,r,f,t); n:nat |] ==> f(n) = t O r(n)"; +by (Blast_tac 1); +qed "mediating_eq"; Goal (* lub_universal_mediating *) "[| lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \