src/Tools/isac/ProgLang/ListC.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Fri, 26 Aug 2016 12:25:03 +0200
changeset 59234 d12736878a81
parent 59206 ebf4a8a63371
child 59244 6870ee668115
permissions -rw-r--r--
separate 'type xlist' for Lucas-Interpretation

Note: xlist is not yet used, only tested with insertion sort.
TODO: switch 'list' and 'xlist'; intermediately outcomment InsSort.thy, inssort.sml
neuper@52122
     1
(* Title:  functions on lists for Programs
neuper@37965
     2
   Author: Walther Neuper 0108
neuper@37965
     3
   (c) due to copyright terms
neuper@37906
     4
*)
neuper@37906
     5
neuper@52122
     6
theory ListC 
wneuper@59206
     7
imports "~~/src/Tools/isac/KEStore"
neuper@37906
     8
begin
neuper@52122
     9
neuper@52122
    10
ML_file "~~/src/Tools/isac/ProgLang/termC.sml"
neuper@52122
    11
ML_file "~~/src/Tools/isac/ProgLang/calculate.sml"
neuper@52122
    12
ML_file "~~/src/Tools/isac/ProgLang/rewrite.sml"
neuper@37906
    13
wneuper@59234
    14
subsection {* Notes on Isac's programming language *}
wneuper@59234
    15
text {*
wneuper@59234
    16
  Isac's programming language combines tacticals (TRY, etc) and 
wneuper@59234
    17
  tactics (Rewrite, etc) with list processing.
wneuper@59234
    18
wneuper@59234
    19
  In order to distinguish list expressions of the meta (programming)
wneuper@59234
    20
  language from the object language in Lucas-Interpretation, a 
wneuper@59234
    21
  new 'type xlist' is introduced.
wneuper@59234
    22
  TODO: Switch the role of 'xlist' and 'list' (the former only used
wneuper@59234
    23
  by InsSort.thy)
wneuper@59234
    24
wneuper@59234
    25
  Isac's programming language preceeded the function package
neuper@52148
    26
  in 2002. For naming "axiomatization" is used for reasons of uniformity
neuper@52148
    27
  with the other replacements for "axioms".
neuper@52148
    28
  Another reminiscence from Isabelle2002 are Isac-specific numerals,
neuper@52148
    29
  introduced in order to have floating point numerals at a time, 
neuper@52148
    30
  when Isabelle did not consider that requirement. For the sake of uniformity
wneuper@59234
    31
  'nat' from List.thy is replaced by 'real' by 'fun parse',
wneuper@59234
    32
  however, 'fun parseNEW' has started to replace this fix (after finishing
wneuper@59234
    33
  this fix, there will be a 'rename all parseNEW --> parse).
neuper@52148
    34
neuper@52148
    35
  Note: *one* "axiomatization" over all equations caused strange "'a list list"
neuper@52148
    36
  types.
neuper@52148
    37
*}
neuper@37906
    38
wneuper@59234
    39
subsection {* Type 'xlist' for Lucas-Interpretation *}
wneuper@59234
    40
(*TODO: ask for shorter deliminters in xlist *)
wneuper@59234
    41
datatype 'a xlist =
wneuper@59234
    42
    XNil    ("[|| ||]")
wneuper@59234
    43
  | XCons 'a  "'a xlist"    (infixr "@#" 65)
wneuper@59234
    44
wneuper@59234
    45
syntax
wneuper@59234
    46
  -- {* list Enumeration *}
wneuper@59234
    47
  "_xlist" :: "args => 'a xlist"    ("[||(_)||]")
wneuper@59234
    48
wneuper@59234
    49
translations
wneuper@59234
    50
  "[||x, xs||]" == "x@#[||xs||]"
wneuper@59234
    51
  "[||x||]" == "x@#[|| ||]"
wneuper@59234
    52
wneuper@59234
    53
term "[|| ||]"
wneuper@59234
    54
term "[||1,2,3||]"
wneuper@59234
    55
wneuper@59234
    56
subsection {* Functions for 'xlist' *}
wneuper@59234
    57
(* TODO: 
wneuper@59234
    58
(1) revise, if definition of identifiers like LENGTH_NIL are still required.
wneuper@59234
    59
(2) switch the role of 'xlist' and 'list' in the functions below, in particular for
wneuper@59234
    60
  'foldr', 'foldr_Nil', 'foldr_Cons' and 'xfoldr', 'xfoldr_Nil', 'xfoldr_Cons'.
wneuper@59234
    61
  For transition phase just outcomment InsSort.thy and inssort.sml.
wneuper@59234
    62
*)
wneuper@59234
    63
wneuper@59234
    64
primrec xfoldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a xlist \<Rightarrow> 'b \<Rightarrow> 'b" where
wneuper@59234
    65
xfoldr_Nil:  "xfoldr f [|| ||] = id" |
wneuper@59234
    66
xfoldr_Cons: "xfoldr f (x @# xs) = f x \<circ> xfoldr f xs"
wneuper@59234
    67
neuper@37997
    68
primrec LENGTH   :: "'a list => real"
neuper@37906
    69
where
neuper@52148
    70
LENGTH_NIL:	"LENGTH [] = 0" |
neuper@52148
    71
LENGTH_CONS: "LENGTH (x#xs) = 1 + LENGTH xs"
neuper@37906
    72
neuper@52148
    73
consts NTH ::  "[real, 'a list] => 'a"
neuper@52148
    74
axiomatization where
neuper@52148
    75
NTH_NIL: "NTH 1 (x # xs) = x" and
neuper@52148
    76
NTH_CONS: (*NO primrec, fun ..*)"1 < n ==> NTH n (x # xs) = NTH (n + -1) xs"
neuper@37906
    77
neuper@52148
    78
axiomatization where
neuper@52148
    79
hd_thm:	"hd (x # xs) = x"
t@42197
    80
neuper@52148
    81
axiomatization where
neuper@52148
    82
tl_Nil:	"tl []  = []" and
neuper@52148
    83
tl_Cons: "tl (x # xs) = xs"
neuper@37906
    84
neuper@52148
    85
axiomatization where
neuper@52148
    86
LAST:	"last (x # xs) = (if xs = [] then x else last xs)"
neuper@37906
    87
neuper@52148
    88
axiomatization where
neuper@52148
    89
butlast_Nil: "butlast []    = []" and
neuper@52148
    90
butlast_Cons:	"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
neuper@52148
    91
neuper@52148
    92
axiomatization where
neuper@52148
    93
map_Nil:"map f [] = []" and
neuper@52148
    94
map_Cons:	"map f (x # xs) = f x  #  map f xs"
neuper@52148
    95
neuper@52148
    96
axiomatization where
neuper@52148
    97
rev_Nil: "rev [] = []" and
neuper@52148
    98
rev_Cons:	"rev (x # xs) = rev xs @ [x]"
neuper@52148
    99
neuper@52148
   100
axiomatization where
neuper@52148
   101
filter_Nil:	"filter P [] = []" and
neuper@52148
   102
filter_Cons: "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" 
neuper@52148
   103
neuper@52148
   104
axiomatization where
neuper@52148
   105
concat_Nil:	"concat [] = []" and
neuper@52148
   106
concat_Cons: "concat (x # xs) = x @ concat xs"
neuper@52148
   107
neuper@52148
   108
axiomatization where
neuper@52148
   109
takeWhile_Nil: "takeWhile P []     = []" and
neuper@52148
   110
takeWhile_Cons: "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
neuper@52148
   111
neuper@52148
   112
axiomatization where
neuper@52148
   113
dropWhile_Nil: "dropWhile P [] = []" and
neuper@52148
   114
dropWhile_Cons: "dropWhile P (x # xs) = (if P x then dropWhile P xs else x#xs)"
neuper@52148
   115
neuper@52148
   116
axiomatization where
neuper@52148
   117
zip_Nil: "zip xs [] = []" and
neuper@52148
   118
zip_Cons:	"zip xs (y # ys) = (case xs of [] => [] | z # zs => (z,y) # zip zs ys)"
neuper@52148
   119
neuper@52148
   120
axiomatization where
neuper@52148
   121
distinct_Nil:	"distinct []     = True" and
neuper@52148
   122
distinct_Cons:	"distinct (x # xs) = (x ~: set xs & distinct xs)"
neuper@52148
   123
neuper@52148
   124
axiomatization where
neuper@52148
   125
remdups_Nil:	"remdups [] = []" and
neuper@52148
   126
remdups_Cons:	"remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)" 
neuper@37906
   127
neuper@37906
   128
(** Lexicographic orderings on lists ...!!!**)
neuper@37906
   129
neuper@37947
   130
ML{* (*the former ListC.ML*)
neuper@37906
   131
(** rule set for evaluating listexpr in scripts **)
neuper@37906
   132
val list_rls = 
neuper@52139
   133
  Rls{id = "list_rls", preconds = [], rew_ord = ("dummy_ord", dummy_ord), 
neuper@52139
   134
    erls = Erls, srls = Erls, calc = [], errpatts = [],
neuper@52139
   135
    rules = [Thm ("refl", num_str @{thm refl}),  (*'a<>b -> FALSE' by fun eval_equal*)
neuper@37906
   136
       Thm ("o_apply", num_str @{thm o_apply}),
neuper@37906
   137
neuper@37906
   138
       Thm ("NTH_CONS",num_str @{thm NTH_CONS}),(*erls for cond. in Atools.ML*)
neuper@37906
   139
       Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
neuper@37906
   140
       Thm ("append_Cons",num_str @{thm append_Cons}),
neuper@37906
   141
       Thm ("append_Nil",num_str @{thm append_Nil}),
neuper@52148
   142
(*       Thm ("butlast_Cons",num_str @{thm butlast_Cons}),
neuper@52148
   143
       Thm ("butlast_Nil",num_str @{thm butlast_Nil}),*)
neuper@37906
   144
       Thm ("concat_Cons",num_str @{thm concat_Cons}),
neuper@37906
   145
       Thm ("concat_Nil",num_str @{thm concat_Nil}),
neuper@52148
   146
(*       Thm ("del_base",num_str @{thm del_base}),
neuper@52148
   147
       Thm ("del_rec",num_str @{thm del_rec}), *)
neuper@37906
   148
neuper@37906
   149
       Thm ("distinct_Cons",num_str @{thm distinct_Cons}),
neuper@37906
   150
       Thm ("distinct_Nil",num_str @{thm distinct_Nil}),
neuper@37906
   151
       Thm ("dropWhile_Cons",num_str @{thm dropWhile_Cons}),
neuper@37906
   152
       Thm ("dropWhile_Nil",num_str @{thm dropWhile_Nil}),
neuper@37906
   153
       Thm ("filter_Cons",num_str @{thm filter_Cons}),
neuper@37906
   154
       Thm ("filter_Nil",num_str @{thm filter_Nil}),
neuper@37906
   155
       Thm ("foldr_Cons",num_str @{thm foldr_Cons}),
neuper@37906
   156
       Thm ("foldr_Nil",num_str @{thm foldr_Nil}),
neuper@37906
   157
       Thm ("hd_thm",num_str @{thm hd_thm}),
neuper@37906
   158
       Thm ("LAST",num_str @{thm LAST}),
neuper@37906
   159
       Thm ("LENGTH_CONS",num_str @{thm LENGTH_CONS}),
neuper@37906
   160
       Thm ("LENGTH_NIL",num_str @{thm LENGTH_NIL}),
neuper@52148
   161
(*       Thm ("list_diff_def",num_str @{thm list_diff_def}),*)
neuper@37906
   162
       Thm ("map_Cons",num_str @{thm map_Cons}),
neuper@37906
   163
       Thm ("map_Nil",num_str @{thm map_Cons}),
neuper@52148
   164
(*       Thm ("mem_Cons",num_str @{thm mem_Cons}),
neuper@52148
   165
       Thm ("mem_Nil",num_str @{thm mem_Nil}), *)
neuper@52148
   166
(*       Thm ("null_Cons",num_str @{thm null_Cons}),
neuper@52148
   167
       Thm ("null_Nil",num_str @{thm null_Nil}),*)
neuper@37906
   168
       Thm ("remdups_Cons",num_str @{thm remdups_Cons}),
neuper@37906
   169
       Thm ("remdups_Nil",num_str @{thm remdups_Nil}),
neuper@37906
   170
       Thm ("rev_Cons",num_str @{thm rev_Cons}),
neuper@37906
   171
       Thm ("rev_Nil",num_str @{thm rev_Nil}),
neuper@37906
   172
       Thm ("take_Nil",num_str @{thm take_Nil}),
neuper@37906
   173
       Thm ("take_Cons",num_str @{thm take_Cons}),
neuper@37906
   174
       Thm ("tl_Cons",num_str @{thm tl_Cons}),
neuper@37906
   175
       Thm ("tl_Nil",num_str @{thm tl_Nil}),
neuper@37906
   176
       Thm ("zip_Cons",num_str @{thm zip_Cons}),
neuper@52139
   177
       Thm ("zip_Nil",num_str @{thm zip_Nil})],
neuper@52139
   178
    scr = EmptyScr}:rls;
neuper@37906
   179
*}
neuper@52122
   180
setup {* KEStore_Elems.add_rlss [("list_rls", (Context.theory_name @{theory}, list_rls))] *}
neuper@52122
   181
neuper@37906
   182
end