src/Tools/isac/Knowledge/DiophantEq.thy
author Walther Neuper <walther.neuper@jku.at>
Wed, 02 Oct 2019 16:02:17 +0200
changeset 59637 8881c5d28f82
parent 59635 9fc1bb69813c
child 59850 f3cac3053e7b
permissions -rw-r--r--
lucin: use #> in Program in analogy to #> in Isabelle/ML
     1 (* Title:  Knowledge/DiophantEq.thy
     2    Author: Mathias Lehnfeld 2011
     3    (c) due to copyright terms
     4 12345678901234567890123456789012345678901234567890123456789012345678901234567890
     5         10        20        30        40        50        60        70        80
     6 *)
     7 
     8 theory DiophantEq imports Base_Tools Equation Test
     9 begin
    10 
    11 axiomatization where
    12   int_isolate_add: "(bdv + a = b) = (bdv = b + (-1)*(a::int))"
    13 
    14 ML \<open>val thy = @{theory}\<close>
    15 
    16 text \<open>problemclass for the usecase\<close>
    17 setup \<open>KEStore_Elems.add_pbts
    18   [(Specify.prep_pbt thy "pbl_equ_dio" [] Celem.e_pblID
    19       (["diophantine","equation"],
    20         [("#Given" ,["boolTestGiven e_e","intTestGiven (v_v::int)"]),
    21           (*                                      TODO: drop ^^^^^*)
    22           ("#Where" ,[]),
    23           ("#Find"  ,["boolTestFind s_s"])],
    24         Rule.e_rls, SOME "solve (e_e::bool, v_v::int)", [["LinEq","solve_lineq_equation"]]))]\<close>
    25 
    26 text \<open>method solving the usecase\<close>
    27 
    28 partial_function (tailrec) diophant_equation :: "bool => int => bool"
    29   where
    30 "diophant_equation e_e v_v = (
    31   Repeat (
    32     (Try (Rewrite_Inst [(''bdv'', v_v)] ''int_isolate_add'' )) #>
    33     (Try (Calculate ''PLUS'')) #>
    34     (Try (Calculate ''TIMES''))) e_e)"
    35 setup \<open>KEStore_Elems.add_mets
    36     [Specify.prep_met thy "met_test_diophant" [] Celem.e_metID
    37       (["Test","solve_diophant"],
    38         [("#Given" ,["boolTestGiven e_e","intTestGiven (v_v::int)"]),
    39           (*                                      TODO: drop ^^^^^*)
    40           ("#Where" ,[]),
    41           ("#Find"  ,["boolTestFind s_s"])],
    42         {rew_ord' = "e_rew_ord", rls' = tval_rls, srls = Rule.e_rls, prls = Rule.e_rls, calc = [],
    43           crls = tval_rls, errpats = [], nrls = Test_simplify},
    44         @{thm diophant_equation.simps})]
    45 \<close>
    46 
    47 end