| author | desharna | 
| Thu, 29 Jul 2021 17:08:46 +0200 | |
| changeset 74079 | 180ee02eb075 | 
| parent 60770 | 240563fbf41d | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/ex/Flag.thy | 
| 1474 | 2 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>Dutch national flag program -- except that the point of Dijkstra's example was to use | 
| 7 | arrays and this uses lists.\<close> | |
| 17456 | 8 | |
| 9 | theory Flag | |
| 10 | imports List | |
| 11 | begin | |
| 0 | 12 | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 13 | definition Colour :: "i set" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 14 | where "Colour == Unit + Unit + Unit" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 15 | |
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 16 | definition red :: "i" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 17 | where "red == inl(one)" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 18 | |
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 19 | definition white :: "i" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 20 | where "white == inr(inl(one))" | 
| 0 | 21 | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 22 | definition blue :: "i" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 23 | where "blue == inr(inr(one))" | 
| 0 | 24 | |
| 58977 | 25 | definition ccase :: "[i,i,i,i]\<Rightarrow>i" | 
| 26 | where "ccase(c,r,w,b) == when(c, \<lambda>x. r, \<lambda>wb. when(wb, \<lambda>x. w, \<lambda>x. b))" | |
| 0 | 27 | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 28 | definition flag :: "i" | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 29 | where | 
| 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 30 | "flag == lam l. letrec | 
| 17456 | 31 | flagx l be lcase(l,<[],<[],[]>>, | 
| 58977 | 32 | \<lambda>h t. split(flagx(t), \<lambda>lr p. split(p, \<lambda>lw lb. | 
| 17456 | 33 | ccase(h, <red$lr,<lw,lb>>, | 
| 34 | <lr,<white$lw,lb>>, | |
| 35 | <lr,<lw,blue$lb>>)))) | |
| 36 | in flagx(l)" | |
| 0 | 37 | |
| 58977 | 38 | axiomatization Perm :: "i \<Rightarrow> i \<Rightarrow> o" | 
| 39 | definition Flag :: "i \<Rightarrow> i \<Rightarrow> o" where | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 40 | "Flag(l,x) == ALL lr:List(Colour).ALL lw:List(Colour).ALL lb:List(Colour). | 
| 58977 | 41 | x = <lr,<lw,lb>> \<longrightarrow> | 
| 42 | (ALL c:Colour.(c mem lr = true \<longrightarrow> c=red) \<and> | |
| 43 | (c mem lw = true \<longrightarrow> c=white) \<and> | |
| 44 | (c mem lb = true \<longrightarrow> c=blue)) \<and> | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 45 | Perm(l,lr @ lw @ lb)" | 
| 0 | 46 | |
| 20140 | 47 | |
| 48 | lemmas flag_defs = Colour_def red_def white_def blue_def ccase_def | |
| 49 | ||
| 58977 | 50 | lemma ColourXH: "a : Colour \<longleftrightarrow> (a=red | a=white | a=blue)" | 
| 20140 | 51 | unfolding simp_type_defs flag_defs by blast | 
| 52 | ||
| 53 | lemma redT: "red : Colour" | |
| 54 | and whiteT: "white : Colour" | |
| 55 | and blueT: "blue : Colour" | |
| 56 | unfolding ColourXH by blast+ | |
| 57 | ||
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
28272diff
changeset | 58 | lemma ccaseT: | 
| 58977 | 59 | "\<lbrakk>c:Colour; c=red \<Longrightarrow> r : C(red); c=white \<Longrightarrow> w : C(white); c=blue \<Longrightarrow> b : C(blue)\<rbrakk> | 
| 60 | \<Longrightarrow> ccase(c,r,w,b) : C(c)" | |
| 32153 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 wenzelm parents: 
28272diff
changeset | 61 | unfolding flag_defs by ncanT | 
| 20140 | 62 | |
| 63 | lemma "flag : List(Colour)->List(Colour)*List(Colour)*List(Colour)" | |
| 64 | apply (unfold flag_def) | |
| 58971 | 65 | apply (typechk redT whiteT blueT ccaseT) | 
| 66 | apply clean_ccs | |
| 20140 | 67 | apply (erule ListPRI [THEN ListPR_wf [THEN wfI]]) | 
| 68 | apply assumption | |
| 69 | done | |
| 70 | ||
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
32153diff
changeset | 71 | lemma "flag : PROD l:List(Colour).{x:List(Colour)*List(Colour)*List(Colour).Flag(x,l)}"
 | 
| 20140 | 72 | apply (unfold flag_def) | 
| 58971 | 73 | apply (gen_ccs redT whiteT blueT ccaseT) | 
| 20140 | 74 | oops | 
| 0 | 75 | |
| 17456 | 76 | end |