src/HOL/MiniML/Maybe.ML
author paulson
Mon, 26 May 1997 12:36:16 +0200
changeset 3339 cfa72a70f2b5
parent 2625 69c1b8a493de
child 3919 c036caebfc75
permissions -rw-r--r--
Tidying and a couple of useful lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     1
(* Title:     HOL/MiniML/Maybe.ML
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     2
   ID:        $Id$
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     3
   Author:    Wolfgang Naraschewski and Tobias Nipkow
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     4
   Copyright  1996 TU Muenchen
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     5
*)
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     6
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     7
(* constructor laws for option_bind *)
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
     8
goalw thy [option_bind_def] "option_bind (Some s) f = (f s)";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
     9
by (Simp_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    10
qed "option_bind_Some";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    11
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    12
goalw thy [option_bind_def] "option_bind None f = None";
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    13
by (Simp_tac 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    14
qed "option_bind_None";
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    15
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    16
Addsimps [option_bind_Some,option_bind_None];
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    17
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    18
(* expansion of option_bind *)
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    19
goal thy
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    20
  "P(option_bind res f) = ((res = None --> P None) & (!s. res = Some s --> P(f s)))";
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    21
by (option.induct_tac "res" 1);
1300
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    22
by (fast_tac (HOL_cs addss !simpset) 1);
c7a8f374339b New theory: type inference for let-free MiniML
nipkow
parents:
diff changeset
    23
by (Asm_simp_tac 1);
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    24
qed "expand_option_bind";
1751
946efd210837 Added thm I_complete_wrt_W to I.
nipkow
parents: 1300
diff changeset
    25
2625
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    26
goal thy
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    27
  "((option_bind m f) = None) = ((m=None) | (? p. m = Some p & f p = None))";
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    28
by(simp_tac (!simpset setloop (split_tac [expand_option_bind])) 1);
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    29
qed "option_bind_eq_None";
1757
f7a573c46611 Added the second half of the W/I correspondence.
nipkow
parents: 1751
diff changeset
    30
2525
477c05586286 The new version of MiniML including "let".
nipkow
parents: 2058
diff changeset
    31
Addsimps [option_bind_eq_None];
2625
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    32
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    33
(* auxiliary lemma *)
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    34
goal Maybe.thy "(y = Some x) = (Some x = y)";
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    35
by( simp_tac (!simpset addsimps [eq_sym_conv]) 1);
69c1b8a493de Some lemmas changed to valuesd
narasche
parents: 2525
diff changeset
    36
qed "rotate_Some";