src/Tools/isac/Knowledge/LogExp.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 01 Sep 2010 16:43:58 +0200
branchisac-update-Isa09-2
changeset 37972 66fc615a1e89
parent 37954 4022d670753c
child 37981 b2877b9d455a
permissions -rw-r--r--
fixed @{theory} in all Knowledge/*.thy
     1 (* all outcommented in order to demonstrate authoring:
     2    WN071203
     3 *)
     4 
     5 theory LogExp imports PolyEq begin
     6 
     7 consts
     8 
     9   ln     :: "real => real"
    10   exp    :: "real => real"         ("E'_ ^^^ _" 80)
    11 
    12 (*--------------------------------------------------*) 
    13   alog   :: "[real, real] => real" ("_ log _" 90)
    14 
    15   (*Script-names*)
    16   Solve'_log    :: "[bool,real,        bool list]  
    17 				    => bool list"
    18                   ("((Script Solve'_log (_ _=))//(_))" 9)
    19 
    20 rules
    21 
    22   equality_pow    "0 < a ==> (l = r) = (a^^^l = a^^^r)"
    23   (* this is what students   ^^^^^^^... are told to do *)
    24   equality_power  "((a log b) = c) = (a^^^(a log b) = a^^^c)"
    25   exp_invers_log  "a^^^(a log b) = b"
    26 
    27 ML {*
    28 val thy = @{theory};
    29 
    30 (** problems **)
    31 store_pbt
    32  (prep_pbt thy "pbl_test_equ_univ_log" [] e_pblID
    33  (["logarithmic","univariate","equation"],
    34   [("#Given",["equality e_","solveFor v_"]),
    35    ("#Where",["matches ((?a log ?v_) = ?b) e_"]),
    36    ("#Find" ,["solutions v_i_"]),
    37    ("#With" ,["||(lhs (Subst (v_i_,v_) e_) - " ^
    38 	      "  (rhs (Subst (v_i_,v_) e_) || < eps)"])
    39    ],
    40   PolyEq_prls, SOME "solve (e_::bool, v_)",
    41   [["Equation","solve_log"]]));
    42 
    43 (** methods **)
    44 store_met
    45  (prep_met thy "met_equ_log" [] e_metID
    46  (["Equation","solve_log"],
    47   [("#Given" ,["equality e_","solveFor v_"]),
    48    ("#Where" ,["matches ((?a log ?v_) = ?b) e_"]),
    49    ("#Find"  ,["solutions v_i_"])
    50   ],
    51    {rew_ord'="termlessI",rls'=PolyEq_erls,srls=e_rls,prls=PolyEq_prls,
    52     calc=[],crls=PolyEq_crls, nrls=norm_Rational},
    53     "Script Solve_log (e_::bool) (v_::real) =     " ^
    54     "(let e_ = ((Rewrite equality_power False) @@ " ^
    55     "           (Rewrite exp_invers_log False) @@ " ^
    56     "           (Rewrite_Set norm_Poly False)) e_ " ^
    57     " in [e_])"
    58    ));
    59 *}
    60 
    61 end