src/HOL/ex/Puzzle.ML
changeset 5456 d2ab1264afd4
parent 5069 3ea049f7979d
child 5479 5a5dfb0f0d7d
equal deleted inserted replaced
5455:6712d95c8113 5456:d2ab1264afd4
     1 (*  Title:      HOL/ex/puzzle.ML
     1 (*  Title:      HOL/ex/puzzle.ML
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Tobias Nipkow
     3     Author:     Tobias Nipkow
     4     Copyright   1993 TU Muenchen
     4     Copyright   1993 TU Muenchen
     5 
     5 
     6 For puzzle.thy.  A question from "Bundeswettbewerb Mathematik"
     6 A question from "Bundeswettbewerb Mathematik"
     7 
     7 
     8 Proof due to Herbert Ehler
     8 Proof due to Herbert Ehler
     9 *)
     9 *)
       
    10 
       
    11 AddSIs [Puzzle.f_ax];
    10 
    12 
    11 (*specialized form of induction needed below*)
    13 (*specialized form of induction needed below*)
    12 val prems = goal Nat.thy "[| P(0); !!n. P(Suc(n)) |] ==> !n. P(n)";
    14 val prems = goal Nat.thy "[| P(0); !!n. P(Suc(n)) |] ==> !n. P(n)";
    13 by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]);
    15 by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]);
    14 qed "nat_exh";
    16 qed "nat_exh";
    19 by (Simp_tac 1);
    21 by (Simp_tac 1);
    20 by (rtac impI 1);
    22 by (rtac impI 1);
    21 by (rtac classical 1);
    23 by (rtac classical 1);
    22 by (dtac not_leE 1);
    24 by (dtac not_leE 1);
    23 by (subgoal_tac "f(na) <= f(f(na))" 1);
    25 by (subgoal_tac "f(na) <= f(f(na))" 1);
    24 by (fast_tac (claset() addIs [Puzzle.f_ax]) 2);
    26 by (Blast_tac 2);
    25 by (rtac Suc_leI 1);
    27 by (rtac Suc_leI 1);
    26 by (fast_tac (claset() delrules [order_refl] 
    28 by (blast_tac (claset() addSDs [spec] addIs [le_less_trans]) 1);
    27                       addIs [Puzzle.f_ax, le_less_trans]) 1);
       
    28 val lemma = result() RS spec RS mp;
    29 val lemma = result() RS spec RS mp;
    29 
    30 
    30 Goal "n <= f(n)";
    31 Goal "n <= f(n)";
    31 by (fast_tac (claset() addIs [lemma]) 1);
    32 by (fast_tac (claset() addIs [lemma]) 1);
    32 qed "lemma1";
    33 qed "lemma1";
    33 
    34 
    34 Goal "f(n) < f(Suc(n))";
    35 Goal "f(n) < f(Suc(n))";
    35 by (deepen_tac (claset() addIs [Puzzle.f_ax, le_less_trans, lemma1]) 0 1);
    36 by (blast_tac (claset() addIs [le_less_trans, lemma1]) 1);
    36 qed "lemma2";
    37 qed "lemma2";
    37 
    38 
    38 val prems = goal Puzzle.thy "(!!n. f(n) <= f(Suc(n))) ==> m<n --> f(m) <= f(n)";
    39 val prems = Goal "(!!n. f(n) <= f(Suc(n))) ==> m<n --> f(m) <= f(n)";
    39 by (res_inst_tac[("n","n")]nat_induct 1);
    40 by (res_inst_tac[("n","n")]nat_induct 1);
    40 by (Simp_tac 1);
    41 by (Simp_tac 1);
    41 by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
    42 by (simp_tac (simpset() addsimps [less_Suc_eq]) 1);
       
    43   (*Blast_tac: PROOF FAILED.  Perhaps remove DETERM for unsafe tactics,
       
    44     or stop rotating prems for recursive rules: the le-assumptions
       
    45     have got out of order.*)
    42 by (best_tac (claset() addIs (le_trans::prems)) 1);
    46 by (best_tac (claset() addIs (le_trans::prems)) 1);
    43 qed_spec_mp "mono_lemma1";
    47 qed_spec_mp "mono_lemma1";
    44 
    48 
    45 val [p1,p2] = goal Puzzle.thy
    49 val [p1,p2] = goal Puzzle.thy
    46     "[| !! n. f(n)<=f(Suc(n));  m<=n |] ==> f(m) <= f(n)";
    50     "[| !! n. f(n)<=f(Suc(n));  m<=n |] ==> f(m) <= f(n)";