src/Tools/isac/Knowledge/Base_Tools.thy
author wneuper <walther.neuper@jku.at>
Sat, 03 Jul 2021 16:21:07 +0200
changeset 60318 e6e7a9b9ced7
parent 60278 343efa173023
child 60331 40eb8aa2b0d6
permissions -rw-r--r--
//test/../rewrite.sml,poly.sml WORK
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@60318
     9
lemma real_unari_minus:   "Not (a is_const) ==> - 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>
wneuper@59602
    28
setup \<open>KEStore_Elems.add_calcs
walther@60278
    29
  [ ("occurs_in", ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in "#occurs_in_")),
walther@60278
    30
    ("some_occur_in", ("Prog_Expr.some_occur_in", Prog_Expr.eval_some_occur_in "#some_occur_in_")),
walther@60278
    31
    ("is_atom", ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_")),
walther@60278
    32
    ("is_even", ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_")),
walther@60278
    33
    ("is_const", ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_")),
walther@59603
    34
    ("le", ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_")),
walther@59603
    35
    ("leq", ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_")),
walther@59603
    36
    ("ident", ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_")),
walther@59603
    37
    ("equal", ("HOL.eq", Prog_Expr.eval_equal "#equal_")),
walther@59603
    38
    ("PLUS", ("Groups.plus_class.plus", (**)eval_binop "#add_")),
walther@59603
    39
    ("MINUS", ("Groups.minus_class.minus", (**)eval_binop "#sub_")),
walther@59603
    40
    ("TIMES", ("Groups.times_class.times", (**)eval_binop "#mult_")),
walther@59603
    41
    ("DIVIDE", ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")),
walther@60275
    42
    ("POWER",("Transcendental.powr", (**)eval_binop "#power_")),
walther@59603
    43
    ("boollist2sum", ("Prog_Expr.boollist2sum", Prog_Expr.eval_boollist2sum ""))]\<close>
wneuper@59602
    44
wneuper@59602
    45
subsection \<open>rewrite-order for rule-sets\<close>
wneuper@59602
    46
ML \<open>
wneuper@59602
    47
\<close> ML \<open>
wneuper@59602
    48
local
wneuper@59602
    49
  open Term;
wneuper@59602
    50
in
walther@59910
    51
  fun termlessI (_: subst) uv = LibraryC.termless uv;
walther@59910
    52
  fun term_ordI (_: subst) uv = Term_Ord.term_ord uv;
wneuper@59602
    53
end;
wneuper@59602
    54
\<close> ML \<open>
walther@59857
    55
(*TODO.WN0509 reduce ids: tless_true = e_rew_ord' = Rewrite_Ord.e_rew_ord = Rewrite_Ord.dummy_ord*)
walther@59857
    56
val tless_true = Rewrite_Ord.dummy_ord;
walther@59887
    57
Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (*<<<---- use Know_Store.xxx here, too*)
wneuper@59602
    58
			[("tless_true", tless_true),
wneuper@59602
    59
			 ("e_rew_ord'", tless_true),
walther@59857
    60
			 ("dummy_ord", Rewrite_Ord.dummy_ord)]);
wneuper@59602
    61
\<close>
wneuper@59602
    62
wneuper@59602
    63
subsection \<open>rule-sets\<close>
wneuper@59602
    64
ML \<open>
wneuper@59602
    65
\<close> ML \<open>
walther@59852
    66
val Atools_erls = Rule_Set.append_rules "Atools_erls" Rule_Set.empty
walther@59878
    67
  [ Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59871
    68
    Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
wneuper@59602
    69
		(*"(~ True) = False"*)
walther@59871
    70
		Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
wneuper@59602
    71
		(*"(~ False) = True"*)
walther@59871
    72
		Rule.Thm ("and_true", ThmC.numerals_to_Free @{thm and_true}),
wneuper@59602
    73
		(*"(?a & True) = ?a"*)
walther@59871
    74
		Rule.Thm ("and_false", ThmC.numerals_to_Free @{thm and_false}),
wneuper@59602
    75
		(*"(?a & False) = False"*)
walther@59871
    76
		Rule.Thm ("or_true", ThmC.numerals_to_Free @{thm or_true}),
wneuper@59602
    77
		(*"(?a | True) = True"*)
walther@59871
    78
		Rule.Thm ("or_false", ThmC.numerals_to_Free @{thm or_false}),
wneuper@59602
    79
		(*"(?a | False) = ?a"*)
wneuper@59602
    80
               
walther@59871
    81
		Rule.Thm ("rat_leq1", ThmC.numerals_to_Free @{thm rat_leq1}),
walther@59871
    82
		Rule.Thm ("rat_leq2", ThmC.numerals_to_Free @{thm rat_leq2}),
walther@59871
    83
		Rule.Thm ("rat_leq3", ThmC.numerals_to_Free @{thm rat_leq3}),
walther@59871
    84
      Rule.Thm ("refl", ThmC.numerals_to_Free @{thm refl}),
walther@59871
    85
		Rule.Thm ("order_refl", ThmC.numerals_to_Free @{thm order_refl}),
walther@59871
    86
		Rule.Thm ("radd_left_cancel_le", ThmC.numerals_to_Free @{thm radd_left_cancel_le}),
wneuper@59602
    87
		
walther@59878
    88
		Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
walther@59878
    89
		Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"),
wneuper@59602
    90
		
walther@59878
    91
		Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),    
walther@60278
    92
		Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"),
walther@60278
    93
		Rule.Eval ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in ""),    
walther@59878
    94
		Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")];
wneuper@59602
    95
\<close>
wneuper@59602
    96
wneuper@59602
    97
ML \<open>
walther@59852
    98
val Atools_crls = Rule_Set.append_rules "Atools_crls" Rule_Set.empty
walther@59878
    99
  [ Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59871
   100
    Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
walther@59871
   101
		Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
walther@59871
   102
		Rule.Thm ("and_true", ThmC.numerals_to_Free @{thm and_true}),
walther@59871
   103
		Rule.Thm ("and_false", ThmC.numerals_to_Free @{thm and_false}),
walther@59871
   104
		Rule.Thm ("or_true", ThmC.numerals_to_Free @{thm or_true}),
walther@59871
   105
		Rule.Thm ("or_false", ThmC.numerals_to_Free @{thm or_false}),
wneuper@59602
   106
               
walther@59871
   107
		Rule.Thm ("rat_leq1", ThmC.numerals_to_Free @{thm rat_leq1}),
walther@59871
   108
		Rule.Thm ("rat_leq2", ThmC.numerals_to_Free @{thm rat_leq2}),
walther@59871
   109
		Rule.Thm ("rat_leq3", ThmC.numerals_to_Free @{thm rat_leq3}),
walther@59871
   110
		Rule.Thm ("refl", ThmC.numerals_to_Free @{thm refl}),
walther@59871
   111
		Rule.Thm ("order_refl", ThmC.numerals_to_Free @{thm order_refl}),
walther@59871
   112
		Rule.Thm ("radd_left_cancel_le", ThmC.numerals_to_Free @{thm radd_left_cancel_le}),
wneuper@59602
   113
		
walther@59878
   114
		Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
walther@59878
   115
		Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"),
wneuper@59602
   116
		
walther@59878
   117
		Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),    
walther@60278
   118
		Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"),
walther@60278
   119
		Rule.Eval ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in ""),    
walther@59878
   120
		Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "")];
wneuper@59602
   121
\<close>
wneuper@59602
   122
wneuper@59602
   123
subsection \<open>ONCE AGAIN extend rule-set for evaluating pre-conditions and program-expressions\<close>
wneuper@59602
   124
text \<open>requires "eval_binop" from above\<close>
wneuper@59602
   125
ML \<open>
walther@59852
   126
val prog_expr = Rule_Set.append_rules "prog_expr" prog_expr
walther@59878
   127
	[ Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"),
walther@59878
   128
		Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"), 
walther@59878
   129
		Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
walther@59878
   130
		Rule.Eval ("Orderings.ord_class.less_eq", Prog_Expr.eval_equ "#less_equal_"),
walther@59878
   131
		Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),
walther@59878
   132
		Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),(*atom <> atom -> False*)
wneuper@59602
   133
       
walther@59878
   134
		Rule.Eval ("Prog_Expr.Vars",Prog_Expr.eval_var "#Vars_"),
wneuper@59602
   135
		
walther@59871
   136
		Rule.Thm ("if_True",ThmC.numerals_to_Free @{thm if_True}),
walther@59871
   137
		Rule.Thm ("if_False",ThmC.numerals_to_Free @{thm if_False})];
wneuper@59602
   138
walther@59852
   139
val prog_expr = Auto_Prog.prep_rls @{theory} (Rule_Set.merge "list_erls"
walther@59851
   140
	(Rule_Def.Repeat {id = "replaced", preconds = [], rew_ord = ("termlessI", termlessI),
walther@59851
   141
    erls = Rule_Def.Repeat {id = "list_elrs", preconds = [], rew_ord = ("termlessI", termlessI), 
walther@59852
   142
    erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59878
   143
    rules = [Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_"),
walther@59878
   144
      Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_")
wneuper@59602
   145
      (*    ~~~~~~ for nth_Cons_*)],
walther@59878
   146
    scr = Rule.Empty_Prog},
walther@59851
   147
    srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59878
   148
    rules = [], scr = Rule.Empty_Prog})
walther@59801
   149
  prog_expr);
wneuper@59602
   150
\<close>
wneuper@59602
   151
subsection \<open>setup for extended rule-set\<close>
wneuper@59602
   152
setup \<open>KEStore_Elems.add_rlss
walther@59801
   153
  [("prog_expr", (Context.theory_name @{theory}, Auto_Prog.prep_rls @{theory} prog_expr))]\<close>
wneuper@59602
   154
walther@60278
   155
ML \<open>
walther@60278
   156
\<close> ML \<open>
walther@60278
   157
\<close> ML \<open>
walther@60278
   158
\<close>
walther@60077
   159
end