src/Tools/isac/ProgLang/ListC.thy
changeset 59866 3b194392ea71
parent 59857 cbb3fae0381d
child 59871 82428ca0d23e
equal deleted inserted replaced
59865:75a9d629ea53 59866:3b194392ea71
     1 (* Title:  functions on lists for Programs
     1 (* Title:  functions on lists for Programs
     2    Author: Walther Neuper 0108
     2    Author: Walther Neuper 0108
     3    (c) due to copyright terms
     3    (c) due to copyright terms
     4 *)
     4 *)
     5 
     5 
     6 theory ListC imports "~~/src/Tools/isac/CalcElements/CalcElements"
     6 theory ListC imports "~~/src/Tools/isac/BaseDefinitions/BaseDefinitions"
     7 
     7 
     8 begin
     8 begin
     9 
     9 
    10 text \<open>Abstract:
    10 text \<open>Abstract:
    11   There are examples in the Knowledge, which demonstrate list processing, e.g. InsSort.thy.
    11   There are examples in the Knowledge, which demonstrate list processing, e.g. InsSort.thy.