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