src/HOL/Option.ML
author paulson
Fri, 29 Nov 1996 18:03:21 +0100
changeset 2284 80ebd1a213fd
parent 2056 93c093620c28
child 2935 998cb95fdd43
permissions -rw-r--r--
Swapped arguments of Crypt (for clarity and because it is conventional)

(*  Title:      Option.ML
    ID:         $Id$
    Author:     Tobias Nipkow
    Copyright   1996  TU Muenchen

Derived rules
*)

open Option;

val [prem] = goal Option.thy "P(opt) ==> P(None) | (? x. P(Some(x)))";
 br (prem RS rev_mp) 1;
 by (option.induct_tac "opt" 1);
 by (ALLGOALS(Fast_tac));
bind_thm("optionE", standard(result() RS disjE));
(*
goal Option.thy "opt=None | (? x.opt=Some(x))"; 
by (option.induct_tac "opt" 1);
by (Simp_tac 1);
by (rtac disjI2 1);
by (rtac exI 1);
by (Asm_full_simp_tac 1);
qed"option_cases";
*)
goal Option.thy "P(case opt of None => a | Some(x) => b(x)) = \
\                ((opt = None --> P a) & (!x. opt = Some x --> P(b(x))))";
by (option.induct_tac "opt" 1);
by (Simp_tac 1);
by (Asm_full_simp_tac 1);
qed"expand_option_case";