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