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