src/Tools/isac/Knowledge/Base_Tools.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60506 145e45cd7a0f
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
     1 theory Base_Tools
     2   imports Interpret.Interpret
     3 (** )"../BridgeJEdit/BridgeJEdit"                    ( *activate after devel.of BridgeJEdit*)
     4 (**) "../BridgeLibisabelle/BridgeLibisabelle"           (*remove after devel.of BridgeJEdit*)
     5                       (*  ^^^ for KEStore_Elems.add_thes *)
     6 begin
     7 subsection \<open>theorems for Base_Tools\<close>
     8 
     9 lemma real_unari_minus:   "Not (a is_num) ==> - a = (-1) * (a::real)" by auto
    10 (*lemma real_unari_minus:   "- a = (-1) * (a::real)" by auto LOOPS WITH NUMERALS*)
    11 (*Semiring_Normalization.comm_ring_1_class.ring_normalization_rules(1)*)
    12 
    13 (* should be in Rational.thy, but needed for asms in e.g. d2_pqformula1 in PolyEq, RootEq... *)
    14 lemma radd_left_cancel_le: "((k::real) + m \<le> k + n) = (m \<le> n)" by auto
    15 (*Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left*)
    16 
    17 thm "Fields.linordered_field_class.mult_imp_le_div_pos" (*0 < ?y \<Longrightarrow> ?z * ?y \<le> ?x \<Longrightarrow> ?z \<le> ?x / ?y*)
    18 
    19 axiomatization where (*for evaluating the assumptions of conditional rules*)
    20   (* should be in Rational.thy, but needed for asms in e.g. d2_pqformula1 in PolyEq, RootEq... *)
    21   rat_leq1: "0 \<noteq> b \<Longrightarrow> 0 \<noteq> d \<Longrightarrow> (a / b \<le> c / d) = (a * d \<le> b * c)"(*Quickcheck found a counterexample*) and
    22   rat_leq2:	          "0 \<noteq> d \<Longrightarrow> (a     \<le> c / d) = (a * d \<le>     c)"(*Quickcheck found a counterexample*) and
    23   rat_leq3:	"0 \<noteq> b           \<Longrightarrow> (a / b \<le> c    ) = (a     \<le> b * c)"(*Quickcheck found a counterexample*)
    24 
    25 
    26 subsection \<open>setup for ML-functions\<close>
    27 text \<open>required by "eval_binop" below\<close>
    28 calculation occurs_in = \<open>Prog_Expr.eval_occurs_in "#occurs_in_"\<close>
    29 calculation some_occur_in = \<open>Prog_Expr.eval_some_occur_in "#some_occur_in_"\<close>
    30 calculation is_atom = \<open>Prog_Expr.eval_is_atom "#is_atom_"\<close>
    31 calculation is_even = \<open>Prog_Expr.eval_is_even "#is_even_"\<close>
    32 calculation le (less) = \<open>Prog_Expr.eval_equ "#less_"\<close>
    33 calculation leq (less_eq) = \<open>Prog_Expr.eval_equ "#less_equal_"\<close>
    34 calculation ident = \<open>Prog_Expr.eval_ident "#ident_"\<close>
    35 calculation equal (HOL.eq) = \<open>Prog_Expr.eval_equal "#equal_"\<close>
    36 calculation PLUS (plus) = \<open>(**)eval_binop "#add_"\<close>
    37 calculation MINUS (minus) = \<open>(**)eval_binop "#sub_"\<close>
    38 calculation TIMES (times) = \<open>(**)eval_binop "#mult_"\<close>
    39 calculation DIVIDE (divide) = \<open>Prog_Expr.eval_cancel "#divide_e"\<close>
    40 calculation POWER (realpow) = \<open>(**)eval_binop "#power_"\<close>
    41 calculation boollist2sum = \<open>Prog_Expr.eval_boollist2sum ""\<close>
    42 
    43 subsection \<open>rewrite-order for rule-sets\<close>
    44 ML \<open>
    45 \<close> ML \<open>
    46 local
    47   open Term;
    48 in
    49   fun termlessI (_: subst) uv = LibraryC.termless uv;
    50   fun term_ordI (_: subst) uv = Term_Ord.term_ord uv;
    51 end;
    52 \<close> ML \<open>
    53 (*TODO.WN0509 reduce ids: tless_true = e_rew_ord' = Rewrite_Ord.function_empty = Rewrite_Ord.function_empty*)
    54 val tless_true = Rewrite_Ord.function_empty;
    55 \<close> ML \<open>
    56 \<close> 
    57 setup \<open>KEStore_Elems.add_rew_ord [
    58   ("tless_true", tless_true),
    59 	("Rewrite_Ord.id_empty", tless_true),
    60 	("e_rew_ord'", tless_true),
    61 	("dummy_ord", Rewrite_Ord.function_empty)]\<close>
    62 
    63 subsection \<open>rule-sets\<close>
    64 ML \<open>
    65 \<close> ML \<open>
    66 val Atools_erls = Rule_Set.append_rules "Atools_erls" Rule_Set.empty [
    67   \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
    68   \<^rule_thm>\<open>not_true\<close>, (*"(~ True) = False"*)
    69 	\<^rule_thm>\<open>not_false\<close>, (*"(~ False) = True"*)
    70 	\<^rule_thm>\<open>and_true\<close>, (*"(?a & True) = ?a"*)
    71 	\<^rule_thm>\<open>and_false\<close>, (*"(?a & False) = False"*)
    72 	\<^rule_thm>\<open>or_true\<close>, (*"(?a | True) = True"*)
    73 	\<^rule_thm>\<open>or_false\<close>, (*"(?a | False) = ?a"*)
    74              
    75 	\<^rule_thm>\<open>rat_leq1\<close>,
    76 	\<^rule_thm>\<open>rat_leq2\<close>,
    77 	\<^rule_thm>\<open>rat_leq3\<close>,
    78   \<^rule_thm>\<open>refl\<close>,
    79 	\<^rule_thm>\<open>order_refl\<close>,
    80 	\<^rule_thm>\<open>radd_left_cancel_le\<close>,
    81 	
    82 	\<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
    83 	\<^rule_eval>\<open>less_eq\<close> (Prog_Expr.eval_equ "#less_equal_"),
    84 	
    85 	\<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
    86 	\<^rule_eval>\<open>Prog_Expr.is_num\<close> (Prog_Expr.eval_is_num "#is_num_"),
    87 	\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),
    88 	\<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches "")];
    89 \<close>
    90 
    91 ML \<open>
    92 val Atools_crls = Rule_Set.append_rules "Atools_crls" Rule_Set.empty [
    93   \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
    94   \<^rule_thm>\<open>not_true\<close>,
    95  	\<^rule_thm>\<open>not_false\<close>,
    96  	\<^rule_thm>\<open>and_true\<close>,
    97  	\<^rule_thm>\<open>and_false\<close>,
    98  	\<^rule_thm>\<open>or_true\<close>,
    99  	\<^rule_thm>\<open>or_false\<close>,
   100  
   101  	\<^rule_thm>\<open>rat_leq1\<close>,
   102  	\<^rule_thm>\<open>rat_leq2\<close>,
   103  	\<^rule_thm>\<open>rat_leq3\<close>,
   104  	\<^rule_thm>\<open>refl\<close>,
   105  	\<^rule_thm>\<open>order_refl\<close>,
   106  	\<^rule_thm>\<open>radd_left_cancel_le\<close>,
   107  	
   108  	\<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   109  	\<^rule_eval>\<open>less_eq\<close> (Prog_Expr.eval_equ "#less_equal_"),
   110  	
   111  	\<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
   112 	\<^rule_eval>\<open>Prog_Expr.is_num\<close> (Prog_Expr.eval_is_num "#is_num_"),
   113  	\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),
   114  	\<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches "")];
   115 \<close>
   116 
   117 subsection \<open>ONCE AGAIN extend rule-set for evaluating pre-conditions and program-expressions\<close>
   118 text \<open>requires "eval_binop" from above\<close>
   119 ML \<open>
   120 val prog_expr = Rule_Set.append_rules "prog_expr" prog_expr [
   121   \<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_"),
   122 	\<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"), 
   123 	\<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   124 	\<^rule_eval>\<open>less_eq\<close> (Prog_Expr.eval_equ "#less_equal_"),
   125 	\<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
   126 	\<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),(*atom <> atom -> False*)
   127      
   128 	\<^rule_eval>\<open>Prog_Expr.Vars\<close> (Prog_Expr.eval_var "#Vars_"),
   129 	
   130 	\<^rule_thm>\<open>if_True\<close>,
   131 	\<^rule_thm>\<open>if_False\<close>];
   132 
   133 val prog_expr = Auto_Prog.prep_rls @{theory} (Rule_Set.merge "list_erls"
   134 	(Rule_Def.Repeat {id = "replaced", preconds = [], rew_ord = ("termlessI", termlessI),
   135     erls = Rule_Def.Repeat {id = "list_elrs", preconds = [], rew_ord = ("termlessI", termlessI), 
   136     erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   137     rules = [
   138       \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   139       \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_") (*..for nth_Cons_*)],
   140     scr = Rule.Empty_Prog},
   141     srls = Rule_Set.Empty, calc = [], errpatts = [],
   142     rules = [], scr = Rule.Empty_Prog})
   143   prog_expr);
   144 \<close>
   145 
   146 subsection \<open>setup for extended rule-set\<close>
   147 
   148 rule_set_knowledge prog_expr = \<open>Auto_Prog.prep_rls @{theory} prog_expr\<close>
   149 
   150 ML \<open>
   151 \<close> ML \<open>
   152 \<close> ML \<open>
   153 \<close>
   154 end