src/Tools/isac/Knowledge/Delete.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 31 Aug 2010 15:36:57 +0200
branchisac-update-Isa09-2
changeset 37965 9c11005c33b8
child 37974 ececb894db9c
permissions -rw-r--r--
updated Knowledge/Atools.thy + some changes + changes ahead

# replaced thms ahead by ./thms-replace-Isa02-Isa09-2.sml
# Knowledge/Delete.thy takes intermediate code:
* fun calc, fun term_of_float,fun var_op_float, fun float_op_var
for Float, which have already been deleted
* thms which are available only with long.identifiers
which cannot be handled in isac's Scripts
# added ProgLang/Language.thy collecting all data about Scripts
neuper@37965
     1
(* Title:  quick and dirty for quick update Isabelle2002 --> Isabelle2009-2
neuper@37965
     2
   Author: Walther Neuper 000301
neuper@37965
     3
   (c) due to copyright terms
neuper@37965
     4
*)
neuper@37965
     5
neuper@37965
     6
theory Delete imports "../ProgLang/Language" begin
neuper@37965
     7
neuper@37965
     8
axioms (* theorems which are available only with long names,
neuper@37965
     9
          which can not yet be handled in isac's scripts *)
neuper@37965
    10
neuper@37965
    11
 real_mult_left_commute "z1 * (z2 * z3) = z2 * (z1 * z3)"
neuper@37965
    12
 (*Semiring_Normalization.comm_semiring_1_class.normalizing_semiring_rules(19)*)
neuper@37965
    13
neuper@37965
    14
neuper@37965
    15
ML {* 
neuper@37965
    16
(* TODO rebuild fun calc, fun term_of_float,fun var_op_float, fun float_op_var:
neuper@37965
    17
   Float ((a, b), _:int * int, (c, d), _:int * int) has been deleted already *)
neuper@37965
    18
neuper@37965
    19
(*.used for calculating built in binary operations in Isabelle2002.
neuper@37965
    20
   integer numerals n are ((n,0),(0,0)) i.e. precision is (0,0)*)
neuper@37965
    21
fun calc "op +" ((a, b), _:int * int) ((c, d), _:int * int)  = (*FIXME.WN1008 drop Float.calc, var_op_float, float_op_var, term_of_float*)
neuper@37965
    22
    if b < d 
neuper@37965
    23
    then ((a + c * power 10 (d - b), b), (0, 0))(*FIXXXME precision*)
neuper@37965
    24
    else ((a * power 10 (b - d) + c, d), (0, 0))(*FIXXXME precision*)
neuper@37965
    25
  | calc "op -" ((a, 0), _) ((c, 0), _) =       (*FIXXXME float + prec.*)
neuper@37965
    26
    ((a - c,0),(0,0))
neuper@37965
    27
  | calc "op *" ((a, b), _) ((c, d), _) =       (*FIXXXME precision*)
neuper@37965
    28
    ((a * c, b + d), (0, 0))
neuper@37965
    29
  | calc "HOL.divide" ((a, 0), _) ((c, 0), _) = (*FIXXXME float + prec.*)
neuper@37965
    30
    ((a div c, 0), (0, 0))
neuper@37965
    31
  | calc "Atools.pow" ((a, b), _) ((c, d), _) = (*FIXXXME float + prec.*)
neuper@37965
    32
    ((power a c, 0), (0, 0))
neuper@37965
    33
  | calc op_ ((a, b), (p11, p12)) ((c, d), (p21, p22)) = 
neuper@37965
    34
    raise error ("calc: not impl. for Float (("^
neuper@37965
    35
		 (string_of_int a  )^","^(string_of_int b  )^"), ("^
neuper@37965
    36
		 (string_of_int p11)^","^(string_of_int p12)^")) "^op_^" (("^
neuper@37965
    37
		 (string_of_int c  )^","^(string_of_int d  )^"), ("^
neuper@37965
    38
		 (string_of_int p21)^","^(string_of_int p22)^"))");
neuper@37965
    39
(*> calc "op +" ((~1,0),(0,0)) ((2,0),(0,0)); 
neuper@37965
    40
val it = ((1,0),(0,0))*)
neuper@37965
    41
neuper@37965
    42
(*.convert internal floatingpoint prepresentation to int and float.*)
neuper@37965
    43
fun term_of_float T ((val1,    0), (         0,          0)) =
neuper@37965
    44
    term_of_num T val1
neuper@37965
    45
  | term_of_float T ((val1, val2), (precision1, precision2)) =
neuper@37965
    46
    let val pT = pairT T T
neuper@37965
    47
    in Const ("Float.Float", (pairT pT pT) --> T)
neuper@37965
    48
	     $ (pairt (pairt (Free (str_of_int val1, T))
neuper@37965
    49
			     (Free (str_of_int val2, T)))
neuper@37965
    50
		      (pairt (Free (str_of_int precision1, T))
neuper@37965
    51
			     (Free (str_of_int precision2, T))))
neuper@37965
    52
    end;
neuper@37965
    53
(*> val t = str2term "Float ((1,2),(0,0))";
neuper@37965
    54
> val Const ("Float.Float", fT) $ _ = t;
neuper@37965
    55
> atomtyp fT;
neuper@37965
    56
> val ffT = (pairT (pairT HOLogic.realT HOLogic.realT) 
neuper@37965
    57
> 	     (pairT HOLogic.realT HOLogic.realT)) --> HOLogic.realT;
neuper@37965
    58
> atomtyp ffT;
neuper@37965
    59
> fT = ffT;
neuper@37965
    60
val it = true : bool
neuper@37965
    61
neuper@37965
    62
t = float_term_of_num HOLogic.realT ((1,2),(0,0));
neuper@37965
    63
val it = true : bool*)
neuper@37965
    64
neuper@37965
    65
(*.assoc. convert internal floatingpoint prepresentation to int and float.*)
neuper@37965
    66
fun var_op_float v op_ optype ntyp ((v1, 0), (0, 0)) =
neuper@37965
    67
    var_op_num v op_ optype ntyp v1
neuper@37965
    68
  | var_op_float v op_ optype T ((v1, v2), (p1, p2)) =
neuper@37965
    69
    let val pT = pairT T T
neuper@37965
    70
    in Const (op_, optype) $ v $ 
neuper@37965
    71
	     (Const ("Float.Float", (pairT pT pT) --> T)
neuper@37965
    72
		    $ (pairt (pairt (Free (str_of_int v1, T))
neuper@37965
    73
				    (Free (str_of_int v2, T)))
neuper@37965
    74
			     (pairt (Free (str_of_int p1, T))
neuper@37965
    75
				    (Free (str_of_int p2, T)))))
neuper@37965
    76
    end;
neuper@37965
    77
(*> val t = str2term "a + b";
neuper@37965
    78
> val Const ("op +", optype) $ _ $ _ = t;
neuper@37965
    79
> val t = str2term "v + Float ((11,-1),(0,0))";val v = str2term "v";
neuper@37965
    80
> t = var_op_float v "op +" optype HOLogic.realT ((11,~1),(0,0));
neuper@37965
    81
val it = true : bool*)
neuper@37965
    82
neuper@37965
    83
(*.assoc. convert internal floatingpoint prepresentation to int and float.*)
neuper@37965
    84
fun float_op_var v op_ optype ntyp ((v1, 0), (0, 0)) =
neuper@37965
    85
    num_op_var v op_ optype ntyp v1
neuper@37965
    86
  | float_op_var v op_ optype T ((v1, v2), (p1, p2)) =
neuper@37965
    87
    let val pT = pairT T T
neuper@37965
    88
    in Const (op_,optype) $ 
neuper@37965
    89
	     (Const ("Float.Float", (pairT pT pT) --> T)
neuper@37965
    90
		    $ (pairt (pairt (Free (str_of_int v1, T))
neuper@37965
    91
				    (Free (str_of_int v2, T)))
neuper@37965
    92
			     (pairt (Free (str_of_int p1, T))
neuper@37965
    93
				    (Free (str_of_int p2, T))))) $ v
neuper@37965
    94
    end;
neuper@37965
    95
(*> val t = str2term "a + b";
neuper@37965
    96
> val Const ("op +", optype) $ _ $ _ = t;
neuper@37965
    97
> val t = str2term "Float ((11,-1),(0,0)) + v";val v = str2term "v";
neuper@37965
    98
> t = float_op_var v "op +" optype HOLogic.realT ((11,~1),(0,0));
neuper@37965
    99
val it = true : bool*)
neuper@37965
   100
*}
neuper@37965
   101
neuper@37965
   102
end