src/Tools/isac/Knowledge/DiffApp.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 22 Dec 2016 11:36:20 +0100
changeset 59279 255c853ea2f0
parent 59269 1da53d1540fe
child 59334 690f0822e102
permissions -rw-r--r--
renamed Ctree.ptree --> Ctree.ctree
neuper@37906
     1
(* application of differential calculus
neuper@37954
     2
   Walther Neuper 2002
neuper@37954
     3
   (c) due to copyright terms
neuper@37906
     4
*)
neuper@37906
     5
neuper@37954
     6
theory DiffApp imports Diff begin
neuper@37906
     7
neuper@37906
     8
consts
neuper@37906
     9
neuper@37906
    10
  Maximum'_value
neuper@37954
    11
             :: "[bool list,real,bool list,real,real set,bool, 
neuper@37954
    12
		    bool list] => bool list"
neuper@37906
    13
               ("((Script Maximum'_value (_ _ _ _ _ _ =))// (_))" 9)
neuper@37906
    14
  
neuper@37906
    15
  Make'_fun'_by'_new'_variable
neuper@37954
    16
             :: "[real,real,bool list,  
neuper@37954
    17
		    bool] => bool"
neuper@37954
    18
               ("((Script Make'_fun'_by'_new'_variable (_ _ _ =))//  
neuper@37954
    19
		   (_))" 9)
neuper@37906
    20
  Make'_fun'_by'_explicit
neuper@37954
    21
             :: "[real,real,bool list,  
neuper@37954
    22
		    bool] => bool"
neuper@37954
    23
               ("((Script Make'_fun'_by'_explicit (_ _ _ =))//  
neuper@37954
    24
		   (_))" 9)
neuper@37906
    25
neuper@37906
    26
  dummy :: real
neuper@37906
    27
neuper@37906
    28
(*for script Maximum_value*)
neuper@37906
    29
  filterVar :: "[real, 'a list] => 'a list"
neuper@37906
    30
neuper@52148
    31
axiomatization where
neuper@52148
    32
  filterVar_Nil:		"filterVar v []     = []" and
neuper@37995
    33
  filterVar_Const:	"filterVar v (x#xs) =                       
neuper@41905
    34
			 (if (v : set (Vars x)) then x#(filterVar v xs)  
neuper@37954
    35
			                    else filterVar v xs)   "
neuper@37954
    36
text {*WN.6.5.03: old decisions in this file partially are being changed
neuper@37954
    37
  in a quick-and-dirty way to make scripts run: Maximum_value,
neuper@37954
    38
  Make_fun_by_new_variable, Make_fun_by_explicit.
neuper@37954
    39
found to be reconsidered:
neuper@37954
    40
- descriptions (Descript.thy)
neuper@37954
    41
- penv: really need term list; or just rerun the whole example with num/var
neuper@37954
    42
- mk_arg, itms2args ... env in script different from penv ?
neuper@37954
    43
- L = SubProblem eq ... show some vars on the worksheet ? (other means for
neuper@37954
    44
  referencing are labels (no on worksheet))
neuper@37954
    45
neuper@37954
    46
WN.6.5.03 quick-and-dirty: mk_arg, itms2args just make most convenient env
neuper@37954
    47
  from penv as is.    
neuper@37954
    48
*}
neuper@37954
    49
neuper@37954
    50
ML {*
neuper@37972
    51
val thy = @{theory};
neuper@37972
    52
s1210629013@55444
    53
val eval_rls = prep_rls' (
neuper@52139
    54
  Rls {id = "eval_rls", preconds = [], rew_ord = ("termlessI", termlessI), 
neuper@52139
    55
    erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@52139
    56
    rules = [Thm ("refl", num_str @{thm refl}),
neuper@52139
    57
      Thm ("order_refl", num_str @{thm order_refl}),
neuper@52139
    58
      Thm ("radd_left_cancel_le", num_str @{thm radd_left_cancel_le}),
neuper@52139
    59
      Thm ("not_true", num_str @{thm not_true}),
neuper@52139
    60
      Thm ("not_false", num_str @{thm not_false}),
neuper@52139
    61
      Thm ("and_true", num_str @{thm and_true}),
neuper@52139
    62
      Thm ("and_false", num_str @{thm and_false}),
neuper@52139
    63
      Thm ("or_true", num_str @{thm or_true}),
neuper@52139
    64
      Thm ("or_false", num_str @{thm or_false}),
neuper@52139
    65
      Thm ("and_commute", num_str @{thm and_commute}),
neuper@52139
    66
      Thm ("or_commute", num_str @{thm or_commute}),
neuper@52139
    67
      
neuper@52139
    68
      Calc ("Orderings.ord_class.less", eval_equ "#less_"),
neuper@52139
    69
      Calc ("Orderings.ord_class.less_eq", eval_equ "#less_equal_"),
neuper@52139
    70
      
neuper@52139
    71
      Calc ("Atools.ident", eval_ident "#ident_"),    
neuper@52139
    72
      Calc ("Atools.is'_const", eval_const "#is_const_"),
neuper@52139
    73
      Calc ("Atools.occurs'_in", eval_occurs_in ""),    
neuper@52139
    74
      Calc ("Tools.matches", eval_matches "")],
wneuper@59186
    75
    scr = Prog ((Thm.term_of o the o (parse thy)) "empty_script")
neuper@52139
    76
    }:rls);
neuper@37995
    77
*}
neuper@52125
    78
setup {* KEStore_Elems.add_rlss [("eval_rls", (Context.theory_name @{theory}, eval_rls))] *}
neuper@37954
    79
neuper@37954
    80
(** problem types **)
s1210629013@55359
    81
setup {* KEStore_Elems.add_pbts
wneuper@59269
    82
  [(Specify.prep_pbt thy "pbl_fun_max" [] e_pblID
s1210629013@55339
    83
      (["maximum_of","function"],
s1210629013@55339
    84
        [("#Given" ,["fixedValues f_ix"]),
s1210629013@55339
    85
          ("#Find"  ,["maximum m_m","valuesFor v_s"]),
s1210629013@55339
    86
          ("#Relate",["relations r_s"])],
s1210629013@55339
    87
        e_rls, NONE, [])),
wneuper@59269
    88
    (Specify.prep_pbt thy "pbl_fun_make" [] e_pblID
s1210629013@55339
    89
      (["make","function"]:pblID,
s1210629013@55339
    90
        [("#Given" ,["functionOf f_f","boundVariable v_v","equalities eqs"]),
s1210629013@55339
    91
          ("#Find"  ,["functionEq f_1"])],
s1210629013@55339
    92
        e_rls, NONE, [])),
wneuper@59269
    93
    (Specify.prep_pbt thy "pbl_fun_max_expl" [] e_pblID
s1210629013@55339
    94
      (["by_explicit","make","function"]:pblID,
s1210629013@55339
    95
        [("#Given" ,["functionOf f_f","boundVariable v_v","equalities eqs"]),
s1210629013@55339
    96
          ("#Find"  ,["functionEq f_1"])],
s1210629013@55339
    97
      e_rls, NONE, [["DiffApp","make_fun_by_explicit"]])),
wneuper@59269
    98
    (Specify.prep_pbt thy "pbl_fun_max_newvar" [] e_pblID
s1210629013@55339
    99
      (["by_new_variable","make","function"]:pblID,
s1210629013@55339
   100
        [("#Given" ,["functionOf f_f","boundVariable v_v","equalities eqs"]),
s1210629013@55339
   101
          (*WN.12.5.03: precond for distinction still missing*)
s1210629013@55339
   102
          ("#Find"  ,["functionEq f_1"])],
s1210629013@55339
   103
      e_rls, NONE, [["DiffApp","make_fun_by_new_variable"]])),
wneuper@59269
   104
    (Specify.prep_pbt thy "pbl_fun_max_interv" [] e_pblID
s1210629013@55339
   105
      (["on_interval","maximum_of","function"]:pblID,
s1210629013@55339
   106
        [("#Given" ,["functionEq t_t","boundVariable v_v","interval i_tv"]),
s1210629013@55339
   107
          (*WN.12.5.03: precond for distinction still missing*)
s1210629013@55339
   108
          ("#Find"  ,["maxArgument v_0"])],
s1210629013@55339
   109
      e_rls, NONE, [])),
wneuper@59269
   110
    (Specify.prep_pbt thy "pbl_tool" [] e_pblID
s1210629013@55339
   111
      (["tool"]:pblID, [], e_rls, NONE, [])),
wneuper@59269
   112
    (Specify.prep_pbt thy "pbl_tool_findvals" [] e_pblID
s1210629013@55339
   113
      (["find_values","tool"]:pblID,
s1210629013@55339
   114
        [("#Given" ,["maxArgument m_ax","functionEq f_f","boundVariable v_v"]),
s1210629013@55339
   115
          ("#Find"  ,["valuesFor v_ls"]),
s1210629013@55339
   116
          ("#Relate",["additionalRels r_s"])],
s1210629013@55339
   117
      e_rls, NONE, []))] *}
s1210629013@55380
   118
neuper@37954
   119
neuper@37954
   120
(** methods, scripts not yet implemented **)
s1210629013@55373
   121
setup {* KEStore_Elems.add_mets
wneuper@59269
   122
  [Specify.prep_met thy "met_diffapp" [] e_metID
s1210629013@55373
   123
      (["DiffApp"], [],
s1210629013@55373
   124
        {rew_ord'="tless_true", rls'=Atools_erls,calc = [], srls = e_rls, prls = e_rls,
s1210629013@55373
   125
          crls = Atools_erls, errpats = [], nrls = norm_Rational (*, asm_rls=[],asm_thm=[]*)},
s1210629013@55373
   126
        "empty_script"),
wneuper@59269
   127
    Specify.prep_met thy "met_diffapp_max" [] e_metID
s1210629013@55373
   128
      (["DiffApp","max_by_calculus"]:metID,
s1210629013@55373
   129
        [("#Given" ,["fixedValues f_ix","maximum m_m","relations r_s", "boundVariable v_v",
s1210629013@55373
   130
              "interval i_tv","errorBound e_rr"]),
s1210629013@55373
   131
          ("#Find"  ,["valuesFor v_s"]),
s1210629013@55373
   132
          ("#Relate",[])],
s1210629013@55373
   133
      {rew_ord'="tless_true",rls'=eval_rls,calc=[],srls=list_rls,prls=e_rls, crls = eval_rls,
s1210629013@55373
   134
        errpats = [], nrls = norm_Rational (*,  asm_rls=[],asm_thm=[]*)},
s1210629013@55373
   135
        "Script Maximum_value(f_ix::bool list)(m_m::real) (r_s::bool list)       " ^
s1210629013@55373
   136
          "      (v_v::real) (itv_v::real set) (e_rr::bool) =                       " ^ 
s1210629013@55373
   137
          " (let e_e = (hd o (filterVar m_m)) r_s;                                  " ^
s1210629013@55373
   138
          "      t_t = (if 1 < LENGTH r_s                                         " ^
s1210629013@55373
   139
          "           then (SubProblem (DiffApp',[make,function],[no_met])        " ^
s1210629013@55373
   140
          "                     [REAL m_m, REAL v_v, BOOL_LIST r_s])             " ^
s1210629013@55373
   141
          "           else (hd r_s));                                             " ^
s1210629013@55373
   142
          "      (m_x::real) =                                                    " ^ 
s1210629013@55373
   143
          "SubProblem(DiffApp',[on_interval,maximum_of,function],                 " ^
s1210629013@55373
   144
          "                                [DiffApp,max_on_interval_by_calculus]) " ^
s1210629013@55373
   145
          "                               [BOOL t_t, REAL v_v, REAL_SET i_tv]    " ^
s1210629013@55373
   146
          " in ((SubProblem (DiffApp',[find_values,tool],[Isac,find_values])      " ^
s1210629013@55373
   147
          "      [REAL m_x, REAL (Rhs t_t), REAL v_v, REAL m_m,                  " ^
s1210629013@55373
   148
          "       BOOL_LIST (dropWhile (ident e_e) r_s)])::bool list))            "),
wneuper@59269
   149
    Specify.prep_met thy "met_diffapp_funnew" [] e_metID
s1210629013@55373
   150
      (["DiffApp","make_fun_by_new_variable"]:metID,
s1210629013@55373
   151
        [("#Given" ,["functionOf f_f","boundVariable v_v","equalities eqs"]),
s1210629013@55373
   152
          ("#Find"  ,["functionEq f_1"])],
s1210629013@55373
   153
        {rew_ord'="tless_true",rls'=eval_rls,srls=list_rls,prls=e_rls, calc=[], crls = eval_rls,
s1210629013@55373
   154
          errpats = [], nrls = norm_Rational(*,asm_rls=[],asm_thm=[]*)},
s1210629013@55373
   155
        "Script Make_fun_by_new_variable (f_f::real) (v_v::real)                " ^
s1210629013@55373
   156
          "      (eqs::bool list) =                                            " ^
s1210629013@55373
   157
          "(let h_h = (hd o (filterVar f_f)) eqs;                                " ^
s1210629013@55373
   158
          "     e_s = dropWhile (ident h_h) eqs;                                " ^
s1210629013@55373
   159
          "     v_s = dropWhile (ident f_f) (Vars h_h);                           " ^
s1210629013@55373
   160
          "     v_1 = NTH 1 v_s;                                               " ^
s1210629013@55373
   161
          "     v_2 = NTH 2 v_s;                                               " ^
s1210629013@55373
   162
          "     e_1 = (hd o (filterVar v_1)) e_s;                               " ^
s1210629013@55373
   163
          "     e_2 = (hd o (filterVar v_2)) e_s;                               " ^
s1210629013@55373
   164
          "  (s_1::bool list) =                                                 " ^
s1210629013@55373
   165
          "                (SubProblem (DiffApp',[univariate,equation],[no_met])" ^
s1210629013@55373
   166
          "                    [BOOL e_1, REAL v_1]);                         " ^
s1210629013@55373
   167
          "  (s_2::bool list) =                                                 " ^
s1210629013@55373
   168
          "                (SubProblem (DiffApp',[univariate,equation],[no_met])" ^
s1210629013@55373
   169
          "                    [BOOL e_2, REAL v_2])" ^
s1210629013@55373
   170
          "in Substitute [(v_1 = (rhs o hd) s_1),(v_2 = (rhs o hd) s_2)] h_h)"),
wneuper@59269
   171
    Specify.prep_met thy "met_diffapp_funexp" [] e_metID
s1210629013@55373
   172
      (["DiffApp","make_fun_by_explicit"]:metID,
s1210629013@55373
   173
        [("#Given" ,["functionOf f_f","boundVariable v_v","equalities eqs"]),
s1210629013@55373
   174
          ("#Find"  ,["functionEq f_1"])],
s1210629013@55373
   175
        {rew_ord'="tless_true",rls'=eval_rls,calc=[],srls=list_rls,prls=e_rls, crls = eval_rls,
s1210629013@55373
   176
          errpats = [], nrls = norm_Rational (*, asm_rls=[],asm_thm=[]*)},
s1210629013@55373
   177
        "Script Make_fun_by_explicit (f_f::real) (v_v::real)                 " ^
s1210629013@55373
   178
          "      (eqs::bool list) =                                         " ^
s1210629013@55373
   179
          " (let h_h = (hd o (filterVar f_f)) eqs;                            " ^
s1210629013@55373
   180
          "      e_1 = hd (dropWhile (ident h_h) eqs);                       " ^
s1210629013@55373
   181
          "      v_s = dropWhile (ident f_f) (Vars h_h);                       " ^
s1210629013@55373
   182
          "      v_1 = hd (dropWhile (ident v_v) v_s);                        " ^
s1210629013@55373
   183
          "      (s_1::bool list)=                                           " ^ 
s1210629013@55373
   184
          "              (SubProblem(DiffApp',[univariate,equation],[no_met])" ^
s1210629013@55373
   185
          "                          [BOOL e_1, REAL v_1])                 " ^
s1210629013@55373
   186
          " in Substitute [(v_1 = (rhs o hd) s_1)] h_h)                       "),
wneuper@59269
   187
    Specify.prep_met thy "met_diffapp_max_oninterval" [] e_metID
s1210629013@55373
   188
      (["DiffApp","max_on_interval_by_calculus"]:metID,
s1210629013@55373
   189
        [("#Given" ,["functionEq t_t","boundVariable v_v","interval i_tv"(*, "errorBound e_rr"*)]),
s1210629013@55373
   190
          ("#Find"  ,["maxArgument v_0"])],
s1210629013@55373
   191
      {rew_ord'="tless_true",rls'=eval_rls,calc=[],srls = e_rls,prls=e_rls, crls = eval_rls,
s1210629013@55373
   192
        errpats = [], nrls = norm_Rational (*, asm_rls=[],asm_thm=[]*)},
s1210629013@55373
   193
      "empty_script"),
wneuper@59269
   194
    Specify.prep_met thy "met_diffapp_findvals" [] e_metID
s1210629013@55373
   195
      (["DiffApp","find_values"]:metID, [],
s1210629013@55373
   196
        {rew_ord'="tless_true",rls'=eval_rls,calc=[],srls = e_rls,prls=e_rls, crls = eval_rls,
s1210629013@55373
   197
          errpats = [], nrls = norm_Rational(*, asm_rls = [], asm_thm = []*)},
s1210629013@55373
   198
        "empty_script")]
s1210629013@55373
   199
*}
s1210629013@55373
   200
ML {*
neuper@37954
   201
val list_rls = append_rls "list_rls" list_rls
neuper@52139
   202
  [Thm ("filterVar_Const", num_str @{thm filterVar_Const}),
neuper@52139
   203
   Thm ("filterVar_Nil", num_str @{thm filterVar_Nil})];
neuper@37954
   204
*}
neuper@52125
   205
setup {* KEStore_Elems.add_rlss [("list_rls", (Context.theory_name @{theory}, list_rls))] *}
neuper@37906
   206
neuper@37906
   207
end