src/HOL/HOL.thy
author kleing
Mon, 05 Nov 2007 22:50:00 +0100
changeset 25297 a5d689d04426
parent 25062 af5ef0d4d655
child 25388 5cd130251825
permissions -rw-r--r--
move itself into HOL types
     1 (*  Title:      HOL/HOL.thy
     2     ID:         $Id$
     3     Author:     Tobias Nipkow, Markus Wenzel, and Larry Paulson
     4 *)
     5 
     6 header {* The basis of Higher-Order Logic *}
     7 
     8 theory HOL
     9 imports CPure
    10 uses
    11   ("hologic.ML")
    12   "~~/src/Tools/IsaPlanner/zipper.ML"
    13   "~~/src/Tools/IsaPlanner/isand.ML"
    14   "~~/src/Tools/IsaPlanner/rw_tools.ML"
    15   "~~/src/Tools/IsaPlanner/rw_inst.ML"
    16   "~~/src/Provers/project_rule.ML"
    17   "~~/src/Provers/hypsubst.ML"
    18   "~~/src/Provers/splitter.ML"
    19   "~~/src/Provers/classical.ML"
    20   "~~/src/Provers/blast.ML"
    21   "~~/src/Provers/clasimp.ML"
    22   "~~/src/Provers/eqsubst.ML"
    23   "~~/src/Provers/quantifier1.ML"
    24   ("simpdata.ML")
    25   "~~/src/Tools/induct.ML"
    26   "~~/src/Tools/code/code_name.ML"
    27   "~~/src/Tools/code/code_funcgr.ML"
    28   "~~/src/Tools/code/code_thingol.ML"
    29   "~~/src/Tools/code/code_target.ML"
    30   "~~/src/Tools/code/code_package.ML"
    31   "~~/src/Tools/nbe.ML"
    32 begin
    33 
    34 subsection {* Primitive logic *}
    35 
    36 subsubsection {* Core syntax *}
    37 
    38 classes type
    39 defaultsort type
    40 
    41 global
    42 
    43 typedecl bool
    44 
    45 arities
    46   bool :: type
    47   "fun" :: (type, type) type
    48 
    49   itself :: (type) type
    50 
    51 judgment
    52   Trueprop      :: "bool => prop"                   ("(_)" 5)
    53 
    54 consts
    55   Not           :: "bool => bool"                   ("~ _" [40] 40)
    56   True          :: bool
    57   False         :: bool
    58   arbitrary     :: 'a
    59 
    60   The           :: "('a => bool) => 'a"
    61   All           :: "('a => bool) => bool"           (binder "ALL " 10)
    62   Ex            :: "('a => bool) => bool"           (binder "EX " 10)
    63   Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
    64   Let           :: "['a, 'a => 'b] => 'b"
    65 
    66   "op ="        :: "['a, 'a] => bool"               (infixl "=" 50)
    67   "op &"        :: "[bool, bool] => bool"           (infixr "&" 35)
    68   "op |"        :: "[bool, bool] => bool"           (infixr "|" 30)
    69   "op -->"      :: "[bool, bool] => bool"           (infixr "-->" 25)
    70 
    71 local
    72 
    73 consts
    74   If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
    75 
    76 
    77 subsubsection {* Additional concrete syntax *}
    78 
    79 notation (output)
    80   "op ="  (infix "=" 50)
    81 
    82 abbreviation
    83   not_equal :: "['a, 'a] => bool"  (infixl "~=" 50) where
    84   "x ~= y == ~ (x = y)"
    85 
    86 notation (output)
    87   not_equal  (infix "~=" 50)
    88 
    89 notation (xsymbols)
    90   Not  ("\<not> _" [40] 40) and
    91   "op &"  (infixr "\<and>" 35) and
    92   "op |"  (infixr "\<or>" 30) and
    93   "op -->"  (infixr "\<longrightarrow>" 25) and
    94   not_equal  (infix "\<noteq>" 50)
    95 
    96 notation (HTML output)
    97   Not  ("\<not> _" [40] 40) and
    98   "op &"  (infixr "\<and>" 35) and
    99   "op |"  (infixr "\<or>" 30) and
   100   not_equal  (infix "\<noteq>" 50)
   101 
   102 abbreviation (iff)
   103   iff :: "[bool, bool] => bool"  (infixr "<->" 25) where
   104   "A <-> B == A = B"
   105 
   106 notation (xsymbols)
   107   iff  (infixr "\<longleftrightarrow>" 25)
   108 
   109 
   110 nonterminals
   111   letbinds  letbind
   112   case_syn  cases_syn
   113 
   114 syntax
   115   "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
   116 
   117   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
   118   ""            :: "letbind => letbinds"                 ("_")
   119   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
   120   "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
   121 
   122   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
   123   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
   124   ""            :: "case_syn => cases_syn"               ("_")
   125   "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
   126 
   127 translations
   128   "THE x. P"              == "The (%x. P)"
   129   "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
   130   "let x = a in e"        == "Let a (%x. e)"
   131 
   132 print_translation {*
   133 (* To avoid eta-contraction of body: *)
   134 [("The", fn [Abs abs] =>
   135      let val (x,t) = atomic_abs_tr' abs
   136      in Syntax.const "_The" $ x $ t end)]
   137 *}
   138 
   139 syntax (xsymbols)
   140   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
   141 
   142 notation (xsymbols)
   143   All  (binder "\<forall>" 10) and
   144   Ex  (binder "\<exists>" 10) and
   145   Ex1  (binder "\<exists>!" 10)
   146 
   147 notation (HTML output)
   148   All  (binder "\<forall>" 10) and
   149   Ex  (binder "\<exists>" 10) and
   150   Ex1  (binder "\<exists>!" 10)
   151 
   152 notation (HOL)
   153   All  (binder "! " 10) and
   154   Ex  (binder "? " 10) and
   155   Ex1  (binder "?! " 10)
   156 
   157 
   158 subsubsection {* Axioms and basic definitions *}
   159 
   160 axioms
   161   eq_reflection:  "(x=y) ==> (x==y)"
   162 
   163   refl:           "t = (t::'a)"
   164 
   165   ext:            "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"
   166     -- {*Extensionality is built into the meta-logic, and this rule expresses
   167          a related property.  It is an eta-expanded version of the traditional
   168          rule, and similar to the ABS rule of HOL*}
   169 
   170   the_eq_trivial: "(THE x. x = a) = (a::'a)"
   171 
   172   impI:           "(P ==> Q) ==> P-->Q"
   173   mp:             "[| P-->Q;  P |] ==> Q"
   174 
   175 
   176 defs
   177   True_def:     "True      == ((%x::bool. x) = (%x. x))"
   178   All_def:      "All(P)    == (P = (%x. True))"
   179   Ex_def:       "Ex(P)     == !Q. (!x. P x --> Q) --> Q"
   180   False_def:    "False     == (!P. P)"
   181   not_def:      "~ P       == P-->False"
   182   and_def:      "P & Q     == !R. (P-->Q-->R) --> R"
   183   or_def:       "P | Q     == !R. (P-->R) --> (Q-->R) --> R"
   184   Ex1_def:      "Ex1(P)    == ? x. P(x) & (! y. P(y) --> y=x)"
   185 
   186 axioms
   187   iff:          "(P-->Q) --> (Q-->P) --> (P=Q)"
   188   True_or_False:  "(P=True) | (P=False)"
   189 
   190 defs
   191   Let_def:      "Let s f == f(s)"
   192   if_def:       "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)"
   193 
   194 finalconsts
   195   "op ="
   196   "op -->"
   197   The
   198   arbitrary
   199 
   200 axiomatization
   201   undefined :: 'a
   202 
   203 axiomatization where
   204   undefined_fun: "undefined x = undefined"
   205 
   206 
   207 subsubsection {* Generic classes and algebraic operations *}
   208 
   209 class default = type +
   210   fixes default :: 'a
   211 
   212 class zero = type + 
   213   fixes zero :: 'a  ("0")
   214 
   215 class one = type +
   216   fixes one  :: 'a  ("1")
   217 
   218 hide (open) const zero one
   219 
   220 class plus = type +
   221   fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "+" 65)
   222 
   223 class minus = type +
   224   fixes uminus :: "'a \<Rightarrow> 'a"  ("- _" [81] 80)
   225     and minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "-" 65)
   226 
   227 class times = type +
   228   fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "*" 70)
   229 
   230 class inverse = type +
   231   fixes inverse :: "'a \<Rightarrow> 'a"
   232     and divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "'/" 70)
   233 
   234 class abs = type +
   235   fixes abs :: "'a \<Rightarrow> 'a"
   236 
   237 notation (xsymbols)
   238   abs  ("\<bar>_\<bar>")
   239 notation (HTML output)
   240   abs  ("\<bar>_\<bar>")
   241 
   242 class sgn = type +
   243   fixes sgn :: "'a \<Rightarrow> 'a"
   244 
   245 class ord = type +
   246   fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
   247     and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
   248 begin
   249 
   250 abbreviation (input)
   251   greater_eq  (infix ">=" 50) where
   252   "x >= y \<equiv> less_eq y x"
   253 
   254 abbreviation (input)
   255   greater  (infix ">" 50) where
   256   "x > y \<equiv> less y x"
   257 
   258 definition
   259   Least :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" (binder "LEAST " 10)
   260 where
   261   "Least P == (THE x. P x \<and> (\<forall>y. P y \<longrightarrow> less_eq x y))"
   262 
   263 end
   264 
   265 notation
   266   less_eq  ("op <=") and
   267   less_eq  ("(_/ <= _)" [51, 51] 50) and
   268   less  ("op <") and
   269   less  ("(_/ < _)"  [51, 51] 50)
   270   
   271 notation (xsymbols)
   272   less_eq  ("op \<le>") and
   273   less_eq  ("(_/ \<le> _)"  [51, 51] 50)
   274 
   275 notation (HTML output)
   276   less_eq  ("op \<le>") and
   277   less_eq  ("(_/ \<le> _)"  [51, 51] 50)
   278 
   279 notation (input)
   280   greater_eq  (infix "\<ge>" 50)
   281 
   282 syntax
   283   "_index1"  :: index    ("\<^sub>1")
   284 translations
   285   (index) "\<^sub>1" => (index) "\<^bsub>\<struct>\<^esub>"
   286 
   287 typed_print_translation {*
   288 let
   289   fun tr' c = (c, fn show_sorts => fn T => fn ts =>
   290     if T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match
   291     else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T);
   292 in map tr' [@{const_syntax HOL.one}, @{const_syntax HOL.zero}] end;
   293 *} -- {* show types that are presumably too general *}
   294 
   295 
   296 subsection {* Fundamental rules *}
   297 
   298 subsubsection {* Equality *}
   299 
   300 text {* Thanks to Stephan Merz *}
   301 lemma subst:
   302   assumes eq: "s = t" and p: "P s"
   303   shows "P t"
   304 proof -
   305   from eq have meta: "s \<equiv> t"
   306     by (rule eq_reflection)
   307   from p show ?thesis
   308     by (unfold meta)
   309 qed
   310 
   311 lemma sym: "s = t ==> t = s"
   312   by (erule subst) (rule refl)
   313 
   314 lemma ssubst: "t = s ==> P s ==> P t"
   315   by (drule sym) (erule subst)
   316 
   317 lemma trans: "[| r=s; s=t |] ==> r=t"
   318   by (erule subst)
   319 
   320 lemma meta_eq_to_obj_eq: 
   321   assumes meq: "A == B"
   322   shows "A = B"
   323   by (unfold meq) (rule refl)
   324 
   325 text {* Useful with @{text erule} for proving equalities from known equalities. *}
   326      (* a = b
   327         |   |
   328         c = d   *)
   329 lemma box_equals: "[| a=b;  a=c;  b=d |] ==> c=d"
   330 apply (rule trans)
   331 apply (rule trans)
   332 apply (rule sym)
   333 apply assumption+
   334 done
   335 
   336 text {* For calculational reasoning: *}
   337 
   338 lemma forw_subst: "a = b ==> P b ==> P a"
   339   by (rule ssubst)
   340 
   341 lemma back_subst: "P a ==> a = b ==> P b"
   342   by (rule subst)
   343 
   344 
   345 subsubsection {*Congruence rules for application*}
   346 
   347 (*similar to AP_THM in Gordon's HOL*)
   348 lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)"
   349 apply (erule subst)
   350 apply (rule refl)
   351 done
   352 
   353 (*similar to AP_TERM in Gordon's HOL and FOL's subst_context*)
   354 lemma arg_cong: "x=y ==> f(x)=f(y)"
   355 apply (erule subst)
   356 apply (rule refl)
   357 done
   358 
   359 lemma arg_cong2: "\<lbrakk> a = b; c = d \<rbrakk> \<Longrightarrow> f a c = f b d"
   360 apply (erule ssubst)+
   361 apply (rule refl)
   362 done
   363 
   364 lemma cong: "[| f = g; (x::'a) = y |] ==> f(x) = g(y)"
   365 apply (erule subst)+
   366 apply (rule refl)
   367 done
   368 
   369 
   370 subsubsection {*Equality of booleans -- iff*}
   371 
   372 lemma iffI: assumes "P ==> Q" and "Q ==> P" shows "P=Q"
   373   by (iprover intro: iff [THEN mp, THEN mp] impI assms)
   374 
   375 lemma iffD2: "[| P=Q; Q |] ==> P"
   376   by (erule ssubst)
   377 
   378 lemma rev_iffD2: "[| Q; P=Q |] ==> P"
   379   by (erule iffD2)
   380 
   381 lemma iffD1: "Q = P \<Longrightarrow> Q \<Longrightarrow> P"
   382   by (drule sym) (rule iffD2)
   383 
   384 lemma rev_iffD1: "Q \<Longrightarrow> Q = P \<Longrightarrow> P"
   385   by (drule sym) (rule rev_iffD2)
   386 
   387 lemma iffE:
   388   assumes major: "P=Q"
   389     and minor: "[| P --> Q; Q --> P |] ==> R"
   390   shows R
   391   by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1])
   392 
   393 
   394 subsubsection {*True*}
   395 
   396 lemma TrueI: "True"
   397   unfolding True_def by (rule refl)
   398 
   399 lemma eqTrueI: "P ==> P = True"
   400   by (iprover intro: iffI TrueI)
   401 
   402 lemma eqTrueE: "P = True ==> P"
   403   by (erule iffD2) (rule TrueI)
   404 
   405 
   406 subsubsection {*Universal quantifier*}
   407 
   408 lemma allI: assumes "!!x::'a. P(x)" shows "ALL x. P(x)"
   409   unfolding All_def by (iprover intro: ext eqTrueI assms)
   410 
   411 lemma spec: "ALL x::'a. P(x) ==> P(x)"
   412 apply (unfold All_def)
   413 apply (rule eqTrueE)
   414 apply (erule fun_cong)
   415 done
   416 
   417 lemma allE:
   418   assumes major: "ALL x. P(x)"
   419     and minor: "P(x) ==> R"
   420   shows R
   421   by (iprover intro: minor major [THEN spec])
   422 
   423 lemma all_dupE:
   424   assumes major: "ALL x. P(x)"
   425     and minor: "[| P(x); ALL x. P(x) |] ==> R"
   426   shows R
   427   by (iprover intro: minor major major [THEN spec])
   428 
   429 
   430 subsubsection {* False *}
   431 
   432 text {*
   433   Depends upon @{text spec}; it is impossible to do propositional
   434   logic before quantifiers!
   435 *}
   436 
   437 lemma FalseE: "False ==> P"
   438   apply (unfold False_def)
   439   apply (erule spec)
   440   done
   441 
   442 lemma False_neq_True: "False = True ==> P"
   443   by (erule eqTrueE [THEN FalseE])
   444 
   445 
   446 subsubsection {* Negation *}
   447 
   448 lemma notI:
   449   assumes "P ==> False"
   450   shows "~P"
   451   apply (unfold not_def)
   452   apply (iprover intro: impI assms)
   453   done
   454 
   455 lemma False_not_True: "False ~= True"
   456   apply (rule notI)
   457   apply (erule False_neq_True)
   458   done
   459 
   460 lemma True_not_False: "True ~= False"
   461   apply (rule notI)
   462   apply (drule sym)
   463   apply (erule False_neq_True)
   464   done
   465 
   466 lemma notE: "[| ~P;  P |] ==> R"
   467   apply (unfold not_def)
   468   apply (erule mp [THEN FalseE])
   469   apply assumption
   470   done
   471 
   472 lemma notI2: "(P \<Longrightarrow> \<not> Pa) \<Longrightarrow> (P \<Longrightarrow> Pa) \<Longrightarrow> \<not> P"
   473   by (erule notE [THEN notI]) (erule meta_mp)
   474 
   475 
   476 subsubsection {*Implication*}
   477 
   478 lemma impE:
   479   assumes "P-->Q" "P" "Q ==> R"
   480   shows "R"
   481 by (iprover intro: assms mp)
   482 
   483 (* Reduces Q to P-->Q, allowing substitution in P. *)
   484 lemma rev_mp: "[| P;  P --> Q |] ==> Q"
   485 by (iprover intro: mp)
   486 
   487 lemma contrapos_nn:
   488   assumes major: "~Q"
   489       and minor: "P==>Q"
   490   shows "~P"
   491 by (iprover intro: notI minor major [THEN notE])
   492 
   493 (*not used at all, but we already have the other 3 combinations *)
   494 lemma contrapos_pn:
   495   assumes major: "Q"
   496       and minor: "P ==> ~Q"
   497   shows "~P"
   498 by (iprover intro: notI minor major notE)
   499 
   500 lemma not_sym: "t ~= s ==> s ~= t"
   501   by (erule contrapos_nn) (erule sym)
   502 
   503 lemma eq_neq_eq_imp_neq: "[| x = a ; a ~= b; b = y |] ==> x ~= y"
   504   by (erule subst, erule ssubst, assumption)
   505 
   506 (*still used in HOLCF*)
   507 lemma rev_contrapos:
   508   assumes pq: "P ==> Q"
   509       and nq: "~Q"
   510   shows "~P"
   511 apply (rule nq [THEN contrapos_nn])
   512 apply (erule pq)
   513 done
   514 
   515 subsubsection {*Existential quantifier*}
   516 
   517 lemma exI: "P x ==> EX x::'a. P x"
   518 apply (unfold Ex_def)
   519 apply (iprover intro: allI allE impI mp)
   520 done
   521 
   522 lemma exE:
   523   assumes major: "EX x::'a. P(x)"
   524       and minor: "!!x. P(x) ==> Q"
   525   shows "Q"
   526 apply (rule major [unfolded Ex_def, THEN spec, THEN mp])
   527 apply (iprover intro: impI [THEN allI] minor)
   528 done
   529 
   530 
   531 subsubsection {*Conjunction*}
   532 
   533 lemma conjI: "[| P; Q |] ==> P&Q"
   534 apply (unfold and_def)
   535 apply (iprover intro: impI [THEN allI] mp)
   536 done
   537 
   538 lemma conjunct1: "[| P & Q |] ==> P"
   539 apply (unfold and_def)
   540 apply (iprover intro: impI dest: spec mp)
   541 done
   542 
   543 lemma conjunct2: "[| P & Q |] ==> Q"
   544 apply (unfold and_def)
   545 apply (iprover intro: impI dest: spec mp)
   546 done
   547 
   548 lemma conjE:
   549   assumes major: "P&Q"
   550       and minor: "[| P; Q |] ==> R"
   551   shows "R"
   552 apply (rule minor)
   553 apply (rule major [THEN conjunct1])
   554 apply (rule major [THEN conjunct2])
   555 done
   556 
   557 lemma context_conjI:
   558   assumes "P" "P ==> Q" shows "P & Q"
   559 by (iprover intro: conjI assms)
   560 
   561 
   562 subsubsection {*Disjunction*}
   563 
   564 lemma disjI1: "P ==> P|Q"
   565 apply (unfold or_def)
   566 apply (iprover intro: allI impI mp)
   567 done
   568 
   569 lemma disjI2: "Q ==> P|Q"
   570 apply (unfold or_def)
   571 apply (iprover intro: allI impI mp)
   572 done
   573 
   574 lemma disjE:
   575   assumes major: "P|Q"
   576       and minorP: "P ==> R"
   577       and minorQ: "Q ==> R"
   578   shows "R"
   579 by (iprover intro: minorP minorQ impI
   580                  major [unfolded or_def, THEN spec, THEN mp, THEN mp])
   581 
   582 
   583 subsubsection {*Classical logic*}
   584 
   585 lemma classical:
   586   assumes prem: "~P ==> P"
   587   shows "P"
   588 apply (rule True_or_False [THEN disjE, THEN eqTrueE])
   589 apply assumption
   590 apply (rule notI [THEN prem, THEN eqTrueI])
   591 apply (erule subst)
   592 apply assumption
   593 done
   594 
   595 lemmas ccontr = FalseE [THEN classical, standard]
   596 
   597 (*notE with premises exchanged; it discharges ~R so that it can be used to
   598   make elimination rules*)
   599 lemma rev_notE:
   600   assumes premp: "P"
   601       and premnot: "~R ==> ~P"
   602   shows "R"
   603 apply (rule ccontr)
   604 apply (erule notE [OF premnot premp])
   605 done
   606 
   607 (*Double negation law*)
   608 lemma notnotD: "~~P ==> P"
   609 apply (rule classical)
   610 apply (erule notE)
   611 apply assumption
   612 done
   613 
   614 lemma contrapos_pp:
   615   assumes p1: "Q"
   616       and p2: "~P ==> ~Q"
   617   shows "P"
   618 by (iprover intro: classical p1 p2 notE)
   619 
   620 
   621 subsubsection {*Unique existence*}
   622 
   623 lemma ex1I:
   624   assumes "P a" "!!x. P(x) ==> x=a"
   625   shows "EX! x. P(x)"
   626 by (unfold Ex1_def, iprover intro: assms exI conjI allI impI)
   627 
   628 text{*Sometimes easier to use: the premises have no shared variables.  Safe!*}
   629 lemma ex_ex1I:
   630   assumes ex_prem: "EX x. P(x)"
   631       and eq: "!!x y. [| P(x); P(y) |] ==> x=y"
   632   shows "EX! x. P(x)"
   633 by (iprover intro: ex_prem [THEN exE] ex1I eq)
   634 
   635 lemma ex1E:
   636   assumes major: "EX! x. P(x)"
   637       and minor: "!!x. [| P(x);  ALL y. P(y) --> y=x |] ==> R"
   638   shows "R"
   639 apply (rule major [unfolded Ex1_def, THEN exE])
   640 apply (erule conjE)
   641 apply (iprover intro: minor)
   642 done
   643 
   644 lemma ex1_implies_ex: "EX! x. P x ==> EX x. P x"
   645 apply (erule ex1E)
   646 apply (rule exI)
   647 apply assumption
   648 done
   649 
   650 
   651 subsubsection {*THE: definite description operator*}
   652 
   653 lemma the_equality:
   654   assumes prema: "P a"
   655       and premx: "!!x. P x ==> x=a"
   656   shows "(THE x. P x) = a"
   657 apply (rule trans [OF _ the_eq_trivial])
   658 apply (rule_tac f = "The" in arg_cong)
   659 apply (rule ext)
   660 apply (rule iffI)
   661  apply (erule premx)
   662 apply (erule ssubst, rule prema)
   663 done
   664 
   665 lemma theI:
   666   assumes "P a" and "!!x. P x ==> x=a"
   667   shows "P (THE x. P x)"
   668 by (iprover intro: assms the_equality [THEN ssubst])
   669 
   670 lemma theI': "EX! x. P x ==> P (THE x. P x)"
   671 apply (erule ex1E)
   672 apply (erule theI)
   673 apply (erule allE)
   674 apply (erule mp)
   675 apply assumption
   676 done
   677 
   678 (*Easier to apply than theI: only one occurrence of P*)
   679 lemma theI2:
   680   assumes "P a" "!!x. P x ==> x=a" "!!x. P x ==> Q x"
   681   shows "Q (THE x. P x)"
   682 by (iprover intro: assms theI)
   683 
   684 lemma the1I2: assumes "EX! x. P x" "\<And>x. P x \<Longrightarrow> Q x" shows "Q (THE x. P x)"
   685 by(iprover intro:assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)]
   686            elim:allE impE)
   687 
   688 lemma the1_equality [elim?]: "[| EX!x. P x; P a |] ==> (THE x. P x) = a"
   689 apply (rule the_equality)
   690 apply  assumption
   691 apply (erule ex1E)
   692 apply (erule all_dupE)
   693 apply (drule mp)
   694 apply  assumption
   695 apply (erule ssubst)
   696 apply (erule allE)
   697 apply (erule mp)
   698 apply assumption
   699 done
   700 
   701 lemma the_sym_eq_trivial: "(THE y. x=y) = x"
   702 apply (rule the_equality)
   703 apply (rule refl)
   704 apply (erule sym)
   705 done
   706 
   707 
   708 subsubsection {*Classical intro rules for disjunction and existential quantifiers*}
   709 
   710 lemma disjCI:
   711   assumes "~Q ==> P" shows "P|Q"
   712 apply (rule classical)
   713 apply (iprover intro: assms disjI1 disjI2 notI elim: notE)
   714 done
   715 
   716 lemma excluded_middle: "~P | P"
   717 by (iprover intro: disjCI)
   718 
   719 text {*
   720   case distinction as a natural deduction rule.
   721   Note that @{term "~P"} is the second case, not the first
   722 *}
   723 lemma case_split_thm:
   724   assumes prem1: "P ==> Q"
   725       and prem2: "~P ==> Q"
   726   shows "Q"
   727 apply (rule excluded_middle [THEN disjE])
   728 apply (erule prem2)
   729 apply (erule prem1)
   730 done
   731 lemmas case_split = case_split_thm [case_names True False]
   732 
   733 (*Classical implies (-->) elimination. *)
   734 lemma impCE:
   735   assumes major: "P-->Q"
   736       and minor: "~P ==> R" "Q ==> R"
   737   shows "R"
   738 apply (rule excluded_middle [of P, THEN disjE])
   739 apply (iprover intro: minor major [THEN mp])+
   740 done
   741 
   742 (*This version of --> elimination works on Q before P.  It works best for
   743   those cases in which P holds "almost everywhere".  Can't install as
   744   default: would break old proofs.*)
   745 lemma impCE':
   746   assumes major: "P-->Q"
   747       and minor: "Q ==> R" "~P ==> R"
   748   shows "R"
   749 apply (rule excluded_middle [of P, THEN disjE])
   750 apply (iprover intro: minor major [THEN mp])+
   751 done
   752 
   753 (*Classical <-> elimination. *)
   754 lemma iffCE:
   755   assumes major: "P=Q"
   756       and minor: "[| P; Q |] ==> R"  "[| ~P; ~Q |] ==> R"
   757   shows "R"
   758 apply (rule major [THEN iffE])
   759 apply (iprover intro: minor elim: impCE notE)
   760 done
   761 
   762 lemma exCI:
   763   assumes "ALL x. ~P(x) ==> P(a)"
   764   shows "EX x. P(x)"
   765 apply (rule ccontr)
   766 apply (iprover intro: assms exI allI notI notE [of "\<exists>x. P x"])
   767 done
   768 
   769 
   770 subsubsection {* Intuitionistic Reasoning *}
   771 
   772 lemma impE':
   773   assumes 1: "P --> Q"
   774     and 2: "Q ==> R"
   775     and 3: "P --> Q ==> P"
   776   shows R
   777 proof -
   778   from 3 and 1 have P .
   779   with 1 have Q by (rule impE)
   780   with 2 show R .
   781 qed
   782 
   783 lemma allE':
   784   assumes 1: "ALL x. P x"
   785     and 2: "P x ==> ALL x. P x ==> Q"
   786   shows Q
   787 proof -
   788   from 1 have "P x" by (rule spec)
   789   from this and 1 show Q by (rule 2)
   790 qed
   791 
   792 lemma notE':
   793   assumes 1: "~ P"
   794     and 2: "~ P ==> P"
   795   shows R
   796 proof -
   797   from 2 and 1 have P .
   798   with 1 show R by (rule notE)
   799 qed
   800 
   801 lemma TrueE: "True ==> P ==> P" .
   802 lemma notFalseE: "~ False ==> P ==> P" .
   803 
   804 lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE
   805   and [Pure.intro!] = iffI conjI impI TrueI notI allI refl
   806   and [Pure.elim 2] = allE notE' impE'
   807   and [Pure.intro] = exI disjI2 disjI1
   808 
   809 lemmas [trans] = trans
   810   and [sym] = sym not_sym
   811   and [Pure.elim?] = iffD1 iffD2 impE
   812 
   813 use "hologic.ML"
   814 
   815 
   816 subsubsection {* Atomizing meta-level connectives *}
   817 
   818 lemma atomize_all [atomize]: "(!!x. P x) == Trueprop (ALL x. P x)"
   819 proof
   820   assume "!!x. P x"
   821   then show "ALL x. P x" ..
   822 next
   823   assume "ALL x. P x"
   824   then show "!!x. P x" by (rule allE)
   825 qed
   826 
   827 lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)"
   828 proof
   829   assume r: "A ==> B"
   830   show "A --> B" by (rule impI) (rule r)
   831 next
   832   assume "A --> B" and A
   833   then show B by (rule mp)
   834 qed
   835 
   836 lemma atomize_not: "(A ==> False) == Trueprop (~A)"
   837 proof
   838   assume r: "A ==> False"
   839   show "~A" by (rule notI) (rule r)
   840 next
   841   assume "~A" and A
   842   then show False by (rule notE)
   843 qed
   844 
   845 lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)"
   846 proof
   847   assume "x == y"
   848   show "x = y" by (unfold `x == y`) (rule refl)
   849 next
   850   assume "x = y"
   851   then show "x == y" by (rule eq_reflection)
   852 qed
   853 
   854 lemma atomize_conj [atomize]:
   855   includes meta_conjunction_syntax
   856   shows "(A && B) == Trueprop (A & B)"
   857 proof
   858   assume conj: "A && B"
   859   show "A & B"
   860   proof (rule conjI)
   861     from conj show A by (rule conjunctionD1)
   862     from conj show B by (rule conjunctionD2)
   863   qed
   864 next
   865   assume conj: "A & B"
   866   show "A && B"
   867   proof -
   868     from conj show A ..
   869     from conj show B ..
   870   qed
   871 qed
   872 
   873 lemmas [symmetric, rulify] = atomize_all atomize_imp
   874   and [symmetric, defn] = atomize_all atomize_imp atomize_eq
   875 
   876 
   877 subsection {* Package setup *}
   878 
   879 subsubsection {* Classical Reasoner setup *}
   880 
   881 lemma thin_refl:
   882   "\<And>X. \<lbrakk> x=x; PROP W \<rbrakk> \<Longrightarrow> PROP W" .
   883 
   884 ML {*
   885 structure Hypsubst = HypsubstFun(
   886 struct
   887   structure Simplifier = Simplifier
   888   val dest_eq = HOLogic.dest_eq
   889   val dest_Trueprop = HOLogic.dest_Trueprop
   890   val dest_imp = HOLogic.dest_imp
   891   val eq_reflection = @{thm HOL.eq_reflection}
   892   val rev_eq_reflection = @{thm HOL.meta_eq_to_obj_eq}
   893   val imp_intr = @{thm HOL.impI}
   894   val rev_mp = @{thm HOL.rev_mp}
   895   val subst = @{thm HOL.subst}
   896   val sym = @{thm HOL.sym}
   897   val thin_refl = @{thm thin_refl};
   898 end);
   899 open Hypsubst;
   900 
   901 structure Classical = ClassicalFun(
   902 struct
   903   val mp = @{thm HOL.mp}
   904   val not_elim = @{thm HOL.notE}
   905   val classical = @{thm HOL.classical}
   906   val sizef = Drule.size_of_thm
   907   val hyp_subst_tacs = [Hypsubst.hyp_subst_tac]
   908 end);
   909 
   910 structure BasicClassical: BASIC_CLASSICAL = Classical; 
   911 open BasicClassical;
   912 
   913 ML_Context.value_antiq "claset"
   914   (Scan.succeed ("claset", "Classical.local_claset_of (ML_Context.the_local_context ())"));
   915 
   916 structure ResAtpset = NamedThmsFun(val name = "atp" val description = "ATP rules");
   917 
   918 structure ResBlacklist = NamedThmsFun(val name = "noatp" val description = "Theorems blacklisted for ATP");
   919 *}
   920 
   921 (*ResBlacklist holds theorems blacklisted to sledgehammer. 
   922   These theorems typically produce clauses that are prolific (match too many equality or
   923   membership literals) and relate to seldom-used facts. Some duplicate other rules.*)
   924 
   925 setup {*
   926 let
   927   (*prevent substitution on bool*)
   928   fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso
   929     Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("bool", []) | _ => false)
   930       (nth (Thm.prems_of thm) (i - 1)) then Hypsubst.hyp_subst_tac i thm else no_tac thm;
   931 in
   932   Hypsubst.hypsubst_setup
   933   #> ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)
   934   #> Classical.setup
   935   #> ResAtpset.setup
   936   #> ResBlacklist.setup
   937 end
   938 *}
   939 
   940 declare iffI [intro!]
   941   and notI [intro!]
   942   and impI [intro!]
   943   and disjCI [intro!]
   944   and conjI [intro!]
   945   and TrueI [intro!]
   946   and refl [intro!]
   947 
   948 declare iffCE [elim!]
   949   and FalseE [elim!]
   950   and impCE [elim!]
   951   and disjE [elim!]
   952   and conjE [elim!]
   953   and conjE [elim!]
   954 
   955 declare ex_ex1I [intro!]
   956   and allI [intro!]
   957   and the_equality [intro]
   958   and exI [intro]
   959 
   960 declare exE [elim!]
   961   allE [elim]
   962 
   963 ML {* val HOL_cs = @{claset} *}
   964 
   965 lemma contrapos_np: "~ Q ==> (~ P ==> Q) ==> P"
   966   apply (erule swap)
   967   apply (erule (1) meta_mp)
   968   done
   969 
   970 declare ex_ex1I [rule del, intro! 2]
   971   and ex1I [intro]
   972 
   973 lemmas [intro?] = ext
   974   and [elim?] = ex1_implies_ex
   975 
   976 (*Better then ex1E for classical reasoner: needs no quantifier duplication!*)
   977 lemma alt_ex1E [elim!]:
   978   assumes major: "\<exists>!x. P x"
   979       and prem: "\<And>x. \<lbrakk> P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y' \<rbrakk> \<Longrightarrow> R"
   980   shows R
   981 apply (rule ex1E [OF major])
   982 apply (rule prem)
   983 apply (tactic {* ares_tac @{thms allI} 1 *})+
   984 apply (tactic {* etac (Classical.dup_elim @{thm allE}) 1 *})
   985 apply iprover
   986 done
   987 
   988 ML {*
   989 structure Blast = BlastFun(
   990 struct
   991   type claset = Classical.claset
   992   val equality_name = @{const_name "op ="}
   993   val not_name = @{const_name Not}
   994   val notE = @{thm HOL.notE}
   995   val ccontr = @{thm HOL.ccontr}
   996   val contr_tac = Classical.contr_tac
   997   val dup_intr = Classical.dup_intr
   998   val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
   999   val claset = Classical.claset
  1000   val rep_cs = Classical.rep_cs
  1001   val cla_modifiers = Classical.cla_modifiers
  1002   val cla_meth' = Classical.cla_meth'
  1003 end);
  1004 val Blast_tac = Blast.Blast_tac;
  1005 val blast_tac = Blast.blast_tac;
  1006 *}
  1007 
  1008 setup Blast.setup
  1009 
  1010 
  1011 subsubsection {* Simplifier *}
  1012 
  1013 lemma eta_contract_eq: "(%s. f s) = f" ..
  1014 
  1015 lemma simp_thms:
  1016   shows not_not: "(~ ~ P) = P"
  1017   and Not_eq_iff: "((~P) = (~Q)) = (P = Q)"
  1018   and
  1019     "(P ~= Q) = (P = (~Q))"
  1020     "(P | ~P) = True"    "(~P | P) = True"
  1021     "(x = x) = True"
  1022   and not_True_eq_False: "(\<not> True) = False"
  1023   and not_False_eq_True: "(\<not> False) = True"
  1024   and
  1025     "(~P) ~= P"  "P ~= (~P)"
  1026     "(True=P) = P"
  1027   and eq_True: "(P = True) = P"
  1028   and "(False=P) = (~P)"
  1029   and eq_False: "(P = False) = (\<not> P)"
  1030   and
  1031     "(True --> P) = P"  "(False --> P) = True"
  1032     "(P --> True) = True"  "(P --> P) = True"
  1033     "(P --> False) = (~P)"  "(P --> ~P) = (~P)"
  1034     "(P & True) = P"  "(True & P) = P"
  1035     "(P & False) = False"  "(False & P) = False"
  1036     "(P & P) = P"  "(P & (P & Q)) = (P & Q)"
  1037     "(P & ~P) = False"    "(~P & P) = False"
  1038     "(P | True) = True"  "(True | P) = True"
  1039     "(P | False) = P"  "(False | P) = P"
  1040     "(P | P) = P"  "(P | (P | Q)) = (P | Q)" and
  1041     "(ALL x. P) = P"  "(EX x. P) = P"  "EX x. x=t"  "EX x. t=x"
  1042     -- {* needed for the one-point-rule quantifier simplification procs *}
  1043     -- {* essential for termination!! *} and
  1044     "!!P. (EX x. x=t & P(x)) = P(t)"
  1045     "!!P. (EX x. t=x & P(x)) = P(t)"
  1046     "!!P. (ALL x. x=t --> P(x)) = P(t)"
  1047     "!!P. (ALL x. t=x --> P(x)) = P(t)"
  1048   by (blast, blast, blast, blast, blast, iprover+)
  1049 
  1050 lemma disj_absorb: "(A | A) = A"
  1051   by blast
  1052 
  1053 lemma disj_left_absorb: "(A | (A | B)) = (A | B)"
  1054   by blast
  1055 
  1056 lemma conj_absorb: "(A & A) = A"
  1057   by blast
  1058 
  1059 lemma conj_left_absorb: "(A & (A & B)) = (A & B)"
  1060   by blast
  1061 
  1062 lemma eq_ac:
  1063   shows eq_commute: "(a=b) = (b=a)"
  1064     and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))"
  1065     and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (iprover, blast+)
  1066 lemma neq_commute: "(a~=b) = (b~=a)" by iprover
  1067 
  1068 lemma conj_comms:
  1069   shows conj_commute: "(P&Q) = (Q&P)"
  1070     and conj_left_commute: "(P&(Q&R)) = (Q&(P&R))" by iprover+
  1071 lemma conj_assoc: "((P&Q)&R) = (P&(Q&R))" by iprover
  1072 
  1073 lemmas conj_ac = conj_commute conj_left_commute conj_assoc
  1074 
  1075 lemma disj_comms:
  1076   shows disj_commute: "(P|Q) = (Q|P)"
  1077     and disj_left_commute: "(P|(Q|R)) = (Q|(P|R))" by iprover+
  1078 lemma disj_assoc: "((P|Q)|R) = (P|(Q|R))" by iprover
  1079 
  1080 lemmas disj_ac = disj_commute disj_left_commute disj_assoc
  1081 
  1082 lemma conj_disj_distribL: "(P&(Q|R)) = (P&Q | P&R)" by iprover
  1083 lemma conj_disj_distribR: "((P|Q)&R) = (P&R | Q&R)" by iprover
  1084 
  1085 lemma disj_conj_distribL: "(P|(Q&R)) = ((P|Q) & (P|R))" by iprover
  1086 lemma disj_conj_distribR: "((P&Q)|R) = ((P|R) & (Q|R))" by iprover
  1087 
  1088 lemma imp_conjR: "(P --> (Q&R)) = ((P-->Q) & (P-->R))" by iprover
  1089 lemma imp_conjL: "((P&Q) -->R)  = (P --> (Q --> R))" by iprover
  1090 lemma imp_disjL: "((P|Q) --> R) = ((P-->R)&(Q-->R))" by iprover
  1091 
  1092 text {* These two are specialized, but @{text imp_disj_not1} is useful in @{text "Auth/Yahalom"}. *}
  1093 lemma imp_disj_not1: "(P --> Q | R) = (~Q --> P --> R)" by blast
  1094 lemma imp_disj_not2: "(P --> Q | R) = (~R --> P --> Q)" by blast
  1095 
  1096 lemma imp_disj1: "((P-->Q)|R) = (P--> Q|R)" by blast
  1097 lemma imp_disj2: "(Q|(P-->R)) = (P--> Q|R)" by blast
  1098 
  1099 lemma imp_cong: "(P = P') ==> (P' ==> (Q = Q')) ==> ((P --> Q) = (P' --> Q'))"
  1100   by iprover
  1101 
  1102 lemma de_Morgan_disj: "(~(P | Q)) = (~P & ~Q)" by iprover
  1103 lemma de_Morgan_conj: "(~(P & Q)) = (~P | ~Q)" by blast
  1104 lemma not_imp: "(~(P --> Q)) = (P & ~Q)" by blast
  1105 lemma not_iff: "(P~=Q) = (P = (~Q))" by blast
  1106 lemma disj_not1: "(~P | Q) = (P --> Q)" by blast
  1107 lemma disj_not2: "(P | ~Q) = (Q --> P)"  -- {* changes orientation :-( *}
  1108   by blast
  1109 lemma imp_conv_disj: "(P --> Q) = ((~P) | Q)" by blast
  1110 
  1111 lemma iff_conv_conj_imp: "(P = Q) = ((P --> Q) & (Q --> P))" by iprover
  1112 
  1113 
  1114 lemma cases_simp: "((P --> Q) & (~P --> Q)) = Q"
  1115   -- {* Avoids duplication of subgoals after @{text split_if}, when the true and false *}
  1116   -- {* cases boil down to the same thing. *}
  1117   by blast
  1118 
  1119 lemma not_all: "(~ (! x. P(x))) = (? x.~P(x))" by blast
  1120 lemma imp_all: "((! x. P x) --> Q) = (? x. P x --> Q)" by blast
  1121 lemma not_ex: "(~ (? x. P(x))) = (! x.~P(x))" by iprover
  1122 lemma imp_ex: "((? x. P x) --> Q) = (! x. P x --> Q)" by iprover
  1123 lemma all_not_ex: "(ALL x. P x) = (~ (EX x. ~ P x ))" by blast
  1124 
  1125 declare All_def [noatp]
  1126 
  1127 lemma ex_disj_distrib: "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))" by iprover
  1128 lemma all_conj_distrib: "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))" by iprover
  1129 
  1130 text {*
  1131   \medskip The @{text "&"} congruence rule: not included by default!
  1132   May slow rewrite proofs down by as much as 50\% *}
  1133 
  1134 lemma conj_cong:
  1135     "(P = P') ==> (P' ==> (Q = Q')) ==> ((P & Q) = (P' & Q'))"
  1136   by iprover
  1137 
  1138 lemma rev_conj_cong:
  1139     "(Q = Q') ==> (Q' ==> (P = P')) ==> ((P & Q) = (P' & Q'))"
  1140   by iprover
  1141 
  1142 text {* The @{text "|"} congruence rule: not included by default! *}
  1143 
  1144 lemma disj_cong:
  1145     "(P = P') ==> (~P' ==> (Q = Q')) ==> ((P | Q) = (P' | Q'))"
  1146   by blast
  1147 
  1148 
  1149 text {* \medskip if-then-else rules *}
  1150 
  1151 lemma if_True: "(if True then x else y) = x"
  1152   by (unfold if_def) blast
  1153 
  1154 lemma if_False: "(if False then x else y) = y"
  1155   by (unfold if_def) blast
  1156 
  1157 lemma if_P: "P ==> (if P then x else y) = x"
  1158   by (unfold if_def) blast
  1159 
  1160 lemma if_not_P: "~P ==> (if P then x else y) = y"
  1161   by (unfold if_def) blast
  1162 
  1163 lemma split_if: "P (if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))"
  1164   apply (rule case_split [of Q])
  1165    apply (simplesubst if_P)
  1166     prefer 3 apply (simplesubst if_not_P, blast+)
  1167   done
  1168 
  1169 lemma split_if_asm: "P (if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))"
  1170 by (simplesubst split_if, blast)
  1171 
  1172 lemmas if_splits [noatp] = split_if split_if_asm
  1173 
  1174 lemma if_cancel: "(if c then x else x) = x"
  1175 by (simplesubst split_if, blast)
  1176 
  1177 lemma if_eq_cancel: "(if x = y then y else x) = x"
  1178 by (simplesubst split_if, blast)
  1179 
  1180 lemma if_bool_eq_conj: "(if P then Q else R) = ((P-->Q) & (~P-->R))"
  1181   -- {* This form is useful for expanding @{text "if"}s on the RIGHT of the @{text "==>"} symbol. *}
  1182   by (rule split_if)
  1183 
  1184 lemma if_bool_eq_disj: "(if P then Q else R) = ((P&Q) | (~P&R))"
  1185   -- {* And this form is useful for expanding @{text "if"}s on the LEFT. *}
  1186   apply (simplesubst split_if, blast)
  1187   done
  1188 
  1189 lemma Eq_TrueI: "P ==> P == True" by (unfold atomize_eq) iprover
  1190 lemma Eq_FalseI: "~P ==> P == False" by (unfold atomize_eq) iprover
  1191 
  1192 text {* \medskip let rules for simproc *}
  1193 
  1194 lemma Let_folded: "f x \<equiv> g x \<Longrightarrow>  Let x f \<equiv> Let x g"
  1195   by (unfold Let_def)
  1196 
  1197 lemma Let_unfold: "f x \<equiv> g \<Longrightarrow>  Let x f \<equiv> g"
  1198   by (unfold Let_def)
  1199 
  1200 text {*
  1201   The following copy of the implication operator is useful for
  1202   fine-tuning congruence rules.  It instructs the simplifier to simplify
  1203   its premise.
  1204 *}
  1205 
  1206 constdefs
  1207   simp_implies :: "[prop, prop] => prop"  (infixr "=simp=>" 1)
  1208   "simp_implies \<equiv> op ==>"
  1209 
  1210 lemma simp_impliesI:
  1211   assumes PQ: "(PROP P \<Longrightarrow> PROP Q)"
  1212   shows "PROP P =simp=> PROP Q"
  1213   apply (unfold simp_implies_def)
  1214   apply (rule PQ)
  1215   apply assumption
  1216   done
  1217 
  1218 lemma simp_impliesE:
  1219   assumes PQ:"PROP P =simp=> PROP Q"
  1220   and P: "PROP P"
  1221   and QR: "PROP Q \<Longrightarrow> PROP R"
  1222   shows "PROP R"
  1223   apply (rule QR)
  1224   apply (rule PQ [unfolded simp_implies_def])
  1225   apply (rule P)
  1226   done
  1227 
  1228 lemma simp_implies_cong:
  1229   assumes PP' :"PROP P == PROP P'"
  1230   and P'QQ': "PROP P' ==> (PROP Q == PROP Q')"
  1231   shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')"
  1232 proof (unfold simp_implies_def, rule equal_intr_rule)
  1233   assume PQ: "PROP P \<Longrightarrow> PROP Q"
  1234   and P': "PROP P'"
  1235   from PP' [symmetric] and P' have "PROP P"
  1236     by (rule equal_elim_rule1)
  1237   then have "PROP Q" by (rule PQ)
  1238   with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1)
  1239 next
  1240   assume P'Q': "PROP P' \<Longrightarrow> PROP Q'"
  1241   and P: "PROP P"
  1242   from PP' and P have P': "PROP P'" by (rule equal_elim_rule1)
  1243   then have "PROP Q'" by (rule P'Q')
  1244   with P'QQ' [OF P', symmetric] show "PROP Q"
  1245     by (rule equal_elim_rule1)
  1246 qed
  1247 
  1248 lemma uncurry:
  1249   assumes "P \<longrightarrow> Q \<longrightarrow> R"
  1250   shows "P \<and> Q \<longrightarrow> R"
  1251   using assms by blast
  1252 
  1253 lemma iff_allI:
  1254   assumes "\<And>x. P x = Q x"
  1255   shows "(\<forall>x. P x) = (\<forall>x. Q x)"
  1256   using assms by blast
  1257 
  1258 lemma iff_exI:
  1259   assumes "\<And>x. P x = Q x"
  1260   shows "(\<exists>x. P x) = (\<exists>x. Q x)"
  1261   using assms by blast
  1262 
  1263 lemma all_comm:
  1264   "(\<forall>x y. P x y) = (\<forall>y x. P x y)"
  1265   by blast
  1266 
  1267 lemma ex_comm:
  1268   "(\<exists>x y. P x y) = (\<exists>y x. P x y)"
  1269   by blast
  1270 
  1271 use "simpdata.ML"
  1272 ML {* open Simpdata *}
  1273 
  1274 setup {*
  1275   Simplifier.method_setup Splitter.split_modifiers
  1276   #> (fn thy => (change_simpset_of thy (fn _ => Simpdata.simpset_simprocs); thy))
  1277   #> Splitter.setup
  1278   #> Clasimp.setup
  1279   #> EqSubst.setup
  1280 *}
  1281 
  1282 text {* Simproc for proving @{text "(y = x) == False"} from premise @{text "~(x = y)"}: *}
  1283 
  1284 simproc_setup neq ("x = y") = {* fn _ =>
  1285 let
  1286   val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI};
  1287   fun is_neq eq lhs rhs thm =
  1288     (case Thm.prop_of thm of
  1289       _ $ (Not $ (eq' $ l' $ r')) =>
  1290         Not = HOLogic.Not andalso eq' = eq andalso
  1291         r' aconv lhs andalso l' aconv rhs
  1292     | _ => false);
  1293   fun proc ss ct =
  1294     (case Thm.term_of ct of
  1295       eq $ lhs $ rhs =>
  1296         (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of_ss ss) of
  1297           SOME thm => SOME (thm RS neq_to_EQ_False)
  1298         | NONE => NONE)
  1299      | _ => NONE);
  1300 in proc end;
  1301 *}
  1302 
  1303 simproc_setup let_simp ("Let x f") = {*
  1304 let
  1305   val (f_Let_unfold, x_Let_unfold) =
  1306     let val [(_$(f$x)$_)] = prems_of @{thm Let_unfold}
  1307     in (cterm_of @{theory} f, cterm_of @{theory} x) end
  1308   val (f_Let_folded, x_Let_folded) =
  1309     let val [(_$(f$x)$_)] = prems_of @{thm Let_folded}
  1310     in (cterm_of @{theory} f, cterm_of @{theory} x) end;
  1311   val g_Let_folded =
  1312     let val [(_$_$(g$_))] = prems_of @{thm Let_folded} in cterm_of @{theory} g end;
  1313 
  1314   fun proc _ ss ct =
  1315     let
  1316       val ctxt = Simplifier.the_context ss;
  1317       val thy = ProofContext.theory_of ctxt;
  1318       val t = Thm.term_of ct;
  1319       val ([t'], ctxt') = Variable.import_terms false [t] ctxt;
  1320     in Option.map (hd o Variable.export ctxt' ctxt o single)
  1321       (case t' of Const ("Let",_) $ x $ f => (* x and f are already in normal form *)
  1322         if is_Free x orelse is_Bound x orelse is_Const x
  1323         then SOME @{thm Let_def}
  1324         else
  1325           let
  1326             val n = case f of (Abs (x,_,_)) => x | _ => "x";
  1327             val cx = cterm_of thy x;
  1328             val {T=xT,...} = rep_cterm cx;
  1329             val cf = cterm_of thy f;
  1330             val fx_g = Simplifier.rewrite ss (Thm.capply cf cx);
  1331             val (_$_$g) = prop_of fx_g;
  1332             val g' = abstract_over (x,g);
  1333           in (if (g aconv g')
  1334                then
  1335                   let
  1336                     val rl =
  1337                       cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] @{thm Let_unfold};
  1338                   in SOME (rl OF [fx_g]) end
  1339                else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*)
  1340                else let
  1341                      val abs_g'= Abs (n,xT,g');
  1342                      val g'x = abs_g'$x;
  1343                      val g_g'x = symmetric (beta_conversion false (cterm_of thy g'x));
  1344                      val rl = cterm_instantiate
  1345                                [(f_Let_folded,cterm_of thy f),(x_Let_folded,cx),
  1346                                 (g_Let_folded,cterm_of thy abs_g')]
  1347                                @{thm Let_folded};
  1348                    in SOME (rl OF [transitive fx_g g_g'x])
  1349                    end)
  1350           end
  1351       | _ => NONE)
  1352     end
  1353 in proc end *}
  1354 
  1355 
  1356 lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
  1357 proof
  1358   assume "True \<Longrightarrow> PROP P"
  1359   from this [OF TrueI] show "PROP P" .
  1360 next
  1361   assume "PROP P"
  1362   then show "PROP P" .
  1363 qed
  1364 
  1365 lemma ex_simps:
  1366   "!!P Q. (EX x. P x & Q)   = ((EX x. P x) & Q)"
  1367   "!!P Q. (EX x. P & Q x)   = (P & (EX x. Q x))"
  1368   "!!P Q. (EX x. P x | Q)   = ((EX x. P x) | Q)"
  1369   "!!P Q. (EX x. P | Q x)   = (P | (EX x. Q x))"
  1370   "!!P Q. (EX x. P x --> Q) = ((ALL x. P x) --> Q)"
  1371   "!!P Q. (EX x. P --> Q x) = (P --> (EX x. Q x))"
  1372   -- {* Miniscoping: pushing in existential quantifiers. *}
  1373   by (iprover | blast)+
  1374 
  1375 lemma all_simps:
  1376   "!!P Q. (ALL x. P x & Q)   = ((ALL x. P x) & Q)"
  1377   "!!P Q. (ALL x. P & Q x)   = (P & (ALL x. Q x))"
  1378   "!!P Q. (ALL x. P x | Q)   = ((ALL x. P x) | Q)"
  1379   "!!P Q. (ALL x. P | Q x)   = (P | (ALL x. Q x))"
  1380   "!!P Q. (ALL x. P x --> Q) = ((EX x. P x) --> Q)"
  1381   "!!P Q. (ALL x. P --> Q x) = (P --> (ALL x. Q x))"
  1382   -- {* Miniscoping: pushing in universal quantifiers. *}
  1383   by (iprover | blast)+
  1384 
  1385 lemmas [simp] =
  1386   triv_forall_equality (*prunes params*)
  1387   True_implies_equals  (*prune asms `True'*)
  1388   if_True
  1389   if_False
  1390   if_cancel
  1391   if_eq_cancel
  1392   imp_disjL
  1393   (*In general it seems wrong to add distributive laws by default: they
  1394     might cause exponential blow-up.  But imp_disjL has been in for a while
  1395     and cannot be removed without affecting existing proofs.  Moreover,
  1396     rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the
  1397     grounds that it allows simplification of R in the two cases.*)
  1398   conj_assoc
  1399   disj_assoc
  1400   de_Morgan_conj
  1401   de_Morgan_disj
  1402   imp_disj1
  1403   imp_disj2
  1404   not_imp
  1405   disj_not1
  1406   not_all
  1407   not_ex
  1408   cases_simp
  1409   the_eq_trivial
  1410   the_sym_eq_trivial
  1411   ex_simps
  1412   all_simps
  1413   simp_thms
  1414 
  1415 lemmas [cong] = imp_cong simp_implies_cong
  1416 lemmas [split] = split_if
  1417 
  1418 ML {* val HOL_ss = @{simpset} *}
  1419 
  1420 text {* Simplifies x assuming c and y assuming ~c *}
  1421 lemma if_cong:
  1422   assumes "b = c"
  1423       and "c \<Longrightarrow> x = u"
  1424       and "\<not> c \<Longrightarrow> y = v"
  1425   shows "(if b then x else y) = (if c then u else v)"
  1426   unfolding if_def using assms by simp
  1427 
  1428 text {* Prevents simplification of x and y:
  1429   faster and allows the execution of functional programs. *}
  1430 lemma if_weak_cong [cong]:
  1431   assumes "b = c"
  1432   shows "(if b then x else y) = (if c then x else y)"
  1433   using assms by (rule arg_cong)
  1434 
  1435 text {* Prevents simplification of t: much faster *}
  1436 lemma let_weak_cong:
  1437   assumes "a = b"
  1438   shows "(let x = a in t x) = (let x = b in t x)"
  1439   using assms by (rule arg_cong)
  1440 
  1441 text {* To tidy up the result of a simproc.  Only the RHS will be simplified. *}
  1442 lemma eq_cong2:
  1443   assumes "u = u'"
  1444   shows "(t \<equiv> u) \<equiv> (t \<equiv> u')"
  1445   using assms by simp
  1446 
  1447 lemma if_distrib:
  1448   "f (if c then x else y) = (if c then f x else f y)"
  1449   by simp
  1450 
  1451 text {* This lemma restricts the effect of the rewrite rule u=v to the left-hand
  1452   side of an equality.  Used in @{text "{Integ,Real}/simproc.ML"} *}
  1453 lemma restrict_to_left:
  1454   assumes "x = y"
  1455   shows "(x = z) = (y = z)"
  1456   using assms by simp
  1457 
  1458 
  1459 subsubsection {* Generic cases and induction *}
  1460 
  1461 text {* Rule projections: *}
  1462 
  1463 ML {*
  1464 structure ProjectRule = ProjectRuleFun
  1465 (struct
  1466   val conjunct1 = @{thm conjunct1};
  1467   val conjunct2 = @{thm conjunct2};
  1468   val mp = @{thm mp};
  1469 end)
  1470 *}
  1471 
  1472 constdefs
  1473   induct_forall where "induct_forall P == \<forall>x. P x"
  1474   induct_implies where "induct_implies A B == A \<longrightarrow> B"
  1475   induct_equal where "induct_equal x y == x = y"
  1476   induct_conj where "induct_conj A B == A \<and> B"
  1477 
  1478 lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\<lambda>x. P x))"
  1479   by (unfold atomize_all induct_forall_def)
  1480 
  1481 lemma induct_implies_eq: "(A ==> B) == Trueprop (induct_implies A B)"
  1482   by (unfold atomize_imp induct_implies_def)
  1483 
  1484 lemma induct_equal_eq: "(x == y) == Trueprop (induct_equal x y)"
  1485   by (unfold atomize_eq induct_equal_def)
  1486 
  1487 lemma induct_conj_eq:
  1488   includes meta_conjunction_syntax
  1489   shows "(A && B) == Trueprop (induct_conj A B)"
  1490   by (unfold atomize_conj induct_conj_def)
  1491 
  1492 lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq
  1493 lemmas induct_rulify [symmetric, standard] = induct_atomize
  1494 lemmas induct_rulify_fallback =
  1495   induct_forall_def induct_implies_def induct_equal_def induct_conj_def
  1496 
  1497 
  1498 lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) =
  1499     induct_conj (induct_forall A) (induct_forall B)"
  1500   by (unfold induct_forall_def induct_conj_def) iprover
  1501 
  1502 lemma induct_implies_conj: "induct_implies C (induct_conj A B) =
  1503     induct_conj (induct_implies C A) (induct_implies C B)"
  1504   by (unfold induct_implies_def induct_conj_def) iprover
  1505 
  1506 lemma induct_conj_curry: "(induct_conj A B ==> PROP C) == (A ==> B ==> PROP C)"
  1507 proof
  1508   assume r: "induct_conj A B ==> PROP C" and A B
  1509   show "PROP C" by (rule r) (simp add: induct_conj_def `A` `B`)
  1510 next
  1511   assume r: "A ==> B ==> PROP C" and "induct_conj A B"
  1512   show "PROP C" by (rule r) (simp_all add: `induct_conj A B` [unfolded induct_conj_def])
  1513 qed
  1514 
  1515 lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry
  1516 
  1517 hide const induct_forall induct_implies induct_equal induct_conj
  1518 
  1519 text {* Method setup. *}
  1520 
  1521 ML {*
  1522   structure Induct = InductFun
  1523   (
  1524     val cases_default = @{thm case_split}
  1525     val atomize = @{thms induct_atomize}
  1526     val rulify = @{thms induct_rulify}
  1527     val rulify_fallback = @{thms induct_rulify_fallback}
  1528   );
  1529 *}
  1530 
  1531 setup Induct.setup
  1532 
  1533 
  1534 subsection {* Other simple lemmas and lemma duplicates *}
  1535 
  1536 lemma Let_0 [simp]: "Let 0 f = f 0"
  1537   unfolding Let_def ..
  1538 
  1539 lemma Let_1 [simp]: "Let 1 f = f 1"
  1540   unfolding Let_def ..
  1541 
  1542 lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x"
  1543   by blast+
  1544 
  1545 lemma choice_eq: "(ALL x. EX! y. P x y) = (EX! f. ALL x. P x (f x))"
  1546   apply (rule iffI)
  1547   apply (rule_tac a = "%x. THE y. P x y" in ex1I)
  1548   apply (fast dest!: theI')
  1549   apply (fast intro: ext the1_equality [symmetric])
  1550   apply (erule ex1E)
  1551   apply (rule allI)
  1552   apply (rule ex1I)
  1553   apply (erule spec)
  1554   apply (erule_tac x = "%z. if z = x then y else f z" in allE)
  1555   apply (erule impE)
  1556   apply (rule allI)
  1557   apply (rule_tac P = "xa = x" in case_split_thm)
  1558   apply (drule_tac [3] x = x in fun_cong, simp_all)
  1559   done
  1560 
  1561 lemma mk_left_commute:
  1562   fixes f (infix "\<otimes>" 60)
  1563   assumes a: "\<And>x y z. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)" and
  1564           c: "\<And>x y. x \<otimes> y = y \<otimes> x"
  1565   shows "x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)"
  1566   by (rule trans [OF trans [OF c a] arg_cong [OF c, of "f y"]])
  1567 
  1568 lemmas eq_sym_conv = eq_commute
  1569 
  1570 lemma nnf_simps:
  1571   "(\<not>(P \<and> Q)) = (\<not> P \<or> \<not> Q)" "(\<not> (P \<or> Q)) = (\<not> P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" 
  1572   "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not>(P = Q)) = ((P \<and> \<not> Q) \<or> (\<not>P \<and> Q))" 
  1573   "(\<not> \<not>(P)) = P"
  1574 by blast+
  1575 
  1576 
  1577 subsection {* Basic ML bindings *}
  1578 
  1579 ML {*
  1580 val FalseE = @{thm FalseE}
  1581 val Let_def = @{thm Let_def}
  1582 val TrueI = @{thm TrueI}
  1583 val allE = @{thm allE}
  1584 val allI = @{thm allI}
  1585 val all_dupE = @{thm all_dupE}
  1586 val arg_cong = @{thm arg_cong}
  1587 val box_equals = @{thm box_equals}
  1588 val ccontr = @{thm ccontr}
  1589 val classical = @{thm classical}
  1590 val conjE = @{thm conjE}
  1591 val conjI = @{thm conjI}
  1592 val conjunct1 = @{thm conjunct1}
  1593 val conjunct2 = @{thm conjunct2}
  1594 val disjCI = @{thm disjCI}
  1595 val disjE = @{thm disjE}
  1596 val disjI1 = @{thm disjI1}
  1597 val disjI2 = @{thm disjI2}
  1598 val eq_reflection = @{thm eq_reflection}
  1599 val ex1E = @{thm ex1E}
  1600 val ex1I = @{thm ex1I}
  1601 val ex1_implies_ex = @{thm ex1_implies_ex}
  1602 val exE = @{thm exE}
  1603 val exI = @{thm exI}
  1604 val excluded_middle = @{thm excluded_middle}
  1605 val ext = @{thm ext}
  1606 val fun_cong = @{thm fun_cong}
  1607 val iffD1 = @{thm iffD1}
  1608 val iffD2 = @{thm iffD2}
  1609 val iffI = @{thm iffI}
  1610 val impE = @{thm impE}
  1611 val impI = @{thm impI}
  1612 val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq}
  1613 val mp = @{thm mp}
  1614 val notE = @{thm notE}
  1615 val notI = @{thm notI}
  1616 val not_all = @{thm not_all}
  1617 val not_ex = @{thm not_ex}
  1618 val not_iff = @{thm not_iff}
  1619 val not_not = @{thm not_not}
  1620 val not_sym = @{thm not_sym}
  1621 val refl = @{thm refl}
  1622 val rev_mp = @{thm rev_mp}
  1623 val spec = @{thm spec}
  1624 val ssubst = @{thm ssubst}
  1625 val subst = @{thm subst}
  1626 val sym = @{thm sym}
  1627 val trans = @{thm trans}
  1628 *}
  1629 
  1630 
  1631 subsection {* Code generator basic setup -- see further @{text Code_Setup.thy} *}
  1632 
  1633 setup "CodeName.setup #> CodeTarget.setup #> Nbe.setup"
  1634 
  1635 class eq (attach "op =") = type
  1636 
  1637 code_datatype True False
  1638 
  1639 lemma [code func]:
  1640   shows "False \<and> x \<longleftrightarrow> False"
  1641     and "True \<and> x \<longleftrightarrow> x"
  1642     and "x \<and> False \<longleftrightarrow> False"
  1643     and "x \<and> True \<longleftrightarrow> x" by simp_all
  1644 
  1645 lemma [code func]:
  1646   shows "False \<or> x \<longleftrightarrow> x"
  1647     and "True \<or> x \<longleftrightarrow> True"
  1648     and "x \<or> False \<longleftrightarrow> x"
  1649     and "x \<or> True \<longleftrightarrow> True" by simp_all
  1650 
  1651 lemma [code func]:
  1652   shows "\<not> True \<longleftrightarrow> False"
  1653     and "\<not> False \<longleftrightarrow> True" by (rule HOL.simp_thms)+
  1654 
  1655 instance bool :: eq ..
  1656 
  1657 lemma [code func]:
  1658   shows "False = P \<longleftrightarrow> \<not> P"
  1659     and "True = P \<longleftrightarrow> P" 
  1660     and "P = False \<longleftrightarrow> \<not> P" 
  1661     and "P = True \<longleftrightarrow> P" by simp_all
  1662 
  1663 code_datatype Trueprop "prop"
  1664 
  1665 code_datatype "TYPE('a)"
  1666 
  1667 lemma Let_case_cert:
  1668   assumes "CASE \<equiv> (\<lambda>x. Let x f)"
  1669   shows "CASE x \<equiv> f x"
  1670   using assms by simp_all
  1671 
  1672 lemma If_case_cert:
  1673   includes meta_conjunction_syntax
  1674   assumes "CASE \<equiv> (\<lambda>b. If b f g)"
  1675   shows "(CASE True \<equiv> f) && (CASE False \<equiv> g)"
  1676   using assms by simp_all
  1677 
  1678 setup {*
  1679   Code.add_case @{thm Let_case_cert}
  1680   #> Code.add_case @{thm If_case_cert}
  1681   #> Code.add_undefined @{const_name undefined}
  1682 *}
  1683 
  1684 
  1685 subsection {* Legacy tactics and ML bindings *}
  1686 
  1687 ML {*
  1688 fun strip_tac i = REPEAT (resolve_tac [impI, allI] i);
  1689 
  1690 (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *)
  1691 local
  1692   fun wrong_prem (Const ("All", _) $ (Abs (_, _, t))) = wrong_prem t
  1693     | wrong_prem (Bound _) = true
  1694     | wrong_prem _ = false;
  1695   val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of);
  1696 in
  1697   fun smp i = funpow i (fn m => filter_right ([spec] RL m)) ([mp]);
  1698   fun smp_tac j = EVERY'[dresolve_tac (smp j), atac];
  1699 end;
  1700 
  1701 val all_conj_distrib = thm "all_conj_distrib";
  1702 val all_simps = thms "all_simps";
  1703 val atomize_not = thm "atomize_not";
  1704 val case_split = thm "case_split";
  1705 val case_split_thm = thm "case_split_thm"
  1706 val cases_simp = thm "cases_simp";
  1707 val choice_eq = thm "choice_eq"
  1708 val cong = thm "cong"
  1709 val conj_comms = thms "conj_comms";
  1710 val conj_cong = thm "conj_cong";
  1711 val de_Morgan_conj = thm "de_Morgan_conj";
  1712 val de_Morgan_disj = thm "de_Morgan_disj";
  1713 val disj_assoc = thm "disj_assoc";
  1714 val disj_comms = thms "disj_comms";
  1715 val disj_cong = thm "disj_cong";
  1716 val eq_ac = thms "eq_ac";
  1717 val eq_cong2 = thm "eq_cong2"
  1718 val Eq_FalseI = thm "Eq_FalseI";
  1719 val Eq_TrueI = thm "Eq_TrueI";
  1720 val Ex1_def = thm "Ex1_def"
  1721 val ex_disj_distrib = thm "ex_disj_distrib";
  1722 val ex_simps = thms "ex_simps";
  1723 val if_cancel = thm "if_cancel";
  1724 val if_eq_cancel = thm "if_eq_cancel";
  1725 val if_False = thm "if_False";
  1726 val iff_conv_conj_imp = thm "iff_conv_conj_imp";
  1727 val iff = thm "iff"
  1728 val if_splits = thms "if_splits";
  1729 val if_True = thm "if_True";
  1730 val if_weak_cong = thm "if_weak_cong"
  1731 val imp_all = thm "imp_all";
  1732 val imp_cong = thm "imp_cong";
  1733 val imp_conjL = thm "imp_conjL";
  1734 val imp_conjR = thm "imp_conjR";
  1735 val imp_conv_disj = thm "imp_conv_disj";
  1736 val simp_implies_def = thm "simp_implies_def";
  1737 val simp_thms = thms "simp_thms";
  1738 val split_if = thm "split_if";
  1739 val the1_equality = thm "the1_equality"
  1740 val theI = thm "theI"
  1741 val theI' = thm "theI'"
  1742 val True_implies_equals = thm "True_implies_equals";
  1743 val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps simp_thms @ @{thms "nnf_simps"})
  1744 
  1745 *}
  1746 
  1747 end