wneuper@59424: theory Base_Tools walther@60149: imports Interpret.Interpret walther@60149: (** )"../BridgeJEdit/BridgeJEdit" ( *activate after devel.of BridgeJEdit*) walther@60149: (**) "../BridgeLibisabelle/BridgeLibisabelle" (*remove after devel.of BridgeJEdit*) walther@60149: (* ^^^ for KEStore_Elems.add_thes *) wneuper@59424: begin wneuper@59602: subsection \theorems for Base_Tools\ walther@60273: walther@60273: lemma real_unari_minus: "- a = (-1) * (a::real)" by auto walther@60273: (*Semiring_Normalization.comm_ring_1_class.ring_normalization_rules(1)*) walther@60273: walther@60273: (* should be in Rational.thy, but needed for asms in e.g. d2_pqformula1 in PolyEq, RootEq... *) walther@60273: lemma radd_left_cancel_le: "((k::real) + m <= k + n) = (m <= n)" by auto walther@60273: (*Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left*) walther@60273: wneuper@59602: axiomatization where (*for evaluating the assumptions of conditional rules*) walther@60273: (* should be in Rational.thy, but needed for asms in e.g. d2_pqformula1 in PolyEq, RootEq... *) walther@60273: rat_leq1: "[| 0 \ b; 0 \ d |] ==> (a / b <= c / d) = (a * d <= b * c)"(*Quickcheck found a counterexample*) and walther@60273: rat_leq2: "0 \ d ==> (a <= c / d) = (a * d <= c)" (*Quickcheck found a counterexample*) and walther@60273: rat_leq3: "0 \ b ==> (a / b <= c ) = (a <= b * c)" (*Quickcheck found a counterexample*) wneuper@59602: wneuper@59602: wneuper@59602: subsection \setup for ML-functions\ wneuper@59602: text \required by "eval_binop" below\ wneuper@59602: setup \KEStore_Elems.add_calcs walther@59603: [ ("occurs_in", ("Prog_Expr.occurs'_in", Prog_Expr.eval_occurs_in "#occurs_in_")), walther@59603: ("some_occur_in", ("Prog_Expr.some'_occur'_in", Prog_Expr.eval_some_occur_in "#some_occur_in_")), walther@59603: ("is_atom", ("Prog_Expr.is'_atom", Prog_Expr.eval_is_atom "#is_atom_")), walther@59603: ("is_even", ("Prog_Expr.is'_even", Prog_Expr.eval_is_even "#is_even_")), walther@59603: ("is_const", ("Prog_Expr.is'_const", Prog_Expr.eval_const "#is_const_")), walther@59603: ("le", ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_")), walther@59603: ("leq", ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_")), walther@59603: ("ident", ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_")), walther@59603: ("equal", ("HOL.eq", Prog_Expr.eval_equal "#equal_")), walther@59603: ("PLUS", ("Groups.plus_class.plus", (**)eval_binop "#add_")), walther@59603: ("MINUS", ("Groups.minus_class.minus", (**)eval_binop "#sub_")), walther@59603: ("TIMES", ("Groups.times_class.times", (**)eval_binop "#mult_")), walther@59603: ("DIVIDE", ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")), walther@59603: ("POWER",("Prog_Expr.pow", (**)eval_binop "#power_")), walther@59603: ("boollist2sum", ("Prog_Expr.boollist2sum", Prog_Expr.eval_boollist2sum ""))]\ wneuper@59602: wneuper@59602: subsection \rewrite-order for rule-sets\ wneuper@59602: ML \ wneuper@59602: \ ML \ wneuper@59602: local wneuper@59602: open Term; wneuper@59602: in walther@59910: fun termlessI (_: subst) uv = LibraryC.termless uv; walther@59910: fun term_ordI (_: subst) uv = Term_Ord.term_ord uv; wneuper@59602: end; wneuper@59602: \ ML \ walther@59857: (*TODO.WN0509 reduce ids: tless_true = e_rew_ord' = Rewrite_Ord.e_rew_ord = Rewrite_Ord.dummy_ord*) walther@59857: val tless_true = Rewrite_Ord.dummy_ord; walther@59887: Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (*<<<---- use Know_Store.xxx here, too*) wneuper@59602: [("tless_true", tless_true), wneuper@59602: ("e_rew_ord'", tless_true), walther@59857: ("dummy_ord", Rewrite_Ord.dummy_ord)]); wneuper@59602: \ wneuper@59602: wneuper@59602: subsection \rule-sets\ wneuper@59602: ML \ wneuper@59602: \ ML \ walther@59852: val Atools_erls = Rule_Set.append_rules "Atools_erls" Rule_Set.empty walther@59878: [ Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"), walther@59871: Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}), wneuper@59602: (*"(~ True) = False"*) walther@59871: Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}), wneuper@59602: (*"(~ False) = True"*) walther@59871: Rule.Thm ("and_true", ThmC.numerals_to_Free @{thm and_true}), wneuper@59602: (*"(?a & True) = ?a"*) walther@59871: Rule.Thm ("and_false", ThmC.numerals_to_Free @{thm and_false}), wneuper@59602: (*"(?a & False) = False"*) walther@59871: Rule.Thm ("or_true", ThmC.numerals_to_Free @{thm or_true}), wneuper@59602: (*"(?a | True) = True"*) walther@59871: Rule.Thm ("or_false", ThmC.numerals_to_Free @{thm or_false}), wneuper@59602: (*"(?a | False) = ?a"*) wneuper@59602: walther@59871: Rule.Thm ("rat_leq1", ThmC.numerals_to_Free @{thm rat_leq1}), walther@59871: Rule.Thm ("rat_leq2", ThmC.numerals_to_Free @{thm rat_leq2}), walther@59871: Rule.Thm ("rat_leq3", ThmC.numerals_to_Free @{thm rat_leq3}), walther@59871: Rule.Thm ("refl", ThmC.numerals_to_Free @{thm refl}), walther@59871: Rule.Thm ("order_refl", ThmC.numerals_to_Free @{thm order_refl}), walther@59871: Rule.Thm ("radd_left_cancel_le", ThmC.numerals_to_Free @{thm radd_left_cancel_le}), wneuper@59602: walther@59878: Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"), walther@59878: Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"), wneuper@59602: walther@59878: Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"), walther@59878: Rule.Eval ("Prog_Expr.is'_const", Prog_Expr.eval_const "#is_const_"), walther@59878: Rule.Eval ("Prog_Expr.occurs'_in", Prog_Expr.eval_occurs_in ""), walther@59878: Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")]; wneuper@59602: \ wneuper@59602: wneuper@59602: ML \ walther@59852: val Atools_crls = Rule_Set.append_rules "Atools_crls" Rule_Set.empty walther@59878: [ Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"), walther@59871: Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}), walther@59871: Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}), walther@59871: Rule.Thm ("and_true", ThmC.numerals_to_Free @{thm and_true}), walther@59871: Rule.Thm ("and_false", ThmC.numerals_to_Free @{thm and_false}), walther@59871: Rule.Thm ("or_true", ThmC.numerals_to_Free @{thm or_true}), walther@59871: Rule.Thm ("or_false", ThmC.numerals_to_Free @{thm or_false}), wneuper@59602: walther@59871: Rule.Thm ("rat_leq1", ThmC.numerals_to_Free @{thm rat_leq1}), walther@59871: Rule.Thm ("rat_leq2", ThmC.numerals_to_Free @{thm rat_leq2}), walther@59871: Rule.Thm ("rat_leq3", ThmC.numerals_to_Free @{thm rat_leq3}), walther@59871: Rule.Thm ("refl", ThmC.numerals_to_Free @{thm refl}), walther@59871: Rule.Thm ("order_refl", ThmC.numerals_to_Free @{thm order_refl}), walther@59871: Rule.Thm ("radd_left_cancel_le", ThmC.numerals_to_Free @{thm radd_left_cancel_le}), wneuper@59602: walther@59878: Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"), walther@59878: Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"), wneuper@59602: walther@59878: Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"), walther@59878: Rule.Eval ("Prog_Expr.is'_const", Prog_Expr.eval_const "#is_const_"), walther@59878: Rule.Eval ("Prog_Expr.occurs'_in", Prog_Expr.eval_occurs_in ""), walther@59878: Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")]; wneuper@59602: \ wneuper@59602: wneuper@59602: subsection \ONCE AGAIN extend rule-set for evaluating pre-conditions and program-expressions\ wneuper@59602: text \requires "eval_binop" from above\ wneuper@59602: ML \ walther@59852: val prog_expr = Rule_Set.append_rules "prog_expr" prog_expr walther@59878: [ Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"), walther@59878: Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"), walther@59878: Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"), walther@59878: Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"), walther@59878: Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"), walther@59878: Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),(*atom <> atom -> False*) wneuper@59602: walther@59878: Rule.Eval ("Prog_Expr.Vars",Prog_Expr.eval_var "#Vars_"), wneuper@59602: walther@59871: Rule.Thm ("if_True",ThmC.numerals_to_Free @{thm if_True}), walther@59871: Rule.Thm ("if_False",ThmC.numerals_to_Free @{thm if_False})]; wneuper@59602: walther@59852: val prog_expr = Auto_Prog.prep_rls @{theory} (Rule_Set.merge "list_erls" walther@59851: (Rule_Def.Repeat {id = "replaced", preconds = [], rew_ord = ("termlessI", termlessI), walther@59851: erls = Rule_Def.Repeat {id = "list_elrs", preconds = [], rew_ord = ("termlessI", termlessI), walther@59852: erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@59878: rules = [Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"), walther@59878: Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_") wneuper@59602: (* ~~~~~~ for nth_Cons_*)], walther@59878: scr = Rule.Empty_Prog}, walther@59851: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@59878: rules = [], scr = Rule.Empty_Prog}) walther@59801: prog_expr); wneuper@59602: \ wneuper@59602: subsection \setup for extended rule-set\ wneuper@59602: setup \KEStore_Elems.add_rlss walther@59801: [("prog_expr", (Context.theory_name @{theory}, Auto_Prog.prep_rls @{theory} prog_expr))]\ wneuper@59602: walther@60077: end