| author | wenzelm | 
| Fri, 24 May 2024 16:15:27 +0200 | |
| changeset 80188 | 3956e8b6a9c9 | 
| parent 60770 | 240563fbf41d | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/ex/List.thy | 
| 1474 | 2 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>Programs defined over lists\<close> | 
| 17456 | 7 | |
| 8 | theory List | |
| 9 | imports Nat | |
| 10 | begin | |
| 0 | 11 | |
| 58977 | 12 | definition map :: "[i\<Rightarrow>i,i]\<Rightarrow>i" | 
| 13 | where "map(f,l) == lrec(l, [], \<lambda>x xs g. f(x)$g)" | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 14 | |
| 58977 | 15 | definition comp :: "[i\<Rightarrow>i,i\<Rightarrow>i]\<Rightarrow>i\<Rightarrow>i" (infixr "\<circ>" 55) | 
| 16 | where "f \<circ> g == (\<lambda>x. f(g(x)))" | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 17 | |
| 58977 | 18 | definition append :: "[i,i]\<Rightarrow>i" (infixr "@" 55) | 
| 19 | where "l @ m == lrec(l, m, \<lambda>x xs g. x$g)" | |
| 0 | 20 | |
| 58977 | 21 | axiomatization member :: "[i,i]\<Rightarrow>i" (infixr "mem" 55) (* FIXME dangling eq *) | 
| 22 | where member_ax: "a mem l == lrec(l, false, \<lambda>h t g. if eq(a,h) then true else g)" | |
| 0 | 23 | |
| 58977 | 24 | definition filter :: "[i,i]\<Rightarrow>i" | 
| 25 | where "filter(f,l) == lrec(l, [], \<lambda>x xs g. if f`x then x$g else g)" | |
| 0 | 26 | |
| 58977 | 27 | definition flat :: "i\<Rightarrow>i" | 
| 28 | where "flat(l) == lrec(l, [], \<lambda>h t g. h @ g)" | |
| 0 | 29 | |
| 58977 | 30 | definition partition :: "[i,i]\<Rightarrow>i" where | 
| 31 | "partition(f,l) == letrec part l a b be lcase(l, <a,b>, \<lambda>x xs. | |
| 17456 | 32 | if f`x then part(xs,x$a,b) else part(xs,a,x$b)) | 
| 1149 | 33 | in part(l,[],[])" | 
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 34 | |
| 58977 | 35 | definition insert :: "[i,i,i]\<Rightarrow>i" | 
| 36 | where "insert(f,a,l) == lrec(l, a$[], \<lambda>h t g. if f`a`h then a$h$t else h$g)" | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 37 | |
| 58977 | 38 | definition isort :: "i\<Rightarrow>i" | 
| 39 | where "isort(f) == lam l. lrec(l, [], \<lambda>h t g. insert(f,h,g))" | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 40 | |
| 58977 | 41 | definition qsort :: "i\<Rightarrow>i" where | 
| 42 | "qsort(f) == lam l. letrec qsortx l be lcase(l, [], \<lambda>h t. | |
| 17456 | 43 | let p be partition(f`h,t) | 
| 58977 | 44 | in split(p, \<lambda>x y. qsortx(x) @ h$qsortx(y))) | 
| 1149 | 45 | in qsortx(l)" | 
| 0 | 46 | |
| 20140 | 47 | lemmas list_defs = map_def comp_def append_def filter_def flat_def | 
| 48 | insert_def isort_def partition_def qsort_def | |
| 49 | ||
| 50 | lemma listBs [simp]: | |
| 58977 | 51 | "\<And>f g. (f \<circ> g) = (\<lambda>a. f(g(a)))" | 
| 52 | "\<And>a f g. (f \<circ> g)(a) = f(g(a))" | |
| 53 | "\<And>f. map(f,[]) = []" | |
| 54 | "\<And>f x xs. map(f,x$xs) = f(x)$map(f,xs)" | |
| 55 | "\<And>m. [] @ m = m" | |
| 56 | "\<And>x xs m. x$xs @ m = x$(xs @ m)" | |
| 57 | "\<And>f. filter(f,[]) = []" | |
| 58 | "\<And>f x xs. filter(f,x$xs) = if f`x then x$filter(f,xs) else filter(f,xs)" | |
| 20140 | 59 | "flat([]) = []" | 
| 58977 | 60 | "\<And>x xs. flat(x$xs) = x @ flat(xs)" | 
| 61 | "\<And>a f. insert(f,a,[]) = a$[]" | |
| 62 | "\<And>a f xs. insert(f,a,x$xs) = if f`a`x then a$x$xs else x$insert(f,a,xs)" | |
| 20140 | 63 | by (simp_all add: list_defs) | 
| 64 | ||
| 58977 | 65 | lemma nmapBnil: "n:Nat \<Longrightarrow> map(f) ^ n ` [] = []" | 
| 20140 | 66 | apply (erule Nat_ind) | 
| 67 | apply simp_all | |
| 68 | done | |
| 69 | ||
| 58977 | 70 | lemma nmapBcons: "n:Nat \<Longrightarrow> map(f)^n`(x$xs) = (f^n`x)$(map(f)^n`xs)" | 
| 20140 | 71 | apply (erule Nat_ind) | 
| 72 | apply simp_all | |
| 73 | done | |
| 74 | ||
| 75 | ||
| 58977 | 76 | lemma mapT: "\<lbrakk>\<And>x. x:A \<Longrightarrow> f(x):B; l : List(A)\<rbrakk> \<Longrightarrow> map(f,l) : List(B)" | 
| 20140 | 77 | apply (unfold map_def) | 
| 58971 | 78 | apply typechk | 
| 20140 | 79 | apply blast | 
| 80 | done | |
| 81 | ||
| 58977 | 82 | lemma appendT: "\<lbrakk>l : List(A); m : List(A)\<rbrakk> \<Longrightarrow> l @ m : List(A)" | 
| 20140 | 83 | apply (unfold append_def) | 
| 58971 | 84 | apply typechk | 
| 20140 | 85 | done | 
| 86 | ||
| 87 | lemma appendTS: | |
| 58977 | 88 |   "\<lbrakk>l : {l:List(A). m : {m:List(A).P(l @ m)}}\<rbrakk> \<Longrightarrow> l @ m : {x:List(A). P(x)}"
 | 
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 89 | by (blast intro!: appendT) | 
| 20140 | 90 | |
| 58977 | 91 | lemma filterT: "\<lbrakk>f:A->Bool; l : List(A)\<rbrakk> \<Longrightarrow> filter(f,l) : List(A)" | 
| 20140 | 92 | apply (unfold filter_def) | 
| 58971 | 93 | apply typechk | 
| 20140 | 94 | done | 
| 95 | ||
| 58977 | 96 | lemma flatT: "l : List(List(A)) \<Longrightarrow> flat(l) : List(A)" | 
| 20140 | 97 | apply (unfold flat_def) | 
| 58971 | 98 | apply (typechk appendT) | 
| 20140 | 99 | done | 
| 100 | ||
| 58977 | 101 | lemma insertT: "\<lbrakk>f : A->A->Bool; a:A; l : List(A)\<rbrakk> \<Longrightarrow> insert(f,a,l) : List(A)" | 
| 20140 | 102 | apply (unfold insert_def) | 
| 58971 | 103 | apply typechk | 
| 20140 | 104 | done | 
| 105 | ||
| 106 | lemma insertTS: | |
| 58977 | 107 |   "\<lbrakk>f : {f:A->A->Bool. a : {a:A. l : {l:List(A).P(insert(f,a,l))}}} \<rbrakk> \<Longrightarrow>
 | 
| 20140 | 108 |    insert(f,a,l)  : {x:List(A). P(x)}"
 | 
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
35762diff
changeset | 109 | by (blast intro!: insertT) | 
| 20140 | 110 | |
| 58977 | 111 | lemma partitionT: "\<lbrakk>f:A->Bool; l : List(A)\<rbrakk> \<Longrightarrow> partition(f,l) : List(A)*List(A)" | 
| 20140 | 112 | apply (unfold partition_def) | 
| 58971 | 113 | apply typechk | 
| 114 | apply clean_ccs | |
| 20140 | 115 | apply (rule ListPRI [THEN wfstI, THEN ListPR_wf [THEN wmap_wf, THEN wfI]]) | 
| 116 | apply assumption+ | |
| 117 | apply (rule ListPRI [THEN wfstI, THEN ListPR_wf [THEN wmap_wf, THEN wfI]]) | |
| 118 | apply assumption+ | |
| 119 | done | |
| 17456 | 120 | |
| 0 | 121 | end |