src/HOLCF/ex/Hoare.thy
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 12036 49f6c49454c2
child 17291 94f6113fe9ed
permissions -rw-r--r--
Merged in license change from Isabelle2004
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
     1
(*  Title:      HOLCF/ex/hoare.thy
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
     3
    Author:     Franz Regensburger
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     4
12036
wenzelm
parents: 10835
diff changeset
     5
Theory for an example by C.A.R. Hoare
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     6
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
     7
p x = if b1 x 
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
     8
         then p (g x)
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     9
         else x fi
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    10
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    11
q x = if b1 x orelse b2 x 
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    12
         then q (g x)
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    13
         else x fi
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    14
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    15
Prove: for all b1 b2 g . 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    16
            q o p  = q 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    17
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    18
In order to get a nice notation we fix the functions b1,b2 and g in the
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    19
signature of this example
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    20
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    21
*)
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    22
2380
90280b3a538b Dummy change to document the change in revision 1.5:
sandnerr
parents: 2379
diff changeset
    23
Hoare = HOLCF + 
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    24
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    25
consts
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    26
        b1:: "'a -> tr"
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    27
        b2:: "'a -> tr"
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    28
         g:: "'a -> 'a"
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    29
        p :: "'a -> 'a"
21eb5e156d91 expanded tabs
clasohm
parents: 1168
diff changeset
    30
        q :: "'a -> 'a"
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    31
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    32
defs
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    33
10835
nipkow
parents: 2380
diff changeset
    34
  p_def  "p == fix$(LAM f. LAM x.
nipkow
parents: 2380
diff changeset
    35
                 If b1$x then f$(g$x) else x fi)"
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    36
10835
nipkow
parents: 2380
diff changeset
    37
  q_def  "q == fix$(LAM f. LAM x.
nipkow
parents: 2380
diff changeset
    38
                 If b1$x orelse b2$x then f$(g$x) else x fi)"
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    39
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    40
end
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    41