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