src/Tools/isac/Knowledge/RootRat.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 15 Mar 2018 12:42:04 +0100
changeset 59406 509d70b507e5
parent 59389 627d25067f2f
child 59416 229e5c9cf78b
permissions -rw-r--r--
separate structure Celem: CALC_ELEMENT, finished on src/
neuper@42318
     1
(* Title:  collecting all knowledge for Root and Rational
neuper@42318
     2
   Author: Richard Lang 02.10 last change by: rlang date: 02.10.20
neuper@42318
     3
   (c) due to copyright terms
neuper@37906
     4
*)
neuper@37906
     5
neuper@37954
     6
theory RootRat imports Root Rational begin
neuper@37906
     7
neuper@37954
     8
ML {*
neuper@37972
     9
val thy = @{theory};
neuper@37972
    10
neuper@37954
    11
val rootrat_erls = 
wneuper@59406
    12
  Celem.merge_rls "rootrat_erls" Root_erls
wneuper@59406
    13
    (Celem.merge_rls "" rational_erls
wneuper@59406
    14
      (Celem.append_rls "" Celem.e_rls []));
neuper@37954
    15
neuper@37954
    16
(*.calculate numeral groundterms.*)
neuper@37954
    17
val calculate_RootRat = 
wneuper@59406
    18
  Celem.append_rls "calculate_RootRat" calculate_Rational
wneuper@59406
    19
	    [Celem.Thm ("distrib_left",TermC.num_str @{thm distrib_left}),
neuper@42318
    20
		      (* w*(z1.0 + z2.0) = w * z1.0 + w * z2.0 *)
wneuper@59406
    21
		     Celem.Thm ("mult_1_left",TermC.num_str @{thm mult_1_left}),
neuper@42318
    22
		      (* 1 * z = z *)
wneuper@59406
    23
		     Celem.Thm ("sym_real_mult_minus1",TermC.num_str (@{thm real_mult_minus1} RS @{thm sym})),
neuper@42318
    24
		       (* "- z1 = -1 * z1"  *)
wneuper@59406
    25
		     Celem.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_")
neuper@42318
    26
		     ];
s1210629013@55444
    27
wneuper@59374
    28
val prep_rls' = LTool.prep_rls @{theory};
neuper@37954
    29
*}
neuper@52125
    30
setup {* KEStore_Elems.add_rlss 
s1210629013@55444
    31
  [("rootrat_erls", (Context.theory_name @{theory}, prep_rls' rootrat_erls)), 
s1210629013@55444
    32
  ("calculate_RootRat", (Context.theory_name @{theory}, prep_rls' calculate_RootRat))] *}
neuper@37906
    33
neuper@37906
    34
end