src/Tools/isac/ProgLang/ListC.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60417 00ba9518dc35
child 60586 007ef64dbb08
permissions -rw-r--r--
polish naming in Rewrite_Order
     1 (* Title:  functions on lists for Programs
     2    Author: Walther Neuper 0108
     3    (c) due to copyright terms
     4 *)
     5 
     6 theory ListC imports "$ISABELLE_ISAC/BaseDefinitions/BaseDefinitions"
     7 
     8 begin
     9 
    10 text \<open>Abstract:
    11   There are examples in the Knowledge, which demonstrate list processing, e.g. InsSort.thy.
    12 
    13   Lists in example calculations must have a different type for list processing in Isac programs.
    14   Presently the lists in examples have a new type xlist with the unfamiliar list
    15   notation "{||a, b, c||}". One could expect math-authors to use such unfamiliar notation 
    16   and expose students to familiar notation in calculations, and not vice versa as done presently.
    17 
    18   Thus this theory resides in ProgLang so far.
    19 \<close>
    20 
    21 subsection \<open>General constants\<close>
    22 consts
    23   EmptyList :: "bool list" 
    24   UniversalList :: "bool list" 
    25 
    26 ML \<open>
    27 \<close> ML \<open>
    28 val UniversalList = TermC.parseNEW''\<^theory> "UniversalList";
    29 val EmptyList = TermC.parseNEW''\<^theory> "[]::bool list";     
    30 \<close> ML \<open>
    31 \<close>
    32 
    33 subsection \<open>Type 'xlist' for Lucas-Interpretation | for InsSort.thy\<close>
    34 (* cp fom ~~/src/HOL/List.thy
    35    TODO: ask for shorter deliminters in xlist *)
    36 datatype  'a xlist =
    37     XNil  ("{|| ||}")
    38   | XCons (xhd: 'a)  (xtl: "'a xlist")  (infixr "@#" 65)
    39 
    40 syntax
    41    \<comment> \<open>list Enumeration\<close>
    42   "_xlist" :: "args => 'a xlist"    ("{|| (_) ||}")
    43 
    44 translations
    45   "{||x, xs||}" == "x@#{||xs||}"
    46   "{||x||}" == "x@#{|| ||}"
    47 
    48 term "{|| ||}"
    49 term "{||1,2,3||}"
    50 
    51 subsection \<open>Functions for 'xlist'\<close>
    52 (* TODO: 
    53 (1) revise, if definition of identifiers like LENGTH_NIL are still required.
    54 (2) switch the role of 'xlist' and 'list' in the functions below, in particular for
    55   'foldr', 'foldr_Nil', 'foldr_Cons' and 'xfoldr', 'xfoldr_Nil', 'xfoldr_Cons'.
    56   For transition phase just outcomment InsSort.thy and inssort.sml.
    57 *)
    58 
    59 primrec xfoldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a xlist \<Rightarrow> 'b \<Rightarrow> 'b" where
    60 xfoldr_Nil:  "xfoldr f {|| ||} = id" |
    61 xfoldr_Cons: "xfoldr f (x @# xs) = f x \<circ> xfoldr f xs"
    62 
    63 primrec Length   :: "'a list => real"
    64 where
    65 LENGTH_NIL:	"Length [] = 0" |
    66 LENGTH_CONS: "Length (x#xs) = 1 + Length xs"
    67 
    68 subsection \<open>Functions for 'list'\<close>
    69 
    70 consts NTH ::  "[real, 'a list] => 'a"
    71 axiomatization where
    72 NTH_NIL: "NTH 1 (x # xs) = x" and
    73 NTH_CONS: (*NO primrec, fun ..*)"1 < n ==> NTH n (x # xs) = NTH (n + -1) xs"
    74 
    75 (* redefine together with identifiers (still) required for Know_Store ..*)
    76 axiomatization where
    77 hd_thm:	"hd (x # xs) = x"
    78 
    79 axiomatization where
    80 tl_Nil:	"tl []  = []" and
    81 tl_Cons: "tl (x # xs) = xs"
    82 
    83 axiomatization where
    84 LAST:	"last (x # xs) = (if xs = [] then x else last xs)"
    85 
    86 axiomatization where
    87 butlast_Nil: "butlast []    = []" and
    88 butlast_Cons:	"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
    89 
    90 axiomatization where
    91 map_Nil:"map f [] = []" and
    92 map_Cons:	"map f (x # xs) = f x  #  map f xs"
    93 
    94 axiomatization where
    95 rev_Nil: "rev [] = []" and
    96 rev_Cons:	"rev (x # xs) = rev xs @ [x]"
    97 
    98 axiomatization where
    99 filter_Nil:	"filter P [] = []" and
   100 filter_Cons: "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)" 
   101 
   102 axiomatization where
   103 concat_Nil:	"concat [] = []" and
   104 concat_Cons: "concat (x # xs) = x @ concat xs"
   105 
   106 axiomatization where
   107 takeWhile_Nil: "takeWhile P []     = []" and
   108 takeWhile_Cons: "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
   109 
   110 axiomatization where
   111 dropWhile_Nil: "dropWhile P [] = []" and
   112 dropWhile_Cons: "dropWhile P (x # xs) = (if P x then dropWhile P xs else x#xs)"
   113 
   114 axiomatization where
   115 zip_Nil: "zip xs [] = []" and
   116 zip_Cons:	"zip xs (y # ys) = (case xs of [] => [] | z # zs => (z,y) # zip zs ys)"
   117 
   118 axiomatization where
   119 distinct_Nil:	"distinct []     = True" and
   120 distinct_Cons:	"distinct (x # xs) = (x ~: set xs & distinct xs)"
   121 
   122 axiomatization where
   123 remdups_Nil:	"remdups [] = []" and
   124 remdups_Cons:	"remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)" 
   125 
   126 consts lastI :: "'a list \<Rightarrow> 'a"
   127 axiomatization where
   128 last_thmI: "lastI (x#xs) = (if xs = [] then x else lastI xs)"
   129 
   130 subsection \<open>rule set for evaluating Porg_Expr, will be extended in several thys\<close>
   131 ML \<open>
   132 \<close> ML \<open>
   133 val prog_expr =
   134   Rule_Def.Repeat {id = "prog_expr", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty), 
   135     erls = Rule_Def.Empty, srls = Rule_Def.Empty, calc = [], errpatts = [],
   136     rules = [Rule_Def.Thm ("refl",  @{thm refl}),  (*'a<>b -> FALSE' by fun eval_equal*)
   137        Rule_Def.Thm ("o_apply",  @{thm o_apply}),
   138 
   139        Rule_Def.Thm ("NTH_CONS", @{thm NTH_CONS}),(*erls for cond. in Atools.ML*)
   140        Rule_Def.Thm ("NTH_NIL", @{thm NTH_NIL}),
   141        Rule_Def.Thm ("append_Cons", @{thm append_Cons}),
   142        Rule_Def.Thm ("append_Nil", @{thm append_Nil}),
   143 (*       Thm ("butlast_Cons",num_str @{thm butlast_Cons}),
   144        Thm ("butlast_Nil",num_str @{thm butlast_Nil}),*)
   145        Rule_Def.Thm ("concat_Cons", @{thm concat_Cons}),
   146        Rule_Def.Thm ("concat_Nil", @{thm concat_Nil}),
   147 (*       Rule_Def.Thm ("del_base",num_str @{thm del_base}),
   148        Rule_Def.Thm ("del_rec",num_str @{thm del_rec}), *)
   149 
   150        Rule_Def.Thm ("distinct_Cons", @{thm distinct_Cons}),
   151        Rule_Def.Thm ("distinct_Nil", @{thm distinct_Nil}),
   152        Rule_Def.Thm ("dropWhile_Cons", @{thm dropWhile_Cons}),
   153        Rule_Def.Thm ("dropWhile_Nil", @{thm dropWhile_Nil}),
   154        Rule_Def.Thm ("filter_Cons", @{thm filter_Cons}),
   155        Rule_Def.Thm ("filter_Nil", @{thm filter_Nil}),
   156        Rule_Def.Thm ("foldr_Cons", @{thm foldr_Cons}),
   157        Rule_Def.Thm ("foldr_Nil", @{thm foldr_Nil}),
   158        Rule_Def.Thm ("hd_thm", @{thm hd_thm}),
   159        Rule_Def.Thm ("LAST", @{thm LAST}),
   160        Rule_Def.Thm ("LENGTH_CONS", @{thm LENGTH_CONS}),
   161        Rule_Def.Thm ("LENGTH_NIL", @{thm LENGTH_NIL}),
   162 (*       Rule_Def.Thm ("list_diff_def",num_str @{thm list_diff_def}),*)
   163        Rule_Def.Thm ("map_Cons", @{thm map_Cons}),
   164        Rule_Def.Thm ("map_Nil", @{thm map_Cons}),
   165 (*       Rule_Def.Thm ("mem_Cons", @{thm mem_Cons}),
   166        Rule_Def.Thm ("mem_Nil", @{thm mem_Nil}), *)
   167 (*       Rule_Def.Thm ("null_Cons", @{thm null_Cons}),
   168        Rule_Def.Thm ("null_Nil", @{thm null_Nil}),*)
   169        Rule_Def.Thm ("remdups_Cons", @{thm remdups_Cons}),
   170        Rule_Def.Thm ("remdups_Nil", @{thm remdups_Nil}),
   171        Rule_Def.Thm ("rev_Cons", @{thm rev_Cons}),
   172        Rule_Def.Thm ("rev_Nil", @{thm rev_Nil}),
   173        Rule_Def.Thm ("take_Nil", @{thm take_Nil}),
   174        Rule_Def.Thm ("take_Cons", @{thm take_Cons}),
   175        Rule_Def.Thm ("tl_Cons", @{thm tl_Cons}),
   176        Rule_Def.Thm ("tl_Nil", @{thm tl_Nil}),
   177        Rule_Def.Thm ("zip_Cons", @{thm zip_Cons}),
   178        Rule_Def.Thm ("zip_Nil", @{thm zip_Nil})],
   179     scr = Rule_Def.Empty_Prog}: Rule_Set.T;
   180 \<close>
   181 rule_set_knowledge prog_expr = prog_expr
   182 
   183 ML \<open>
   184 \<close> ML \<open>
   185 \<close> ML \<open>
   186 \<close>
   187 end