src/Tools/isac/ProgLang/ListC.thy
author Thomas Leh <t.leh@gmx.at>
Wed, 27 Jul 2011 09:30:15 +0200
branchdecompose-isar
changeset 42211 51c3c007d7fd
parent 42197 7497ff20f1e8
child 42451 bc03b5d60547
permissions -rw-r--r--
prepared "axioms" (Isabelle2002) --> "axiomatization" (Isabelle2011)

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