src/Tools/isac/Knowledge/LogExp.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 14 Sep 2010 12:12:42 +0200
branchisac-update-Isa09-2
changeset 38009 b49723351533
parent 37992 351a9e94c38d
child 38010 a37a3ab989f4
permissions -rw-r--r--
adapted is_copy_named from v___ to v'''

Unclear comment for 'fun is_copy_named_generating',
thus probably still broken; waiting for further tests.
Tests are still run from Build_Isac.
neuper@37906
     1
(* all outcommented in order to demonstrate authoring:
neuper@37906
     2
   WN071203
neuper@37906
     3
*)
neuper@37954
     4
neuper@37954
     5
theory LogExp imports PolyEq begin
neuper@37906
     6
neuper@37906
     7
consts
neuper@37906
     8
neuper@37906
     9
  ln     :: "real => real"
neuper@37906
    10
  exp    :: "real => real"         ("E'_ ^^^ _" 80)
neuper@37906
    11
neuper@37906
    12
(*--------------------------------------------------*) 
neuper@37906
    13
  alog   :: "[real, real] => real" ("_ log _" 90)
neuper@37906
    14
neuper@37906
    15
  (*Script-names*)
neuper@37954
    16
  Solve'_log    :: "[bool,real,        bool list]  
neuper@37954
    17
				    => bool list"
neuper@37906
    18
                  ("((Script Solve'_log (_ _=))//(_))" 9)
neuper@37906
    19
neuper@37983
    20
axioms
neuper@37906
    21
neuper@37983
    22
  equality_pow:    "0 < a ==> (l = r) = (a^^^l = a^^^r)"
neuper@37906
    23
  (* this is what students   ^^^^^^^... are told to do *)
neuper@37983
    24
  equality_power:  "((a log b) = c) = (a^^^(a log b) = a^^^c)"
neuper@37983
    25
  exp_invers_log:  "a^^^(a log b) = b"
neuper@37954
    26
neuper@37954
    27
ML {*
neuper@37972
    28
val thy = @{theory};
neuper@37972
    29
neuper@37954
    30
(** problems **)
neuper@37954
    31
store_pbt
neuper@37972
    32
 (prep_pbt thy "pbl_test_equ_univ_log" [] e_pblID
neuper@37954
    33
 (["logarithmic","univariate","equation"],
neuper@37981
    34
  [("#Given",["equality e_e","solveFor v_v"]),
neuper@37991
    35
   ("#Where",["matches ((?a log ?v_v) = ?b) e_e"]),
neuper@38009
    36
   ("#Find" ,["solutions v_i'''"]),
neuper@38009
    37
   ("#With" ,["||(lhs (Subst (v_i''', v_v) e_e) - " ^
neuper@38009
    38
	      "  (rhs (Subst (v_i''', v_v) e_e) || < eps)"])
neuper@37954
    39
   ],
neuper@37981
    40
  PolyEq_prls, SOME "solve (e_e::bool, v_v)",
neuper@37954
    41
  [["Equation","solve_log"]]));
neuper@37992
    42
*}
neuper@37992
    43
ML {*
neuper@37954
    44
(** methods **)
neuper@37954
    45
store_met
neuper@37972
    46
 (prep_met thy "met_equ_log" [] e_metID
neuper@37954
    47
 (["Equation","solve_log"],
neuper@37981
    48
  [("#Given" ,["equality e_e","solveFor v_v"]),
neuper@37991
    49
   ("#Where" ,["matches ((?a log ?v_v) = ?b) e_e"]),
neuper@38009
    50
   ("#Find"  ,["solutions v_i'''"])
neuper@37954
    51
  ],
neuper@37954
    52
   {rew_ord'="termlessI",rls'=PolyEq_erls,srls=e_rls,prls=PolyEq_prls,
neuper@37954
    53
    calc=[],crls=PolyEq_crls, nrls=norm_Rational},
neuper@37982
    54
    "Script Solve_log (e_e::bool) (v_v::real) =     " ^
neuper@37981
    55
    "(let e_e = ((Rewrite equality_power False) @@ " ^
neuper@37954
    56
    "           (Rewrite exp_invers_log False) @@ " ^
neuper@37981
    57
    "           (Rewrite_Set norm_Poly False)) e_e " ^
neuper@37992
    58
    " in [e_e])"
neuper@37954
    59
   ));
neuper@37954
    60
*}
neuper@37906
    61
neuper@37906
    62
end