neuper@41931: (* Title: Knowledge/DiophantEq.thy neuper@41931: Author: Mathias Lehnfeld 2011 neuper@41931: (c) due to copyright terms neuper@41931: 12345678901234567890123456789012345678901234567890123456789012345678901234567890 neuper@41931: 10 20 30 40 50 60 70 80 neuper@41931: *) neuper@41931: wneuper@59424: theory DiophantEq imports Base_Tools Equation Test neuper@41931: begin neuper@41921: neuper@52148: axiomatization where neuper@41931: int_isolate_add: "(bdv + a = b) = (bdv = b + (-1)*(a::int))" neuper@41921: wneuper@59472: text \problemclass for the usecase\ wneuper@59472: setup \KEStore_Elems.add_pbts wenzelm@60290: [(Problem.prep_input @{theory} "pbl_equ_dio" [] Problem.id_empty walther@59997: (["diophantine", "equation"], walther@59997: [("#Given" ,["boolTestGiven e_e", "intTestGiven (v_v::int)"]), s1210629013@55339: (* TODO: drop ^^^^^*) s1210629013@55339: ("#Where" ,[]), s1210629013@55339: ("#Find" ,["boolTestFind s_s"])], walther@59997: Rule_Set.empty, SOME "solve (e_e::bool, v_v::int)", [["LinEq", "solve_lineq_equation"]]))]\ neuper@41921: wneuper@59472: text \method solving the usecase\ wneuper@59545: wneuper@59504: partial_function (tailrec) diophant_equation :: "bool => int => bool" wneuper@59504: where walther@59635: "diophant_equation e_e v_v = ( walther@59635: Repeat ( walther@59637: (Try (Rewrite_Inst [(''bdv'', v_v)] ''int_isolate_add'' )) #> walther@59637: (Try (Calculate ''PLUS'')) #> walther@59635: (Try (Calculate ''TIMES''))) e_e)" wenzelm@60303: wenzelm@60303: method met_test_diophant : "Test/solve_diophant" = wenzelm@60303: \{rew_ord' = "e_rew_ord", rls' = tval_rls, srls = Rule_Set.empty, prls = Rule_Set.empty, calc = [], wenzelm@60303: crls = tval_rls, errpats = [], nrls = Test_simplify}\ wenzelm@60303: Program: diophant_equation.simps wenzelm@60303: Given: "boolTestGiven e_e" "intTestGiven (v_v::int)" wenzelm@60303: (* TODO: drop ^^^^^*) wenzelm@60303: Where: wenzelm@60303: Find: "boolTestFind s_s" wenzelm@60303: ML \ walther@60278: \ ML \ wneuper@59472: \ neuper@41921: walther@60278: end