src/HOLCF/ex/hoare.thy
author lcp
Thu, 12 Jan 1995 03:00:58 +0100
changeset 850 a744f9749885
parent 244 929fc2c63bd0
permissions -rw-r--r--
Added constants Ord_alt, ++, **
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     1
(*  Title:	HOLCF/ex/hoare.thy
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     2
    ID:         $Id$
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     3
    Author: 	Franz Regensburger
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     4
    Copyright	1993 Technische Universitaet Muenchen
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     5
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     6
Theory for an example by C.A.R. Hoare 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     7
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     8
p x = if b1(x) 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     9
         then p(g(x))
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    10
         else x fi
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    11
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    12
q x = if b1(x) orelse b2(x) 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    13
         then q (g(x))
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    14
         else x fi
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    15
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    16
Prove: for all b1 b2 g . 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    17
            q o p  = q 
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    18
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    19
In order to get a nice notation we fix the functions b1,b2 and g in the
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    20
signature of this example
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    21
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    22
*)
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    23
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    24
Hoare = Tr2 +
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    25
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    26
consts
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    27
	b1:: "'a -> tr"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    28
	b2:: "'a -> tr"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    29
	 g:: "'a -> 'a"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    30
	p :: "'a -> 'a"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    31
	q :: "'a -> 'a"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    32
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    33
rules
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    34
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    35
  p_def  "p == fix[LAM f. LAM x.\
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    36
\                 If b1[x] then f[g[x]] else x fi]"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    37
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    38
  q_def  "q == fix[LAM f. LAM x.\
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    39
\                 If b1[x] orelse b2[x] then f[g[x]] else x fi]"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    40
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    41
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    42
end
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    43