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