src/Tools/isac/ProgLang/ListC.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sun, 29 Sep 2013 18:27:37 +0200
changeset 52139 511fc271f783
parent 52135 9659f65c9df6
child 52148 aabc6c8e930a
permissions -rw-r--r--
collected updates since changeset 9690a8d5f1c
     1 (* Title:  functions on lists for Programs
     2    Author: Walther Neuper 0108
     3    (c) due to copyright terms
     4 *)
     5 
     6 theory ListC 
     7 imports Complex_Main "~~/src/Tools/isac/KEStore"
     8 begin
     9 
    10 ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
    11 ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
    12 ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
    13 
    14 text {* 'nat' in List.thy replaced by 'real' *}
    15 
    16 primrec LENGTH   :: "'a list => real"
    17 where
    18   LENGTH_NIL:	"LENGTH [] = 0"     (*length: 'a list => nat*)
    19 | LENGTH_CONS: "LENGTH (x#xs) = 1 + LENGTH xs"
    20 
    21 primrec del :: "['a list, 'a] => 'a list"
    22 where
    23   del_base: "del [] x = []"
    24 | del_rec:  "del (y#ys) x = (if x = y then ys else y#(del ys x))"
    25 
    26 definition
    27   list_diff :: "['a list, 'a list] => 'a list"         (* as -- bs *)
    28               ("(_ --/ _)" [66, 66] 65)
    29   where "a -- b == foldl del a b"
    30   
    31 consts NTH ::  "[real, 'a list] => 'a"
    32 
    33 axioms(*axiomatization where*)
    34  (*** more than one non-variable in pattern in "nth_ 1 [x] = x"--*)
    35   NTH_NIL:      "NTH 1 (x#xs) = x" (*and*)
    36 (*  NTH_CONS:     "NTH n (x#xs) = NTH (n+ -1) xs"  *)
    37 
    38 (*rewriter does not reach base case   ......    ;
    39   the condition involves another rule set (erls, eval_binop in Atools):*)
    40   NTH_CONS:     "1 < n ==> NTH n (x#xs) = NTH (n+ - 1) xs" (*and*)
    41 
    42 (*primrec from Isabelle/src/HOL/List.thy -- def.twice not allowed*)
    43 (*primrec*)
    44   hd_thm:	"hd(x#xs) = x" (*and*)
    45 (*primrec*)
    46   tl_Nil:	"tl([])   = []" (*and*)
    47   tl_Cons:		"tl(x#xs) = xs" (*and*)
    48 (*primrec*)
    49   null_Nil:	"null([])   = True" (*and*)
    50   null_Cons:	"null(x#xs) = False" (*and*)
    51 (*primrec*)
    52   LAST:	"last(x#xs) = (if xs=[] then x else last xs)" (*and*)
    53 (*primrec*)
    54   butlast_Nil:	"butlast []    = []" (*and*)
    55   butlast_Cons:	"butlast(x#xs) = (if xs=[] then [] else x#butlast xs)" (*and*)
    56 (*primrec
    57   mem_Nil:	"x mem []     = False"
    58   mem_Cons:	"x mem (y#ys) = (if y=x then True else x mem ys)"
    59 *)
    60   mem_Nil:	"x : set []     = False" (*and*)
    61   mem_Cons:	"x : set (y#ys) = (if y=x then True else x : set ys)" (*and*)
    62 (*primrec-------already named---
    63   "set [] = {}"
    64   "set (x#xs) = insert x (set xs)"
    65   primrec
    66   list_all_Nil  "list_all P [] = True"
    67   list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"
    68 ----------------*)
    69 (*primrec*)
    70   map_Nil:	"map f []     = []" (*and*)
    71   map_Cons:	"map f (x#xs) = f(x)#map f xs" (*and*)
    72 (*primrec*)
    73   append_Nil:  "[]    @ys = ys" (*and*)
    74   append_Cons: "(x#xs)@ys = x#(xs@ys)" (*and*)
    75 (*primrec*)
    76   rev_Nil:	"rev([])   = []" (*and*)
    77   rev_Cons:	"rev(x#xs) = rev(xs) @ [x]" (*and*)
    78 (*primrec*)
    79   filter_Nil:	"filter P []     = []" (*and*)
    80   filter_Cons:	"filter P (x#xs) =(if P x then x#filter P xs else filter P xs)" (*and*)
    81 (*primrec-------already named---
    82   foldl_Nil  "foldl f a [] = a"
    83   foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs"
    84 ----------------*)
    85 (*primrec*)
    86   foldr_Nil:	"foldr f [] a     = a" (*and*)
    87   foldr_Cons:	"foldr f (x#xs) a = f x (foldr f xs a)" (*and*)
    88 (*primrec*)
    89   concat_Nil:	"concat([])   = []" (*and*)
    90   concat_Cons:	"concat(x#xs) = x @ concat(xs)" (*and*)
    91 (*primrec-------already named---
    92   drop_Nil  "drop n [] = []"
    93   drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"
    94   (* Warning: simpset does not contain this definition but separate theorems 
    95      for n=0 / n=Suc k*)
    96 (*primrec*)
    97   take_Nil  "take n [] = []"
    98   take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
    99   (* Warning: simpset does not contain this definition but separate theorems 
   100      for n=0 / n=Suc k*)
   101 (*primrec*) 
   102   nth_Cons  "(x#xs)!n = (case n of 0 => x | (Suc k) => xs!k)"
   103   (* Warning: simpset does not contain this definition but separate theorems 
   104      for n=0 / n=Suc k*)
   105 (*primrec*)
   106  "    [][i:=v] = []"
   107  "(x#xs)[i:=v] = (case i of 0     => v # xs 
   108 			  | Suc j => x # xs[j:=v])"
   109 ----------------*)
   110 (*primrec*)
   111   takeWhile_Nil:	"takeWhile P []     = []" (*and*)
   112   takeWhile_Cons:
   113   "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])" (*and*)
   114 (*primrec*)
   115   dropWhile_Nil:	"dropWhile P []     = []" (*and*)
   116   dropWhile_Cons:
   117   "dropWhile P (x#xs) = (if P x then dropWhile P xs else x#xs)" (*and*)
   118 (*primrec*)
   119   zip_Nil:	"zip xs []     = []" (*and*)
   120   zip_Cons:	"zip xs (y#ys) =(case xs of [] => [] | z#zs =>(z,y)#zip zs ys)" (*and*)
   121   (* Warning: simpset does not contain this definition but separate theorems 
   122      for xs=[] / xs=z#zs *)
   123 (*primrec
   124   upt_0   "[i..0(] = []"
   125   upt_Suc "[i..(Suc j)(] = (if i <= j then [i..j(] @ [j] else [])"
   126 *)
   127 (*primrec*)
   128   distinct_Nil:	"distinct []     = True" (*and*)
   129   distinct_Cons:	"distinct (x#xs) = (x ~: set xs & distinct xs)" (*and*)
   130 (*primrec*)
   131   remdups_Nil:	"remdups [] = []" (*and*)
   132   remdups_Cons:	"remdups (x#xs) = 
   133 		 (if x : set xs then remdups xs else x # remdups xs)" 
   134 (*primrec-------already named---
   135   replicate_0   "replicate  0      x = []"
   136   replicate_Suc "replicate (Suc n) x = x # replicate n x"
   137 ----------------*)
   138 
   139 (** Lexicographic orderings on lists ...!!!**)
   140 
   141 ML{* (*the former ListC.ML*)
   142 (** rule set for evaluating listexpr in scripts **)
   143 val list_rls = 
   144   Rls{id = "list_rls", preconds = [], rew_ord = ("dummy_ord", dummy_ord), 
   145     erls = Erls, srls = Erls, calc = [], errpatts = [],
   146     rules = [Thm ("refl", num_str @{thm refl}),  (*'a<>b -> FALSE' by fun eval_equal*)
   147        Thm ("o_apply", num_str @{thm o_apply}),
   148 
   149        Thm ("NTH_CONS",num_str @{thm NTH_CONS}),(*erls for cond. in Atools.ML*)
   150        Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   151        Thm ("append_Cons",num_str @{thm append_Cons}),
   152        Thm ("append_Nil",num_str @{thm append_Nil}),
   153        Thm ("butlast_Cons",num_str @{thm butlast_Cons}),
   154        Thm ("butlast_Nil",num_str @{thm butlast_Nil}),
   155        Thm ("concat_Cons",num_str @{thm concat_Cons}),
   156        Thm ("concat_Nil",num_str @{thm concat_Nil}),
   157        Thm ("del_base",num_str @{thm del_base}),
   158        Thm ("del_rec",num_str @{thm del_rec}),
   159 
   160        Thm ("distinct_Cons",num_str @{thm distinct_Cons}),
   161        Thm ("distinct_Nil",num_str @{thm distinct_Nil}),
   162        Thm ("dropWhile_Cons",num_str @{thm dropWhile_Cons}),
   163        Thm ("dropWhile_Nil",num_str @{thm dropWhile_Nil}),
   164        Thm ("filter_Cons",num_str @{thm filter_Cons}),
   165        Thm ("filter_Nil",num_str @{thm filter_Nil}),
   166        Thm ("foldr_Cons",num_str @{thm foldr_Cons}),
   167        Thm ("foldr_Nil",num_str @{thm foldr_Nil}),
   168        Thm ("hd_thm",num_str @{thm hd_thm}),
   169        Thm ("LAST",num_str @{thm LAST}),
   170        Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
   171        Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
   172        Thm ("list_diff_def",num_str @{thm list_diff_def}),
   173        Thm ("map_Cons",num_str @{thm map_Cons}),
   174        Thm ("map_Nil",num_str @{thm map_Cons}),
   175        Thm ("mem_Cons",num_str @{thm mem_Cons}),
   176        Thm ("mem_Nil",num_str @{thm mem_Nil}),
   177        Thm ("null_Cons",num_str @{thm null_Cons}),
   178        Thm ("null_Nil",num_str @{thm null_Nil}),
   179        Thm ("remdups_Cons",num_str @{thm remdups_Cons}),
   180        Thm ("remdups_Nil",num_str @{thm remdups_Nil}),
   181        Thm ("rev_Cons",num_str @{thm rev_Cons}),
   182        Thm ("rev_Nil",num_str @{thm rev_Nil}),
   183        Thm ("take_Nil",num_str @{thm take_Nil}),
   184        Thm ("take_Cons",num_str @{thm take_Cons}),
   185        Thm ("tl_Cons",num_str @{thm tl_Cons}),
   186        Thm ("tl_Nil",num_str @{thm tl_Nil}),
   187        Thm ("zip_Cons",num_str @{thm zip_Cons}),
   188        Thm ("zip_Nil",num_str @{thm zip_Nil})],
   189     scr = EmptyScr}:rls;
   190 *}
   191 
   192 ML{*
   193 ruleset' := overwritelthy @{theory} (!ruleset', [("list_rls", list_rls)]);
   194 *}
   195 setup {* KEStore_Elems.add_rlss [("list_rls", (Context.theory_name @{theory}, list_rls))] *}
   196 
   197 end