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