src/HOL/Import/HOL/HOL4Base.thy
author wenzelm
Wed, 08 Oct 2008 19:20:29 +0200
changeset 28529 7ff939586e83
parent 23290 c358025ad8db
child 30952 7ab2716dd93b
permissions -rw-r--r--
setmp_noncritical makes it work with future scheduler;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
     1
(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
     2
17566
484ff733f29c new header syntax;
wenzelm
parents: 17188
diff changeset
     3
theory HOL4Base imports "../HOL4Compat" "../HOL4Syntax" begin
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     4
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     5
;setup_theory bool
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     6
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     7
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
     8
  ARB :: "'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
     9
  "ARB == SOME x::'a::type. True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    10
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    11
lemma ARB_DEF: "ARB = (SOME x::'a::type. True)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    12
  by (import bool ARB_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    13
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    14
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
    15
  IN :: "'a => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    16
  "IN == %(x::'a::type) f::'a::type => bool. f x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    17
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    18
lemma IN_DEF: "IN = (%(x::'a::type) f::'a::type => bool. f x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    19
  by (import bool IN_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    20
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    21
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
    22
  RES_FORALL :: "('a => bool) => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    23
  "RES_FORALL ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    24
%(p::'a::type => bool) m::'a::type => bool. ALL x::'a::type. IN x p --> m x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    25
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    26
lemma RES_FORALL_DEF: "RES_FORALL =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    27
(%(p::'a::type => bool) m::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    28
    ALL x::'a::type. IN x p --> m x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    29
  by (import bool RES_FORALL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    30
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    31
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
    32
  RES_EXISTS :: "('a => bool) => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    33
  "RES_EXISTS ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    34
%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    35
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    36
lemma RES_EXISTS_DEF: "RES_EXISTS =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    37
(%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    38
  by (import bool RES_EXISTS_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    39
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    40
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
    41
  RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    42
  "RES_EXISTS_UNIQUE ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    43
%(p::'a::type => bool) m::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    44
   RES_EXISTS p m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    45
   RES_FORALL p
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    46
    (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    47
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    48
lemma RES_EXISTS_UNIQUE_DEF: "RES_EXISTS_UNIQUE =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    49
(%(p::'a::type => bool) m::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    50
    RES_EXISTS p m &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    51
    RES_FORALL p
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    52
     (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    53
  by (import bool RES_EXISTS_UNIQUE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    54
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    55
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
    56
  RES_SELECT :: "('a => bool) => ('a => bool) => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    57
  "RES_SELECT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    58
%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    59
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    60
lemma RES_SELECT_DEF: "RES_SELECT =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    61
(%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    62
  by (import bool RES_SELECT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    63
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    64
lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    65
  by (import bool EXCLUDED_MIDDLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    66
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    67
lemma FORALL_THM: "All (f::'a::type => bool) = All f"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    68
  by (import bool FORALL_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    69
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    70
lemma EXISTS_THM: "Ex (f::'a::type => bool) = Ex f"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    71
  by (import bool EXISTS_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    72
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    73
lemma F_IMP: "ALL t::bool. ~ t --> t --> False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    74
  by (import bool F_IMP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    75
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    76
lemma NOT_AND: "~ ((t::bool) & ~ t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    77
  by (import bool NOT_AND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    78
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    79
lemma AND_CLAUSES: "ALL t::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    80
   (True & t) = t &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    81
   (t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    82
  by (import bool AND_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    83
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    84
lemma OR_CLAUSES: "ALL t::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    85
   (True | t) = True &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    86
   (t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    87
  by (import bool OR_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    88
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    89
lemma IMP_CLAUSES: "ALL t::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    90
   (True --> t) = t &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    91
   (t --> True) = True &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    92
   (False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    93
  by (import bool IMP_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    94
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
    95
lemma NOT_CLAUSES: "(ALL t::bool. (~ ~ t) = t) & (~ True) = False & (~ False) = True"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    96
  by (import bool NOT_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    97
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    98
lemma BOOL_EQ_DISTINCT: "True ~= False & False ~= True"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    99
  by (import bool BOOL_EQ_DISTINCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   100
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   101
lemma EQ_CLAUSES: "ALL t::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   102
   (True = t) = t &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   103
   (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   104
  by (import bool EQ_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   105
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   106
lemma COND_CLAUSES: "ALL (t1::'a::type) t2::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   107
   (if True then t1 else t2) = t1 & (if False then t1 else t2) = t2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   108
  by (import bool COND_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   109
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   110
lemma SELECT_UNIQUE: "ALL (P::'a::type => bool) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   111
   (ALL y::'a::type. P y = (y = x)) --> Eps P = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   112
  by (import bool SELECT_UNIQUE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   113
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   114
lemma BOTH_EXISTS_AND_THM: "ALL (P::bool) Q::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   115
   (EX x::'a::type. P & Q) = ((EX x::'a::type. P) & (EX x::'a::type. Q))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   116
  by (import bool BOTH_EXISTS_AND_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   117
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   118
lemma BOTH_FORALL_OR_THM: "ALL (P::bool) Q::bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   119
   (ALL x::'a::type. P | Q) = ((ALL x::'a::type. P) | (ALL x::'a::type. Q))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   120
  by (import bool BOTH_FORALL_OR_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   121
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   122
lemma BOTH_FORALL_IMP_THM: "ALL (P::bool) Q::bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   123
   (ALL x::'a::type. P --> Q) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   124
   ((EX x::'a::type. P) --> (ALL x::'a::type. Q))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   125
  by (import bool BOTH_FORALL_IMP_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   126
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   127
lemma BOTH_EXISTS_IMP_THM: "ALL (P::bool) Q::bool.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   128
   (EX x::'a::type. P --> Q) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   129
   ((ALL x::'a::type. P) --> (EX x::'a::type. Q))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   130
  by (import bool BOTH_EXISTS_IMP_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   131
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   132
lemma OR_IMP_THM: "ALL (A::bool) B::bool. (A = (B | A)) = (B --> A)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   133
  by (import bool OR_IMP_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   134
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   135
lemma DE_MORGAN_THM: "ALL (A::bool) B::bool. (~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   136
  by (import bool DE_MORGAN_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   137
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   138
lemma IMP_F_EQ_F: "ALL t::bool. (t --> False) = (t = False)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   139
  by (import bool IMP_F_EQ_F)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   140
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   141
lemma EQ_EXPAND: "ALL (t1::bool) t2::bool. (t1 = t2) = (t1 & t2 | ~ t1 & ~ t2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   142
  by (import bool EQ_EXPAND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   143
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   144
lemma COND_RATOR: "ALL (b::bool) (f::'a::type => 'b::type) (g::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   145
   x::'a::type. (if b then f else g) x = (if b then f x else g x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   146
  by (import bool COND_RATOR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   147
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   148
lemma COND_ABS: "ALL (b::bool) (f::'a::type => 'b::type) g::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   149
   (%x::'a::type. if b then f x else g x) = (if b then f else g)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   150
  by (import bool COND_ABS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   151
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   152
lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   153
   (if b then t1 else t2) = ((~ b | t1) & (b | t2))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   154
  by (import bool COND_EXPAND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   155
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   156
lemma ONE_ONE_THM: "ALL f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   157
   inj f = (ALL (x1::'a::type) x2::'a::type. f x1 = f x2 --> x1 = x2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   158
  by (import bool ONE_ONE_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   159
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   160
lemma ABS_REP_THM: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   161
 (%P::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   162
     (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   163
      ((Ex::(('b::type => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   164
        ((TYPE_DEFINITION::('a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   165
                           => ('b::type => 'a::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   166
          P))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   167
      ((Ex::(('b::type => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   168
        (%x::'b::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   169
            (Ex::(('a::type => 'b::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   170
             (%abs::'a::type => 'b::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   171
                 (op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   172
                  ((All::('b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   173
                    (%a::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   174
                        (op =::'b::type => 'b::type => bool) (abs (x a)) a))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   175
                  ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   176
                    (%r::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   177
                        (op =::bool => bool => bool) (P r)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   178
                         ((op =::'a::type => 'a::type => bool) (x (abs r))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   179
                           r)))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   180
  by (import bool ABS_REP_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   181
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   182
lemma LET_RAND: "(P::'b::type => bool) (Let (M::'a::type) (N::'a::type => 'b::type)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   183
(let x::'a::type = M in P (N x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   184
  by (import bool LET_RAND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   185
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   186
lemma LET_RATOR: "Let (M::'a::type) (N::'a::type => 'b::type => 'c::type) (b::'b::type) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   187
(let x::'a::type = M in N x b)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   188
  by (import bool LET_RATOR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   189
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   190
lemma SWAP_FORALL_THM: "ALL P::'a::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   191
   (ALL x::'a::type. All (P x)) = (ALL (y::'b::type) x::'a::type. P x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   192
  by (import bool SWAP_FORALL_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   193
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   194
lemma SWAP_EXISTS_THM: "ALL P::'a::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   195
   (EX x::'a::type. Ex (P x)) = (EX (y::'b::type) x::'a::type. P x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   196
  by (import bool SWAP_EXISTS_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   197
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   198
lemma AND_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   199
   (Q --> P = P') & (P' --> Q = Q') --> (P & Q) = (P' & Q')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   200
  by (import bool AND_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   201
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   202
lemma OR_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   203
   (~ Q --> P = P') & (~ P' --> Q = Q') --> (P | Q) = (P' | Q')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   204
  by (import bool OR_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   205
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   206
lemma COND_CONG: "ALL (P::bool) (Q::bool) (x::'a::type) (x'::'a::type) (y::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   207
   y'::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   208
   P = Q & (Q --> x = x') & (~ Q --> y = y') -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   209
   (if P then x else y) = (if Q then x' else y')"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   210
  by (import bool COND_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   211
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   212
lemma MONO_COND: "((x::bool) --> (y::bool)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   213
((z::bool) --> (w::bool)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   214
(if b::bool then x else z) --> (if b then y else w)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   215
  by (import bool MONO_COND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   216
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   217
lemma SKOLEM_THM: "ALL P::'a::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   218
   (ALL x::'a::type. Ex (P x)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   219
   (EX f::'a::type => 'b::type. ALL x::'a::type. P x (f x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   220
  by (import bool SKOLEM_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   221
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   222
lemma bool_case_thm: "(ALL (e0::'a::type) e1::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   223
    (case True of True => e0 | False => e1) = e0) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   224
(ALL (e0::'a::type) e1::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   225
    (case False of True => e0 | False => e1) = e1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   226
  by (import bool bool_case_thm)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   227
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   228
lemma bool_case_ID: "ALL (x::'a::type) b::bool. (case b of True => x | _ => x) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   229
  by (import bool bool_case_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   230
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   231
lemma boolAxiom: "ALL (e0::'a::type) e1::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   232
   EX x::bool => 'a::type. x True = e0 & x False = e1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   233
  by (import bool boolAxiom)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   234
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   235
lemma UEXISTS_OR_THM: "ALL (P::'a::type => bool) Q::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   236
   (EX! x::'a::type. P x | Q x) --> Ex1 P | Ex1 Q"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   237
  by (import bool UEXISTS_OR_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   238
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   239
lemma UEXISTS_SIMP: "(EX! x::'a::type. (t::bool)) = (t & (ALL x::'a::type. All (op = x)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   240
  by (import bool UEXISTS_SIMP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   241
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   242
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   243
  RES_ABSTRACT :: "('a => bool) => ('a => 'b) => 'a => 'b" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   244
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   245
specification (RES_ABSTRACT) RES_ABSTRACT_DEF: "(ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   246
    IN x p --> RES_ABSTRACT p m x = m x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   247
(ALL (p::'a::type => bool) (m1::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   248
    m2::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   249
    (ALL x::'a::type. IN x p --> m1 x = m2 x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   250
    RES_ABSTRACT p m1 = RES_ABSTRACT p m2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   251
  by (import bool RES_ABSTRACT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   252
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   253
lemma BOOL_FUN_CASES_THM: "ALL f::bool => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   254
   f = (%b::bool. True) |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   255
   f = (%b::bool. False) | f = (%b::bool. b) | f = Not"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   256
  by (import bool BOOL_FUN_CASES_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   257
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   258
lemma BOOL_FUN_INDUCT: "ALL P::(bool => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   259
   P (%b::bool. True) & P (%b::bool. False) & P (%b::bool. b) & P Not -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   260
   All P"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   261
  by (import bool BOOL_FUN_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   262
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   263
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   264
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   265
;setup_theory combin
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   266
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   267
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   268
  K :: "'a => 'b => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   269
  "K == %(x::'a::type) y::'b::type. x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   270
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   271
lemma K_DEF: "K = (%(x::'a::type) y::'b::type. x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   272
  by (import combin K_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   273
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   274
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   275
  S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   276
  "S ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   277
%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   278
   x::'a::type. f x (g x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   279
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   280
lemma S_DEF: "S =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   281
(%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   282
    x::'a::type. f x (g x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   283
  by (import combin S_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   284
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   285
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   286
  I :: "'a => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   287
  "(op ==::('a::type => 'a::type) => ('a::type => 'a::type) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   288
 (I::'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   289
 ((S::('a::type => ('a::type => 'a::type) => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   290
      => ('a::type => 'a::type => 'a::type) => 'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   291
   (K::'a::type => ('a::type => 'a::type) => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   292
   (K::'a::type => 'a::type => 'a::type))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   293
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   294
lemma I_DEF: "(op =::('a::type => 'a::type) => ('a::type => 'a::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   295
 (I::'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   296
 ((S::('a::type => ('a::type => 'a::type) => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   297
      => ('a::type => 'a::type => 'a::type) => 'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   298
   (K::'a::type => ('a::type => 'a::type) => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   299
   (K::'a::type => 'a::type => 'a::type))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   300
  by (import combin I_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   301
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   302
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   303
  C :: "('a => 'b => 'c) => 'b => 'a => 'c" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   304
  "C == %(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   305
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   306
lemma C_DEF: "C =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   307
(%(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   308
  by (import combin C_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   309
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   310
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   311
  W :: "('a => 'a => 'b) => 'a => 'b" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   312
  "W == %(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   313
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   314
lemma W_DEF: "W = (%(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   315
  by (import combin W_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   316
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   317
lemma I_THM: "ALL x::'a::type. I x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   318
  by (import combin I_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   319
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   320
lemma I_o_ID: "ALL f::'a::type => 'b::type. I o f = f & f o I = f"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   321
  by (import combin I_o_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   322
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   323
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   324
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   325
;setup_theory sum
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   326
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   327
lemma ISL_OR_ISR: "ALL x::'a::type + 'b::type. ISL x | ISR x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   328
  by (import sum ISL_OR_ISR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   329
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   330
lemma INL: "ALL x::'a::type + 'b::type. ISL x --> Inl (OUTL x) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   331
  by (import sum INL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   332
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   333
lemma INR: "ALL x::'a::type + 'b::type. ISR x --> Inr (OUTR x) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   334
  by (import sum INR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   335
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   336
lemma sum_case_cong: "ALL (M::'b::type + 'c::type) (M'::'b::type + 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   337
   (f::'b::type => 'a::type) g::'c::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   338
   M = M' &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   339
   (ALL x::'b::type. M' = Inl x --> f x = (f'::'b::type => 'a::type) x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   340
   (ALL y::'c::type. M' = Inr y --> g y = (g'::'c::type => 'a::type) y) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   341
   sum_case f g M = sum_case f' g' M'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   342
  by (import sum sum_case_cong)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   343
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   344
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   345
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   346
;setup_theory one
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   347
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   348
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   349
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   350
;setup_theory option
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   351
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   352
lemma option_CLAUSES: "(op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   353
 ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   354
   (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   355
       (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   356
        (%y::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   357
            (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   358
             ((op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   359
               ((Some::'a::type ~=> 'a::type) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   360
               ((Some::'a::type ~=> 'a::type) y))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   361
             ((op =::'a::type => 'a::type => bool) x y))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   362
 ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   363
   ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   364
     (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   365
         (op =::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   366
          ((the::'a::type option => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   367
            ((Some::'a::type ~=> 'a::type) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   368
          x))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   369
   ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   370
     ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   371
       (%x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   372
           (Not::bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   373
            ((op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   374
              (None::'a::type option) ((Some::'a::type ~=> 'a::type) x))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   375
     ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   376
       ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   377
         (%x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   378
             (Not::bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   379
              ((op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   380
                ((Some::'a::type ~=> 'a::type) x) (None::'a::type option))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   381
       ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   382
         ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   383
           (%x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   384
               (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   385
                ((IS_SOME::'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   386
                  ((Some::'a::type ~=> 'a::type) x))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   387
                (True::bool)))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   388
         ((op &::bool => bool => bool)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   389
           ((op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   390
             ((IS_SOME::'a::type option => bool) (None::'a::type option))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   391
             (False::bool))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   392
           ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   393
             ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   394
               (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   395
                   (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   396
                    ((IS_NONE::'a::type option => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   397
                    ((op =::'a::type option => 'a::type option => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   398
                      (None::'a::type option))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   399
             ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   400
               ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   401
                 (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   402
                     (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   403
                      ((Not::bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   404
                        ((IS_SOME::'a::type option => bool) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   405
                      ((op =::'a::type option => 'a::type option => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   406
                        (None::'a::type option))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   407
               ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   408
                 ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   409
                   (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   410
                       (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   411
                        ((IS_SOME::'a::type option => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   412
                        ((op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   413
                          ((Some::'a::type ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   414
                            ((the::'a::type option => 'a::type) x))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   415
                          x)))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   416
                 ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   417
                   ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   418
                     (%x::'a::type option.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   419
                         (op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   420
                          ((option_case::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   421
   => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   422
                            (None::'a::type option)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   423
                            (Some::'a::type ~=> 'a::type) x)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   424
                          x))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   425
                   ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   426
                     ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   427
                       (%x::'a::type option.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   428
                           (op =::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   429
                                  => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   430
                            ((option_case::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   431
     => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   432
                              x (Some::'a::type ~=> 'a::type) x)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   433
                            x))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   434
                     ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   435
                       ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   436
                         (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   437
                             (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   438
                              ((IS_NONE::'a::type option => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   439
                              ((op =::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   440
                                ((option_case::'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   441
         => ('a::type => 'b::type) => 'a::type option => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   442
                                  (e::'b::type) (f::'a::type => 'b::type) x)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   443
                                e)))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   444
                       ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   445
                         ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   446
                           (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   447
                               (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   448
                                ((IS_SOME::'a::type option => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   449
                                ((op =::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   450
                                  ((option_case::'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   451
           => ('a::type => 'b::type) => 'a::type option => 'b::type)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   452
                                    e f x)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   453
                                  (f ((the::'a::type option => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   454
 x)))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   455
                         ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   456
                           ((All::('a::type option => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   457
                             (%x::'a::type option.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   458
                                 (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   459
                                  ((IS_SOME::'a::type option => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   460
                                  ((op =::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   461
    => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   462
                                    ((option_case::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   463
             => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   464
(ea::'a::type option) (Some::'a::type ~=> 'a::type) x)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   465
                                    x)))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   466
                           ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   467
                             ((All::('b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   468
                               (%u::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   469
                                   (All::(('a::type => 'b::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   470
   => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   471
                                    (%f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   472
  (op =::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   473
   ((option_case::'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   474
                  => ('a::type => 'b::type) => 'a::type option => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   475
     u f (None::'a::type option))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   476
   u)))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   477
                             ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   478
                               ((All::('b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   479
                                 (%u::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   480
                                     (All::(('a::type => 'b::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   481
     => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   482
(%f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   483
    (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   484
     (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   485
         (op =::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   486
          ((option_case::'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   487
                         => ('a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   488
                            => 'a::type option => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   489
            u f ((Some::'a::type ~=> 'a::type) x))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   490
          (f x)))))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   491
                               ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   492
                                 ((All::(('a::type => 'b::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   493
  => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   494
                                   (%f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   495
 (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   496
  (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   497
      (op =::'b::type option => 'b::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   498
       ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   499
         f ((Some::'a::type ~=> 'a::type) x))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   500
       ((Some::'b::type ~=> 'b::type) (f x)))))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   501
                                 ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   502
                                   ((All::(('a::type => 'b::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   503
    => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   504
                                     (%f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   505
   (op =::'b::type option => 'b::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   506
    ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type) f
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   507
      (None::'a::type option))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   508
    (None::'b::type option)))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   509
                                   ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   510
                                     ((op =::'a::type option
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   511
       => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   512
 ((OPTION_JOIN::'a::type option option ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   513
   (None::'a::type option option))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   514
 (None::'a::type option))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   515
                                     ((All::('a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   516
      => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   517
 (%x::'a::type option.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   518
     (op =::'a::type option => 'a::type option => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   519
      ((OPTION_JOIN::'a::type option option ~=> 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   520
        ((Some::'a::type option ~=> 'a::type option) x))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   521
      x))))))))))))))))))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   522
  by (import option option_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   523
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   524
lemma option_case_compute: "option_case (e::'b::type) (f::'a::type => 'b::type) (x::'a::type option) =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   525
(if IS_SOME x then f (the x) else e)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   526
  by (import option option_case_compute)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   527
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   528
lemma OPTION_MAP_EQ_SOME: "ALL (f::'a::type => 'b::type) (x::'a::type option) y::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   529
   (option_map f x = Some y) = (EX z::'a::type. x = Some z & y = f z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   530
  by (import option OPTION_MAP_EQ_SOME)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   531
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   532
lemma OPTION_JOIN_EQ_SOME: "ALL (x::'a::type option option) xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   533
   (OPTION_JOIN x = Some xa) = (x = Some (Some xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   534
  by (import option OPTION_JOIN_EQ_SOME)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   535
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   536
lemma option_case_cong: "ALL (M::'a::type option) (M'::'a::type option) (u::'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   537
   f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   538
   M = M' &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   539
   (M' = None --> u = (u'::'b::type)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   540
   (ALL x::'a::type. M' = Some x --> f x = (f'::'a::type => 'b::type) x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   541
   option_case u f M = option_case u' f' M'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   542
  by (import option option_case_cong)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   543
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   544
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   545
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   546
;setup_theory marker
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   547
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   548
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   549
  stmarker :: "'a => 'a" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   550
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   551
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   552
  stmarker_primdef: "stmarker == %x::'a::type. x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   553
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   554
lemma stmarker_def: "ALL x::'a::type. stmarker x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   555
  by (import marker stmarker_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   556
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   557
lemma move_left_conj: "ALL (x::bool) (xa::bool) xb::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   558
   (x & stmarker xb) = (stmarker xb & x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   559
   ((stmarker xb & x) & xa) = (stmarker xb & x & xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   560
   (x & stmarker xb & xa) = (stmarker xb & x & xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   561
  by (import marker move_left_conj)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   562
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   563
lemma move_right_conj: "ALL (x::bool) (xa::bool) xb::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   564
   (stmarker xb & x) = (x & stmarker xb) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   565
   (x & xa & stmarker xb) = ((x & xa) & stmarker xb) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   566
   ((x & stmarker xb) & xa) = ((x & xa) & stmarker xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   567
  by (import marker move_right_conj)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   568
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   569
lemma move_left_disj: "ALL (x::bool) (xa::bool) xb::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   570
   (x | stmarker xb) = (stmarker xb | x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   571
   ((stmarker xb | x) | xa) = (stmarker xb | x | xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   572
   (x | stmarker xb | xa) = (stmarker xb | x | xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   573
  by (import marker move_left_disj)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   574
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   575
lemma move_right_disj: "ALL (x::bool) (xa::bool) xb::bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   576
   (stmarker xb | x) = (x | stmarker xb) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   577
   (x | xa | stmarker xb) = ((x | xa) | stmarker xb) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   578
   ((x | stmarker xb) | xa) = ((x | xa) | stmarker xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   579
  by (import marker move_right_disj)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   580
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   581
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   582
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   583
;setup_theory relation
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   584
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   585
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   586
  TC :: "('a => 'a => bool) => 'a => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   587
  "TC ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   588
%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   589
   ALL P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   590
      (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   591
      (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   592
          P x y & P y z --> P x z) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   593
      P a b"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   594
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   595
lemma TC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   596
   TC R a b =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   597
   (ALL P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   598
       (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   599
       (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   600
           P x y & P y z --> P x z) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   601
       P a b)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   602
  by (import relation TC_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   603
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   604
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   605
  RTC :: "('a => 'a => bool) => 'a => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   606
  "RTC ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   607
%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   608
   ALL P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   609
      (ALL x::'a::type. P x x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   610
      (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   611
          R x y & P y z --> P x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   612
      P a b"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   613
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   614
lemma RTC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   615
   RTC R a b =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   616
   (ALL P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   617
       (ALL x::'a::type. P x x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   618
       (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   619
           R x y & P y z --> P x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   620
       P a b)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   621
  by (import relation RTC_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   622
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   623
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   624
  RC :: "('a => 'a => bool) => 'a => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   625
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   626
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   627
  RC_primdef: "RC ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   628
%(R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. x = y | R x y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   629
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   630
lemma RC_def: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   631
   RC R x y = (x = y | R x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   632
  by (import relation RC_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   633
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   634
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   635
  transitive :: "('a => 'a => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   636
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   637
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   638
  transitive_primdef: "transitive ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   639
%R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   640
   ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   641
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   642
lemma transitive_def: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   643
   transitive R =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   644
   (ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   645
  by (import relation transitive_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   646
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   647
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   648
  pred_reflexive :: "('a => 'a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   649
  "pred_reflexive == %R::'a::type => 'a::type => bool. ALL x::'a::type. R x x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   650
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   651
lemma reflexive_def: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   652
   pred_reflexive R = (ALL x::'a::type. R x x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   653
  by (import relation reflexive_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   654
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   655
lemma TC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (TC x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   656
  by (import relation TC_TRANSITIVE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   657
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   658
lemma RTC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   659
   (ALL x::'a::type. xa x x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   660
   (ALL (xb::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   661
       x xb y & xa y z --> xa xb z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   662
   (ALL (xb::'a::type) xc::'a::type. RTC x xb xc --> xa xb xc)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   663
  by (import relation RTC_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   664
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   665
lemma TC_RULES: "ALL x::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   666
   (ALL (xa::'a::type) xb::'a::type. x xa xb --> TC x xa xb) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   667
   (ALL (xa::'a::type) (xb::'a::type) xc::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   668
       TC x xa xb & TC x xb xc --> TC x xa xc)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   669
  by (import relation TC_RULES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   670
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   671
lemma RTC_RULES: "ALL x::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   672
   (ALL xa::'a::type. RTC x xa xa) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   673
   (ALL (xa::'a::type) (xb::'a::type) xc::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   674
       x xa xb & RTC x xb xc --> RTC x xa xc)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   675
  by (import relation RTC_RULES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   676
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   677
lemma RTC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   678
   (ALL x::'a::type. P x x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   679
   (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   680
       R x y & RTC R y z & P y z --> P x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   681
   (ALL (x::'a::type) y::'a::type. RTC R x y --> P x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   682
  by (import relation RTC_STRONG_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   683
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   684
lemma RTC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   685
   RTC R x y --> (ALL z::'a::type. RTC R y z --> RTC R x z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   686
  by (import relation RTC_RTC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   687
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   688
lemma RTC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (RTC x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   689
  by (import relation RTC_TRANSITIVE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   690
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   691
lemma RTC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RTC R)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   692
  by (import relation RTC_REFLEXIVE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   693
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   694
lemma RC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RC R)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   695
  by (import relation RC_REFLEXIVE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   696
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   697
lemma TC_SUBSET: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   698
   x xa xb --> TC x xa xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   699
  by (import relation TC_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   700
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   701
lemma RTC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   702
   R x y --> RTC R x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   703
  by (import relation RTC_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   704
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   705
lemma RC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   706
   R x y --> RC R x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   707
  by (import relation RC_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   708
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   709
lemma RC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   710
   RC R x y --> RTC R x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   711
  by (import relation RC_RTC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   712
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   713
lemma TC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   714
   (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   715
   (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   716
       xa x y & xa y z --> xa x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   717
   (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   718
  by (import relation TC_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   719
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   720
lemma TC_INDUCT_LEFT1: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   721
   (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   722
   (ALL (xb::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   723
       x xb y & xa y z --> xa xb z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   724
   (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   725
  by (import relation TC_INDUCT_LEFT1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   726
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   727
lemma TC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   728
   (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   729
   (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   730
       P x y & P y z & TC R x y & TC R y z --> P x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   731
   (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   732
  by (import relation TC_STRONG_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   733
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   734
lemma TC_STRONG_INDUCT_LEFT1: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   735
   (ALL (x::'a::type) y::'a::type. R x y --> P x y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   736
   (ALL (x::'a::type) (y::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   737
       R x y & P y z & TC R y z --> P x z) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   738
   (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   739
  by (import relation TC_STRONG_INDUCT_LEFT1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   740
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   741
lemma TC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   742
   TC R x y --> RTC R x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   743
  by (import relation TC_RTC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   744
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   745
lemma RTC_TC_RC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   746
   RTC R x y --> RC R x y | TC R x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   747
  by (import relation RTC_TC_RC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   748
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   749
lemma TC_RC_EQNS: "ALL R::'a::type => 'a::type => bool. RC (TC R) = RTC R & TC (RC R) = RTC R"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   750
  by (import relation TC_RC_EQNS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   751
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   752
lemma RC_IDEM: "ALL R::'a::type => 'a::type => bool. RC (RC R) = RC R"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   753
  by (import relation RC_IDEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   754
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   755
lemma TC_IDEM: "ALL R::'a::type => 'a::type => bool. TC (TC R) = TC R"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   756
  by (import relation TC_IDEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   757
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   758
lemma RTC_IDEM: "ALL R::'a::type => 'a::type => bool. RTC (RTC R) = RTC R"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   759
  by (import relation RTC_IDEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   760
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   761
lemma RTC_CASES1: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   762
   RTC x xa xb = (xa = xb | (EX u::'a::type. x xa u & RTC x u xb))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   763
  by (import relation RTC_CASES1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   764
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   765
lemma RTC_CASES2: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   766
   RTC x xa xb = (xa = xb | (EX u::'a::type. RTC x xa u & x u xb))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   767
  by (import relation RTC_CASES2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   768
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   769
lemma RTC_CASES_RTC_TWICE: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   770
   RTC x xa xb = (EX u::'a::type. RTC x xa u & RTC x u xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   771
  by (import relation RTC_CASES_RTC_TWICE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   772
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   773
lemma TC_CASES1: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   774
   TC R x z --> R x z | (EX y::'a::type. R x y & TC R y z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   775
  by (import relation TC_CASES1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   776
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   777
lemma TC_CASES2: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   778
   TC R x z --> R x z | (EX y::'a::type. TC R x y & R y z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   779
  by (import relation TC_CASES2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   780
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   781
lemma TC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   782
   (ALL (x::'a::type) y::'a::type. R x y --> Q x y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   783
   (ALL (x::'a::type) y::'a::type. TC R x y --> TC Q x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   784
  by (import relation TC_MONOTONE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   785
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   786
lemma RTC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   787
   (ALL (x::'a::type) y::'a::type. R x y --> Q x y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   788
   (ALL (x::'a::type) y::'a::type. RTC R x y --> RTC Q x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   789
  by (import relation RTC_MONOTONE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   790
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   791
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   792
  WF :: "('a => 'a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   793
  "WF ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   794
%R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   795
   ALL B::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   796
      Ex B -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   797
      (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   798
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   799
lemma WF_DEF: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   800
   WF R =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   801
   (ALL B::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   802
       Ex B -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   803
       (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   804
  by (import relation WF_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   805
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   806
lemma WF_INDUCTION_THM: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   807
   WF R -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   808
   (ALL P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   809
       (ALL x::'a::type. (ALL y::'a::type. R y x --> P y) --> P x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   810
       All P)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   811
  by (import relation WF_INDUCTION_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   812
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   813
lemma WF_NOT_REFL: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   814
   WF x --> x xa xb --> xa ~= xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   815
  by (import relation WF_NOT_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   816
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   817
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   818
  EMPTY_REL :: "'a => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   819
  "EMPTY_REL == %(x::'a::type) y::'a::type. False"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   820
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   821
lemma EMPTY_REL_DEF: "ALL (x::'a::type) y::'a::type. EMPTY_REL x y = False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   822
  by (import relation EMPTY_REL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   823
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   824
lemma WF_EMPTY_REL: "WF EMPTY_REL"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   825
  by (import relation WF_EMPTY_REL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   826
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   827
lemma WF_SUBSET: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   828
   WF x & (ALL (xb::'a::type) y::'a::type. xa xb y --> x xb y) --> WF xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   829
  by (import relation WF_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   830
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   831
lemma WF_TC: "ALL R::'a::type => 'a::type => bool. WF R --> WF (TC R)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   832
  by (import relation WF_TC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   833
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   834
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   835
  inv_image :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   836
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   837
defs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   838
  inv_image_primdef: "relation.inv_image ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   839
%(R::'b::type => 'b::type => bool) (f::'a::type => 'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   840
   y::'a::type. R (f x) (f y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   841
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   842
lemma inv_image_def: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   843
   relation.inv_image R f = (%(x::'a::type) y::'a::type. R (f x) (f y))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   844
  by (import relation inv_image_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   845
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   846
lemma WF_inv_image: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   847
   WF R --> WF (relation.inv_image R f)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   848
  by (import relation WF_inv_image)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   849
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   850
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   851
  RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   852
  "RESTRICT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   853
%(f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   854
   y::'a::type. if R y x then f y else ARB"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   855
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   856
lemma RESTRICT_DEF: "ALL (f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   857
   RESTRICT f R x = (%y::'a::type. if R y x then f y else ARB)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   858
  by (import relation RESTRICT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   859
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   860
lemma RESTRICT_LEMMA: "ALL (x::'a::type => 'b::type) (xa::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   861
   (xb::'a::type) xc::'a::type. xa xb xc --> RESTRICT x xa xc xb = x xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   862
  by (import relation RESTRICT_LEMMA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   863
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   864
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   865
  approx :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => ('a => 'b) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   866
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   867
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   868
  approx_primdef: "approx ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   869
%(R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   870
   (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   871
   f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   872
   f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   873
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   874
lemma approx_def: "ALL (R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   875
   (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   876
   f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   877
   approx R M x f = (f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   878
  by (import relation approx_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   879
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   880
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   881
  the_fun :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'a => 'b" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   882
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   883
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   884
  the_fun_primdef: "the_fun ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   885
%(R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   886
   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   887
   Eps (approx R M x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   888
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   889
lemma the_fun_def: "ALL (R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   890
   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   891
   the_fun R M x = Eps (approx R M x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   892
  by (import relation the_fun_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   893
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   894
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   895
  WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   896
  "WFREC ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   897
%(R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   898
   (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   899
   M (RESTRICT
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   900
       (the_fun (TC R)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   901
         (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   902
       R x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   903
    x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   904
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   905
lemma WFREC_DEF: "ALL (R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   906
   M::('a::type => 'b::type) => 'a::type => 'b::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   907
   WFREC R M =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   908
   (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   909
       M (RESTRICT
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   910
           (the_fun (TC R)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   911
             (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   912
             x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   913
           R x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   914
        x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   915
  by (import relation WFREC_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   916
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   917
lemma WFREC_THM: "ALL (R::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   918
   M::('a::type => 'b::type) => 'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   919
   WF R --> (ALL x::'a::type. WFREC R M x = M (RESTRICT (WFREC R M) R x) x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   920
  by (import relation WFREC_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   921
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   922
lemma WFREC_COROLLARY: "ALL (M::('a::type => 'b::type) => 'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   923
   (R::'a::type => 'a::type => bool) f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   924
   f = WFREC R M --> WF R --> (ALL x::'a::type. f x = M (RESTRICT f R x) x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   925
  by (import relation WFREC_COROLLARY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   926
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   927
lemma WF_RECURSION_THM: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   928
   WF R -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   929
   (ALL M::('a::type => 'b::type) => 'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   930
       EX! f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   931
          ALL x::'a::type. f x = M (RESTRICT f R x) x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   932
  by (import relation WF_RECURSION_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   933
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   934
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   935
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   936
;setup_theory pair
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   937
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   938
lemma CURRY_ONE_ONE_THM: "(curry (f::'a::type * 'b::type => 'c::type) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   939
 curry (g::'a::type * 'b::type => 'c::type)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   940
(f = g)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   941
  by (import pair CURRY_ONE_ONE_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   942
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
   943
lemma UNCURRY_ONE_ONE_THM: "(op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   944
 ((op =::('a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   945
         => ('a::type * 'b::type => 'c::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   946
   ((split::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   947
            => 'a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   948
     (f::'a::type => 'b::type => 'c::type))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   949
   ((split::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   950
            => 'a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   951
     (g::'a::type => 'b::type => 'c::type)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   952
 ((op =::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   953
         => ('a::type => 'b::type => 'c::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   954
   f g)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   955
  by (import pair UNCURRY_ONE_ONE_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   956
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   957
lemma pair_Axiom: "ALL f::'a::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   958
   EX x::'a::type * 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   959
      ALL (xa::'a::type) y::'b::type. x (xa, y) = f xa y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   960
  by (import pair pair_Axiom)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   961
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   962
lemma UNCURRY_CONG: "ALL (M::'a::type * 'b::type) (M'::'a::type * 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   963
   f::'a::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   964
   M = M' &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   965
   (ALL (x::'a::type) y::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   966
       M' = (x, y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   967
       f x y = (f'::'a::type => 'b::type => 'c::type) x y) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   968
   split f M = split f' M'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   969
  by (import pair UNCURRY_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   970
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   971
lemma ELIM_PEXISTS: "(EX p::'a::type * 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   972
    (P::'a::type => 'b::type => bool) (fst p) (snd p)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   973
(EX p1::'a::type. Ex (P p1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   974
  by (import pair ELIM_PEXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   975
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   976
lemma ELIM_PFORALL: "(ALL p::'a::type * 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   977
    (P::'a::type => 'b::type => bool) (fst p) (snd p)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   978
(ALL p1::'a::type. All (P p1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   979
  by (import pair ELIM_PFORALL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   980
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   981
lemma PFORALL_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   982
 (%x::'a::type => 'b::type => bool.
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
   983
     (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   984
      ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   985
        (%xa::'a::type. (All::('b::type => bool) => bool) (x xa)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   986
      ((All::('a::type * 'b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   987
        ((split::('a::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   988
                 => 'a::type * 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   989
          x)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   990
  by (import pair PFORALL_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   991
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   992
lemma PEXISTS_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   993
 (%x::'a::type => 'b::type => bool.
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
   994
     (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   995
      ((Ex::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   996
        (%xa::'a::type. (Ex::('b::type => bool) => bool) (x xa)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   997
      ((Ex::('a::type * 'b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   998
        ((split::('a::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
   999
                 => 'a::type * 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1000
          x)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1001
  by (import pair PEXISTS_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1002
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1003
lemma LET2_RAND: "(All::(('c::type => 'd::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1004
 (%x::'c::type => 'd::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1005
     (All::('a::type * 'b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1006
      (%xa::'a::type * 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1007
          (All::(('a::type => 'b::type => 'c::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1008
           (%xb::'a::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1009
               (op =::'d::type => 'd::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1010
                (x ((Let::'a::type * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1011
                          => ('a::type * 'b::type => 'c::type) => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1012
                     xa ((split::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1013
                                 => 'a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1014
                          xb)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1015
                ((Let::'a::type * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1016
                       => ('a::type * 'b::type => 'd::type) => 'd::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1017
                  xa ((split::('a::type => 'b::type => 'd::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1018
                              => 'a::type * 'b::type => 'd::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1019
                       (%(xa::'a::type) y::'b::type. x (xb xa y)))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1020
  by (import pair LET2_RAND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1021
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1022
lemma LET2_RATOR: "(All::('a1::type * 'a2::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1023
 (%x::'a1::type * 'a2::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1024
     (All::(('a1::type => 'a2::type => 'b::type => 'c::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1025
           => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1026
      (%xa::'a1::type => 'a2::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1027
          (All::('b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1028
           (%xb::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1029
               (op =::'c::type => 'c::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1030
                ((Let::'a1::type * 'a2::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1031
                       => ('a1::type * 'a2::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1032
                          => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1033
                  x ((split::('a1::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1034
                              => 'a2::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1035
                             => 'a1::type * 'a2::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1036
                                => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1037
                      xa)
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
  1038
                  xb)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1039
                ((Let::'a1::type * 'a2::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1040
                       => ('a1::type * 'a2::type => 'c::type) => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1041
                  x ((split::('a1::type => 'a2::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1042
                             => 'a1::type * 'a2::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1043
                      (%(x::'a1::type) y::'a2::type. xa x y xb))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1044
  by (import pair LET2_RATOR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1045
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1046
lemma pair_case_cong: "ALL (x::'a::type * 'b::type) (xa::'a::type * 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1047
   xb::'a::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1048
   x = xa &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1049
   (ALL (x::'a::type) y::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1050
       xa = (x, y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1051
       xb x y = (f'::'a::type => 'b::type => 'c::type) x y) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1052
   split xb x = split f' xa"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1053
  by (import pair pair_case_cong)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1054
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1055
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1056
  LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1057
  "LEX ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1058
%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1059
   (s::'a::type, t::'b::type) (u::'a::type, v::'b::type).
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1060
   R1 s u | s = u & R2 t v"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1061
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1062
lemma LEX_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1063
   LEX R1 R2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1064
   (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type).
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1065
       R1 s u | s = u & R2 t v)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1066
  by (import pair LEX_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1067
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1068
lemma WF_LEX: "ALL (x::'a::type => 'a::type => bool) xa::'b::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1069
   WF x & WF xa --> WF (LEX x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1070
  by (import pair WF_LEX)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1071
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1072
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1073
  RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1074
  "RPROD ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1075
%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1076
   (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1077
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1078
lemma RPROD_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1079
   RPROD R1 R2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1080
   (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1081
  by (import pair RPROD_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1082
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1083
lemma WF_RPROD: "ALL (R::'a::type => 'a::type => bool) Q::'b::type => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1084
   WF R & WF Q --> WF (RPROD R Q)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1085
  by (import pair WF_RPROD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1086
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1087
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1088
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1089
;setup_theory num
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1090
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1091
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1092
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1093
;setup_theory prim_rec
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1094
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1095
lemma LESS_0_0: "0 < Suc 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1096
  by (import prim_rec LESS_0_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1097
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1098
lemma LESS_LEMMA1: "ALL (x::nat) xa::nat. x < Suc xa --> x = xa | x < xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1099
  by (import prim_rec LESS_LEMMA1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1100
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1101
lemma LESS_LEMMA2: "ALL (m::nat) n::nat. m = n | m < n --> m < Suc n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1102
  by (import prim_rec LESS_LEMMA2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1103
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1104
lemma LESS_THM: "ALL (m::nat) n::nat. (m < Suc n) = (m = n | m < n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1105
  by (import prim_rec LESS_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1106
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1107
lemma LESS_SUC_IMP: "ALL (x::nat) xa::nat. x < Suc xa --> x ~= xa --> x < xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1108
  by (import prim_rec LESS_SUC_IMP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1109
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1110
lemma EQ_LESS: "ALL n::nat. Suc (m::nat) = n --> m < n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1111
  by (import prim_rec EQ_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1112
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1113
lemma NOT_LESS_EQ: "ALL (m::nat) n::nat. m = n --> ~ m < n"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1114
  by (import prim_rec NOT_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1115
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1116
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1117
  SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1118
  "(op ==::((nat => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1119
         => 'a::type => ('a::type => 'a::type) => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1120
        => ((nat => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1121
            => 'a::type => ('a::type => 'a::type) => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1122
           => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1123
 (SIMP_REC_REL::(nat => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1124
                => 'a::type => ('a::type => 'a::type) => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1125
 (%(fun::nat => 'a::type) (x::'a::type) (f::'a::type => 'a::type) n::nat.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1126
     (op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1127
      ((op =::'a::type => 'a::type => bool) (fun (0::nat)) x)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1128
      ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1129
        (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1130
            (op -->::bool => bool => bool) ((op <::nat => nat => bool) m n)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1131
             ((op =::'a::type => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1132
               (fun ((Suc::nat => nat) m)) (f (fun m))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1133
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1134
lemma SIMP_REC_REL: "(All::((nat => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1135
 (%fun::nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1136
     (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1137
      (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1138
          (All::(('a::type => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1139
           (%f::'a::type => 'a::type.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1140
               (All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1141
                (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1142
                    (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1143
                     ((SIMP_REC_REL::(nat => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1144
                                     => 'a::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1145
  => ('a::type => 'a::type) => nat => bool)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1146
                       fun x f n)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1147
                     ((op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1148
                       ((op =::'a::type => 'a::type => bool) (fun (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1149
                         x)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1150
                       ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1151
                         (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1152
                             (op -->::bool => bool => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1153
                              ((op <::nat => nat => bool) m n)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1154
                              ((op =::'a::type => 'a::type => bool)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1155
                                (fun ((Suc::nat => nat) m))
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1156
                                (f (fun m))))))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1157
  by (import prim_rec SIMP_REC_REL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1158
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1159
lemma SIMP_REC_EXISTS: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1160
   EX fun::nat => 'a::type. SIMP_REC_REL fun x f n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1161
  by (import prim_rec SIMP_REC_EXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1162
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1163
lemma SIMP_REC_REL_UNIQUE: "ALL (x::'a::type) (xa::'a::type => 'a::type) (xb::nat => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1164
   (xc::nat => 'a::type) (xd::nat) xe::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1165
   SIMP_REC_REL xb x xa xd & SIMP_REC_REL xc x xa xe -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1166
   (ALL n::nat. n < xd & n < xe --> xb n = xc n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1167
  by (import prim_rec SIMP_REC_REL_UNIQUE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1168
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1169
lemma SIMP_REC_REL_UNIQUE_RESULT: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1170
   EX! y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1171
      EX g::nat => 'a::type. SIMP_REC_REL g x f (Suc n) & y = g n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1172
  by (import prim_rec SIMP_REC_REL_UNIQUE_RESULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1173
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1174
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1175
  SIMP_REC :: "'a => ('a => 'a) => nat => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1176
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1177
specification (SIMP_REC) SIMP_REC: "ALL (x::'a::type) (f'::'a::type => 'a::type) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1178
   EX g::nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1179
      SIMP_REC_REL g x f' (Suc n) & SIMP_REC x f' n = g n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1180
  by (import prim_rec SIMP_REC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1181
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1182
lemma LESS_SUC_SUC: "ALL m::nat. m < Suc m & m < Suc (Suc m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1183
  by (import prim_rec LESS_SUC_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1184
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1185
lemma SIMP_REC_THM: "ALL (x::'a::type) f::'a::type => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1186
   SIMP_REC x f 0 = x &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1187
   (ALL m::nat. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1188
  by (import prim_rec SIMP_REC_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1189
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1190
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1191
  PRE :: "nat => nat" 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1192
  "PRE == %m::nat. if m = 0 then 0 else SOME n::nat. m = Suc n"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1193
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1194
lemma PRE_DEF: "ALL m::nat. PRE m = (if m = 0 then 0 else SOME n::nat. m = Suc n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1195
  by (import prim_rec PRE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1196
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1197
lemma PRE: "PRE 0 = 0 & (ALL m::nat. PRE (Suc m) = m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1198
  by (import prim_rec PRE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1199
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1200
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1201
  PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1202
  "PRIM_REC_FUN ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1203
%(x::'a::type) f::'a::type => nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1204
   SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1205
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1206
lemma PRIM_REC_FUN: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1207
   PRIM_REC_FUN x f =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1208
   SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1209
  by (import prim_rec PRIM_REC_FUN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1210
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1211
lemma PRIM_REC_EQN: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1212
   (ALL n::nat. PRIM_REC_FUN x f 0 n = x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1213
   (ALL (m::nat) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1214
       PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1215
  by (import prim_rec PRIM_REC_EQN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1216
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1217
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1218
  PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1219
  "PRIM_REC ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1220
%(x::'a::type) (f::'a::type => nat => 'a::type) m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1221
   PRIM_REC_FUN x f m (PRE m)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1222
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1223
lemma PRIM_REC: "ALL (x::'a::type) (f::'a::type => nat => 'a::type) m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1224
   PRIM_REC x f m = PRIM_REC_FUN x f m (PRE m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1225
  by (import prim_rec PRIM_REC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1226
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1227
lemma PRIM_REC_THM: "ALL (x::'a::type) f::'a::type => nat => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1228
   PRIM_REC x f 0 = x &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1229
   (ALL m::nat. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1230
  by (import prim_rec PRIM_REC_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1231
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1232
lemma DC: "ALL (P::'a::type => bool) (R::'a::type => 'a::type => bool) a::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1233
   P a & (ALL x::'a::type. P x --> (EX y::'a::type. P y & R x y)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1234
   (EX x::nat => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1235
       x 0 = a & (ALL n::nat. P (x n) & R (x n) (x (Suc n))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1236
  by (import prim_rec DC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1237
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1238
lemma num_Axiom_old: "ALL (e::'a::type) f::'a::type => nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1239
   EX! fn1::nat => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1240
      fn1 0 = e & (ALL n::nat. fn1 (Suc n) = f (fn1 n) n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1241
  by (import prim_rec num_Axiom_old)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1242
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1243
lemma num_Axiom: "ALL (e::'a::type) f::nat => 'a::type => 'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1244
   EX x::nat => 'a::type. x 0 = e & (ALL n::nat. x (Suc n) = f n (x n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1245
  by (import prim_rec num_Axiom)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1246
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1247
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1248
  wellfounded :: "('a => 'a => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1249
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1250
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1251
  wellfounded_primdef: "wellfounded ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1252
%R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1253
   ~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1254
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1255
lemma wellfounded_def: "ALL R::'a::type => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1256
   wellfounded R =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1257
   (~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1258
  by (import prim_rec wellfounded_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1259
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1260
lemma WF_IFF_WELLFOUNDED: "ALL R::'a::type => 'a::type => bool. WF R = wellfounded R"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1261
  by (import prim_rec WF_IFF_WELLFOUNDED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1262
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1263
lemma WF_PRED: "WF (%(x::nat) y::nat. y = Suc x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1264
  by (import prim_rec WF_PRED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1265
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1266
lemma WF_LESS: "(WF::(nat => nat => bool) => bool) (op <::nat => nat => bool)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1267
  by (import prim_rec WF_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1268
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1269
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1270
  measure :: "('a => nat) => 'a => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1271
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1272
defs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1273
  measure_primdef: "prim_rec.measure == relation.inv_image op <"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1274
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1275
lemma measure_def: "prim_rec.measure = relation.inv_image op <"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1276
  by (import prim_rec measure_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1277
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1278
lemma WF_measure: "ALL x::'a::type => nat. WF (prim_rec.measure x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1279
  by (import prim_rec WF_measure)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1280
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1281
lemma measure_thm: "ALL (x::'a::type => nat) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1282
   prim_rec.measure x xa xb = (x xa < x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1283
  by (import prim_rec measure_thm)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1284
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1285
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1286
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1287
;setup_theory arithmetic
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1288
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1289
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1290
  nat_elim__magic :: "nat => nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1291
  "nat_elim__magic == %n::nat. n"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1292
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1293
lemma nat_elim__magic: "ALL n::nat. nat_elim__magic n = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1294
  by (import arithmetic nat_elim__magic)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1295
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1296
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1297
  EVEN :: "nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1298
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1299
specification (EVEN) EVEN: "EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1300
  by (import arithmetic EVEN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1301
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1302
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1303
  ODD :: "nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1304
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1305
specification (ODD) ODD: "ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1306
  by (import arithmetic ODD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1307
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1308
lemma TWO: "2 = Suc 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1309
  by (import arithmetic TWO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1310
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1311
lemma NORM_0: "(op =::nat => nat => bool) (0::nat) (0::nat)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1312
  by (import arithmetic NORM_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1313
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1314
lemma num_case_compute: "ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1315
   nat_case (f::'a::type) (g::nat => 'a::type) n =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1316
   (if n = 0 then f else g (PRE n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1317
  by (import arithmetic num_case_compute)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1318
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1319
lemma ADD_CLAUSES: "0 + (m::nat) = m &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1320
m + 0 = m & Suc m + (n::nat) = Suc (m + n) & m + Suc n = Suc (m + n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1321
  by (import arithmetic ADD_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1322
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1323
lemma LESS_ADD: "ALL (m::nat) n::nat. n < m --> (EX p::nat. p + n = m)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1324
  by (import arithmetic LESS_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1325
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1326
lemma LESS_ANTISYM: "ALL (m::nat) n::nat. ~ (m < n & n < m)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1327
  by (import arithmetic LESS_ANTISYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1328
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1329
lemma LESS_LESS_SUC: "ALL (x::nat) xa::nat. ~ (x < xa & xa < Suc x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1330
  by (import arithmetic LESS_LESS_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1331
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1332
lemma FUN_EQ_LEMMA: "ALL (f::'a::type => bool) (x1::'a::type) x2::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1333
   f x1 & ~ f x2 --> x1 ~= x2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1334
  by (import arithmetic FUN_EQ_LEMMA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1335
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1336
lemma LESS_NOT_SUC: "ALL (m::nat) n::nat. m < n & n ~= Suc m --> Suc m < n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1337
  by (import arithmetic LESS_NOT_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1338
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1339
lemma LESS_0_CASES: "ALL m::nat. 0 = m | 0 < m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1340
  by (import arithmetic LESS_0_CASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1341
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1342
lemma LESS_CASES_IMP: "ALL (m::nat) n::nat. ~ m < n & m ~= n --> n < m"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1343
  by (import arithmetic LESS_CASES_IMP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1344
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1345
lemma LESS_CASES: "ALL (m::nat) n::nat. m < n | n <= m"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1346
  by (import arithmetic LESS_CASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1347
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1348
lemma LESS_EQ_SUC_REFL: "ALL m::nat. m <= Suc m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1349
  by (import arithmetic LESS_EQ_SUC_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1350
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1351
lemma LESS_ADD_NONZERO: "ALL (m::nat) n::nat. n ~= 0 --> m < m + n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1352
  by (import arithmetic LESS_ADD_NONZERO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1353
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1354
lemma LESS_EQ_ANTISYM: "ALL (x::nat) xa::nat. ~ (x < xa & xa <= x)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1355
  by (import arithmetic LESS_EQ_ANTISYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1356
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1357
lemma SUB_0: "ALL m::nat. 0 - m = 0 & m - 0 = m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1358
  by (import arithmetic SUB_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1359
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1360
lemma SUC_SUB1: "ALL m::nat. Suc m - 1 = m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1361
  by (import arithmetic SUC_SUB1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1362
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1363
lemma PRE_SUB1: "ALL m::nat. PRE m = m - 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1364
  by (import arithmetic PRE_SUB1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1365
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1366
lemma MULT_CLAUSES: "ALL (x::nat) xa::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1367
   0 * x = 0 &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1368
   x * 0 = 0 &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1369
   1 * x = x &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1370
   x * 1 = x & Suc x * xa = x * xa + xa & x * Suc xa = x + x * xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1371
  by (import arithmetic MULT_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1372
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1373
lemma PRE_SUB: "ALL (m::nat) n::nat. PRE (m - n) = PRE m - n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1374
  by (import arithmetic PRE_SUB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1375
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1376
lemma ADD_EQ_1: "ALL (m::nat) n::nat. (m + n = 1) = (m = 1 & n = 0 | m = 0 & n = 1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1377
  by (import arithmetic ADD_EQ_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1378
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1379
lemma ADD_INV_0_EQ: "ALL (m::nat) n::nat. (m + n = m) = (n = 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1380
  by (import arithmetic ADD_INV_0_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1381
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1382
lemma PRE_SUC_EQ: "ALL (m::nat) n::nat. 0 < n --> (m = PRE n) = (Suc m = n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1383
  by (import arithmetic PRE_SUC_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1384
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1385
lemma INV_PRE_EQ: "ALL (m::nat) n::nat. 0 < m & 0 < n --> (PRE m = PRE n) = (m = n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1386
  by (import arithmetic INV_PRE_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1387
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1388
lemma LESS_SUC_NOT: "ALL (m::nat) n::nat. m < n --> ~ n < Suc m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1389
  by (import arithmetic LESS_SUC_NOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1390
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1391
lemma ADD_EQ_SUB: "ALL (m::nat) (n::nat) p::nat. n <= p --> (m + n = p) = (m = p - n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1392
  by (import arithmetic ADD_EQ_SUB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1393
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1394
lemma LESS_ADD_1: "ALL (x::nat) xa::nat. xa < x --> (EX xb::nat. x = xa + (xb + 1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1395
  by (import arithmetic LESS_ADD_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1396
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1397
lemma NOT_ODD_EQ_EVEN: "ALL (n::nat) m::nat. Suc (n + n) ~= m + m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1398
  by (import arithmetic NOT_ODD_EQ_EVEN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1399
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1400
lemma MULT_SUC_EQ: "ALL (p::nat) (m::nat) n::nat. (n * Suc p = m * Suc p) = (n = m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1401
  by (import arithmetic MULT_SUC_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1402
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1403
lemma MULT_EXP_MONO: "ALL (p::nat) (q::nat) (n::nat) m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1404
   (n * Suc q ^ p = m * Suc q ^ p) = (n = m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1405
  by (import arithmetic MULT_EXP_MONO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1406
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1407
lemma LESS_ADD_SUC: "ALL (m::nat) n::nat. m < m + Suc n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1408
  by (import arithmetic LESS_ADD_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1409
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1410
lemma LESS_OR_EQ_ADD: "ALL (n::nat) m::nat. n < m | (EX p::nat. n = p + m)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1411
  by (import arithmetic LESS_OR_EQ_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1412
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1413
lemma WOP: "(All::((nat => bool) => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1414
 (%P::nat => bool.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1415
     (op -->::bool => bool => bool) ((Ex::(nat => bool) => bool) P)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1416
      ((Ex::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1417
        (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1418
            (op &::bool => bool => bool) (P n)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1419
             ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1420
               (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1421
                   (op -->::bool => bool => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1422
                    ((op <::nat => nat => bool) m n)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  1423
                    ((Not::bool => bool) (P m)))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1424
  by (import arithmetic WOP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1425
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1426
lemma INV_PRE_LESS: "ALL m>0. ALL n::nat. (PRE m < PRE n) = (m < n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1427
  by (import arithmetic INV_PRE_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1428
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1429
lemma INV_PRE_LESS_EQ: "ALL n>0. ALL m::nat. (PRE m <= PRE n) = (m <= n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1430
  by (import arithmetic INV_PRE_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1431
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1432
lemma SUB_EQ_EQ_0: "ALL (m::nat) n::nat. (m - n = m) = (m = 0 | n = 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1433
  by (import arithmetic SUB_EQ_EQ_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1434
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1435
lemma SUB_LESS_OR: "ALL (m::nat) n::nat. n < m --> n <= m - 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1436
  by (import arithmetic SUB_LESS_OR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1437
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1438
lemma LESS_SUB_ADD_LESS: "ALL (n::nat) (m::nat) i::nat. i < n - m --> i + m < n"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1439
  by (import arithmetic LESS_SUB_ADD_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1440
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1441
lemma LESS_EQ_SUB_LESS: "ALL (x::nat) xa::nat. xa <= x --> (ALL c::nat. (x - xa < c) = (x < xa + c))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1442
  by (import arithmetic LESS_EQ_SUB_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1443
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1444
lemma NOT_SUC_LESS_EQ: "ALL (x::nat) xa::nat. (~ Suc x <= xa) = (xa <= x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1445
  by (import arithmetic NOT_SUC_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1446
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1447
lemma SUB_LESS_EQ_ADD: "ALL (m::nat) p::nat. m <= p --> (ALL n::nat. (p - m <= n) = (p <= m + n))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1448
  by (import arithmetic SUB_LESS_EQ_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1449
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1450
lemma SUB_CANCEL: "ALL (x::nat) (xa::nat) xb::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1451
   xa <= x & xb <= x --> (x - xa = x - xb) = (xa = xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1452
  by (import arithmetic SUB_CANCEL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1453
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1454
lemma NOT_EXP_0: "ALL (m::nat) n::nat. Suc n ^ m ~= 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1455
  by (import arithmetic NOT_EXP_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1456
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1457
lemma ZERO_LESS_EXP: "ALL (m::nat) n::nat. 0 < Suc n ^ m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1458
  by (import arithmetic ZERO_LESS_EXP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1459
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1460
lemma ODD_OR_EVEN: "ALL x::nat. EX xa::nat. x = Suc (Suc 0) * xa | x = Suc (Suc 0) * xa + 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1461
  by (import arithmetic ODD_OR_EVEN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1462
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1463
lemma LESS_EXP_SUC_MONO: "ALL (n::nat) m::nat. Suc (Suc m) ^ n < Suc (Suc m) ^ Suc n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1464
  by (import arithmetic LESS_EXP_SUC_MONO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1465
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1466
lemma LESS_LESS_CASES: "ALL (m::nat) n::nat. m = n | m < n | n < m"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1467
  by (import arithmetic LESS_LESS_CASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1468
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1469
lemma LESS_EQUAL_ADD: "ALL (m::nat) n::nat. m <= n --> (EX p::nat. n = m + p)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1470
  by (import arithmetic LESS_EQUAL_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1471
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1472
lemma MULT_EQ_1: "ALL (x::nat) y::nat. (x * y = 1) = (x = 1 & y = 1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1473
  by (import arithmetic MULT_EQ_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1474
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1475
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1476
  FACT :: "nat => nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1477
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1478
specification (FACT) FACT: "FACT 0 = 1 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1479
  by (import arithmetic FACT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1480
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1481
lemma FACT_LESS: "ALL n::nat. 0 < FACT n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1482
  by (import arithmetic FACT_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1483
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1484
lemma EVEN_ODD: "ALL n::nat. EVEN n = (~ ODD n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1485
  by (import arithmetic EVEN_ODD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1486
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1487
lemma ODD_EVEN: "ALL x::nat. ODD x = (~ EVEN x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1488
  by (import arithmetic ODD_EVEN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1489
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1490
lemma EVEN_OR_ODD: "ALL x::nat. EVEN x | ODD x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1491
  by (import arithmetic EVEN_OR_ODD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1492
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1493
lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1494
  by (import arithmetic EVEN_AND_ODD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1495
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1496
lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1497
  by (import arithmetic EVEN_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1498
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1499
lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1500
  by (import arithmetic EVEN_MULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1501
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1502
lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1503
  by (import arithmetic ODD_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1504
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1505
lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1506
  by (import arithmetic ODD_MULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1507
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1508
lemma EVEN_DOUBLE: "ALL n::nat. EVEN (2 * n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1509
  by (import arithmetic EVEN_DOUBLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1510
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1511
lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (2 * x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1512
  by (import arithmetic ODD_DOUBLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1513
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1514
lemma EVEN_ODD_EXISTS: "ALL x::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1515
   (EVEN x --> (EX m::nat. x = 2 * m)) &
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1516
   (ODD x --> (EX m::nat. x = Suc (2 * m)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1517
  by (import arithmetic EVEN_ODD_EXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1518
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1519
lemma EVEN_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = 2 * m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1520
  by (import arithmetic EVEN_EXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1521
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1522
lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc (2 * m))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1523
  by (import arithmetic ODD_EXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1524
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1525
lemma NOT_SUC_LESS_EQ_0: "ALL x::nat. ~ Suc x <= 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1526
  by (import arithmetic NOT_SUC_LESS_EQ_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1527
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1528
lemma NOT_LEQ: "ALL (x::nat) xa::nat. (~ x <= xa) = (Suc xa <= x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1529
  by (import arithmetic NOT_LEQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1530
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1531
lemma NOT_NUM_EQ: "ALL (x::nat) xa::nat. (x ~= xa) = (Suc x <= xa | Suc xa <= x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1532
  by (import arithmetic NOT_NUM_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1533
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1534
lemma NOT_GREATER_EQ: "ALL (x::nat) xa::nat. (~ xa <= x) = (Suc x <= xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1535
  by (import arithmetic NOT_GREATER_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1536
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1537
lemma SUC_ADD_SYM: "ALL (m::nat) n::nat. Suc (m + n) = Suc n + m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1538
  by (import arithmetic SUC_ADD_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1539
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1540
lemma NOT_SUC_ADD_LESS_EQ: "ALL (m::nat) n::nat. ~ Suc (m + n) <= m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1541
  by (import arithmetic NOT_SUC_ADD_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1542
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1543
lemma SUB_LEFT_ADD: "ALL (m::nat) (n::nat) p::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1544
   m + (n - p) = (if n <= p then m else m + n - p)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1545
  by (import arithmetic SUB_LEFT_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1546
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1547
lemma SUB_RIGHT_ADD: "ALL (m::nat) (n::nat) p::nat. m - n + p = (if m <= n then p else m + p - n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1548
  by (import arithmetic SUB_RIGHT_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1549
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1550
lemma SUB_LEFT_SUB: "ALL (m::nat) (n::nat) p::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1551
   m - (n - p) = (if n <= p then m else m + p - n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1552
  by (import arithmetic SUB_LEFT_SUB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1553
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1554
lemma SUB_LEFT_SUC: "ALL (m::nat) n::nat. Suc (m - n) = (if m <= n then Suc 0 else Suc m - n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1555
  by (import arithmetic SUB_LEFT_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1556
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1557
lemma SUB_LEFT_LESS_EQ: "ALL (m::nat) (n::nat) p::nat. (m <= n - p) = (m + p <= n | m <= 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1558
  by (import arithmetic SUB_LEFT_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1559
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1560
lemma SUB_RIGHT_LESS_EQ: "ALL (m::nat) (n::nat) p::nat. (m - n <= p) = (m <= n + p)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1561
  by (import arithmetic SUB_RIGHT_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1562
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1563
lemma SUB_RIGHT_LESS: "ALL (m::nat) (n::nat) p::nat. (m - n < p) = (m < n + p & 0 < p)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1564
  by (import arithmetic SUB_RIGHT_LESS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1565
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1566
lemma SUB_RIGHT_GREATER_EQ: "ALL (m::nat) (n::nat) p::nat. (p <= m - n) = (n + p <= m | p <= 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1567
  by (import arithmetic SUB_RIGHT_GREATER_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1568
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1569
lemma SUB_LEFT_GREATER: "ALL (m::nat) (n::nat) p::nat. (n - p < m) = (n < m + p & 0 < m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1570
  by (import arithmetic SUB_LEFT_GREATER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1571
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1572
lemma SUB_RIGHT_GREATER: "ALL (m::nat) (n::nat) p::nat. (p < m - n) = (n + p < m)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1573
  by (import arithmetic SUB_RIGHT_GREATER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1574
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1575
lemma SUB_LEFT_EQ: "ALL (m::nat) (n::nat) p::nat. (m = n - p) = (m + p = n | m <= 0 & n <= p)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1576
  by (import arithmetic SUB_LEFT_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1577
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1578
lemma SUB_RIGHT_EQ: "ALL (m::nat) (n::nat) p::nat. (m - n = p) = (m = n + p | m <= n & p <= 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1579
  by (import arithmetic SUB_RIGHT_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1580
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1581
lemma LE: "(ALL n::nat. (n <= 0) = (n = 0)) &
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1582
(ALL (m::nat) n::nat. (m <= Suc n) = (m = Suc n | m <= n))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1583
  by (import arithmetic LE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1584
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1585
lemma DA: "ALL (k::nat) n::nat. 0 < n --> (EX (x::nat) q::nat. k = q * n + x & x < n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1586
  by (import arithmetic DA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1587
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1588
lemma DIV_LESS_EQ: "ALL n>0. ALL k::nat. k div n <= k"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1589
  by (import arithmetic DIV_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1590
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1591
lemma DIV_UNIQUE: "ALL (n::nat) (k::nat) q::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1592
   (EX r::nat. k = q * n + r & r < n) --> k div n = q"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1593
  by (import arithmetic DIV_UNIQUE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1594
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1595
lemma MOD_UNIQUE: "ALL (n::nat) (k::nat) r::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1596
   (EX q::nat. k = q * n + r & r < n) --> k mod n = r"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1597
  by (import arithmetic MOD_UNIQUE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1598
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1599
lemma DIV_MULT: "ALL (n::nat) r::nat. r < n --> (ALL q::nat. (q * n + r) div n = q)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1600
  by (import arithmetic DIV_MULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1601
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1602
lemma MOD_EQ_0: "ALL n>0. ALL k::nat. k * n mod n = 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1603
  by (import arithmetic MOD_EQ_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1604
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1605
lemma ZERO_MOD: "(All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1606
 (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1607
     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1608
      ((op =::nat => nat => bool) ((op mod::nat => nat => nat) (0::nat) n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1609
        (0::nat)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1610
  by (import arithmetic ZERO_MOD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1611
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1612
lemma ZERO_DIV: "(All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1613
 (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1614
     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1615
      ((op =::nat => nat => bool) ((op div::nat => nat => nat) (0::nat) n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1616
        (0::nat)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1617
  by (import arithmetic ZERO_DIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1618
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1619
lemma MOD_MULT: "ALL (n::nat) r::nat. r < n --> (ALL q::nat. (q * n + r) mod n = r)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1620
  by (import arithmetic MOD_MULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1621
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1622
lemma MOD_TIMES: "ALL n>0. ALL (q::nat) r::nat. (q * n + r) mod n = r mod n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1623
  by (import arithmetic MOD_TIMES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1624
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1625
lemma MOD_PLUS: "ALL n>0. ALL (j::nat) k::nat. (j mod n + k mod n) mod n = (j + k) mod n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1626
  by (import arithmetic MOD_PLUS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1627
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1628
lemma MOD_MOD: "ALL n>0. ALL k::nat. k mod n mod n = k mod n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1629
  by (import arithmetic MOD_MOD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1630
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1631
lemma ADD_DIV_ADD_DIV: "ALL x>0. ALL (xa::nat) r::nat. (xa * x + r) div x = xa + r div x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1632
  by (import arithmetic ADD_DIV_ADD_DIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1633
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1634
lemma MOD_MULT_MOD: "ALL (m::nat) n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1635
   0 < n & 0 < m --> (ALL x::nat. x mod (n * m) mod n = x mod n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1636
  by (import arithmetic MOD_MULT_MOD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1637
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1638
lemma DIVMOD_ID: "(All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1639
 (%n::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1640
     (op -->::bool => bool => bool) ((op <::nat => nat => bool) (0::nat) n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1641
      ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1642
        ((op =::nat => nat => bool) ((op div::nat => nat => nat) n n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1643
          (1::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1644
        ((op =::nat => nat => bool) ((op mod::nat => nat => nat) n n)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1645
          (0::nat))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1646
  by (import arithmetic DIVMOD_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1647
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1648
lemma DIV_DIV_DIV_MULT: "ALL (x::nat) xa::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1649
   0 < x & 0 < xa --> (ALL xb::nat. xb div x div xa = xb div (x * xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1650
  by (import arithmetic DIV_DIV_DIV_MULT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1651
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1652
lemma DIV_P: "ALL (P::nat => bool) (p::nat) q::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1653
   0 < q --> P (p div q) = (EX (k::nat) r::nat. p = k * q + r & r < q & P k)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1654
  by (import arithmetic DIV_P)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1655
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1656
lemma MOD_P: "ALL (P::nat => bool) (p::nat) q::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1657
   0 < q --> P (p mod q) = (EX (k::nat) r::nat. p = k * q + r & r < q & P r)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1658
  by (import arithmetic MOD_P)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1659
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1660
lemma MOD_TIMES2: "ALL n>0. ALL (j::nat) k::nat. j mod n * (k mod n) mod n = j * k mod n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1661
  by (import arithmetic MOD_TIMES2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1662
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1663
lemma MOD_COMMON_FACTOR: "ALL (n::nat) (p::nat) q::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1664
   0 < n & 0 < q --> n * (p mod q) = n * p mod (n * q)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1665
  by (import arithmetic MOD_COMMON_FACTOR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1666
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1667
lemma num_case_cong: "ALL (M::nat) (M'::nat) (b::'a::type) f::nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1668
   M = M' &
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1669
   (M' = 0 --> b = (b'::'a::type)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1670
   (ALL n::nat. M' = Suc n --> f n = (f'::nat => 'a::type) n) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1671
   nat_case b f M = nat_case b' f' M'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1672
  by (import arithmetic num_case_cong)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1673
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1674
lemma SUC_ELIM_THM: "ALL P::nat => nat => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1675
   (ALL n::nat. P (Suc n) n) = (ALL n>0. P n (n - 1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1676
  by (import arithmetic SUC_ELIM_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1677
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1678
lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1679
(ALL x::nat. (b = a + x --> P 0) & (a = b + x --> P x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1680
  by (import arithmetic SUB_ELIM_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1681
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1682
lemma PRE_ELIM_THM: "(P::nat => bool) (PRE (n::nat)) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1683
(ALL m::nat. (n = 0 --> P 0) & (n = Suc m --> P m))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1684
  by (import arithmetic PRE_ELIM_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1685
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1686
lemma MULT_INCREASES: "ALL (m::nat) n::nat. 1 < m & 0 < n --> Suc n <= m * n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1687
  by (import arithmetic MULT_INCREASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1688
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1689
lemma EXP_ALWAYS_BIG_ENOUGH: "ALL b>1. ALL n::nat. EX m::nat. n <= b ^ m"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1690
  by (import arithmetic EXP_ALWAYS_BIG_ENOUGH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1691
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1692
lemma EXP_EQ_0: "ALL (n::nat) m::nat. (n ^ m = 0) = (n = 0 & 0 < m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1693
  by (import arithmetic EXP_EQ_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1694
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1695
lemma EXP_1: "(All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1696
 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1697
     (op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1698
      ((op =::nat => nat => bool) ((op ^::nat => nat => nat) (1::nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1699
        (1::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1700
      ((op =::nat => nat => bool) ((op ^::nat => nat => nat) x (1::nat)) x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1701
  by (import arithmetic EXP_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1702
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1703
lemma EXP_EQ_1: "ALL (n::nat) m::nat. (n ^ m = 1) = (n = 1 | m = 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1704
  by (import arithmetic EXP_EQ_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1705
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1706
lemma MIN_MAX_EQ: "ALL (x::nat) xa::nat. (min x xa = max x xa) = (x = xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1707
  by (import arithmetic MIN_MAX_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1708
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1709
lemma MIN_MAX_LT: "ALL (x::nat) xa::nat. (min x xa < max x xa) = (x ~= xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1710
  by (import arithmetic MIN_MAX_LT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1711
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1712
lemma MIN_MAX_PRED: "ALL (P::nat => bool) (m::nat) n::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1713
   P m & P n --> P (min m n) & P (max m n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1714
  by (import arithmetic MIN_MAX_PRED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1715
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1716
lemma MIN_LT: "ALL (x::nat) xa::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1717
   (min xa x < xa) = (xa ~= x & min xa x = x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1718
   (min xa x < x) = (xa ~= x & min xa x = xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1719
   (xa < min xa x) = False & (x < min xa x) = False"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1720
  by (import arithmetic MIN_LT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1721
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1722
lemma MAX_LT: "ALL (x::nat) xa::nat.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1723
   (xa < max xa x) = (xa ~= x & max xa x = x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1724
   (x < max xa x) = (xa ~= x & max xa x = xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1725
   (max xa x < xa) = False & (max xa x < x) = False"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1726
  by (import arithmetic MAX_LT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1727
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1728
lemma MIN_LE: "ALL (x::nat) xa::nat. min xa x <= xa & min xa x <= x"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1729
  by (import arithmetic MIN_LE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1730
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1731
lemma MAX_LE: "ALL (x::nat) xa::nat. xa <= max xa x & x <= max xa x"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1732
  by (import arithmetic MAX_LE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1733
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1734
lemma MIN_0: "ALL x::nat. min x 0 = 0 & min 0 x = 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1735
  by (import arithmetic MIN_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1736
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1737
lemma MAX_0: "ALL x::nat. max x 0 = x & max 0 x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1738
  by (import arithmetic MAX_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1739
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1740
lemma EXISTS_GREATEST: "ALL P::nat => bool.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1741
   (Ex P & (EX x::nat. ALL y::nat. x < y --> ~ P y)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1742
   (EX x::nat. P x & (ALL y::nat. x < y --> ~ P y))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1743
  by (import arithmetic EXISTS_GREATEST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1744
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1745
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1746
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1747
;setup_theory hrat
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1748
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1749
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1750
  trat_1 :: "nat * nat" 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1751
  "trat_1 == (0, 0)"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1752
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1753
lemma trat_1: "trat_1 = (0, 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1754
  by (import hrat trat_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1755
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1756
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1757
  trat_inv :: "nat * nat => nat * nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1758
  "trat_inv == %(x::nat, y::nat). (y, x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1759
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1760
lemma trat_inv: "ALL (x::nat) y::nat. trat_inv (x, y) = (y, x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1761
  by (import hrat trat_inv)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1762
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1763
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1764
  trat_add :: "nat * nat => nat * nat => nat * nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1765
  "trat_add ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1766
%(x::nat, y::nat) (x'::nat, y'::nat).
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1767
   (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1768
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1769
lemma trat_add: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1770
   trat_add (x, y) (x', y') =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1771
   (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1772
  by (import hrat trat_add)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1773
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1774
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1775
  trat_mul :: "nat * nat => nat * nat => nat * nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1776
  "trat_mul ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1777
%(x::nat, y::nat) (x'::nat, y'::nat).
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1778
   (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1779
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1780
lemma trat_mul: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1781
   trat_mul (x, y) (x', y') = (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1782
  by (import hrat trat_mul)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1783
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1784
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1785
  trat_sucint :: "nat => nat * nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1786
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1787
specification (trat_sucint) trat_sucint: "trat_sucint 0 = trat_1 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1788
(ALL n::nat. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1789
  by (import hrat trat_sucint)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1790
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1791
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1792
  trat_eq :: "nat * nat => nat * nat => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1793
  "trat_eq ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1794
%(x::nat, y::nat) (x'::nat, y'::nat). Suc x * Suc y' = Suc x' * Suc y"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1795
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1796
lemma trat_eq: "ALL (x::nat) (y::nat) (x'::nat) y'::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1797
   trat_eq (x, y) (x', y') = (Suc x * Suc y' = Suc x' * Suc y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1798
  by (import hrat trat_eq)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1799
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1800
lemma TRAT_EQ_REFL: "ALL p::nat * nat. trat_eq p p"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1801
  by (import hrat TRAT_EQ_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1802
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1803
lemma TRAT_EQ_SYM: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = trat_eq q p"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1804
  by (import hrat TRAT_EQ_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1805
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1806
lemma TRAT_EQ_TRANS: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1807
   trat_eq p q & trat_eq q r --> trat_eq p r"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1808
  by (import hrat TRAT_EQ_TRANS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1809
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1810
lemma TRAT_EQ_AP: "ALL (p::nat * nat) q::nat * nat. p = q --> trat_eq p q"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1811
  by (import hrat TRAT_EQ_AP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1812
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1813
lemma TRAT_ADD_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_add h i = trat_add i h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1814
  by (import hrat TRAT_ADD_SYM_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1815
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1816
lemma TRAT_MUL_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_mul h i = trat_mul i h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1817
  by (import hrat TRAT_MUL_SYM_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1818
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1819
lemma TRAT_INV_WELLDEFINED: "ALL (p::nat * nat) q::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1820
   trat_eq p q --> trat_eq (trat_inv p) (trat_inv q)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1821
  by (import hrat TRAT_INV_WELLDEFINED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1822
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1823
lemma TRAT_ADD_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1824
   trat_eq p q --> trat_eq (trat_add p r) (trat_add q r)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1825
  by (import hrat TRAT_ADD_WELLDEFINED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1826
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1827
lemma TRAT_ADD_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1828
   trat_eq p1 p2 & trat_eq q1 q2 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1829
   trat_eq (trat_add p1 q1) (trat_add p2 q2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1830
  by (import hrat TRAT_ADD_WELLDEFINED2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1831
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1832
lemma TRAT_MUL_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1833
   trat_eq p q --> trat_eq (trat_mul p r) (trat_mul q r)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1834
  by (import hrat TRAT_MUL_WELLDEFINED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1835
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1836
lemma TRAT_MUL_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1837
   trat_eq p1 p2 & trat_eq q1 q2 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1838
   trat_eq (trat_mul p1 q1) (trat_mul p2 q2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1839
  by (import hrat TRAT_MUL_WELLDEFINED2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1840
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1841
lemma TRAT_ADD_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_add h i) (trat_add i h)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1842
  by (import hrat TRAT_ADD_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1843
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1844
lemma TRAT_ADD_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1845
   trat_eq (trat_add h (trat_add i j)) (trat_add (trat_add h i) j)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1846
  by (import hrat TRAT_ADD_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1847
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1848
lemma TRAT_MUL_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_mul h i) (trat_mul i h)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1849
  by (import hrat TRAT_MUL_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1850
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1851
lemma TRAT_MUL_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1852
   trat_eq (trat_mul h (trat_mul i j)) (trat_mul (trat_mul h i) j)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1853
  by (import hrat TRAT_MUL_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1854
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1855
lemma TRAT_LDISTRIB: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1856
   trat_eq (trat_mul h (trat_add i j))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1857
    (trat_add (trat_mul h i) (trat_mul h j))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1858
  by (import hrat TRAT_LDISTRIB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1859
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1860
lemma TRAT_MUL_LID: "ALL h::nat * nat. trat_eq (trat_mul trat_1 h) h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1861
  by (import hrat TRAT_MUL_LID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1862
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1863
lemma TRAT_MUL_LINV: "ALL h::nat * nat. trat_eq (trat_mul (trat_inv h) h) trat_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1864
  by (import hrat TRAT_MUL_LINV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1865
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1866
lemma TRAT_NOZERO: "ALL (h::nat * nat) i::nat * nat. ~ trat_eq (trat_add h i) h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1867
  by (import hrat TRAT_NOZERO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1868
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1869
lemma TRAT_ADD_TOTAL: "ALL (h::nat * nat) i::nat * nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1870
   trat_eq h i |
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1871
   (EX d::nat * nat. trat_eq h (trat_add i d)) |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1872
   (EX d::nat * nat. trat_eq i (trat_add h d))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1873
  by (import hrat TRAT_ADD_TOTAL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1874
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1875
lemma TRAT_SUCINT_0: "ALL n::nat. trat_eq (trat_sucint n) (n, 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1876
  by (import hrat TRAT_SUCINT_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1877
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1878
lemma TRAT_ARCH: "ALL h::nat * nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1879
   EX (n::nat) d::nat * nat. trat_eq (trat_sucint n) (trat_add h d)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1880
  by (import hrat TRAT_ARCH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1881
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1882
lemma TRAT_SUCINT: "trat_eq (trat_sucint 0) trat_1 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1883
(ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1884
    trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1885
  by (import hrat TRAT_SUCINT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1886
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1887
lemma TRAT_EQ_EQUIV: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = (trat_eq p = trat_eq q)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1888
  by (import hrat TRAT_EQ_EQUIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1889
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1890
typedef (open) hrat = "{x::nat * nat => bool. EX xa::nat * nat. x = trat_eq xa}" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1891
  by (rule typedef_helper,import hrat hrat_TY_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1892
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1893
lemmas hrat_TY_DEF = typedef_hol2hol4 [OF type_definition_hrat]
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1894
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1895
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1896
  mk_hrat :: "(nat * nat => bool) => hrat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1897
  dest_hrat :: "hrat => nat * nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1898
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1899
specification (dest_hrat mk_hrat) hrat_tybij: "(ALL a::hrat. mk_hrat (dest_hrat a) = a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1900
(ALL r::nat * nat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1901
    (EX x::nat * nat. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1902
  by (import hrat hrat_tybij)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1903
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1904
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1905
  hrat_1 :: "hrat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1906
  "hrat_1 == mk_hrat (trat_eq trat_1)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1907
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1908
lemma hrat_1: "hrat_1 = mk_hrat (trat_eq trat_1)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1909
  by (import hrat hrat_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1910
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1911
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1912
  hrat_inv :: "hrat => hrat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1913
  "hrat_inv == %T1::hrat. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1914
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1915
lemma hrat_inv: "ALL T1::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1916
   hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1917
  by (import hrat hrat_inv)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1918
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1919
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1920
  hrat_add :: "hrat => hrat => hrat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1921
  "hrat_add ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1922
%(T1::hrat) T2::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1923
   mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1924
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1925
lemma hrat_add: "ALL (T1::hrat) T2::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1926
   hrat_add T1 T2 =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1927
   mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1928
  by (import hrat hrat_add)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1929
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1930
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1931
  hrat_mul :: "hrat => hrat => hrat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1932
  "hrat_mul ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1933
%(T1::hrat) T2::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1934
   mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1935
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1936
lemma hrat_mul: "ALL (T1::hrat) T2::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1937
   hrat_mul T1 T2 =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1938
   mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1939
  by (import hrat hrat_mul)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1940
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1941
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1942
  hrat_sucint :: "nat => hrat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1943
  "hrat_sucint == %T1::nat. mk_hrat (trat_eq (trat_sucint T1))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1944
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1945
lemma hrat_sucint: "ALL T1::nat. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1946
  by (import hrat hrat_sucint)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1947
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1948
lemma HRAT_ADD_SYM: "ALL (h::hrat) i::hrat. hrat_add h i = hrat_add i h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1949
  by (import hrat HRAT_ADD_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1950
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1951
lemma HRAT_ADD_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1952
   hrat_add h (hrat_add i j) = hrat_add (hrat_add h i) j"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1953
  by (import hrat HRAT_ADD_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1954
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1955
lemma HRAT_MUL_SYM: "ALL (h::hrat) i::hrat. hrat_mul h i = hrat_mul i h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1956
  by (import hrat HRAT_MUL_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1957
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1958
lemma HRAT_MUL_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1959
   hrat_mul h (hrat_mul i j) = hrat_mul (hrat_mul h i) j"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1960
  by (import hrat HRAT_MUL_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1961
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1962
lemma HRAT_LDISTRIB: "ALL (h::hrat) (i::hrat) j::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1963
   hrat_mul h (hrat_add i j) = hrat_add (hrat_mul h i) (hrat_mul h j)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1964
  by (import hrat HRAT_LDISTRIB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1965
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1966
lemma HRAT_MUL_LID: "ALL h::hrat. hrat_mul hrat_1 h = h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1967
  by (import hrat HRAT_MUL_LID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1968
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1969
lemma HRAT_MUL_LINV: "ALL h::hrat. hrat_mul (hrat_inv h) h = hrat_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1970
  by (import hrat HRAT_MUL_LINV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1971
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1972
lemma HRAT_NOZERO: "ALL (h::hrat) i::hrat. hrat_add h i ~= h"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1973
  by (import hrat HRAT_NOZERO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1974
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1975
lemma HRAT_ADD_TOTAL: "ALL (h::hrat) i::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1976
   h = i | (EX x::hrat. h = hrat_add i x) | (EX x::hrat. i = hrat_add h x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1977
  by (import hrat HRAT_ADD_TOTAL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1978
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1979
lemma HRAT_ARCH: "ALL h::hrat. EX (x::nat) xa::hrat. hrat_sucint x = hrat_add h xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1980
  by (import hrat HRAT_ARCH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1981
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  1982
lemma HRAT_SUCINT: "hrat_sucint 0 = hrat_1 &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1983
(ALL x::nat. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1984
  by (import hrat HRAT_SUCINT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1985
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1986
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1987
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1988
;setup_theory hreal
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1989
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1990
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1991
  hrat_lt :: "hrat => hrat => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1992
  "hrat_lt == %(x::hrat) y::hrat. EX d::hrat. y = hrat_add x d"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1993
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1994
lemma hrat_lt: "ALL (x::hrat) y::hrat. hrat_lt x y = (EX d::hrat. y = hrat_add x d)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1995
  by (import hreal hrat_lt)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1996
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  1997
lemma HRAT_LT_REFL: "ALL x::hrat. ~ hrat_lt x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1998
  by (import hreal HRAT_LT_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1999
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2000
lemma HRAT_LT_TRANS: "ALL (x::hrat) (y::hrat) z::hrat. hrat_lt x y & hrat_lt y z --> hrat_lt x z"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2001
  by (import hreal HRAT_LT_TRANS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2002
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2003
lemma HRAT_LT_ANTISYM: "ALL (x::hrat) y::hrat. ~ (hrat_lt x y & hrat_lt y x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2004
  by (import hreal HRAT_LT_ANTISYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2005
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2006
lemma HRAT_LT_TOTAL: "ALL (x::hrat) y::hrat. x = y | hrat_lt x y | hrat_lt y x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2007
  by (import hreal HRAT_LT_TOTAL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2008
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2009
lemma HRAT_MUL_RID: "ALL x::hrat. hrat_mul x hrat_1 = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2010
  by (import hreal HRAT_MUL_RID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2011
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2012
lemma HRAT_MUL_RINV: "ALL x::hrat. hrat_mul x (hrat_inv x) = hrat_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2013
  by (import hreal HRAT_MUL_RINV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2014
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2015
lemma HRAT_RDISTRIB: "ALL (x::hrat) (y::hrat) z::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2016
   hrat_mul (hrat_add x y) z = hrat_add (hrat_mul x z) (hrat_mul y z)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2017
  by (import hreal HRAT_RDISTRIB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2018
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2019
lemma HRAT_LT_ADDL: "ALL (x::hrat) y::hrat. hrat_lt x (hrat_add x y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2020
  by (import hreal HRAT_LT_ADDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2021
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2022
lemma HRAT_LT_ADDR: "ALL (x::hrat) xa::hrat. hrat_lt xa (hrat_add x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2023
  by (import hreal HRAT_LT_ADDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2024
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2025
lemma HRAT_LT_GT: "ALL (x::hrat) y::hrat. hrat_lt x y --> ~ hrat_lt y x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2026
  by (import hreal HRAT_LT_GT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2027
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2028
lemma HRAT_LT_NE: "ALL (x::hrat) y::hrat. hrat_lt x y --> x ~= y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2029
  by (import hreal HRAT_LT_NE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2030
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2031
lemma HRAT_EQ_LADD: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_add x y = hrat_add x z) = (y = z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2032
  by (import hreal HRAT_EQ_LADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2033
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2034
lemma HRAT_EQ_LMUL: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_mul x y = hrat_mul x z) = (y = z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2035
  by (import hreal HRAT_EQ_LMUL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2036
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2037
lemma HRAT_LT_ADD2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2038
   hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_add u v) (hrat_add x y)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2039
  by (import hreal HRAT_LT_ADD2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2040
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2041
lemma HRAT_LT_LADD: "ALL (x::hrat) (y::hrat) z::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2042
   hrat_lt (hrat_add z x) (hrat_add z y) = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2043
  by (import hreal HRAT_LT_LADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2044
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2045
lemma HRAT_LT_RADD: "ALL (x::hrat) (y::hrat) z::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2046
   hrat_lt (hrat_add x z) (hrat_add y z) = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2047
  by (import hreal HRAT_LT_RADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2048
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2049
lemma HRAT_LT_MUL2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2050
   hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_mul u v) (hrat_mul x y)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2051
  by (import hreal HRAT_LT_MUL2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2052
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2053
lemma HRAT_LT_LMUL: "ALL (x::hrat) (y::hrat) z::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2054
   hrat_lt (hrat_mul z x) (hrat_mul z y) = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2055
  by (import hreal HRAT_LT_LMUL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2056
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2057
lemma HRAT_LT_RMUL: "ALL (x::hrat) (y::hrat) z::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2058
   hrat_lt (hrat_mul x z) (hrat_mul y z) = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2059
  by (import hreal HRAT_LT_RMUL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2060
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2061
lemma HRAT_LT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2062
  by (import hreal HRAT_LT_LMUL1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2063
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2064
lemma HRAT_LT_RMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2065
  by (import hreal HRAT_LT_RMUL1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2066
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2067
lemma HRAT_GT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2068
  by (import hreal HRAT_GT_LMUL1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2069
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2070
lemma HRAT_LT_L1: "ALL (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2071
   hrat_lt (hrat_mul (hrat_inv x) y) hrat_1 = hrat_lt y x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2072
  by (import hreal HRAT_LT_L1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2073
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2074
lemma HRAT_LT_R1: "ALL (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2075
   hrat_lt (hrat_mul x (hrat_inv y)) hrat_1 = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2076
  by (import hreal HRAT_LT_R1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2077
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2078
lemma HRAT_GT_L1: "ALL (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2079
   hrat_lt hrat_1 (hrat_mul (hrat_inv x) y) = hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2080
  by (import hreal HRAT_GT_L1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2081
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2082
lemma HRAT_INV_MUL: "ALL (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2083
   hrat_inv (hrat_mul x y) = hrat_mul (hrat_inv x) (hrat_inv y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2084
  by (import hreal HRAT_INV_MUL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2085
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2086
lemma HRAT_UP: "ALL x::hrat. Ex (hrat_lt x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2087
  by (import hreal HRAT_UP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2088
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2089
lemma HRAT_DOWN: "ALL x::hrat. EX xa::hrat. hrat_lt xa x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2090
  by (import hreal HRAT_DOWN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2091
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2092
lemma HRAT_DOWN2: "ALL (x::hrat) y::hrat. EX xa::hrat. hrat_lt xa x & hrat_lt xa y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2093
  by (import hreal HRAT_DOWN2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2094
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2095
lemma HRAT_MEAN: "ALL (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2096
   hrat_lt x y --> (EX xa::hrat. hrat_lt x xa & hrat_lt xa y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2097
  by (import hreal HRAT_MEAN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2098
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2099
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2100
  isacut :: "(hrat => bool) => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2101
  "isacut ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2102
%C::hrat => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2103
   Ex C &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2104
   (EX x::hrat. ~ C x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2105
   (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2106
   (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2107
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2108
lemma isacut: "ALL C::hrat => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2109
   isacut C =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2110
   (Ex C &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2111
    (EX x::hrat. ~ C x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2112
    (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2113
    (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2114
  by (import hreal isacut)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2115
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2116
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2117
  cut_of_hrat :: "hrat => hrat => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2118
  "cut_of_hrat == %(x::hrat) y::hrat. hrat_lt y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2119
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2120
lemma cut_of_hrat: "ALL x::hrat. cut_of_hrat x = (%y::hrat. hrat_lt y x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2121
  by (import hreal cut_of_hrat)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2122
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2123
lemma ISACUT_HRAT: "ALL h::hrat. isacut (cut_of_hrat h)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2124
  by (import hreal ISACUT_HRAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2125
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2126
typedef (open) hreal = "Collect isacut" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2127
  by (rule typedef_helper,import hreal hreal_TY_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2128
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2129
lemmas hreal_TY_DEF = typedef_hol2hol4 [OF type_definition_hreal]
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2130
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2131
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2132
  hreal :: "(hrat => bool) => hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2133
  cut :: "hreal => hrat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2134
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2135
specification (cut hreal) hreal_tybij: "(ALL a::hreal. hreal (hreal.cut a) = a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2136
(ALL r::hrat => bool. isacut r = (hreal.cut (hreal r) = r))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2137
  by (import hreal hreal_tybij)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2138
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2139
lemma EQUAL_CUTS: "ALL (X::hreal) Y::hreal. hreal.cut X = hreal.cut Y --> X = Y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2140
  by (import hreal EQUAL_CUTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2141
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2142
lemma CUT_ISACUT: "ALL x::hreal. isacut (hreal.cut x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2143
  by (import hreal CUT_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2144
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2145
lemma CUT_NONEMPTY: "ALL x::hreal. Ex (hreal.cut x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2146
  by (import hreal CUT_NONEMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2147
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2148
lemma CUT_BOUNDED: "ALL x::hreal. EX xa::hrat. ~ hreal.cut x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2149
  by (import hreal CUT_BOUNDED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2150
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2151
lemma CUT_DOWN: "ALL (x::hreal) (xa::hrat) xb::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2152
   hreal.cut x xa & hrat_lt xb xa --> hreal.cut x xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2153
  by (import hreal CUT_DOWN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2154
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2155
lemma CUT_UP: "ALL (x::hreal) xa::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2156
   hreal.cut x xa --> (EX y::hrat. hreal.cut x y & hrat_lt xa y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2157
  by (import hreal CUT_UP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2158
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2159
lemma CUT_UBOUND: "ALL (x::hreal) (xa::hrat) xb::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2160
   ~ hreal.cut x xa & hrat_lt xa xb --> ~ hreal.cut x xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2161
  by (import hreal CUT_UBOUND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2162
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2163
lemma CUT_STRADDLE: "ALL (X::hreal) (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2164
   hreal.cut X x & ~ hreal.cut X y --> hrat_lt x y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2165
  by (import hreal CUT_STRADDLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2166
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2167
lemma CUT_NEARTOP_ADD: "ALL (X::hreal) e::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2168
   EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_add x e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2169
  by (import hreal CUT_NEARTOP_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2170
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2171
lemma CUT_NEARTOP_MUL: "ALL (X::hreal) u::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2172
   hrat_lt hrat_1 u -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2173
   (EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2174
  by (import hreal CUT_NEARTOP_MUL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2175
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2176
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2177
  hreal_1 :: "hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2178
  "hreal_1 == hreal (cut_of_hrat hrat_1)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2179
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2180
lemma hreal_1: "hreal_1 = hreal (cut_of_hrat hrat_1)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2181
  by (import hreal hreal_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2182
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2183
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2184
  hreal_add :: "hreal => hreal => hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2185
  "hreal_add ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2186
%(X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2187
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2188
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2189
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2190
           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2191
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2192
lemma hreal_add: "ALL (X::hreal) Y::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2193
   hreal_add X Y =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2194
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2195
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2196
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2197
           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2198
  by (import hreal hreal_add)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2199
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2200
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2201
  hreal_mul :: "hreal => hreal => hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2202
  "hreal_mul ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2203
%(X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2204
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2205
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2206
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2207
           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2208
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2209
lemma hreal_mul: "ALL (X::hreal) Y::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2210
   hreal_mul X Y =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2211
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2212
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2213
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2214
           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2215
  by (import hreal hreal_mul)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2216
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2217
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2218
  hreal_inv :: "hreal => hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2219
  "hreal_inv ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2220
%X::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2221
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2222
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2223
        EX d::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2224
           hrat_lt d hrat_1 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2225
           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2226
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2227
lemma hreal_inv: "ALL X::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2228
   hreal_inv X =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2229
   hreal
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2230
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2231
        EX d::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2232
           hrat_lt d hrat_1 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2233
           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2234
  by (import hreal hreal_inv)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2235
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2236
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2237
  hreal_sup :: "(hreal => bool) => hreal" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2238
  "hreal_sup ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2239
%P::hreal => bool. hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2240
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2241
lemma hreal_sup: "ALL P::hreal => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2242
   hreal_sup P = hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2243
  by (import hreal hreal_sup)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2244
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2245
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2246
  hreal_lt :: "hreal => hreal => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2247
  "hreal_lt ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2248
%(X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2249
   X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2250
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2251
lemma hreal_lt: "ALL (X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2252
   hreal_lt X Y = (X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2253
  by (import hreal hreal_lt)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2254
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2255
lemma HREAL_INV_ISACUT: "ALL X::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2256
   isacut
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2257
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2258
        EX d::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2259
           hrat_lt d hrat_1 &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2260
           (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2261
  by (import hreal HREAL_INV_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2262
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2263
lemma HREAL_ADD_ISACUT: "ALL (X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2264
   isacut
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2265
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2266
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2267
           w = hrat_add x y & hreal.cut X x & hreal.cut Y y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2268
  by (import hreal HREAL_ADD_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2269
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2270
lemma HREAL_MUL_ISACUT: "ALL (X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2271
   isacut
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2272
    (%w::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2273
        EX (x::hrat) y::hrat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2274
           w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2275
  by (import hreal HREAL_MUL_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2276
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2277
lemma HREAL_ADD_SYM: "ALL (X::hreal) Y::hreal. hreal_add X Y = hreal_add Y X"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2278
  by (import hreal HREAL_ADD_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2279
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2280
lemma HREAL_MUL_SYM: "ALL (X::hreal) Y::hreal. hreal_mul X Y = hreal_mul Y X"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2281
  by (import hreal HREAL_MUL_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2282
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2283
lemma HREAL_ADD_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2284
   hreal_add X (hreal_add Y Z) = hreal_add (hreal_add X Y) Z"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2285
  by (import hreal HREAL_ADD_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2286
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2287
lemma HREAL_MUL_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2288
   hreal_mul X (hreal_mul Y Z) = hreal_mul (hreal_mul X Y) Z"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2289
  by (import hreal HREAL_MUL_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2290
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2291
lemma HREAL_LDISTRIB: "ALL (X::hreal) (Y::hreal) Z::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2292
   hreal_mul X (hreal_add Y Z) = hreal_add (hreal_mul X Y) (hreal_mul X Z)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2293
  by (import hreal HREAL_LDISTRIB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2294
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2295
lemma HREAL_MUL_LID: "ALL X::hreal. hreal_mul hreal_1 X = X"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2296
  by (import hreal HREAL_MUL_LID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2297
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2298
lemma HREAL_MUL_LINV: "ALL X::hreal. hreal_mul (hreal_inv X) X = hreal_1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2299
  by (import hreal HREAL_MUL_LINV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2300
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2301
lemma HREAL_NOZERO: "ALL (X::hreal) Y::hreal. hreal_add X Y ~= X"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2302
  by (import hreal HREAL_NOZERO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2303
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2304
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2305
  hreal_sub :: "hreal => hreal => hreal" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2306
  "hreal_sub ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2307
%(Y::hreal) X::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2308
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2309
    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2310
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2311
lemma hreal_sub: "ALL (Y::hreal) X::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2312
   hreal_sub Y X =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2313
   hreal
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2314
    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2315
  by (import hreal hreal_sub)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2316
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2317
lemma HREAL_LT_LEMMA: "ALL (X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2318
   hreal_lt X Y --> (EX x::hrat. ~ hreal.cut X x & hreal.cut Y x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2319
  by (import hreal HREAL_LT_LEMMA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2320
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2321
lemma HREAL_SUB_ISACUT: "ALL (X::hreal) Y::hreal.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2322
   hreal_lt X Y -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2323
   isacut
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2324
    (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2325
  by (import hreal HREAL_SUB_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2326
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2327
lemma HREAL_SUB_ADD: "ALL (X::hreal) Y::hreal. hreal_lt X Y --> hreal_add (hreal_sub Y X) X = Y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2328
  by (import hreal HREAL_SUB_ADD)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2329
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2330
lemma HREAL_LT_TOTAL: "ALL (X::hreal) Y::hreal. X = Y | hreal_lt X Y | hreal_lt Y X"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2331
  by (import hreal HREAL_LT_TOTAL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2332
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2333
lemma HREAL_LT: "ALL (X::hreal) Y::hreal. hreal_lt X Y = (EX D::hreal. Y = hreal_add X D)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2334
  by (import hreal HREAL_LT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2335
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2336
lemma HREAL_ADD_TOTAL: "ALL (X::hreal) Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2337
   X = Y |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2338
   (EX D::hreal. Y = hreal_add X D) | (EX D::hreal. X = hreal_add Y D)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2339
  by (import hreal HREAL_ADD_TOTAL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2340
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2341
lemma HREAL_SUP_ISACUT: "ALL P::hreal => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2342
   Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2343
   isacut (%w::hrat. EX X::hreal. P X & hreal.cut X w)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2344
  by (import hreal HREAL_SUP_ISACUT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2345
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2346
lemma HREAL_SUP: "ALL P::hreal => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2347
   Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2348
   (ALL Y::hreal.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2349
       (EX X::hreal. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2350
  by (import hreal HREAL_SUP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2351
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2352
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2353
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2354
;setup_theory numeral
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2355
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2356
lemma numeral_suc: "Suc ALT_ZERO = NUMERAL_BIT1 ALT_ZERO &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2357
(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2358
(ALL x::nat. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2359
  by (import numeral numeral_suc)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2360
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2361
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2362
  iZ :: "nat => nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2363
  "iZ == %x::nat. x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2364
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2365
lemma iZ: "ALL x::nat. iZ x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2366
  by (import numeral iZ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2367
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2368
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2369
  iiSUC :: "nat => nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2370
  "iiSUC == %n::nat. Suc (Suc n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2371
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2372
lemma iiSUC: "ALL n::nat. iiSUC n = Suc (Suc n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2373
  by (import numeral iiSUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2374
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2375
lemma numeral_distrib: "(op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2376
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2377
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2378
       (op =::nat => nat => bool) ((op +::nat => nat => nat) (0::nat) x) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2379
 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2380
   ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2381
     (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2382
         (op =::nat => nat => bool) ((op +::nat => nat => nat) x (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2383
          x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2384
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2385
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2386
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2387
           (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2388
            (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2389
                (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2390
                 ((op +::nat => nat => nat) ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2391
                   ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2392
                 ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2393
                   ((iZ::nat => nat) ((op +::nat => nat => nat) x xa))))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2394
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2395
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2396
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2397
             (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2398
              ((op *::nat => nat => nat) (0::nat) x) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2399
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2400
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2401
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2402
               (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2403
                ((op *::nat => nat => nat) x (0::nat)) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2404
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2405
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2406
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2407
                 (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2408
                  (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2409
                      (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2410
                       ((op *::nat => nat => nat) ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2411
                         ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2412
                       ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2413
                         ((op *::nat => nat => nat) x xa)))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2414
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2415
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2416
               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2417
                   (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2418
                    ((op -::nat => nat => nat) (0::nat) x) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2419
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2420
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2421
                 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2422
                     (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2423
                      ((op -::nat => nat => nat) x (0::nat)) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2424
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2425
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2426
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2427
                       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2428
                        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2429
                            (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2430
                             ((op -::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2431
                               ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2432
                               ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2433
                             ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2434
                               ((op -::nat => nat => nat) x xa)))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2435
                 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2436
                   ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2437
                     (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2438
                         (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2439
                          ((op ^::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2440
                            ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2441
                              ((NUMERAL_BIT1::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2442
                          (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2443
                   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2444
                     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2445
                       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2446
                           (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2447
                            ((op ^::nat => nat => nat) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2448
                              ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2449
                                ((NUMERAL_BIT2::nat => nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2450
                            (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2451
                     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2452
                       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2453
                         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2454
                             (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2455
                              ((op ^::nat => nat => nat) x (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2456
                              (1::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2457
                       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2458
                         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2459
                           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2460
                               (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2461
                                (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2462
                                    (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2463
                                     ((op ^::nat => nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2464
 ((NUMERAL::nat => nat) x) ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2465
                                     ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2466
 ((op ^::nat => nat => nat) x xa)))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2467
                         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2468
                           ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2469
                             ((Suc::nat => nat) (0::nat)) (1::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2470
                           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2471
                             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2472
                               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2473
                                   (op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2474
                                    ((Suc::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2475
((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2476
                                    ((NUMERAL::nat => nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2477
((Suc::nat => nat) x))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2478
                             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2479
                               ((op =::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2480
                                 ((PRE::nat => nat) (0::nat)) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2481
                               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2482
                                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2483
                                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2484
 (op =::nat => nat => bool) ((PRE::nat => nat) ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2485
  ((NUMERAL::nat => nat) ((PRE::nat => nat) x))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2486
                                 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2487
                                   ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2488
                                     (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2489
   (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2490
    ((op =::nat => nat => bool) ((NUMERAL::nat => nat) x) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2491
    ((op =::nat => nat => bool) x (ALT_ZERO::nat))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2492
                                   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2493
                                     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2494
 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2495
     (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2496
      ((op =::nat => nat => bool) (0::nat) ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2497
      ((op =::nat => nat => bool) x (ALT_ZERO::nat))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2498
                                     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2499
 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2500
   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2501
       (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2502
        (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2503
            (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2504
             ((op =::nat => nat => bool) ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2505
               ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2506
             ((op =::nat => nat => bool) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2507
 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2508
   ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2509
     (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2510
         (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2511
          ((op <::nat => nat => bool) x (0::nat)) (False::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2512
   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2513
     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2514
       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2515
           (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2516
            ((op <::nat => nat => bool) (0::nat) ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2517
            ((op <::nat => nat => bool) (ALT_ZERO::nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2518
     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2519
       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2520
         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2521
             (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2522
              (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2523
                  (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2524
                   ((op <::nat => nat => bool) ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2525
                     ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2526
                   ((op <::nat => nat => bool) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2527
       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2528
         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2529
           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2530
               (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2531
                ((op <::nat => nat => bool) x (0::nat)) (False::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2532
         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2533
           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2534
             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2535
                 (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2536
                  ((op <::nat => nat => bool) (0::nat)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2537
                    ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2538
                  ((op <::nat => nat => bool) (ALT_ZERO::nat) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2539
           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2540
             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2541
               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2542
                   (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2543
                    (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2544
                        (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2545
                         ((op <::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2546
                           ((NUMERAL::nat => nat) xa)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2547
                           ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2548
                         ((op <::nat => nat => bool) xa x))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2549
             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2550
               ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2551
                 (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2552
                     (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2553
                      ((op <=::nat => nat => bool) (0::nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2554
                      (True::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2555
               ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2556
                 ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2557
                   (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2558
                       (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2559
                        ((op <=::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2560
                          ((NUMERAL::nat => nat) x) (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2561
                        ((op <=::nat => nat => bool) x (ALT_ZERO::nat))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2562
                 ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2563
                   ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2564
                     (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2565
                         (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2566
                          (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2567
                              (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2568
                               ((op <=::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2569
                                 ((NUMERAL::nat => nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2570
                                 ((NUMERAL::nat => nat) xa))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2571
                               ((op <=::nat => nat => bool) x xa))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2572
                   ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2573
                     ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2574
                       (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2575
                           (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2576
                            ((op <=::nat => nat => bool) (0::nat) x)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2577
                            (True::bool)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2578
                     ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2579
                       ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2580
                         (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2581
                             (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2582
                              ((op <=::nat => nat => bool) x (0::nat))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2583
                              ((op =::nat => nat => bool) x (0::nat))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2584
                       ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2585
                         ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2586
                           (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2587
                               (All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2588
                                (%xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2589
                                    (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2590
                                     ((op <=::nat => nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2591
 ((NUMERAL::nat => nat) xa) ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2592
                                     ((op <=::nat => nat => bool) xa x))))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2593
                         ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2594
                           ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2595
                             (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2596
                                 (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2597
                                  ((ODD::nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2598
                                    ((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2599
                                  ((ODD::nat => bool) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2600
                           ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2601
                             ((All::(nat => bool) => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2602
                               (%x::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2603
                                   (op =::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2604
                                    ((EVEN::nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2605
((NUMERAL::nat => nat) x))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2606
                                    ((EVEN::nat => bool) x)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2607
                             ((op &::bool => bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2608
                               ((Not::bool => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2609
                                 ((ODD::nat => bool) (0::nat)))
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2610
                               ((EVEN::nat => bool)
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2611
                                 (0::nat))))))))))))))))))))))))))))))))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2612
  by (import numeral numeral_distrib)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2613
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2614
lemma numeral_iisuc: "iiSUC ALT_ZERO = NUMERAL_BIT2 ALT_ZERO &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2615
iiSUC (NUMERAL_BIT1 (n::nat)) = NUMERAL_BIT1 (Suc n) &
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2616
iiSUC (NUMERAL_BIT2 n) = NUMERAL_BIT2 (Suc n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2617
  by (import numeral numeral_iisuc)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2618
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2619
lemma numeral_add: "ALL (x::nat) xa::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2620
   iZ (ALT_ZERO + x) = x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2621
   iZ (x + ALT_ZERO) = x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2622
   iZ (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (iZ (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2623
   iZ (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2624
   iZ (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2625
   iZ (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2626
   Suc (ALT_ZERO + x) = Suc x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2627
   Suc (x + ALT_ZERO) = Suc x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2628
   Suc (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT1 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2629
   Suc (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2630
   Suc (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2631
   Suc (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT1 (iiSUC (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2632
   iiSUC (ALT_ZERO + x) = iiSUC x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2633
   iiSUC (x + ALT_ZERO) = iiSUC x &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2634
   iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (Suc (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2635
   iiSUC (NUMERAL_BIT1 x + NUMERAL_BIT2 xa) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2636
   NUMERAL_BIT1 (iiSUC (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2637
   iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT1 xa) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2638
   NUMERAL_BIT1 (iiSUC (x + xa)) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2639
   iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (iiSUC (x + xa))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2640
  by (import numeral numeral_add)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2641
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2642
lemma numeral_eq: "ALL (x::nat) xa::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2643
   (ALT_ZERO = NUMERAL_BIT1 x) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2644
   (NUMERAL_BIT1 x = ALT_ZERO) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2645
   (ALT_ZERO = NUMERAL_BIT2 x) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2646
   (NUMERAL_BIT2 x = ALT_ZERO) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2647
   (NUMERAL_BIT1 x = NUMERAL_BIT2 xa) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2648
   (NUMERAL_BIT2 x = NUMERAL_BIT1 xa) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2649
   (NUMERAL_BIT1 x = NUMERAL_BIT1 xa) = (x = xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2650
   (NUMERAL_BIT2 x = NUMERAL_BIT2 xa) = (x = xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2651
  by (import numeral numeral_eq)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2652
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2653
lemma numeral_lt: "ALL (x::nat) xa::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2654
   (ALT_ZERO < NUMERAL_BIT1 x) = True &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2655
   (ALT_ZERO < NUMERAL_BIT2 x) = True &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2656
   (x < ALT_ZERO) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2657
   (NUMERAL_BIT1 x < NUMERAL_BIT1 xa) = (x < xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2658
   (NUMERAL_BIT2 x < NUMERAL_BIT2 xa) = (x < xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2659
   (NUMERAL_BIT1 x < NUMERAL_BIT2 xa) = (~ xa < x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2660
   (NUMERAL_BIT2 x < NUMERAL_BIT1 xa) = (x < xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2661
  by (import numeral numeral_lt)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2662
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2663
lemma numeral_lte: "ALL (x::nat) xa::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2664
   (ALT_ZERO <= x) = True &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2665
   (NUMERAL_BIT1 x <= ALT_ZERO) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2666
   (NUMERAL_BIT2 x <= ALT_ZERO) = False &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2667
   (NUMERAL_BIT1 x <= NUMERAL_BIT1 xa) = (x <= xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2668
   (NUMERAL_BIT1 x <= NUMERAL_BIT2 xa) = (x <= xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2669
   (NUMERAL_BIT2 x <= NUMERAL_BIT1 xa) = (~ xa <= x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2670
   (NUMERAL_BIT2 x <= NUMERAL_BIT2 xa) = (x <= xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2671
  by (import numeral numeral_lte)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2672
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2673
lemma numeral_pre: "PRE ALT_ZERO = ALT_ZERO &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2674
PRE (NUMERAL_BIT1 ALT_ZERO) = ALT_ZERO &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2675
(ALL x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2676
    PRE (NUMERAL_BIT1 (NUMERAL_BIT1 x)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2677
    NUMERAL_BIT2 (PRE (NUMERAL_BIT1 x))) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2678
(ALL x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2679
    PRE (NUMERAL_BIT1 (NUMERAL_BIT2 x)) = NUMERAL_BIT2 (NUMERAL_BIT1 x)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2680
(ALL x::nat. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2681
  by (import numeral numeral_pre)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2682
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2683
lemma bit_initiality: "ALL (zf::'a::type) (b1f::nat => 'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2684
   b2f::nat => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2685
   EX x::nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2686
      x ALT_ZERO = zf &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2687
      (ALL n::nat. x (NUMERAL_BIT1 n) = b1f n (x n)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2688
      (ALL n::nat. x (NUMERAL_BIT2 n) = b2f n (x n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2689
  by (import numeral bit_initiality)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2690
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2691
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2692
  iBIT_cases :: "nat => 'a => (nat => 'a) => (nat => 'a) => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2693
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2694
specification (iBIT_cases) iBIT_cases: "(ALL (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2695
    iBIT_cases ALT_ZERO zf bf1 bf2 = zf) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2696
(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2697
    iBIT_cases (NUMERAL_BIT1 n) zf bf1 bf2 = bf1 n) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2698
(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2699
    iBIT_cases (NUMERAL_BIT2 n) zf bf1 bf2 = bf2 n)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2700
  by (import numeral iBIT_cases)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2701
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2702
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2703
  iDUB :: "nat => nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2704
  "iDUB == %x::nat. x + x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2705
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2706
lemma iDUB: "ALL x::nat. iDUB x = x + x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2707
  by (import numeral iDUB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2708
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2709
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2710
  iSUB :: "bool => nat => nat => nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2711
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2712
specification (iSUB) iSUB_DEF: "(ALL (b::bool) x::nat. iSUB b ALT_ZERO x = ALT_ZERO) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2713
(ALL (b::bool) (n::nat) x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2714
    iSUB b (NUMERAL_BIT1 n) x =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2715
    (if b
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2716
     then iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2717
           (%m::nat. NUMERAL_BIT1 (iSUB False n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2718
     else iBIT_cases x (iDUB n) (%m::nat. NUMERAL_BIT1 (iSUB False n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2719
           (%m::nat. iDUB (iSUB False n m)))) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2720
(ALL (b::bool) (n::nat) x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2721
    iSUB b (NUMERAL_BIT2 n) x =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2722
    (if b
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2723
     then iBIT_cases x (NUMERAL_BIT2 n)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2724
           (%m::nat. NUMERAL_BIT1 (iSUB True n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2725
           (%m::nat. iDUB (iSUB True n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2726
     else iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2727
           (%m::nat. NUMERAL_BIT1 (iSUB False n m))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2728
  by (import numeral iSUB_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2729
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2730
lemma bit_induction: "ALL P::nat => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2731
   P ALT_ZERO &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2732
   (ALL n::nat. P n --> P (NUMERAL_BIT1 n)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2733
   (ALL n::nat. P n --> P (NUMERAL_BIT2 n)) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2734
   All P"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2735
  by (import numeral bit_induction)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2736
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2737
lemma iSUB_THM: "ALL (xa::bool) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2738
   iSUB xa ALT_ZERO (x::nat) = ALT_ZERO &
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2739
   iSUB True xb ALT_ZERO = xb &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2740
   iSUB False (NUMERAL_BIT1 xb) ALT_ZERO = iDUB xb &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2741
   iSUB True (NUMERAL_BIT1 xb) (NUMERAL_BIT1 xc) = iDUB (iSUB True xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2742
   iSUB False (NUMERAL_BIT1 xb) (NUMERAL_BIT1 xc) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2743
   NUMERAL_BIT1 (iSUB False xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2744
   iSUB True (NUMERAL_BIT1 xb) (NUMERAL_BIT2 xc) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2745
   NUMERAL_BIT1 (iSUB False xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2746
   iSUB False (NUMERAL_BIT1 xb) (NUMERAL_BIT2 xc) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2747
   iDUB (iSUB False xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2748
   iSUB False (NUMERAL_BIT2 xb) ALT_ZERO = NUMERAL_BIT1 xb &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2749
   iSUB True (NUMERAL_BIT2 xb) (NUMERAL_BIT1 xc) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2750
   NUMERAL_BIT1 (iSUB True xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2751
   iSUB False (NUMERAL_BIT2 xb) (NUMERAL_BIT1 xc) = iDUB (iSUB True xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2752
   iSUB True (NUMERAL_BIT2 xb) (NUMERAL_BIT2 xc) = iDUB (iSUB True xb xc) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2753
   iSUB False (NUMERAL_BIT2 xb) (NUMERAL_BIT2 xc) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2754
   NUMERAL_BIT1 (iSUB False xb xc)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2755
  by (import numeral iSUB_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2756
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2757
lemma numeral_sub: "ALL (x::nat) xa::nat.
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2758
   NUMERAL (x - xa) = (if xa < x then NUMERAL (iSUB True x xa) else 0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2759
  by (import numeral numeral_sub)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2760
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2761
lemma iDUB_removal: "ALL x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2762
   iDUB (NUMERAL_BIT1 x) = NUMERAL_BIT2 (iDUB x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2763
   iDUB (NUMERAL_BIT2 x) = NUMERAL_BIT2 (NUMERAL_BIT1 x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2764
   iDUB ALT_ZERO = ALT_ZERO"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2765
  by (import numeral iDUB_removal)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2766
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2767
lemma numeral_mult: "ALL (x::nat) xa::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2768
   ALT_ZERO * x = ALT_ZERO &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2769
   x * ALT_ZERO = ALT_ZERO &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2770
   NUMERAL_BIT1 x * xa = iZ (iDUB (x * xa) + xa) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2771
   NUMERAL_BIT2 x * xa = iDUB (iZ (x * xa + xa))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2772
  by (import numeral numeral_mult)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2773
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2774
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2775
  iSQR :: "nat => nat" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2776
  "iSQR == %x::nat. x * x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2777
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2778
lemma iSQR: "ALL x::nat. iSQR x = x * x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2779
  by (import numeral iSQR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2780
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2781
lemma numeral_exp: "(ALL x::nat. x ^ ALT_ZERO = NUMERAL_BIT1 ALT_ZERO) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2782
(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT1 xa = x * iSQR (x ^ xa)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2783
(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT2 xa = iSQR x * iSQR (x ^ xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2784
  by (import numeral numeral_exp)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2785
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2786
lemma numeral_evenodd: "ALL x::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2787
   EVEN ALT_ZERO &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2788
   EVEN (NUMERAL_BIT2 x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2789
   ~ EVEN (NUMERAL_BIT1 x) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2790
   ~ ODD ALT_ZERO & ~ ODD (NUMERAL_BIT2 x) & ODD (NUMERAL_BIT1 x)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2791
  by (import numeral numeral_evenodd)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2792
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2793
lemma numeral_fact: "ALL n::nat. FACT n = (if n = 0 then 1 else n * FACT (PRE n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2794
  by (import numeral numeral_fact)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2795
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2796
lemma numeral_funpow: "ALL n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2797
   ((f::'a::type => 'a::type) ^ n) (x::'a::type) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2798
   (if n = 0 then x else (f ^ (n - 1)) (f x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2799
  by (import numeral numeral_funpow)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2800
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2801
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2802
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2803
;setup_theory ind_type
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2804
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2805
lemma INJ_INVERSE2: "ALL P::'A::type => 'B::type => 'C::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2806
   (ALL (x1::'A::type) (y1::'B::type) (x2::'A::type) y2::'B::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2807
       (P x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2)) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2808
   (EX (x::'C::type => 'A::type) Y::'C::type => 'B::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2809
       ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2810
  by (import ind_type INJ_INVERSE2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2811
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2812
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2813
  NUMPAIR :: "nat => nat => nat" 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2814
  "NUMPAIR == %(x::nat) y::nat. 2 ^ x * (2 * y + 1)"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2815
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2816
lemma NUMPAIR: "ALL (x::nat) y::nat. NUMPAIR x y = 2 ^ x * (2 * y + 1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2817
  by (import ind_type NUMPAIR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2818
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2819
lemma NUMPAIR_INJ_LEMMA: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2820
   NUMPAIR x xa = NUMPAIR xb xc --> x = xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2821
  by (import ind_type NUMPAIR_INJ_LEMMA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2822
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2823
lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2824
   (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2825
  by (import ind_type NUMPAIR_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2826
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2827
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2828
  NUMSND :: "nat => nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2829
  NUMFST :: "nat => nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2830
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2831
specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2832
  by (import ind_type NUMPAIR_DEST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2833
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2834
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2835
  NUMSUM :: "bool => nat => nat" 
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2836
  "NUMSUM == %(b::bool) x::nat. if b then Suc (2 * x) else 2 * x"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2837
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2838
lemma NUMSUM: "ALL (b::bool) x::nat. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2839
  by (import ind_type NUMSUM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2840
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2841
lemma NUMSUM_INJ: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2842
   (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2843
  by (import ind_type NUMSUM_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2844
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2845
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2846
  NUMRIGHT :: "nat => nat" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2847
  NUMLEFT :: "nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2848
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2849
specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2850
  by (import ind_type NUMSUM_DEST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2851
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2852
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2853
  INJN :: "nat => nat => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2854
  "INJN == %(m::nat) (n::nat) a::'a::type. n = m"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2855
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2856
lemma INJN: "ALL m::nat. INJN m = (%(n::nat) a::'a::type. n = m)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2857
  by (import ind_type INJN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2858
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2859
lemma INJN_INJ: "ALL (n1::nat) n2::nat. (INJN n1 = INJN n2) = (n1 = n2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2860
  by (import ind_type INJN_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2861
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2862
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2863
  INJA :: "'a => nat => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2864
  "INJA == %(a::'a::type) (n::nat) b::'a::type. b = a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2865
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2866
lemma INJA: "ALL a::'a::type. INJA a = (%(n::nat) b::'a::type. b = a)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2867
  by (import ind_type INJA)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2868
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2869
lemma INJA_INJ: "ALL (a1::'a::type) a2::'a::type. (INJA a1 = INJA a2) = (a1 = a2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2870
  by (import ind_type INJA_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2871
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2872
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2873
  INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2874
  "INJF == %(f::nat => nat => 'a::type => bool) n::nat. f (NUMFST n) (NUMSND n)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2875
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2876
lemma INJF: "ALL f::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2877
   INJF f = (%n::nat. f (NUMFST n) (NUMSND n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2878
  by (import ind_type INJF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2879
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2880
lemma INJF_INJ: "ALL (f1::nat => nat => 'a::type => bool) f2::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2881
   (INJF f1 = INJF f2) = (f1 = f2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2882
  by (import ind_type INJF_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2883
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2884
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2885
  INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2886
  "INJP ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2887
%(f1::nat => 'a::type => bool) (f2::nat => 'a::type => bool) (n::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2888
   a::'a::type. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2889
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2890
lemma INJP: "ALL (f1::nat => 'a::type => bool) f2::nat => 'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2891
   INJP f1 f2 =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2892
   (%(n::nat) a::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2893
       if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2894
  by (import ind_type INJP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2895
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2896
lemma INJP_INJ: "ALL (f1::nat => 'a::type => bool) (f1'::nat => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2897
   (f2::nat => 'a::type => bool) f2'::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2898
   (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2899
  by (import ind_type INJP_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2900
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2901
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2902
  ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2903
  "ZCONSTR ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2904
%(c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2905
   INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2906
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2907
lemma ZCONSTR: "ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2908
   ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2909
  by (import ind_type ZCONSTR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2910
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2911
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2912
  ZBOT :: "nat => 'a => bool" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2913
  "ZBOT == INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2914
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2915
lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2916
  by (import ind_type ZBOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2917
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2918
lemma ZCONSTR_ZBOT: "ALL (x::nat) (xa::'a::type) xb::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2919
   ZCONSTR x xa xb ~= ZBOT"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2920
  by (import ind_type ZCONSTR_ZBOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2921
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2922
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2923
  ZRECSPACE :: "(nat => 'a => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2924
  "ZRECSPACE ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2925
%a0::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2926
   ALL ZRECSPACE'::(nat => 'a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2927
      (ALL a0::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2928
          a0 = ZBOT |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2929
          (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2930
              a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2931
          ZRECSPACE' a0) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2932
      ZRECSPACE' a0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2933
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2934
lemma ZRECSPACE: "ZRECSPACE =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2935
(%a0::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2936
    ALL ZRECSPACE'::(nat => 'a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2937
       (ALL a0::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2938
           a0 = ZBOT |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2939
           (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2940
               a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2941
           ZRECSPACE' a0) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2942
       ZRECSPACE' a0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2943
  by (import ind_type ZRECSPACE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2944
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2945
lemma ZRECSPACE_rules: "(op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2946
 ((ZRECSPACE::(nat => 'a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2947
   (ZBOT::nat => 'a::type => bool))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2948
 ((All::(nat => bool) => bool)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2949
   (%c::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2950
       (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2951
        (%i::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2952
            (All::((nat => nat => 'a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2953
             (%r::nat => nat => 'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2954
                 (op -->::bool => bool => bool)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2955
                  ((All::(nat => bool) => bool)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2956
                    (%n::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2957
                        (ZRECSPACE::(nat => 'a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2958
                         (r n)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2959
                  ((ZRECSPACE::(nat => 'a::type => bool) => bool)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2960
                    ((ZCONSTR::nat
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2961
                               => 'a::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2962
                                  => (nat => nat => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2963
                                     => nat => 'a::type => bool)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2964
                      c i r))))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2965
  by (import ind_type ZRECSPACE_rules)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2966
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2967
lemma ZRECSPACE_ind: "ALL x::(nat => 'a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2968
   x ZBOT &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2969
   (ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2970
       (ALL n::nat. x (r n)) --> x (ZCONSTR c i r)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2971
   (ALL a0::nat => 'a::type => bool. ZRECSPACE a0 --> x a0)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2972
  by (import ind_type ZRECSPACE_ind)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2973
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2974
lemma ZRECSPACE_cases: "ALL a0::nat => 'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2975
   ZRECSPACE a0 =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2976
   (a0 = ZBOT |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2977
    (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2978
        a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE (r n))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2979
  by (import ind_type ZRECSPACE_cases)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2980
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2981
typedef (open) ('a) recspace = "(Collect::((nat => 'a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2982
          => (nat => 'a::type => bool) set)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2983
 (ZRECSPACE::(nat => 'a::type => bool) => bool)" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2984
  by (rule typedef_helper,import ind_type recspace_TY_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2985
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2986
lemmas recspace_TY_DEF = typedef_hol2hol4 [OF type_definition_recspace]
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2987
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2988
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2989
  mk_rec :: "(nat => 'a => bool) => 'a recspace" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2990
  dest_rec :: "'a recspace => nat => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2991
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2992
specification (dest_rec mk_rec) recspace_repfns: "(ALL a::'a::type recspace. mk_rec (dest_rec a) = a) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  2993
(ALL r::nat => 'a::type => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2994
  by (import ind_type recspace_repfns)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2995
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2996
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  2997
  BOTTOM :: "'a recspace" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2998
  "BOTTOM == mk_rec ZBOT"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2999
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3000
lemma BOTTOM: "BOTTOM = mk_rec ZBOT"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3001
  by (import ind_type BOTTOM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3002
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3003
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3004
  CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3005
  "CONSTR ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3006
%(c::nat) (i::'a::type) r::nat => 'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3007
   mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3008
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3009
lemma CONSTR: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3010
   CONSTR c i r = mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3011
  by (import ind_type CONSTR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3012
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3013
lemma MK_REC_INJ: "ALL (x::nat => 'a::type => bool) y::nat => 'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3014
   mk_rec x = mk_rec y --> ZRECSPACE x & ZRECSPACE y --> x = y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3015
  by (import ind_type MK_REC_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3016
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3017
lemma DEST_REC_INJ: "ALL (x::'a::type recspace) y::'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3018
   (dest_rec x = dest_rec y) = (x = y)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3019
  by (import ind_type DEST_REC_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3020
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3021
lemma CONSTR_BOT: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3022
   CONSTR c i r ~= BOTTOM"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3023
  by (import ind_type CONSTR_BOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3024
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3025
lemma CONSTR_INJ: "ALL (c1::nat) (i1::'a::type) (r1::nat => 'a::type recspace) (c2::nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3026
   (i2::'a::type) r2::nat => 'a::type recspace.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3027
   (CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3028
  by (import ind_type CONSTR_INJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3029
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3030
lemma CONSTR_IND: "ALL P::'a::type recspace => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3031
   P BOTTOM &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3032
   (ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3033
       (ALL n::nat. P (r n)) --> P (CONSTR c i r)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3034
   All P"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3035
  by (import ind_type CONSTR_IND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3036
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3037
lemma CONSTR_REC: "ALL Fn::nat
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3038
        => 'a::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3039
           => (nat => 'a::type recspace) => (nat => 'b::type) => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3040
   EX f::'a::type recspace => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3041
      ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3042
         f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3043
  by (import ind_type CONSTR_REC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3044
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3045
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3046
  FCONS :: "'a => (nat => 'a) => nat => 'a" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3047
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3048
specification (FCONS) FCONS: "(ALL (a::'a::type) f::nat => 'a::type. FCONS a f 0 = a) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3049
(ALL (a::'a::type) (f::nat => 'a::type) n::nat. FCONS a f (Suc n) = f n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3050
  by (import ind_type FCONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3051
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3052
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3053
  FNIL :: "nat => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3054
  "FNIL == %n::nat. SOME x::'a::type. True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3055
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3056
lemma FNIL: "ALL n::nat. FNIL n = (SOME x::'a::type. True)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3057
  by (import ind_type FNIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3058
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3059
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3060
  ISO :: "('a => 'b) => ('b => 'a) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3061
  "ISO ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3062
%(f::'a::type => 'b::type) g::'b::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3063
   (ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3064
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3065
lemma ISO: "ALL (f::'a::type => 'b::type) g::'b::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3066
   ISO f g =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3067
   ((ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3068
  by (import ind_type ISO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3069
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3070
lemma ISO_REFL: "ISO (%x::'a::type. x) (%x::'a::type. x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3071
  by (import ind_type ISO_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3072
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3073
lemma ISO_FUN: "ISO (f::'a::type => 'c::type) (f'::'c::type => 'a::type) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3074
ISO (g::'b::type => 'd::type) (g'::'d::type => 'b::type) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3075
ISO (%(h::'a::type => 'b::type) a'::'c::type. g (h (f' a')))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3076
 (%(h::'c::type => 'd::type) a::'a::type. g' (h (f a)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3077
  by (import ind_type ISO_FUN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3078
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3079
lemma ISO_USAGE: "ISO (f::'a::type => 'b::type) (g::'b::type => 'a::type) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3080
(ALL P::'a::type => bool. All P = (ALL x::'b::type. P (g x))) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3081
(ALL P::'a::type => bool. Ex P = (EX x::'b::type. P (g x))) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3082
(ALL (a::'a::type) b::'b::type. (a = g b) = (f a = b))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3083
  by (import ind_type ISO_USAGE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3084
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3085
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3086
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3087
;setup_theory divides
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3088
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3089
lemma ONE_DIVIDES_ALL: "(All::(nat => bool) => bool) ((op dvd::nat => nat => bool) (1::nat))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3090
  by (import divides ONE_DIVIDES_ALL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3091
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3092
lemma DIVIDES_ADD_2: "ALL (a::nat) (b::nat) c::nat. a dvd b & a dvd b + c --> a dvd c"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3093
  by (import divides DIVIDES_ADD_2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3094
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3095
lemma DIVIDES_FACT: "ALL b>0. b dvd FACT b"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3096
  by (import divides DIVIDES_FACT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3097
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3098
lemma DIVIDES_MULT_LEFT: "ALL (x::nat) xa::nat. (x * xa dvd xa) = (xa = 0 | x = 1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3099
  by (import divides DIVIDES_MULT_LEFT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3100
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3101
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3102
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3103
;setup_theory prime
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3104
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3105
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3106
  prime :: "nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3107
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3108
defs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3109
  prime_primdef: "prime.prime == %a::nat. a ~= 1 & (ALL b::nat. b dvd a --> b = a | b = 1)"
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3110
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3111
lemma prime_def: "ALL a::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3112
   prime.prime a = (a ~= 1 & (ALL b::nat. b dvd a --> b = a | b = 1))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3113
  by (import prime prime_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3114
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3115
lemma NOT_PRIME_0: "~ prime.prime 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3116
  by (import prime NOT_PRIME_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3117
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3118
lemma NOT_PRIME_1: "~ prime.prime 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3119
  by (import prime NOT_PRIME_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3120
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3121
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3122
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3123
;setup_theory list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3124
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3125
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3126
  EL :: "nat => 'a list => 'a" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3127
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3128
specification (EL) EL: "(ALL l::'a::type list. EL 0 l = hd l) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3129
(ALL (l::'a::type list) n::nat. EL (Suc n) l = EL n (tl l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3130
  by (import list EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3131
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3132
lemma NULL: "(op &::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3133
 ((null::'a::type list => bool) ([]::'a::type list))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3134
 ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3135
   (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3136
       (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3137
        (%xa::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3138
            (Not::bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3139
             ((null::'a::type list => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3140
               ((op #::'a::type => 'a::type list => 'a::type list) x xa)))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3141
  by (import list NULL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3142
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3143
lemma list_case_compute: "ALL l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3144
   list_case (b::'b::type) (f::'a::type => 'a::type list => 'b::type) l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3145
   (if null l then b else f (hd l) (tl l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3146
  by (import list list_case_compute)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3147
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3148
lemma LIST_NOT_EQ: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3149
   l1 ~= l2 --> (ALL (x::'a::type) xa::'a::type. x # l1 ~= xa # l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3150
  by (import list LIST_NOT_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3151
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3152
lemma NOT_EQ_LIST: "ALL (h1::'a::type) h2::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3153
   h1 ~= h2 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3154
   (ALL (x::'a::type list) xa::'a::type list. h1 # x ~= h2 # xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3155
  by (import list NOT_EQ_LIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3156
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3157
lemma EQ_LIST: "ALL (h1::'a::type) h2::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3158
   h1 = h2 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3159
   (ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3160
       l1 = l2 --> h1 # l1 = h2 # l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3161
  by (import list EQ_LIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3162
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3163
lemma CONS: "ALL l::'a::type list. ~ null l --> hd l # tl l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3164
  by (import list CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3165
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3166
lemma MAP_EQ_NIL: "ALL (l::'a::type list) f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3167
   (map f l = []) = (l = []) & ([] = map f l) = (l = [])"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3168
  by (import list MAP_EQ_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3169
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3170
lemma EVERY_EL: "(All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3171
 (%l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3172
     (All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3173
      (%P::'a::type => bool.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3174
          (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3175
           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3176
           ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3177
             (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3178
                 (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3179
                  ((op <::nat => nat => bool) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3180
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3181
                  (P ((EL::nat => 'a::type list => 'a::type) n l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3182
  by (import list EVERY_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3183
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3184
lemma EVERY_CONJ: "ALL l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3185
   list_all
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3186
    (%x::'a::type. (P::'a::type => bool) x & (Q::'a::type => bool) x) l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3187
   (list_all P l & list_all Q l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3188
  by (import list EVERY_CONJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3189
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3190
lemma EVERY_MEM: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3191
   list_all P l = (ALL e::'a::type. e mem l --> P e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3192
  by (import list EVERY_MEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3193
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3194
lemma EXISTS_MEM: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3195
   list_exists P l = (EX e::'a::type. e mem l & P e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3196
  by (import list EXISTS_MEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3197
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3198
lemma MEM_APPEND: "ALL (e::'a::type) (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3199
   e mem l1 @ l2 = (e mem l1 | e mem l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3200
  by (import list MEM_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3201
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3202
lemma EXISTS_APPEND: "ALL (P::'a::type => bool) (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3203
   list_exists P (l1 @ l2) = (list_exists P l1 | list_exists P l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3204
  by (import list EXISTS_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3205
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3206
lemma NOT_EVERY: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3207
   (~ list_all P l) = list_exists (Not o P) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3208
  by (import list NOT_EVERY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3209
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3210
lemma NOT_EXISTS: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3211
   (~ list_exists P l) = list_all (Not o P) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3212
  by (import list NOT_EXISTS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3213
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3214
lemma MEM_MAP: "ALL (l::'a::type list) (f::'a::type => 'b::type) x::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3215
   x mem map f l = (EX y::'a::type. x = f y & y mem l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3216
  by (import list MEM_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3217
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3218
lemma LENGTH_CONS: "ALL (l::'a::type list) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3219
   (length l = Suc n) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3220
   (EX (h::'a::type) l'::'a::type list. length l' = n & l = h # l')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3221
  by (import list LENGTH_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3222
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3223
lemma LENGTH_EQ_CONS: "ALL (P::'a::type list => bool) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3224
   (ALL l::'a::type list. length l = Suc n --> P l) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3225
   (ALL l::'a::type list. length l = n --> (ALL x::'a::type. P (x # l)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3226
  by (import list LENGTH_EQ_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3227
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3228
lemma LENGTH_EQ_NIL: "ALL P::'a::type list => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3229
   (ALL l::'a::type list. length l = 0 --> P l) = P []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3230
  by (import list LENGTH_EQ_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3231
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3232
lemma CONS_ACYCLIC: "ALL (l::'a::type list) x::'a::type. l ~= x # l & x # l ~= l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3233
  by (import list CONS_ACYCLIC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3234
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3235
lemma APPEND_eq_NIL: "(ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3236
    ([] = l1 @ l2) = (l1 = [] & l2 = [])) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3237
(ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3238
    (l1 @ l2 = []) = (l1 = [] & l2 = []))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3239
  by (import list APPEND_eq_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3240
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3241
lemma APPEND_11: "(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3242
    (l1 @ l2 = l1 @ l3) = (l2 = l3)) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3243
(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3244
    (l2 @ l1 = l3 @ l1) = (l2 = l3))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3245
  by (import list APPEND_11)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3246
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3247
lemma EL_compute: "ALL n::nat.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3248
   EL n (l::'a::type list) = (if n = 0 then hd l else EL (PRE n) (tl l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3249
  by (import list EL_compute)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3250
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3251
lemma WF_LIST_PRED: "WF (%(L1::'a::type list) L2::'a::type list. EX h::'a::type. L2 = h # L1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3252
  by (import list WF_LIST_PRED)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3253
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3254
lemma list_size_cong: "ALL (M::'a::type list) (N::'a::type list) (f::'a::type => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3255
   f'::'a::type => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3256
   M = N & (ALL x::'a::type. x mem N --> f x = f' x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3257
   list_size f M = list_size f' N"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3258
  by (import list list_size_cong)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3259
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3260
lemma FOLDR_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3261
   (f::'a::type => 'b::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3262
   f'::'a::type => 'b::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3263
   l = l' &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3264
   b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f x a = f' x a) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3265
   foldr f l b = foldr f' l' b'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3266
  by (import list FOLDR_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3267
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3268
lemma FOLDL_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3269
   (f::'b::type => 'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3270
   f'::'b::type => 'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3271
   l = l' &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3272
   b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f a x = f' a x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3273
   foldl f b l = foldl f' b' l'"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3274
  by (import list FOLDL_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3275
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3276
lemma MAP_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (f::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3277
   f'::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3278
   l1 = l2 & (ALL x::'a::type. x mem l2 --> f x = f' x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3279
   map f l1 = map f' l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3280
  by (import list MAP_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3281
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3282
lemma EXISTS_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3283
   P'::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3284
   l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3285
   list_exists P l1 = list_exists P' l2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3286
  by (import list EXISTS_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3287
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3288
lemma EVERY_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3289
   P'::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3290
   l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3291
   list_all P l1 = list_all P' l2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3292
  by (import list EVERY_CONG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3293
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3294
lemma EVERY_MONOTONIC: "ALL (P::'a::type => bool) Q::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3295
   (ALL x::'a::type. P x --> Q x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3296
   (ALL l::'a::type list. list_all P l --> list_all Q l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3297
  by (import list EVERY_MONOTONIC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3298
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3299
lemma LENGTH_ZIP: "ALL (l1::'a::type list) l2::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3300
   length l1 = length l2 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3301
   length (zip l1 l2) = length l1 & length (zip l1 l2) = length l2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3302
  by (import list LENGTH_ZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3303
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3304
lemma LENGTH_UNZIP: "ALL pl::('a::type * 'b::type) list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3305
   length (fst (unzip pl)) = length pl & length (snd (unzip pl)) = length pl"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3306
  by (import list LENGTH_UNZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3307
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3308
lemma ZIP_UNZIP: "ALL l::('a::type * 'b::type) list. ZIP (unzip l) = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3309
  by (import list ZIP_UNZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3310
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3311
lemma UNZIP_ZIP: "ALL (l1::'a::type list) l2::'b::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3312
   length l1 = length l2 --> unzip (zip l1 l2) = (l1, l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3313
  by (import list UNZIP_ZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3314
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3315
lemma ZIP_MAP: "ALL (l1::'a::type list) (l2::'b::type list) (f1::'a::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3316
   f2::'b::type => 'd::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3317
   length l1 = length l2 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3318
   zip (map f1 l1) l2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3319
   map (%p::'a::type * 'b::type. (f1 (fst p), snd p)) (zip l1 l2) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3320
   zip l1 (map f2 l2) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3321
   map (%p::'a::type * 'b::type. (fst p, f2 (snd p))) (zip l1 l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3322
  by (import list ZIP_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3323
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3324
lemma MEM_ZIP: "(All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3325
 (%l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3326
     (All::('b::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3327
      (%l2::'b::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3328
          (All::('a::type * 'b::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3329
           (%p::'a::type * 'b::type.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3330
               (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3331
                ((op =::nat => nat => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3332
                  ((size::'a::type list => nat) l1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3333
                  ((size::'b::type list => nat) l2))
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3334
                ((op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3335
                  ((op mem::'a::type * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3336
                            => ('a::type * 'b::type) list => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3337
                    p ((zip::'a::type list
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3338
                             => 'b::type list => ('a::type * 'b::type) list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3339
                        l1 l2))
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3340
                  ((Ex::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3341
                    (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3342
                        (op &::bool => bool => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3343
                         ((op <::nat => nat => bool) n
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3344
                           ((size::'a::type list => nat) l1))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3345
                         ((op =::'a::type * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3346
                                 => 'a::type * 'b::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3347
                           p ((Pair::'a::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3348
                                     => 'b::type => 'a::type * 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3349
                               ((EL::nat => 'a::type list => 'a::type) n l1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3350
                               ((EL::nat => 'b::type list => 'b::type) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3351
                                 l2)))))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3352
  by (import list MEM_ZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3353
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3354
lemma EL_ZIP: "ALL (l1::'a::type list) (l2::'b::type list) n::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3355
   length l1 = length l2 & n < length l1 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3356
   EL n (zip l1 l2) = (EL n l1, EL n l2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3357
  by (import list EL_ZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3358
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3359
lemma MAP2_ZIP: "(All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3360
 (%l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3361
     (All::('b::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3362
      (%l2::'b::type list.
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
  3363
          (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3364
           ((op =::nat => nat => bool) ((size::'a::type list => nat) l1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3365
             ((size::'b::type list => nat) l2))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3366
           ((All::(('a::type => 'b::type => 'c::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3367
             (%f::'a::type => 'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3368
                 (op =::'c::type list => 'c::type list => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3369
                  ((map2::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3370
                          => 'a::type list
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3371
                             => 'b::type list => 'c::type list)
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
  3372
                    f l1 l2)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3373
                  ((map::('a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3374
                         => ('a::type * 'b::type) list => 'c::type list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3375
                    ((split::('a::type => 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3376
                             => 'a::type * 'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3377
                      f)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3378
                    ((zip::'a::type list
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3379
                           => 'b::type list => ('a::type * 'b::type) list)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3380
                      l1 l2))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3381
  by (import list MAP2_ZIP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3382
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3383
lemma MEM_EL: "(All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3384
 (%l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3385
     (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3386
      (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3387
          (op =::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3388
           ((op mem::'a::type => 'a::type list => bool) x l)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3389
           ((Ex::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3390
             (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  3391
                 (op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3392
                  ((op <::nat => nat => bool) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3393
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3394
                  ((op =::'a::type => 'a::type => bool) x
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3395
                    ((EL::nat => 'a::type list => 'a::type) n l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3396
  by (import list MEM_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3397
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3398
lemma LAST_CONS: "(ALL x::'a::type. last [x] = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3399
(ALL (x::'a::type) (xa::'a::type) xb::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3400
    last (x # xa # xb) = last (xa # xb))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3401
  by (import list LAST_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3402
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3403
lemma FRONT_CONS: "(ALL x::'a::type. butlast [x] = []) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3404
(ALL (x::'a::type) (xa::'a::type) xb::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3405
    butlast (x # xa # xb) = x # butlast (xa # xb))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3406
  by (import list FRONT_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3407
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3408
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3409
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3410
;setup_theory pred_set
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3411
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3412
lemma EXTENSION: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3413
   (s = t) = (ALL x::'a::type. IN x s = IN x t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3414
  by (import pred_set EXTENSION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3415
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3416
lemma NOT_EQUAL_SETS: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3417
   (x ~= xa) = (EX xb::'a::type. IN xb xa = (~ IN xb x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3418
  by (import pred_set NOT_EQUAL_SETS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3419
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3420
lemma NUM_SET_WOP: "ALL s::nat => bool.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3421
   (EX n::nat. IN n s) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3422
   (EX n::nat. IN n s & (ALL m::nat. IN m s --> n <= m))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3423
  by (import pred_set NUM_SET_WOP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3424
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3425
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3426
  GSPEC :: "('b => 'a * bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3427
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3428
specification (GSPEC) GSPECIFICATION: "ALL (f::'b::type => 'a::type * bool) v::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3429
   IN v (GSPEC f) = (EX x::'b::type. (v, True) = f x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3430
  by (import pred_set GSPECIFICATION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3431
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3432
lemma SET_MINIMUM: "ALL (s::'a::type => bool) M::'a::type => nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3433
   (EX x::'a::type. IN x s) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3434
   (EX x::'a::type. IN x s & (ALL y::'a::type. IN y s --> M x <= M y))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3435
  by (import pred_set SET_MINIMUM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3436
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3437
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3438
  EMPTY :: "'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3439
  "EMPTY == %x::'a::type. False"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3440
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3441
lemma EMPTY_DEF: "EMPTY = (%x::'a::type. False)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3442
  by (import pred_set EMPTY_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3443
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3444
lemma NOT_IN_EMPTY: "ALL x::'a::type. ~ IN x EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3445
  by (import pred_set NOT_IN_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3446
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3447
lemma MEMBER_NOT_EMPTY: "ALL x::'a::type => bool. (EX xa::'a::type. IN xa x) = (x ~= EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3448
  by (import pred_set MEMBER_NOT_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3449
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3450
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3451
  UNIV :: "'a => bool" 
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3452
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3453
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3454
  UNIV_def: "pred_set.UNIV == %x::'a::type. True"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3455
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3456
lemma UNIV_DEF: "pred_set.UNIV = (%x::'a::type. True)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3457
  by (import pred_set UNIV_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3458
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3459
lemma IN_UNIV: "ALL x::'a::type. IN x pred_set.UNIV"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3460
  by (import pred_set IN_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3461
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3462
lemma UNIV_NOT_EMPTY: "pred_set.UNIV ~= EMPTY"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3463
  by (import pred_set UNIV_NOT_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3464
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3465
lemma EMPTY_NOT_UNIV: "EMPTY ~= pred_set.UNIV"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3466
  by (import pred_set EMPTY_NOT_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3467
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3468
lemma EQ_UNIV: "(ALL x::'a::type. IN x (s::'a::type => bool)) = (s = pred_set.UNIV)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3469
  by (import pred_set EQ_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3470
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3471
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3472
  SUBSET :: "('a => bool) => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3473
  "SUBSET ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3474
%(s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3475
   ALL x::'a::type. IN x s --> IN x t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3476
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3477
lemma SUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3478
   SUBSET s t = (ALL x::'a::type. IN x s --> IN x t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3479
  by (import pred_set SUBSET_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3480
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3481
lemma SUBSET_TRANS: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3482
   SUBSET x xa & SUBSET xa xb --> SUBSET x xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3483
  by (import pred_set SUBSET_TRANS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3484
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3485
lemma SUBSET_REFL: "ALL x::'a::type => bool. SUBSET x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3486
  by (import pred_set SUBSET_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3487
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3488
lemma SUBSET_ANTISYM: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3489
   SUBSET x xa & SUBSET xa x --> x = xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3490
  by (import pred_set SUBSET_ANTISYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3491
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3492
lemma EMPTY_SUBSET: "All (SUBSET EMPTY)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3493
  by (import pred_set EMPTY_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3494
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3495
lemma SUBSET_EMPTY: "ALL x::'a::type => bool. SUBSET x EMPTY = (x = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3496
  by (import pred_set SUBSET_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3497
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3498
lemma SUBSET_UNIV: "ALL x::'a::type => bool. SUBSET x pred_set.UNIV"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3499
  by (import pred_set SUBSET_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3500
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3501
lemma UNIV_SUBSET: "ALL x::'a::type => bool. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3502
  by (import pred_set UNIV_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3503
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3504
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3505
  PSUBSET :: "('a => bool) => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3506
  "PSUBSET == %(s::'a::type => bool) t::'a::type => bool. SUBSET s t & s ~= t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3507
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3508
lemma PSUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3509
   PSUBSET s t = (SUBSET s t & s ~= t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3510
  by (import pred_set PSUBSET_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3511
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3512
lemma PSUBSET_TRANS: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3513
   PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3514
  by (import pred_set PSUBSET_TRANS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3515
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3516
lemma PSUBSET_IRREFL: "ALL x::'a::type => bool. ~ PSUBSET x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3517
  by (import pred_set PSUBSET_IRREFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3518
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3519
lemma NOT_PSUBSET_EMPTY: "ALL x::'a::type => bool. ~ PSUBSET x EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3520
  by (import pred_set NOT_PSUBSET_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3521
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3522
lemma NOT_UNIV_PSUBSET: "ALL x::'a::type => bool. ~ PSUBSET pred_set.UNIV x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3523
  by (import pred_set NOT_UNIV_PSUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3524
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3525
lemma PSUBSET_UNIV: "ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3526
   PSUBSET x pred_set.UNIV = (EX xa::'a::type. ~ IN xa x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3527
  by (import pred_set PSUBSET_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3528
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3529
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3530
  UNION :: "('a => bool) => ('a => bool) => 'a => bool" 
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3531
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3532
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3533
  UNION_def: "pred_set.UNION ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3534
%(s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3535
   GSPEC (%x::'a::type. (x, IN x s | IN x t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3536
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3537
lemma UNION_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3538
   pred_set.UNION s t = GSPEC (%x::'a::type. (x, IN x s | IN x t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3539
  by (import pred_set UNION_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3540
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3541
lemma IN_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3542
   IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3543
  by (import pred_set IN_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3544
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3545
lemma UNION_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3546
   pred_set.UNION x (pred_set.UNION xa xb) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3547
   pred_set.UNION (pred_set.UNION x xa) xb"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3548
  by (import pred_set UNION_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3549
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3550
lemma UNION_IDEMPOT: "ALL x::'a::type => bool. pred_set.UNION x x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3551
  by (import pred_set UNION_IDEMPOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3552
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3553
lemma UNION_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3554
   pred_set.UNION x xa = pred_set.UNION xa x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3555
  by (import pred_set UNION_COMM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3556
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3557
lemma SUBSET_UNION: "(ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3558
    SUBSET x (pred_set.UNION x xa)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3559
(ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3560
    SUBSET x (pred_set.UNION xa x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3561
  by (import pred_set SUBSET_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3562
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3563
lemma UNION_SUBSET: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3564
   SUBSET (pred_set.UNION s t) u = (SUBSET s u & SUBSET t u)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3565
  by (import pred_set UNION_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3566
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3567
lemma SUBSET_UNION_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3568
   SUBSET x xa = (pred_set.UNION x xa = xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3569
  by (import pred_set SUBSET_UNION_ABSORPTION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3570
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3571
lemma UNION_EMPTY: "(ALL x::'a::type => bool. pred_set.UNION EMPTY x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3572
(ALL x::'a::type => bool. pred_set.UNION x EMPTY = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3573
  by (import pred_set UNION_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3574
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3575
lemma UNION_UNIV: "(ALL x::'a::type => bool. pred_set.UNION pred_set.UNIV x = pred_set.UNIV) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3576
(ALL x::'a::type => bool. pred_set.UNION x pred_set.UNIV = pred_set.UNIV)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3577
  by (import pred_set UNION_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3578
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3579
lemma EMPTY_UNION: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3580
   (pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3581
  by (import pred_set EMPTY_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3582
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3583
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3584
  INTER :: "('a => bool) => ('a => bool) => 'a => bool" 
14684
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3585
d796124e435c removed 'constdefs' hack;
wenzelm
parents: 14516
diff changeset
  3586
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3587
  INTER_def: "pred_set.INTER ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3588
%(s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3589
   GSPEC (%x::'a::type. (x, IN x s & IN x t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3590
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3591
lemma INTER_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3592
   pred_set.INTER s t = GSPEC (%x::'a::type. (x, IN x s & IN x t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3593
  by (import pred_set INTER_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3594
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3595
lemma IN_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3596
   IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3597
  by (import pred_set IN_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3598
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3599
lemma INTER_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3600
   pred_set.INTER x (pred_set.INTER xa xb) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3601
   pred_set.INTER (pred_set.INTER x xa) xb"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3602
  by (import pred_set INTER_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3603
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3604
lemma INTER_IDEMPOT: "ALL x::'a::type => bool. pred_set.INTER x x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3605
  by (import pred_set INTER_IDEMPOT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3606
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3607
lemma INTER_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3608
   pred_set.INTER x xa = pred_set.INTER xa x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3609
  by (import pred_set INTER_COMM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3610
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3611
lemma INTER_SUBSET: "(ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3612
    SUBSET (pred_set.INTER x xa) x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3613
(ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3614
    SUBSET (pred_set.INTER xa x) x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3615
  by (import pred_set INTER_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3616
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3617
lemma SUBSET_INTER: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3618
   SUBSET s (pred_set.INTER t u) = (SUBSET s t & SUBSET s u)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3619
  by (import pred_set SUBSET_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3620
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3621
lemma SUBSET_INTER_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3622
   SUBSET x xa = (pred_set.INTER x xa = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3623
  by (import pred_set SUBSET_INTER_ABSORPTION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3624
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3625
lemma INTER_EMPTY: "(ALL x::'a::type => bool. pred_set.INTER EMPTY x = EMPTY) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3626
(ALL x::'a::type => bool. pred_set.INTER x EMPTY = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3627
  by (import pred_set INTER_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3628
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3629
lemma INTER_UNIV: "(ALL x::'a::type => bool. pred_set.INTER pred_set.UNIV x = x) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3630
(ALL x::'a::type => bool. pred_set.INTER x pred_set.UNIV = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3631
  by (import pred_set INTER_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3632
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3633
lemma UNION_OVER_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3634
   pred_set.INTER x (pred_set.UNION xa xb) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3635
   pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER x xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3636
  by (import pred_set UNION_OVER_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3637
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3638
lemma INTER_OVER_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3639
   pred_set.UNION x (pred_set.INTER xa xb) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3640
   pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3641
  by (import pred_set INTER_OVER_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3642
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3643
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3644
  DISJOINT :: "('a => bool) => ('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3645
  "DISJOINT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3646
%(s::'a::type => bool) t::'a::type => bool. pred_set.INTER s t = EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3647
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3648
lemma DISJOINT_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3649
   DISJOINT s t = (pred_set.INTER s t = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3650
  by (import pred_set DISJOINT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3651
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3652
lemma IN_DISJOINT: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3653
   DISJOINT x xa = (~ (EX xb::'a::type. IN xb x & IN xb xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3654
  by (import pred_set IN_DISJOINT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3655
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3656
lemma DISJOINT_SYM: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3657
   DISJOINT x xa = DISJOINT xa x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3658
  by (import pred_set DISJOINT_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3659
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3660
lemma DISJOINT_EMPTY: "ALL x::'a::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3661
  by (import pred_set DISJOINT_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3662
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3663
lemma DISJOINT_EMPTY_REFL: "ALL x::'a::type => bool. (x = EMPTY) = DISJOINT x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3664
  by (import pred_set DISJOINT_EMPTY_REFL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3665
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3666
lemma DISJOINT_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3667
   DISJOINT (pred_set.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3668
  by (import pred_set DISJOINT_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3669
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3670
lemma DISJOINT_UNION_BOTH: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3671
   DISJOINT (pred_set.UNION s t) u = (DISJOINT s u & DISJOINT t u) &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3672
   DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3673
  by (import pred_set DISJOINT_UNION_BOTH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3674
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3675
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3676
  DIFF :: "('a => bool) => ('a => bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3677
  "DIFF ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3678
%(s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3679
   GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3680
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3681
lemma DIFF_DEF: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3682
   DIFF s t = GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3683
  by (import pred_set DIFF_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3684
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3685
lemma IN_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3686
   IN x (DIFF s t) = (IN x s & ~ IN x t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3687
  by (import pred_set IN_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3688
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3689
lemma DIFF_EMPTY: "ALL s::'a::type => bool. DIFF s EMPTY = s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3690
  by (import pred_set DIFF_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3691
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3692
lemma EMPTY_DIFF: "ALL s::'a::type => bool. DIFF EMPTY s = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3693
  by (import pred_set EMPTY_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3694
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3695
lemma DIFF_UNIV: "ALL s::'a::type => bool. DIFF s pred_set.UNIV = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3696
  by (import pred_set DIFF_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3697
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3698
lemma DIFF_DIFF: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3699
   DIFF (DIFF x xa) xa = DIFF x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3700
  by (import pred_set DIFF_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3701
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3702
lemma DIFF_EQ_EMPTY: "ALL x::'a::type => bool. DIFF x x = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3703
  by (import pred_set DIFF_EQ_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3704
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3705
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3706
  INSERT :: "'a => ('a => bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3707
  "INSERT ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3708
%(x::'a::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3709
   GSPEC (%y::'a::type. (y, y = x | IN y s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3710
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3711
lemma INSERT_DEF: "ALL (x::'a::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3712
   INSERT x s = GSPEC (%y::'a::type. (y, y = x | IN y s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3713
  by (import pred_set INSERT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3714
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3715
lemma IN_INSERT: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3716
   IN x (INSERT xa xb) = (x = xa | IN x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3717
  by (import pred_set IN_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3718
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3719
lemma COMPONENT: "ALL (x::'a::type) xa::'a::type => bool. IN x (INSERT x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3720
  by (import pred_set COMPONENT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3721
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3722
lemma SET_CASES: "ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3723
   x = EMPTY |
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3724
   (EX (xa::'a::type) xb::'a::type => bool. x = INSERT xa xb & ~ IN xa xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3725
  by (import pred_set SET_CASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3726
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3727
lemma DECOMPOSITION: "ALL (s::'a::type => bool) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3728
   IN x s = (EX t::'a::type => bool. s = INSERT x t & ~ IN x t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3729
  by (import pred_set DECOMPOSITION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3730
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3731
lemma ABSORPTION: "ALL (x::'a::type) xa::'a::type => bool. IN x xa = (INSERT x xa = xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3732
  by (import pred_set ABSORPTION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3733
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3734
lemma INSERT_INSERT: "ALL (x::'a::type) xa::'a::type => bool. INSERT x (INSERT x xa) = INSERT x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3735
  by (import pred_set INSERT_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3736
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3737
lemma INSERT_COMM: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3738
   INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3739
  by (import pred_set INSERT_COMM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3740
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3741
lemma INSERT_UNIV: "ALL x::'a::type. INSERT x pred_set.UNIV = pred_set.UNIV"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3742
  by (import pred_set INSERT_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3743
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3744
lemma NOT_INSERT_EMPTY: "ALL (x::'a::type) xa::'a::type => bool. INSERT x xa ~= EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3745
  by (import pred_set NOT_INSERT_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3746
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3747
lemma NOT_EMPTY_INSERT: "ALL (x::'a::type) xa::'a::type => bool. EMPTY ~= INSERT x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3748
  by (import pred_set NOT_EMPTY_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3749
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3750
lemma INSERT_UNION: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3751
   pred_set.UNION (INSERT x s) t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3752
   (if IN x t then pred_set.UNION s t else INSERT x (pred_set.UNION s t))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3753
  by (import pred_set INSERT_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3754
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3755
lemma INSERT_UNION_EQ: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3756
   pred_set.UNION (INSERT x s) t = INSERT x (pred_set.UNION s t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3757
  by (import pred_set INSERT_UNION_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3758
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3759
lemma INSERT_INTER: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3760
   pred_set.INTER (INSERT x s) t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3761
   (if IN x t then INSERT x (pred_set.INTER s t) else pred_set.INTER s t)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3762
  by (import pred_set INSERT_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3763
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3764
lemma DISJOINT_INSERT: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3765
   DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3766
  by (import pred_set DISJOINT_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3767
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3768
lemma INSERT_SUBSET: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3769
   SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3770
  by (import pred_set INSERT_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3771
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3772
lemma SUBSET_INSERT: "ALL (x::'a::type) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3773
   ~ IN x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3774
   (ALL xb::'a::type => bool. SUBSET xa (INSERT x xb) = SUBSET xa xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3775
  by (import pred_set SUBSET_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3776
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3777
lemma INSERT_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3778
   DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3779
  by (import pred_set INSERT_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3780
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3781
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3782
  DELETE :: "('a => bool) => 'a => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3783
  "DELETE == %(s::'a::type => bool) x::'a::type. DIFF s (INSERT x EMPTY)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3784
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3785
lemma DELETE_DEF: "ALL (s::'a::type => bool) x::'a::type. DELETE s x = DIFF s (INSERT x EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3786
  by (import pred_set DELETE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3787
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3788
lemma IN_DELETE: "ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3789
   IN xa (DELETE x xb) = (IN xa x & xa ~= xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3790
  by (import pred_set IN_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3791
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3792
lemma DELETE_NON_ELEMENT: "ALL (x::'a::type) xa::'a::type => bool. (~ IN x xa) = (DELETE xa x = xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3793
  by (import pred_set DELETE_NON_ELEMENT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3794
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3795
lemma IN_DELETE_EQ: "ALL (s::'a::type => bool) (x::'a::type) x'::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3796
   (IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3797
  by (import pred_set IN_DELETE_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3798
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3799
lemma EMPTY_DELETE: "ALL x::'a::type. DELETE EMPTY x = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3800
  by (import pred_set EMPTY_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3801
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3802
lemma DELETE_DELETE: "ALL (x::'a::type) xa::'a::type => bool. DELETE (DELETE xa x) x = DELETE xa x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3803
  by (import pred_set DELETE_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3804
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3805
lemma DELETE_COMM: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3806
   DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3807
  by (import pred_set DELETE_COMM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3808
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3809
lemma DELETE_SUBSET: "ALL (x::'a::type) xa::'a::type => bool. SUBSET (DELETE xa x) xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3810
  by (import pred_set DELETE_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3811
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3812
lemma SUBSET_DELETE: "ALL (x::'a::type) (xa::'a::type => bool) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3813
   SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3814
  by (import pred_set SUBSET_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3815
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3816
lemma SUBSET_INSERT_DELETE: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3817
   SUBSET s (INSERT x t) = SUBSET (DELETE s x) t"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3818
  by (import pred_set SUBSET_INSERT_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3819
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3820
lemma DIFF_INSERT: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3821
   DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3822
  by (import pred_set DIFF_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3823
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3824
lemma PSUBSET_INSERT_SUBSET: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3825
   PSUBSET x xa = (EX xb::'a::type. ~ IN xb x & SUBSET (INSERT xb x) xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3826
  by (import pred_set PSUBSET_INSERT_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3827
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3828
lemma PSUBSET_MEMBER: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3829
   PSUBSET s t = (SUBSET s t & (EX y::'a::type. IN y t & ~ IN y s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3830
  by (import pred_set PSUBSET_MEMBER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3831
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3832
lemma DELETE_INSERT: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3833
   DELETE (INSERT x xb) xa =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3834
   (if x = xa then DELETE xb xa else INSERT x (DELETE xb xa))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3835
  by (import pred_set DELETE_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3836
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3837
lemma INSERT_DELETE: "ALL (x::'a::type) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3838
   IN x xa --> INSERT x (DELETE xa x) = xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3839
  by (import pred_set INSERT_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3840
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3841
lemma DELETE_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3842
   pred_set.INTER (DELETE x xb) xa = DELETE (pred_set.INTER x xa) xb"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3843
  by (import pred_set DELETE_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3844
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3845
lemma DISJOINT_DELETE_SYM: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3846
   DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3847
  by (import pred_set DISJOINT_DELETE_SYM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3848
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3849
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3850
  CHOICE :: "('a => bool) => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3851
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3852
specification (CHOICE) CHOICE_DEF: "ALL x::'a::type => bool. x ~= EMPTY --> IN (CHOICE x) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3853
  by (import pred_set CHOICE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3854
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3855
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3856
  REST :: "('a => bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3857
  "REST == %s::'a::type => bool. DELETE s (CHOICE s)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3858
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3859
lemma REST_DEF: "ALL s::'a::type => bool. REST s = DELETE s (CHOICE s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3860
  by (import pred_set REST_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3861
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3862
lemma CHOICE_NOT_IN_REST: "ALL x::'a::type => bool. ~ IN (CHOICE x) (REST x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3863
  by (import pred_set CHOICE_NOT_IN_REST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3864
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3865
lemma CHOICE_INSERT_REST: "ALL s::'a::type => bool. s ~= EMPTY --> INSERT (CHOICE s) (REST s) = s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3866
  by (import pred_set CHOICE_INSERT_REST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3867
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3868
lemma REST_SUBSET: "ALL x::'a::type => bool. SUBSET (REST x) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3869
  by (import pred_set REST_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3870
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3871
lemma REST_PSUBSET: "ALL x::'a::type => bool. x ~= EMPTY --> PSUBSET (REST x) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3872
  by (import pred_set REST_PSUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3873
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3874
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3875
  SING :: "('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3876
  "SING == %s::'a::type => bool. EX x::'a::type. s = INSERT x EMPTY"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3877
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3878
lemma SING_DEF: "ALL s::'a::type => bool. SING s = (EX x::'a::type. s = INSERT x EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3879
  by (import pred_set SING_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3880
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3881
lemma SING: "ALL x::'a::type. SING (INSERT x EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3882
  by (import pred_set SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3883
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3884
lemma IN_SING: "ALL (x::'a::type) xa::'a::type. IN x (INSERT xa EMPTY) = (x = xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3885
  by (import pred_set IN_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3886
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3887
lemma NOT_SING_EMPTY: "ALL x::'a::type. INSERT x EMPTY ~= EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3888
  by (import pred_set NOT_SING_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3889
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3890
lemma NOT_EMPTY_SING: "ALL x::'a::type. EMPTY ~= INSERT x EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3891
  by (import pred_set NOT_EMPTY_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3892
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3893
lemma EQUAL_SING: "ALL (x::'a::type) xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3894
   (INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3895
  by (import pred_set EQUAL_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3896
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3897
lemma DISJOINT_SING_EMPTY: "ALL x::'a::type. DISJOINT (INSERT x EMPTY) EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3898
  by (import pred_set DISJOINT_SING_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3899
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3900
lemma INSERT_SING_UNION: "ALL (x::'a::type => bool) xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3901
   INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3902
  by (import pred_set INSERT_SING_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3903
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3904
lemma SING_DELETE: "ALL x::'a::type. DELETE (INSERT x EMPTY) x = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3905
  by (import pred_set SING_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3906
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3907
lemma DELETE_EQ_SING: "ALL (x::'a::type => bool) xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3908
   IN xa x --> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3909
  by (import pred_set DELETE_EQ_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3910
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3911
lemma CHOICE_SING: "ALL x::'a::type. CHOICE (INSERT x EMPTY) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3912
  by (import pred_set CHOICE_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3913
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3914
lemma REST_SING: "ALL x::'a::type. REST (INSERT x EMPTY) = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3915
  by (import pred_set REST_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3916
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3917
lemma SING_IFF_EMPTY_REST: "ALL x::'a::type => bool. SING x = (x ~= EMPTY & REST x = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3918
  by (import pred_set SING_IFF_EMPTY_REST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3919
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3920
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3921
  IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3922
  "IMAGE ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3923
%(f::'a::type => 'b::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3924
   GSPEC (%x::'a::type. (f x, IN x s))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3925
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3926
lemma IMAGE_DEF: "ALL (f::'a::type => 'b::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3927
   IMAGE f s = GSPEC (%x::'a::type. (f x, IN x s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3928
  by (import pred_set IMAGE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3929
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3930
lemma IN_IMAGE: "ALL (x::'b::type) (xa::'a::type => bool) xb::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3931
   IN x (IMAGE xb xa) = (EX xc::'a::type. x = xb xc & IN xc xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3932
  by (import pred_set IN_IMAGE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3933
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3934
lemma IMAGE_IN: "ALL (x::'a::type) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3935
   IN x xa --> (ALL xb::'a::type => 'b::type. IN (xb x) (IMAGE xb xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3936
  by (import pred_set IMAGE_IN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3937
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3938
lemma IMAGE_EMPTY: "ALL x::'a::type => 'b::type. IMAGE x EMPTY = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3939
  by (import pred_set IMAGE_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3940
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3941
lemma IMAGE_ID: "ALL x::'a::type => bool. IMAGE (%x::'a::type. x) x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3942
  by (import pred_set IMAGE_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3943
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3944
lemma IMAGE_COMPOSE: "ALL (x::'b::type => 'c::type) (xa::'a::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3945
   xb::'a::type => bool. IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3946
  by (import pred_set IMAGE_COMPOSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3947
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3948
lemma IMAGE_INSERT: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3949
   IMAGE x (INSERT xa xb) = INSERT (x xa) (IMAGE x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3950
  by (import pred_set IMAGE_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3951
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3952
lemma IMAGE_EQ_EMPTY: "ALL (s::'a::type => bool) x::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3953
   (IMAGE x s = EMPTY) = (s = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3954
  by (import pred_set IMAGE_EQ_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3955
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3956
lemma IMAGE_DELETE: "ALL (f::'a::type => 'b::type) (x::'a::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3957
   ~ IN x s --> IMAGE f (DELETE s x) = IMAGE f s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3958
  by (import pred_set IMAGE_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3959
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3960
lemma IMAGE_UNION: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3961
   IMAGE x (pred_set.UNION xa xb) = pred_set.UNION (IMAGE x xa) (IMAGE x xb)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3962
  by (import pred_set IMAGE_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3963
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3964
lemma IMAGE_SUBSET: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3965
   SUBSET x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3966
   (ALL xb::'a::type => 'b::type. SUBSET (IMAGE xb x) (IMAGE xb xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3967
  by (import pred_set IMAGE_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3968
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3969
lemma IMAGE_INTER: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3970
   SUBSET (IMAGE f (pred_set.INTER s t))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3971
    (pred_set.INTER (IMAGE f s) (IMAGE f t))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3972
  by (import pred_set IMAGE_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3973
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3974
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  3975
  INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3976
  "INJ ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3977
%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3978
   (ALL x::'a::type. IN x s --> IN (f x) t) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3979
   (ALL (x::'a::type) y::'a::type. IN x s & IN y s --> f x = f y --> x = y)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3980
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3981
lemma INJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3982
   INJ f s t =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3983
   ((ALL x::'a::type. IN x s --> IN (f x) t) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3984
    (ALL (x::'a::type) y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3985
        IN x s & IN y s --> f x = f y --> x = y))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3986
  by (import pred_set INJ_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3987
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3988
lemma INJ_ID: "ALL x::'a::type => bool. INJ (%x::'a::type. x) x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3989
  by (import pred_set INJ_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3990
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3991
lemma INJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3992
   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3993
   INJ x xb xc & INJ xa xc xd --> INJ (xa o x) xb xd"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3994
  by (import pred_set INJ_COMPOSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3995
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3996
lemma INJ_EMPTY: "ALL x::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3997
   All (INJ x EMPTY) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  3998
   (ALL xa::'a::type => bool. INJ x xa EMPTY = (xa = EMPTY))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  3999
  by (import pred_set INJ_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4000
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4001
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4002
  SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4003
  "SURJ ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4004
%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4005
   (ALL x::'a::type. IN x s --> IN (f x) t) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4006
   (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4007
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4008
lemma SURJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4009
   SURJ f s t =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4010
   ((ALL x::'a::type. IN x s --> IN (f x) t) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4011
    (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4012
  by (import pred_set SURJ_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4013
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4014
lemma SURJ_ID: "ALL x::'a::type => bool. SURJ (%x::'a::type. x) x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4015
  by (import pred_set SURJ_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4016
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4017
lemma SURJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4018
   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4019
   SURJ x xb xc & SURJ xa xc xd --> SURJ (xa o x) xb xd"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4020
  by (import pred_set SURJ_COMPOSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4021
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4022
lemma SURJ_EMPTY: "ALL x::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4023
   (ALL xa::'b::type => bool. SURJ x EMPTY xa = (xa = EMPTY)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4024
   (ALL xa::'a::type => bool. SURJ x xa EMPTY = (xa = EMPTY))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4025
  by (import pred_set SURJ_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4026
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4027
lemma IMAGE_SURJ: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4028
   SURJ x xa xb = (IMAGE x xa = xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4029
  by (import pred_set IMAGE_SURJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4030
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4031
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4032
  BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4033
  "BIJ ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4034
%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4035
   INJ f s t & SURJ f s t"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4036
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4037
lemma BIJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4038
   BIJ f s t = (INJ f s t & SURJ f s t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4039
  by (import pred_set BIJ_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4040
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4041
lemma BIJ_ID: "ALL x::'a::type => bool. BIJ (%x::'a::type. x) x x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4042
  by (import pred_set BIJ_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4043
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4044
lemma BIJ_EMPTY: "ALL x::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4045
   (ALL xa::'b::type => bool. BIJ x EMPTY xa = (xa = EMPTY)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4046
   (ALL xa::'a::type => bool. BIJ x xa EMPTY = (xa = EMPTY))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4047
  by (import pred_set BIJ_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4048
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4049
lemma BIJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4050
   (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4051
   BIJ x xb xc & BIJ xa xc xd --> BIJ (xa o x) xb xd"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4052
  by (import pred_set BIJ_COMPOSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4053
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4054
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4055
  LINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4056
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4057
specification (LINV) LINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4058
   INJ f s t --> (ALL x::'a::type. IN x s --> LINV f s (f x) = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4059
  by (import pred_set LINV_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4060
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4061
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4062
  RINV :: "('a => 'b) => ('a => bool) => 'b => 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4063
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4064
specification (RINV) RINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4065
   SURJ f s t --> (ALL x::'b::type. IN x t --> f (RINV f s x) = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4066
  by (import pred_set RINV_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4067
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4068
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4069
  FINITE :: "('a => bool) => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4070
  "FINITE ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4071
%s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4072
   ALL P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4073
      P EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4074
      (ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4075
          P s --> (ALL e::'a::type. P (INSERT e s))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4076
      P s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4077
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4078
lemma FINITE_DEF: "ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4079
   FINITE s =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4080
   (ALL P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4081
       P EMPTY &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4082
       (ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4083
           P s --> (ALL e::'a::type. P (INSERT e s))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4084
       P s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4085
  by (import pred_set FINITE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4086
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4087
lemma FINITE_EMPTY: "FINITE EMPTY"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4088
  by (import pred_set FINITE_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4089
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4090
lemma FINITE_INDUCT: "ALL P::('a::type => bool) => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4091
   P EMPTY &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4092
   (ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4093
       FINITE s & P s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4094
       (ALL e::'a::type. ~ IN e s --> P (INSERT e s))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4095
   (ALL s::'a::type => bool. FINITE s --> P s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4096
  by (import pred_set FINITE_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4097
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4098
lemma FINITE_INSERT: "ALL (x::'a::type) s::'a::type => bool. FINITE (INSERT x s) = FINITE s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4099
  by (import pred_set FINITE_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4100
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4101
lemma FINITE_DELETE: "ALL (x::'a::type) s::'a::type => bool. FINITE (DELETE s x) = FINITE s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4102
  by (import pred_set FINITE_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4103
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4104
lemma FINITE_UNION: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4105
   FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4106
  by (import pred_set FINITE_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4107
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4108
lemma INTER_FINITE: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4109
   FINITE s --> (ALL t::'a::type => bool. FINITE (pred_set.INTER s t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4110
  by (import pred_set INTER_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4111
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4112
lemma SUBSET_FINITE: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4113
   FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> FINITE t)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4114
  by (import pred_set SUBSET_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4115
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4116
lemma PSUBSET_FINITE: "ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4117
   FINITE x --> (ALL xa::'a::type => bool. PSUBSET xa x --> FINITE xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4118
  by (import pred_set PSUBSET_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4119
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4120
lemma FINITE_DIFF: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4121
   FINITE s --> (ALL t::'a::type => bool. FINITE (DIFF s t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4122
  by (import pred_set FINITE_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4123
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4124
lemma FINITE_SING: "ALL x::'a::type. FINITE (INSERT x EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4125
  by (import pred_set FINITE_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4126
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4127
lemma SING_FINITE: "ALL x::'a::type => bool. SING x --> FINITE x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4128
  by (import pred_set SING_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4129
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4130
lemma IMAGE_FINITE: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4131
   FINITE s --> (ALL f::'a::type => 'b::type. FINITE (IMAGE f s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4132
  by (import pred_set IMAGE_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4133
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4134
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4135
  CARD :: "('a => bool) => nat" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4136
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4137
specification (CARD) CARD_DEF: "(op &::bool => bool => bool)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4138
 ((op =::nat => nat => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4139
   ((CARD::('a::type => bool) => nat) (EMPTY::'a::type => bool)) (0::nat))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4140
 ((All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4141
   (%s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4142
       (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4143
        ((FINITE::('a::type => bool) => bool) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4144
        ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4145
          (%x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4146
              (op =::nat => nat => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4147
               ((CARD::('a::type => bool) => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4148
                 ((INSERT::'a::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4149
                           => ('a::type => bool) => 'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4150
                   x s))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4151
               ((If::bool => nat => nat => nat)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4152
                 ((IN::'a::type => ('a::type => bool) => bool) x s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4153
                 ((CARD::('a::type => bool) => nat) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4154
                 ((Suc::nat => nat)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4155
                   ((CARD::('a::type => bool) => nat) s)))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4156
  by (import pred_set CARD_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4157
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4158
lemma CARD_EMPTY: "CARD EMPTY = 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4159
  by (import pred_set CARD_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4160
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4161
lemma CARD_INSERT: "ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4162
   FINITE s -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4163
   (ALL x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4164
       CARD (INSERT x s) = (if IN x s then CARD s else Suc (CARD s)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4165
  by (import pred_set CARD_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4166
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4167
lemma CARD_EQ_0: "ALL s::'a::type => bool. FINITE s --> (CARD s = 0) = (s = EMPTY)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4168
  by (import pred_set CARD_EQ_0)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4169
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4170
lemma CARD_DELETE: "ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4171
   FINITE s -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4172
   (ALL x::'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4173
       CARD (DELETE s x) = (if IN x s then CARD s - 1 else CARD s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4174
  by (import pred_set CARD_DELETE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4175
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4176
lemma CARD_INTER_LESS_EQ: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4177
   FINITE s -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4178
   (ALL t::'a::type => bool. CARD (pred_set.INTER s t) <= CARD s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4179
  by (import pred_set CARD_INTER_LESS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4180
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4181
lemma CARD_UNION: "ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4182
   FINITE s -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4183
   (ALL t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4184
       FINITE t -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4185
       CARD (pred_set.UNION s t) + CARD (pred_set.INTER s t) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4186
       CARD s + CARD t)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4187
  by (import pred_set CARD_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4188
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4189
lemma CARD_SUBSET: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4190
   FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> CARD t <= CARD s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4191
  by (import pred_set CARD_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4192
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4193
lemma CARD_PSUBSET: "ALL s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4194
   FINITE s --> (ALL t::'a::type => bool. PSUBSET t s --> CARD t < CARD s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4195
  by (import pred_set CARD_PSUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4196
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4197
lemma CARD_SING: "ALL x::'a::type. CARD (INSERT x EMPTY) = 1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4198
  by (import pred_set CARD_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4199
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4200
lemma SING_IFF_CARD1: "ALL x::'a::type => bool. SING x = (CARD x = 1 & FINITE x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4201
  by (import pred_set SING_IFF_CARD1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4202
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4203
lemma CARD_DIFF: "ALL t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4204
   FINITE t -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4205
   (ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4206
       FINITE s --> CARD (DIFF s t) = CARD s - CARD (pred_set.INTER s t))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4207
  by (import pred_set CARD_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4208
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4209
lemma LESS_CARD_DIFF: "ALL t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4210
   FINITE t -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4211
   (ALL s::'a::type => bool.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4212
       FINITE s --> CARD t < CARD s --> 0 < CARD (DIFF s t))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4213
  by (import pred_set LESS_CARD_DIFF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4214
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4215
lemma FINITE_COMPLETE_INDUCTION: "ALL P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4216
   (ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4217
       (ALL y::'a::type => bool. PSUBSET y x --> P y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4218
       FINITE x --> P x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4219
   (ALL x::'a::type => bool. FINITE x --> P x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4220
  by (import pred_set FINITE_COMPLETE_INDUCTION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4221
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4222
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4223
  INFINITE :: "('a => bool) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4224
  "INFINITE == %s::'a::type => bool. ~ FINITE s"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4225
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4226
lemma INFINITE_DEF: "ALL s::'a::type => bool. INFINITE s = (~ FINITE s)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4227
  by (import pred_set INFINITE_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4228
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4229
lemma NOT_IN_FINITE: "(op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4230
 ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4231
 ((All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4232
   (%s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4233
       (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4234
        ((FINITE::('a::type => bool) => bool) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4235
        ((Ex::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4236
          (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4237
              (Not::bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4238
               ((IN::'a::type => ('a::type => bool) => bool) x s)))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4239
  by (import pred_set NOT_IN_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4240
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4241
lemma INFINITE_INHAB: "ALL x::'a::type => bool. INFINITE x --> (EX xa::'a::type. IN xa x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4242
  by (import pred_set INFINITE_INHAB)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4243
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4244
lemma IMAGE_11_INFINITE: "ALL f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4245
   (ALL (x::'a::type) y::'a::type. f x = f y --> x = y) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4246
   (ALL s::'a::type => bool. INFINITE s --> INFINITE (IMAGE f s))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4247
  by (import pred_set IMAGE_11_INFINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4248
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4249
lemma INFINITE_SUBSET: "ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4250
   INFINITE x --> (ALL xa::'a::type => bool. SUBSET x xa --> INFINITE xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4251
  by (import pred_set INFINITE_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4252
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4253
lemma IN_INFINITE_NOT_FINITE: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4254
   INFINITE x & FINITE xa --> (EX xb::'a::type. IN xb x & ~ IN xb xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4255
  by (import pred_set IN_INFINITE_NOT_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4256
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4257
lemma INFINITE_UNIV: "(op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4258
 ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4259
 ((Ex::(('a::type => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4260
   (%f::'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4261
       (op &::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4262
        ((All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4263
          (%x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4264
              (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4265
               (%y::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4266
                   (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4267
                    ((op =::'a::type => 'a::type => bool) (f x) (f y))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4268
                    ((op =::'a::type => 'a::type => bool) x y))))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4269
        ((Ex::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4270
          (%y::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4271
              (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4272
               (%x::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4273
                   (Not::bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4274
                    ((op =::'a::type => 'a::type => bool) (f x) y))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4275
  by (import pred_set INFINITE_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4276
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4277
lemma FINITE_PSUBSET_INFINITE: "ALL x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4278
   INFINITE x =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4279
   (ALL xa::'a::type => bool. FINITE xa --> SUBSET xa x --> PSUBSET xa x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4280
  by (import pred_set FINITE_PSUBSET_INFINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4281
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4282
lemma FINITE_PSUBSET_UNIV: "(op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4283
 ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4284
 ((All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4285
   (%s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4286
       (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4287
        ((FINITE::('a::type => bool) => bool) s)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4288
        ((PSUBSET::('a::type => bool) => ('a::type => bool) => bool) s
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4289
          (pred_set.UNIV::'a::type => bool))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4290
  by (import pred_set FINITE_PSUBSET_UNIV)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4291
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4292
lemma INFINITE_DIFF_FINITE: "ALL (s::'a::type => bool) t::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4293
   INFINITE s & FINITE t --> DIFF s t ~= EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4294
  by (import pred_set INFINITE_DIFF_FINITE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4295
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4296
lemma FINITE_ISO_NUM: "ALL s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4297
   FINITE s -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4298
   (EX f::nat => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4299
       (ALL (n::nat) m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4300
           n < CARD s & m < CARD s --> f n = f m --> n = m) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4301
       s = GSPEC (%n::nat. (f n, n < CARD s)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4302
  by (import pred_set FINITE_ISO_NUM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4303
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4304
lemma FINITE_WEAK_ENUMERATE: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4305
 (%x::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4306
     (op =::bool => bool => bool) ((FINITE::('a::type => bool) => bool) x)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4307
      ((Ex::((nat => 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4308
        (%f::nat => 'a::type.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4309
            (Ex::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4310
             (%b::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4311
                 (All::('a::type => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4312
                  (%e::'a::type.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4313
                      (op =::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4314
                       ((IN::'a::type => ('a::type => bool) => bool) e x)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4315
                       ((Ex::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4316
                         (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4317
                             (op &::bool => bool => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  4318
                              ((op <::nat => nat => bool) n b)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4319
                              ((op =::'a::type => 'a::type => bool) e
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4320
                                (f n)))))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4321
  by (import pred_set FINITE_WEAK_ENUMERATE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4322
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4323
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4324
  BIGUNION :: "(('a => bool) => bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4325
  "BIGUNION ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4326
%P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4327
   GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4328
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4329
lemma BIGUNION: "ALL P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4330
   BIGUNION P =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4331
   GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4332
  by (import pred_set BIGUNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4333
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4334
lemma IN_BIGUNION: "ALL (x::'a::type) xa::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4335
   IN x (BIGUNION xa) = (EX s::'a::type => bool. IN x s & IN s xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4336
  by (import pred_set IN_BIGUNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4337
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4338
lemma BIGUNION_EMPTY: "BIGUNION EMPTY = EMPTY"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4339
  by (import pred_set BIGUNION_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4340
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4341
lemma BIGUNION_SING: "ALL x::'a::type => bool. BIGUNION (INSERT x EMPTY) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4342
  by (import pred_set BIGUNION_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4343
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4344
lemma BIGUNION_UNION: "ALL (x::('a::type => bool) => bool) xa::('a::type => bool) => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4345
   BIGUNION (pred_set.UNION x xa) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4346
   pred_set.UNION (BIGUNION x) (BIGUNION xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4347
  by (import pred_set BIGUNION_UNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4348
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4349
lemma DISJOINT_BIGUNION: "(ALL (s::('a::type => bool) => bool) t::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4350
    DISJOINT (BIGUNION s) t =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4351
    (ALL s'::'a::type => bool. IN s' s --> DISJOINT s' t)) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4352
(ALL (x::('a::type => bool) => bool) xa::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4353
    DISJOINT xa (BIGUNION x) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4354
    (ALL xb::'a::type => bool. IN xb x --> DISJOINT xa xb))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4355
  by (import pred_set DISJOINT_BIGUNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4356
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4357
lemma BIGUNION_INSERT: "ALL (x::'a::type => bool) xa::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4358
   BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4359
  by (import pred_set BIGUNION_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4360
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4361
lemma BIGUNION_SUBSET: "ALL (X::'a::type => bool) P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4362
   SUBSET (BIGUNION P) X = (ALL Y::'a::type => bool. IN Y P --> SUBSET Y X)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4363
  by (import pred_set BIGUNION_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4364
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4365
lemma FINITE_BIGUNION: "ALL x::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4366
   FINITE x & (ALL s::'a::type => bool. IN s x --> FINITE s) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4367
   FINITE (BIGUNION x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4368
  by (import pred_set FINITE_BIGUNION)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4369
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4370
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4371
  BIGINTER :: "(('a => bool) => bool) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4372
  "BIGINTER ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4373
%B::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4374
   GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4375
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4376
lemma BIGINTER: "ALL B::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4377
   BIGINTER B =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4378
   GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4379
  by (import pred_set BIGINTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4380
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4381
lemma IN_BIGINTER: "IN (x::'a::type) (BIGINTER (B::('a::type => bool) => bool)) =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4382
(ALL P::'a::type => bool. IN P B --> IN x P)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4383
  by (import pred_set IN_BIGINTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4384
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4385
lemma BIGINTER_INSERT: "ALL (P::'a::type => bool) B::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4386
   BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4387
  by (import pred_set BIGINTER_INSERT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4388
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4389
lemma BIGINTER_EMPTY: "BIGINTER EMPTY = pred_set.UNIV"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4390
  by (import pred_set BIGINTER_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4391
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4392
lemma BIGINTER_INTER: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4393
   BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4394
  by (import pred_set BIGINTER_INTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4395
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4396
lemma BIGINTER_SING: "ALL x::'a::type => bool. BIGINTER (INSERT x EMPTY) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4397
  by (import pred_set BIGINTER_SING)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4398
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4399
lemma SUBSET_BIGINTER: "ALL (X::'a::type => bool) P::('a::type => bool) => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4400
   SUBSET X (BIGINTER P) = (ALL x::'a::type => bool. IN x P --> SUBSET X x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4401
  by (import pred_set SUBSET_BIGINTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4402
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4403
lemma DISJOINT_BIGINTER: "ALL (x::'a::type => bool) (xa::'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4404
   xb::('a::type => bool) => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4405
   IN xa xb & DISJOINT xa x -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4406
   DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4407
  by (import pred_set DISJOINT_BIGINTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4408
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4409
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4410
  CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4411
  "CROSS ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4412
%(P::'a::type => bool) Q::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4413
   GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4414
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4415
lemma CROSS_DEF: "ALL (P::'a::type => bool) Q::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4416
   CROSS P Q =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4417
   GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4418
  by (import pred_set CROSS_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4419
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4420
lemma IN_CROSS: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type * 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4421
   IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4422
  by (import pred_set IN_CROSS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4423
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4424
lemma CROSS_EMPTY: "ALL x::'a::type => bool. CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4425
  by (import pred_set CROSS_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4426
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4427
lemma CROSS_INSERT_LEFT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4428
   CROSS (INSERT xb x) xa =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4429
   pred_set.UNION (CROSS (INSERT xb EMPTY) xa) (CROSS x xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4430
  by (import pred_set CROSS_INSERT_LEFT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4431
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4432
lemma CROSS_INSERT_RIGHT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'b::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4433
   CROSS x (INSERT xb xa) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4434
   pred_set.UNION (CROSS x (INSERT xb EMPTY)) (CROSS x xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4435
  by (import pred_set CROSS_INSERT_RIGHT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4436
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4437
lemma FINITE_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4438
   FINITE x & FINITE xa --> FINITE (CROSS x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4439
  by (import pred_set FINITE_CROSS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4440
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4441
lemma CROSS_SINGS: "ALL (x::'a::type) xa::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4442
   CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4443
  by (import pred_set CROSS_SINGS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4444
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4445
lemma CARD_SING_CROSS: "ALL (x::'a::type) s::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4446
   FINITE s --> CARD (CROSS (INSERT x EMPTY) s) = CARD s"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4447
  by (import pred_set CARD_SING_CROSS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4448
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4449
lemma CARD_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4450
   FINITE x & FINITE xa --> CARD (CROSS x xa) = CARD x * CARD xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4451
  by (import pred_set CARD_CROSS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4452
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4453
lemma CROSS_SUBSET: "ALL (x::'a::type => bool) (xa::'b::type => bool) (xb::'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4454
   xc::'b::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4455
   SUBSET (CROSS xb xc) (CROSS x xa) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4456
   (xb = EMPTY | xc = EMPTY | SUBSET xb x & SUBSET xc xa)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4457
  by (import pred_set CROSS_SUBSET)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4458
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4459
lemma FINITE_CROSS_EQ: "ALL (P::'a::type => bool) Q::'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4460
   FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4461
  by (import pred_set FINITE_CROSS_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4462
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4463
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4464
  COMPL :: "('a => bool) => 'a => bool" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4465
  "COMPL == DIFF pred_set.UNIV"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4466
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4467
lemma COMPL_DEF: "ALL P::'a::type => bool. COMPL P = DIFF pred_set.UNIV P"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4468
  by (import pred_set COMPL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4469
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4470
lemma IN_COMPL: "ALL (x::'a::type) xa::'a::type => bool. IN x (COMPL xa) = (~ IN x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4471
  by (import pred_set IN_COMPL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4472
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4473
lemma COMPL_COMPL: "ALL x::'a::type => bool. COMPL (COMPL x) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4474
  by (import pred_set COMPL_COMPL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4475
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4476
lemma COMPL_CLAUSES: "ALL x::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4477
   pred_set.INTER (COMPL x) x = EMPTY &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4478
   pred_set.UNION (COMPL x) x = pred_set.UNIV"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4479
  by (import pred_set COMPL_CLAUSES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4480
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4481
lemma COMPL_SPLITS: "ALL (x::'a::type => bool) xa::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4482
   pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER (COMPL x) xa) = xa"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4483
  by (import pred_set COMPL_SPLITS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4484
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4485
lemma INTER_UNION_COMPL: "ALL (x::'a::type => bool) xa::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4486
   pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4487
  by (import pred_set INTER_UNION_COMPL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4488
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4489
lemma COMPL_EMPTY: "COMPL EMPTY = pred_set.UNIV"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4490
  by (import pred_set COMPL_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4491
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4492
consts
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4493
  count :: "nat => nat => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4494
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4495
defs
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4496
  count_primdef: "count == %n::nat. GSPEC (%m::nat. (m, m < n))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4497
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4498
lemma count_def: "ALL n::nat. count n = GSPEC (%m::nat. (m, m < n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4499
  by (import pred_set count_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4500
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4501
lemma IN_COUNT: "ALL (m::nat) n::nat. IN m (count n) = (m < n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4502
  by (import pred_set IN_COUNT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4503
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4504
lemma COUNT_ZERO: "count 0 = EMPTY"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4505
  by (import pred_set COUNT_ZERO)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4506
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4507
lemma COUNT_SUC: "ALL n::nat. count (Suc n) = INSERT n (count n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4508
  by (import pred_set COUNT_SUC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4509
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4510
lemma FINITE_COUNT: "ALL n::nat. FINITE (count n)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4511
  by (import pred_set FINITE_COUNT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4512
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4513
lemma CARD_COUNT: "ALL n::nat. CARD (count n) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4514
  by (import pred_set CARD_COUNT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4515
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4516
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4517
  ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" 
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4518
  "ITSET_tupled ==
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4519
%f::'a::type => 'b::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4520
   WFREC
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4521
    (SOME R::('a::type => bool) * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4522
             => ('a::type => bool) * 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4523
        WF R &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4524
        (ALL (b::'b::type) s::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4525
            FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4526
    (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4527
        (v::'a::type => bool, v1::'b::type).
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4528
        if FINITE v
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4529
        then if v = EMPTY then v1
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4530
             else ITSET_tupled (REST v, f (CHOICE v) v1)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4531
        else ARB)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4532
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4533
lemma ITSET_tupled_primitive_def: "ALL f::'a::type => 'b::type => 'b::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4534
   ITSET_tupled f =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4535
   WFREC
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4536
    (SOME R::('a::type => bool) * 'b::type
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4537
             => ('a::type => bool) * 'b::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4538
        WF R &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4539
        (ALL (b::'b::type) s::'a::type => bool.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4540
            FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b)))
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4541
    (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4542
        (v::'a::type => bool, v1::'b::type).
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4543
        if FINITE v
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4544
        then if v = EMPTY then v1
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4545
             else ITSET_tupled (REST v, f (CHOICE v) v1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4546
        else ARB)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4547
  by (import pred_set ITSET_tupled_primitive_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4548
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4549
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4550
  ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4551
  "ITSET ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4552
%(f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) x1::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4553
   ITSET_tupled f (x, x1)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4554
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4555
lemma ITSET_curried_def: "ALL (f::'a::type => 'b::type => 'b::type) (x::'a::type => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4556
   x1::'b::type. ITSET f x x1 = ITSET_tupled f (x, x1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4557
  by (import pred_set ITSET_curried_def)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4558
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4559
lemma ITSET_IND: "ALL P::('a::type => bool) => 'b::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4560
   (ALL (s::'a::type => bool) b::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4561
       (FINITE s & s ~= EMPTY -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4562
        P (REST s) ((f::'a::type => 'b::type => 'b::type) (CHOICE s) b)) -->
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4563
       P s b) -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4564
   (ALL v::'a::type => bool. All (P v))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4565
  by (import pred_set ITSET_IND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4566
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4567
lemma ITSET_THM: "ALL (s::'a::type => bool) (f::'a::type => 'b::type => 'b::type) b::'b::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4568
   FINITE s -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4569
   ITSET f s b =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4570
   (if s = EMPTY then b else ITSET f (REST s) (f (CHOICE s) b))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4571
  by (import pred_set ITSET_THM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4572
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4573
lemma ITSET_EMPTY: "ALL (x::'a::type => 'b::type => 'b::type) xa::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4574
   ITSET x EMPTY xa = xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4575
  by (import pred_set ITSET_EMPTY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4576
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4577
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4578
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4579
;setup_theory operator
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4580
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4581
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4582
  ASSOC :: "('a => 'a => 'a) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4583
  "ASSOC ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4584
%f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4585
   ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4586
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4587
lemma ASSOC_DEF: "ALL f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4588
   ASSOC f =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4589
   (ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4590
  by (import operator ASSOC_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4591
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4592
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4593
  COMM :: "('a => 'a => 'b) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4594
  "COMM ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4595
%f::'a::type => 'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4596
   ALL (x::'a::type) y::'a::type. f x y = f y x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4597
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4598
lemma COMM_DEF: "ALL f::'a::type => 'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4599
   COMM f = (ALL (x::'a::type) y::'a::type. f x y = f y x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4600
  by (import operator COMM_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4601
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4602
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4603
  FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4604
  "FCOMM ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4605
%(f::'a::type => 'b::type => 'a::type) g::'c::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4606
   ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4607
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4608
lemma FCOMM_DEF: "ALL (f::'a::type => 'b::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4609
   g::'c::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4610
   FCOMM f g =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4611
   (ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4612
  by (import operator FCOMM_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4613
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4614
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4615
  RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4616
  "RIGHT_ID ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4617
%(f::'a::type => 'b::type => 'a::type) e::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4618
   ALL x::'a::type. f x e = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4619
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4620
lemma RIGHT_ID_DEF: "ALL (f::'a::type => 'b::type => 'a::type) e::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4621
   RIGHT_ID f e = (ALL x::'a::type. f x e = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4622
  by (import operator RIGHT_ID_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4623
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4624
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4625
  LEFT_ID :: "('a => 'b => 'b) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4626
  "LEFT_ID ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4627
%(f::'a::type => 'b::type => 'b::type) e::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4628
   ALL x::'b::type. f e x = x"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4629
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4630
lemma LEFT_ID_DEF: "ALL (f::'a::type => 'b::type => 'b::type) e::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4631
   LEFT_ID f e = (ALL x::'b::type. f e x = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4632
  by (import operator LEFT_ID_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4633
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4634
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4635
  MONOID :: "('a => 'a => 'a) => 'a => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4636
  "MONOID ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4637
%(f::'a::type => 'a::type => 'a::type) e::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4638
   ASSOC f & RIGHT_ID f e & LEFT_ID f e"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4639
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4640
lemma MONOID_DEF: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4641
   MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4642
  by (import operator MONOID_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4643
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4644
lemma ASSOC_CONJ: "ASSOC op &"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4645
  by (import operator ASSOC_CONJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4646
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4647
lemma ASSOC_DISJ: "ASSOC op |"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4648
  by (import operator ASSOC_DISJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4649
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4650
lemma FCOMM_ASSOC: "ALL x::'a::type => 'a::type => 'a::type. FCOMM x x = ASSOC x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4651
  by (import operator FCOMM_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4652
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4653
lemma MONOID_CONJ_T: "MONOID op & True"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4654
  by (import operator MONOID_CONJ_T)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4655
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4656
lemma MONOID_DISJ_F: "MONOID op | False"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4657
  by (import operator MONOID_DISJ_F)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4658
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4659
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4660
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4661
;setup_theory rich_list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4662
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4663
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4664
  SNOC :: "'a => 'a list => 'a list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4665
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4666
specification (SNOC) SNOC: "(ALL x::'a::type. SNOC x [] = [x]) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4667
(ALL (x::'a::type) (x'::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4668
    SNOC x (x' # l) = x' # SNOC x l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4669
  by (import rich_list SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4670
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4671
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4672
  SCANL :: "('b => 'a => 'b) => 'b => 'a list => 'b list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4673
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4674
specification (SCANL) SCANL: "(ALL (f::'b::type => 'a::type => 'b::type) e::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4675
    SCANL f e [] = [e]) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4676
(ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4677
    l::'a::type list. SCANL f e (x # l) = e # SCANL f (f e x) l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4678
  by (import rich_list SCANL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4679
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4680
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4681
  SCANR :: "('a => 'b => 'b) => 'b => 'a list => 'b list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4682
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4683
specification (SCANR) SCANR: "(ALL (f::'a::type => 'b::type => 'b::type) e::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4684
    SCANR f e [] = [e]) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4685
(ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4686
    l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4687
    SCANR f e (x # l) = f x (hd (SCANR f e l)) # SCANR f e l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4688
  by (import rich_list SCANR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4689
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4690
lemma IS_EL_DEF: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4691
  by (import rich_list IS_EL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4692
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4693
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4694
  AND_EL :: "bool list => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4695
  "AND_EL == list_all I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4696
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4697
lemma AND_EL_DEF: "AND_EL = list_all I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4698
  by (import rich_list AND_EL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4699
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4700
constdefs
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4701
  OR_EL :: "bool list => bool" 
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4702
  "OR_EL == list_exists I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4703
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4704
lemma OR_EL_DEF: "OR_EL = list_exists I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4705
  by (import rich_list OR_EL_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4706
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4707
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4708
  FIRSTN :: "nat => 'a list => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4709
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4710
specification (FIRSTN) FIRSTN: "(ALL l::'a::type list. FIRSTN 0 l = []) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4711
(ALL (n::nat) (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4712
    FIRSTN (Suc n) (x # l) = x # FIRSTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4713
  by (import rich_list FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4714
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4715
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4716
  BUTFIRSTN :: "nat => 'a list => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4717
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4718
specification (BUTFIRSTN) BUTFIRSTN: "(ALL l::'a::type list. BUTFIRSTN 0 l = l) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4719
(ALL (n::nat) (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4720
    BUTFIRSTN (Suc n) (x # l) = BUTFIRSTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4721
  by (import rich_list BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4722
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4723
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4724
  SEG :: "nat => nat => 'a list => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4725
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4726
specification (SEG) SEG: "(ALL (k::nat) l::'a::type list. SEG 0 k l = []) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4727
(ALL (m::nat) (x::'a::type) l::'a::type list.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4728
    SEG (Suc m) 0 (x # l) = x # SEG m 0 l) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4729
(ALL (m::nat) (k::nat) (x::'a::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4730
    SEG (Suc m) (Suc k) (x # l) = SEG (Suc m) k l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4731
  by (import rich_list SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4732
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4733
lemma LAST: "ALL (x::'a::type) l::'a::type list. last (SNOC x l) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4734
  by (import rich_list LAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4735
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4736
lemma BUTLAST: "ALL (x::'a::type) l::'a::type list. butlast (SNOC x l) = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4737
  by (import rich_list BUTLAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4738
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4739
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4740
  LASTN :: "nat => 'a list => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4741
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4742
specification (LASTN) LASTN: "(ALL l::'a::type list. LASTN 0 l = []) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4743
(ALL (n::nat) (x::'a::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4744
    LASTN (Suc n) (SNOC x l) = SNOC x (LASTN n l))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4745
  by (import rich_list LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4746
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4747
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4748
  BUTLASTN :: "nat => 'a list => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4749
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4750
specification (BUTLASTN) BUTLASTN: "(ALL l::'a::type list. BUTLASTN 0 l = l) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4751
(ALL (n::nat) (x::'a::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4752
    BUTLASTN (Suc n) (SNOC x l) = BUTLASTN n l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4753
  by (import rich_list BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4754
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4755
lemma EL: "(ALL x::'a::type list. EL 0 x = hd x) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4756
(ALL (x::nat) xa::'a::type list. EL (Suc x) xa = EL x (tl xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4757
  by (import rich_list EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4758
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4759
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4760
  ELL :: "nat => 'a list => 'a" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4761
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4762
specification (ELL) ELL: "(ALL l::'a::type list. ELL 0 l = last l) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4763
(ALL (n::nat) l::'a::type list. ELL (Suc n) l = ELL n (butlast l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4764
  by (import rich_list ELL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4765
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4766
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4767
  IS_PREFIX :: "'a list => 'a list => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4768
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4769
specification (IS_PREFIX) IS_PREFIX: "(ALL l::'a::type list. IS_PREFIX l [] = True) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4770
(ALL (x::'a::type) l::'a::type list. IS_PREFIX [] (x # l) = False) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4771
(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4772
    IS_PREFIX (x1 # l1) (x2 # l2) = (x1 = x2 & IS_PREFIX l1 l2))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4773
  by (import rich_list IS_PREFIX)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4774
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4775
lemma SNOC_APPEND: "ALL (x::'a::type) l::'a::type list. SNOC x l = l @ [x]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4776
  by (import rich_list SNOC_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4777
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4778
lemma REVERSE: "rev [] = [] &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4779
(ALL (x::'a::type) xa::'a::type list. rev (x # xa) = SNOC x (rev xa))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4780
  by (import rich_list REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4781
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4782
lemma REVERSE_SNOC: "ALL (x::'a::type) l::'a::type list. rev (SNOC x l) = x # rev l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4783
  by (import rich_list REVERSE_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4784
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4785
lemma SNOC_Axiom: "ALL (e::'b::type) f::'a::type => 'a::type list => 'b::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4786
   EX x::'a::type list => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4787
      x [] = e &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4788
      (ALL (xa::'a::type) l::'a::type list. x (SNOC xa l) = f xa l (x l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4789
  by (import rich_list SNOC_Axiom)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4790
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4791
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4792
  IS_SUFFIX :: "'a list => 'a list => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4793
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4794
specification (IS_SUFFIX) IS_SUFFIX: "(ALL l::'a::type list. IS_SUFFIX l [] = True) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4795
(ALL (x::'a::type) l::'a::type list. IS_SUFFIX [] (SNOC x l) = False) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4796
(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4797
    IS_SUFFIX (SNOC x1 l1) (SNOC x2 l2) = (x1 = x2 & IS_SUFFIX l1 l2))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4798
  by (import rich_list IS_SUFFIX)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4799
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4800
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4801
  IS_SUBLIST :: "'a list => 'a list => bool" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4802
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4803
specification (IS_SUBLIST) IS_SUBLIST: "(ALL l::'a::type list. IS_SUBLIST l [] = True) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4804
(ALL (x::'a::type) l::'a::type list. IS_SUBLIST [] (x # l) = False) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4805
(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4806
    IS_SUBLIST (x1 # l1) (x2 # l2) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4807
    (x1 = x2 & IS_PREFIX l1 l2 | IS_SUBLIST l1 (x2 # l2)))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4808
  by (import rich_list IS_SUBLIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4809
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4810
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4811
  SPLITP :: "('a => bool) => 'a list => 'a list * 'a list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4812
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4813
specification (SPLITP) SPLITP: "(ALL P::'a::type => bool. SPLITP P [] = ([], [])) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4814
(ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4815
    SPLITP P (x # l) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4816
    (if P x then ([], x # l) else (x # fst (SPLITP P l), snd (SPLITP P l))))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4817
  by (import rich_list SPLITP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4818
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4819
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4820
  PREFIX :: "('a => bool) => 'a list => 'a list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4821
  "PREFIX == %(P::'a::type => bool) l::'a::type list. fst (SPLITP (Not o P) l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4822
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4823
lemma PREFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4824
   PREFIX P l = fst (SPLITP (Not o P) l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4825
  by (import rich_list PREFIX_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4826
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4827
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4828
  SUFFIX :: "('a => bool) => 'a list => 'a list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4829
  "SUFFIX ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4830
%P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4831
   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else [])
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4832
    []"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4833
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4834
lemma SUFFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4835
   SUFFIX P l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4836
   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else [])
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4837
    [] l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4838
  by (import rich_list SUFFIX_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4839
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4840
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4841
  UNZIP_FST :: "('a * 'b) list => 'a list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4842
  "UNZIP_FST == %l::('a::type * 'b::type) list. fst (unzip l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4843
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4844
lemma UNZIP_FST_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_FST l = fst (unzip l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4845
  by (import rich_list UNZIP_FST_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4846
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4847
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4848
  UNZIP_SND :: "('a * 'b) list => 'b list" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4849
  "UNZIP_SND == %l::('a::type * 'b::type) list. snd (unzip l)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4850
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4851
lemma UNZIP_SND_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_SND l = snd (unzip l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4852
  by (import rich_list UNZIP_SND_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4853
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4854
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4855
  GENLIST :: "(nat => 'a) => nat => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4856
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4857
specification (GENLIST) GENLIST: "(ALL f::nat => 'a::type. GENLIST f 0 = []) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4858
(ALL (f::nat => 'a::type) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4859
    GENLIST f (Suc n) = SNOC (f n) (GENLIST f n))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4860
  by (import rich_list GENLIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4861
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4862
consts
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4863
  REPLICATE :: "nat => 'a => 'a list" 
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4864
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4865
specification (REPLICATE) REPLICATE: "(ALL x::'a::type. REPLICATE 0 x = []) &
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4866
(ALL (n::nat) x::'a::type. REPLICATE (Suc n) x = x # REPLICATE n x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4867
  by (import rich_list REPLICATE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4868
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4869
lemma LENGTH_MAP2: "ALL (l1::'a::type list) l2::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4870
   length l1 = length l2 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4871
   (ALL f::'a::type => 'b::type => 'c::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4872
       length (map2 f l1 l2) = length l1 &
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4873
       length (map2 f l1 l2) = length l2)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4874
  by (import rich_list LENGTH_MAP2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4875
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4876
lemma NULL_EQ_NIL: "ALL l::'a::type list. null l = (l = [])"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4877
  by (import rich_list NULL_EQ_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4878
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4879
lemma LENGTH_EQ: "ALL (x::'a::type list) y::'a::type list. x = y --> length x = length y"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4880
  by (import rich_list LENGTH_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4881
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4882
lemma LENGTH_NOT_NULL: "ALL l::'a::type list. (0 < length l) = (~ null l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4883
  by (import rich_list LENGTH_NOT_NULL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4884
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4885
lemma SNOC_INDUCT: "ALL P::'a::type list => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4886
   P [] &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4887
   (ALL l::'a::type list. P l --> (ALL x::'a::type. P (SNOC x l))) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4888
   All P"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4889
  by (import rich_list SNOC_INDUCT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4890
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4891
lemma SNOC_CASES: "ALL x'::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4892
   x' = [] | (EX (x::'a::type) l::'a::type list. x' = SNOC x l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4893
  by (import rich_list SNOC_CASES)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4894
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4895
lemma LENGTH_SNOC: "ALL (x::'a::type) l::'a::type list. length (SNOC x l) = Suc (length l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4896
  by (import rich_list LENGTH_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4897
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4898
lemma NOT_NIL_SNOC: "ALL (x::'a::type) xa::'a::type list. [] ~= SNOC x xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4899
  by (import rich_list NOT_NIL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4900
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4901
lemma NOT_SNOC_NIL: "ALL (x::'a::type) xa::'a::type list. SNOC x xa ~= []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4902
  by (import rich_list NOT_SNOC_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4903
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4904
lemma SNOC_11: "ALL (x::'a::type) (l::'a::type list) (x'::'a::type) l'::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4905
   (SNOC x l = SNOC x' l') = (x = x' & l = l')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4906
  by (import rich_list SNOC_11)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4907
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4908
lemma SNOC_EQ_LENGTH_EQ: "ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4909
   SNOC x1 l1 = SNOC x2 l2 --> length l1 = length l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4910
  by (import rich_list SNOC_EQ_LENGTH_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4911
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4912
lemma SNOC_REVERSE_CONS: "ALL (x::'a::type) xa::'a::type list. SNOC x xa = rev (x # rev xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4913
  by (import rich_list SNOC_REVERSE_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4914
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4915
lemma MAP_SNOC: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4916
   map x (SNOC xa xb) = SNOC (x xa) (map x xb)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4917
  by (import rich_list MAP_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4918
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4919
lemma FOLDR_SNOC: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4920
   l::'a::type list. foldr f (SNOC x l) e = foldr f l (f x e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4921
  by (import rich_list FOLDR_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4922
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4923
lemma FOLDL_SNOC: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4924
   l::'a::type list. foldl f e (SNOC x l) = f (foldl f e l) x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4925
  by (import rich_list FOLDL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4926
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4927
lemma FOLDR_FOLDL: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4928
   MONOID f e --> (ALL l::'a::type list. foldr f l e = foldl f e l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4929
  by (import rich_list FOLDR_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4930
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4931
lemma LENGTH_FOLDR: "ALL l::'a::type list. length l = foldr (%x::'a::type. Suc) l 0"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4932
  by (import rich_list LENGTH_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4933
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  4934
lemma LENGTH_FOLDL: "ALL l::'a::type list. length l = foldl (%(l'::nat) x::'a::type. Suc l') 0 l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4935
  by (import rich_list LENGTH_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4936
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4937
lemma MAP_FOLDR: "ALL (f::'a::type => 'b::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4938
   map f l = foldr (%x::'a::type. op # (f x)) l []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4939
  by (import rich_list MAP_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4940
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4941
lemma MAP_FOLDL: "ALL (f::'a::type => 'b::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4942
   map f l = foldl (%(l'::'b::type list) x::'a::type. SNOC (f x) l') [] l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4943
  by (import rich_list MAP_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4944
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4945
lemma MAP_o: "ALL (f::'b::type => 'c::type) g::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4946
   map (f o g) = map f o map g"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4947
  by (import rich_list MAP_o)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4948
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4949
lemma FILTER_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4950
   filter P l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4951
   foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else l') l []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4952
  by (import rich_list FILTER_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4953
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4954
lemma FILTER_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4955
   filter P (SNOC x l) = (if P x then SNOC x (filter P l) else filter P l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4956
  by (import rich_list FILTER_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4957
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4958
lemma FILTER_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4959
   filter P l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4960
   foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else l')
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4961
    [] l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4962
  by (import rich_list FILTER_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4963
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4964
lemma FILTER_COMM: "ALL (f1::'a::type => bool) (f2::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4965
   filter f1 (filter f2 l) = filter f2 (filter f1 l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4966
  by (import rich_list FILTER_COMM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4967
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4968
lemma FILTER_IDEM: "ALL (f::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4969
   filter f (filter f l) = filter f l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4970
  by (import rich_list FILTER_IDEM)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4971
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4972
lemma LENGTH_SEG: "ALL (n::nat) (k::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4973
   n + k <= length l --> length (SEG n k l) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4974
  by (import rich_list LENGTH_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4975
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4976
lemma APPEND_NIL: "(ALL l::'a::type list. l @ [] = l) & (ALL x::'a::type list. [] @ x = x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4977
  by (import rich_list APPEND_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4978
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4979
lemma APPEND_SNOC: "ALL (l1::'a::type list) (x::'a::type) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4980
   l1 @ SNOC x l2 = SNOC x (l1 @ l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4981
  by (import rich_list APPEND_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4982
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4983
lemma APPEND_FOLDR: "ALL (l1::'a::type list) l2::'a::type list. l1 @ l2 = foldr op # l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4984
  by (import rich_list APPEND_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4985
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4986
lemma APPEND_FOLDL: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4987
   l1 @ l2 = foldl (%(l'::'a::type list) x::'a::type. SNOC x l') l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4988
  by (import rich_list APPEND_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4989
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4990
lemma CONS_APPEND: "ALL (x::'a::type) l::'a::type list. x # l = [x] @ l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4991
  by (import rich_list CONS_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4992
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4993
lemma ASSOC_APPEND: "ASSOC op @"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4994
  by (import rich_list ASSOC_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4995
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4996
lemma MONOID_APPEND_NIL: "MONOID op @ []"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4997
  by (import rich_list MONOID_APPEND_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  4998
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  4999
lemma APPEND_LENGTH_EQ: "ALL (l1::'a::type list) l1'::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5000
   length l1 = length l1' -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5001
   (ALL (l2::'a::type list) l2'::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5002
       length l2 = length l2' -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5003
       (l1 @ l2 = l1' @ l2') = (l1 = l1' & l2 = l2'))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5004
  by (import rich_list APPEND_LENGTH_EQ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5005
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5006
lemma FLAT_SNOC: "ALL (x::'a::type list) l::'a::type list list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5007
   concat (SNOC x l) = concat l @ x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5008
  by (import rich_list FLAT_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5009
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5010
lemma FLAT_FOLDR: "ALL l::'a::type list list. concat l = foldr op @ l []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5011
  by (import rich_list FLAT_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5012
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5013
lemma FLAT_FOLDL: "ALL l::'a::type list list. concat l = foldl op @ [] l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5014
  by (import rich_list FLAT_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5015
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5016
lemma LENGTH_FLAT: "ALL l::'a::type list list. length (concat l) = sum (map size l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5017
  by (import rich_list LENGTH_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5018
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5019
lemma REVERSE_FOLDR: "ALL l::'a::type list. rev l = foldr SNOC l []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5020
  by (import rich_list REVERSE_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5021
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5022
lemma REVERSE_FOLDL: "ALL l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5023
   rev l = foldl (%(l'::'a::type list) x::'a::type. x # l') [] l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5024
  by (import rich_list REVERSE_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5025
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5026
lemma ALL_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5027
   list_all P (SNOC x l) = (list_all P l & P x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5028
  by (import rich_list ALL_EL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5029
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5030
lemma ALL_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5031
   list_all P (map f l) = list_all (P o f) l"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5032
  by (import rich_list ALL_EL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5033
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5034
lemma SOME_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5035
   list_exists P (SNOC x l) = (P x | list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5036
  by (import rich_list SOME_EL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5037
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5038
lemma IS_EL_SNOC: "ALL (y::'a::type) (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5039
   y mem SNOC x l = (y = x | y mem l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5040
  by (import rich_list IS_EL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5041
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5042
lemma SUM_SNOC: "ALL (x::nat) l::nat list. sum (SNOC x l) = sum l + x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5043
  by (import rich_list SUM_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5044
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5045
lemma SUM_FOLDL: "ALL l::nat list. sum l = foldl op + 0 l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5046
  by (import rich_list SUM_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5047
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5048
lemma IS_PREFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5049
   IS_PREFIX l1 l2 = (EX l::'a::type list. l1 = l2 @ l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5050
  by (import rich_list IS_PREFIX_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5051
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5052
lemma IS_SUFFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5053
   IS_SUFFIX l1 l2 = (EX l::'a::type list. l1 = l @ l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5054
  by (import rich_list IS_SUFFIX_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5055
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5056
lemma IS_SUBLIST_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5057
   IS_SUBLIST l1 l2 =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5058
   (EX (l::'a::type list) l'::'a::type list. l1 = l @ l2 @ l')"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5059
  by (import rich_list IS_SUBLIST_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5060
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5061
lemma IS_PREFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5062
   IS_PREFIX l1 l2 --> IS_SUBLIST l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5063
  by (import rich_list IS_PREFIX_IS_SUBLIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5064
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5065
lemma IS_SUFFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5066
   IS_SUFFIX l1 l2 --> IS_SUBLIST l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5067
  by (import rich_list IS_SUFFIX_IS_SUBLIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5068
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5069
lemma IS_PREFIX_REVERSE: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5070
   IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5071
  by (import rich_list IS_PREFIX_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5072
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5073
lemma IS_SUFFIX_REVERSE: "ALL (l2::'a::type list) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5074
   IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5075
  by (import rich_list IS_SUFFIX_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5076
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5077
lemma IS_SUBLIST_REVERSE: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5078
   IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5079
  by (import rich_list IS_SUBLIST_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5080
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5081
lemma PREFIX_FOLDR: "ALL (P::'a::type => bool) x::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5082
   PREFIX P x =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5083
   foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else []) x []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5084
  by (import rich_list PREFIX_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5085
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5086
lemma PREFIX: "(ALL x::'a::type => bool. PREFIX x [] = []) &
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5087
(ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5088
    PREFIX x (xa # xb) = (if x xa then xa # PREFIX x xb else []))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5089
  by (import rich_list PREFIX)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5090
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5091
lemma IS_PREFIX_PREFIX: "ALL (P::'a::type => bool) l::'a::type list. IS_PREFIX l (PREFIX P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5092
  by (import rich_list IS_PREFIX_PREFIX)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5093
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5094
lemma LENGTH_SCANL: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5095
   length (SCANL f e l) = Suc (length l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5096
  by (import rich_list LENGTH_SCANL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5097
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5098
lemma LENGTH_SCANR: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5099
   length (SCANR f e l) = Suc (length l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5100
  by (import rich_list LENGTH_SCANR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5101
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5102
lemma COMM_MONOID_FOLDL: "ALL x::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5103
   COMM x -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5104
   (ALL xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5105
       MONOID x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5106
       (ALL (e::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5107
           foldl x e l = x e (foldl x xa l)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5108
  by (import rich_list COMM_MONOID_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5109
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5110
lemma COMM_MONOID_FOLDR: "ALL x::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5111
   COMM x -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5112
   (ALL xa::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5113
       MONOID x xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5114
       (ALL (e::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5115
           foldr x l e = x e (foldr x l xa)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5116
  by (import rich_list COMM_MONOID_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5117
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5118
lemma FCOMM_FOLDR_APPEND: "ALL (x::'a::type => 'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5119
   xa::'b::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5120
   FCOMM x xa -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5121
   (ALL xb::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5122
       LEFT_ID x xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5123
       (ALL (l1::'b::type list) l2::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5124
           foldr xa (l1 @ l2) xb = x (foldr xa l1 xb) (foldr xa l2 xb)))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5125
  by (import rich_list FCOMM_FOLDR_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5126
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5127
lemma FCOMM_FOLDL_APPEND: "ALL (x::'a::type => 'b::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5128
   xa::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5129
   FCOMM x xa -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5130
   (ALL xb::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5131
       RIGHT_ID xa xb -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5132
       (ALL (l1::'b::type list) l2::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5133
           foldl x xb (l1 @ l2) = xa (foldl x xb l1) (foldl x xb l2)))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5134
  by (import rich_list FCOMM_FOLDL_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5135
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5136
lemma FOLDL_SINGLE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5137
   foldl x xa [xb] = x xa xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5138
  by (import rich_list FOLDL_SINGLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5139
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5140
lemma FOLDR_SINGLE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5141
   foldr x [xb] xa = x xb xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5142
  by (import rich_list FOLDR_SINGLE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5143
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5144
lemma FOLDR_CONS_NIL: "ALL l::'a::type list. foldr op # l [] = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5145
  by (import rich_list FOLDR_CONS_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5146
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5147
lemma FOLDL_SNOC_NIL: "ALL l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5148
   foldl (%(xs::'a::type list) x::'a::type. SNOC x xs) [] l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5149
  by (import rich_list FOLDL_SNOC_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5150
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5151
lemma FOLDR_REVERSE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5152
   foldr x (rev xb) xa = foldl (%(xa::'b::type) y::'a::type. x y xa) xa xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5153
  by (import rich_list FOLDR_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5154
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5155
lemma FOLDL_REVERSE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5156
   foldl x xa (rev xb) = foldr (%(xa::'b::type) y::'a::type. x y xa) xb xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5157
  by (import rich_list FOLDL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5158
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5159
lemma FOLDR_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5160
   (g::'b::type => 'a::type) l::'b::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5161
   foldr f (map g l) e = foldr (%x::'b::type. f (g x)) l e"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5162
  by (import rich_list FOLDR_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5163
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5164
lemma FOLDL_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5165
   (g::'b::type => 'a::type) l::'b::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5166
   foldl f e (map g l) = foldl (%(x::'a::type) y::'b::type. f x (g y)) e l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5167
  by (import rich_list FOLDL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5168
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5169
lemma ALL_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5170
   list_all P l = foldr (%x::'a::type. op & (P x)) l True"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5171
  by (import rich_list ALL_EL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5172
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5173
lemma ALL_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5174
   list_all P l = foldl (%(l'::bool) x::'a::type. l' & P x) True l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5175
  by (import rich_list ALL_EL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5176
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5177
lemma SOME_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5178
   list_exists P l = foldr (%x::'a::type. op | (P x)) l False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5179
  by (import rich_list SOME_EL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5180
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5181
lemma SOME_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5182
   list_exists P l = foldl (%(l'::bool) x::'a::type. l' | P x) False l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5183
  by (import rich_list SOME_EL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5184
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5185
lemma ALL_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5186
   list_all x xa = foldr op & (map x xa) True"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5187
  by (import rich_list ALL_EL_FOLDR_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5188
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5189
lemma ALL_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5190
   list_all x xa = foldl op & True (map x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5191
  by (import rich_list ALL_EL_FOLDL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5192
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5193
lemma SOME_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5194
   list_exists x xa = foldr op | (map x xa) False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5195
  by (import rich_list SOME_EL_FOLDR_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5196
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5197
lemma SOME_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5198
   list_exists x xa = foldl op | False (map x xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5199
  by (import rich_list SOME_EL_FOLDL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5200
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5201
lemma FOLDR_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5202
   (P::'a::type => bool) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5203
   foldr f (filter P l) e =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5204
   foldr (%(x::'a::type) y::'a::type. if P x then f x y else y) l e"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5205
  by (import rich_list FOLDR_FILTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5206
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5207
lemma FOLDL_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5208
   (P::'a::type => bool) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5209
   foldl f e (filter P l) =
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5210
   foldl (%(x::'a::type) y::'a::type. if P y then f x y else x) e l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5211
  by (import rich_list FOLDL_FILTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5212
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5213
lemma ASSOC_FOLDR_FLAT: "ALL f::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5214
   ASSOC f -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5215
   (ALL e::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5216
       LEFT_ID f e -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5217
       (ALL l::'a::type list list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5218
           foldr f (concat l) e = foldr f (map (FOLDR f e) l) e))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5219
  by (import rich_list ASSOC_FOLDR_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5220
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5221
lemma ASSOC_FOLDL_FLAT: "ALL f::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5222
   ASSOC f -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5223
   (ALL e::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5224
       RIGHT_ID f e -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5225
       (ALL l::'a::type list list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5226
           foldl f e (concat l) = foldl f e (map (foldl f e) l)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5227
  by (import rich_list ASSOC_FOLDL_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5228
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5229
lemma SOME_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5230
   list_exists P (map f l) = list_exists (P o f) l"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5231
  by (import rich_list SOME_EL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5232
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5233
lemma SOME_EL_DISJ: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5234
   list_exists (%x::'a::type. P x | Q x) l =
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5235
   (list_exists P l | list_exists Q l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5236
  by (import rich_list SOME_EL_DISJ)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5237
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5238
lemma IS_EL_FOLDR: "ALL (x::'a::type) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5239
   x mem xa = foldr (%xa::'a::type. op | (x = xa)) xa False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5240
  by (import rich_list IS_EL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5241
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5242
lemma IS_EL_FOLDL: "ALL (x::'a::type) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5243
   x mem xa = foldl (%(l'::bool) xa::'a::type. l' | x = xa) False xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5244
  by (import rich_list IS_EL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5245
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5246
lemma NULL_FOLDR: "ALL l::'a::type list. null l = foldr (%(x::'a::type) l'::bool. False) l True"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5247
  by (import rich_list NULL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5248
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5249
lemma NULL_FOLDL: "ALL l::'a::type list. null l = foldl (%(x::bool) l'::'a::type. False) True l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5250
  by (import rich_list NULL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5251
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5252
lemma SEG_LENGTH_ID: "ALL l::'a::type list. SEG (length l) 0 l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5253
  by (import rich_list SEG_LENGTH_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5254
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5255
lemma SEG_SUC_CONS: "ALL (m::nat) (n::nat) (l::'a::type list) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5256
   SEG m (Suc n) (x # l) = SEG m n l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5257
  by (import rich_list SEG_SUC_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5258
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5259
lemma SEG_0_SNOC: "ALL (m::nat) (l::'a::type list) x::'a::type.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5260
   m <= length l --> SEG m 0 (SNOC x l) = SEG m 0 l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5261
  by (import rich_list SEG_0_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5262
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5263
lemma BUTLASTN_SEG: "ALL (n::nat) l::'a::type list.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5264
   n <= length l --> BUTLASTN n l = SEG (length l - n) 0 l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5265
  by (import rich_list BUTLASTN_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5266
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5267
lemma LASTN_CONS: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5268
   n <= length l --> (ALL x::'a::type. LASTN n (x # l) = LASTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5269
  by (import rich_list LASTN_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5270
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5271
lemma LENGTH_LASTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (LASTN n l) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5272
  by (import rich_list LENGTH_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5273
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5274
lemma LASTN_LENGTH_ID: "ALL l::'a::type list. LASTN (length l) l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5275
  by (import rich_list LASTN_LENGTH_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5276
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5277
lemma LASTN_LASTN: "ALL (l::'a::type list) (n::nat) m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5278
   m <= length l --> n <= m --> LASTN n (LASTN m l) = LASTN n l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5279
  by (import rich_list LASTN_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5280
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5281
lemma FIRSTN_LENGTH_ID: "ALL l::'a::type list. FIRSTN (length l) l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5282
  by (import rich_list FIRSTN_LENGTH_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5283
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5284
lemma FIRSTN_SNOC: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5285
   n <= length l --> (ALL x::'a::type. FIRSTN n (SNOC x l) = FIRSTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5286
  by (import rich_list FIRSTN_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5287
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5288
lemma BUTLASTN_LENGTH_NIL: "ALL l::'a::type list. BUTLASTN (length l) l = []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5289
  by (import rich_list BUTLASTN_LENGTH_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5290
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5291
lemma BUTLASTN_SUC_BUTLAST: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5292
   n < length l --> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5293
  by (import rich_list BUTLASTN_SUC_BUTLAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5294
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5295
lemma BUTLASTN_BUTLAST: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5296
   n < length l --> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5297
  by (import rich_list BUTLASTN_BUTLAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5298
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5299
lemma LENGTH_BUTLASTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5300
   n <= length l --> length (BUTLASTN n l) = length l - n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5301
  by (import rich_list LENGTH_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5302
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5303
lemma BUTLASTN_BUTLASTN: "ALL (m::nat) (n::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5304
   n + m <= length l --> BUTLASTN n (BUTLASTN m l) = BUTLASTN (n + m) l"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5305
  by (import rich_list BUTLASTN_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5306
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5307
lemma APPEND_BUTLASTN_LASTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5308
   n <= length l --> BUTLASTN n l @ LASTN n l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5309
  by (import rich_list APPEND_BUTLASTN_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5310
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5311
lemma APPEND_FIRSTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5312
   m + n = length l --> FIRSTN n l @ LASTN m l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5313
  by (import rich_list APPEND_FIRSTN_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5314
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5315
lemma BUTLASTN_APPEND2: "ALL (n::nat) (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5316
   n <= length l2 --> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5317
  by (import rich_list BUTLASTN_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5318
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5319
lemma BUTLASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5320
   BUTLASTN (length l2) (l1 @ l2) = l1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5321
  by (import rich_list BUTLASTN_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5322
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5323
lemma LASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list. LASTN (length l2) (l1 @ l2) = l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5324
  by (import rich_list LASTN_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5325
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5326
lemma BUTLASTN_CONS: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5327
   n <= length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5328
   (ALL x::'a::type. BUTLASTN n (x # l) = x # BUTLASTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5329
  by (import rich_list BUTLASTN_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5330
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5331
lemma BUTLASTN_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. BUTLASTN (length l) (x # l) = [x]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5332
  by (import rich_list BUTLASTN_LENGTH_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5333
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5334
lemma LAST_LASTN_LAST: "ALL (n::nat) l::'a::type list.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5335
   n <= length l --> 0 < n --> last (LASTN n l) = last l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5336
  by (import rich_list LAST_LASTN_LAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5337
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5338
lemma BUTLASTN_LASTN_NIL: "ALL (n::nat) l::'a::type list. n <= length l --> BUTLASTN n (LASTN n l) = []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5339
  by (import rich_list BUTLASTN_LASTN_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5340
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5341
lemma LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5342
   n + m <= length l -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5343
   LASTN n (BUTLASTN m l) = BUTLASTN m (LASTN (n + m) l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5344
  by (import rich_list LASTN_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5345
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5346
lemma BUTLASTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5347
   m <= n & n <= length l -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5348
   BUTLASTN m (LASTN n l) = LASTN (n - m) (BUTLASTN m l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5349
  by (import rich_list BUTLASTN_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5350
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5351
lemma LASTN_1: "ALL l::'a::type list. l ~= [] --> LASTN 1 l = [last l]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5352
  by (import rich_list LASTN_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5353
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5354
lemma BUTLASTN_1: "ALL l::'a::type list. l ~= [] --> BUTLASTN 1 l = butlast l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5355
  by (import rich_list BUTLASTN_1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5356
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5357
lemma BUTLASTN_APPEND1: "ALL (l2::'a::type list) n::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5358
   length l2 <= n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5359
   (ALL l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5360
       BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5361
  by (import rich_list BUTLASTN_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5362
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5363
lemma LASTN_APPEND2: "ALL (n::nat) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5364
   n <= length l2 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5365
   (ALL l1::'a::type list. LASTN n (l1 @ l2) = LASTN n l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5366
  by (import rich_list LASTN_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5367
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5368
lemma LASTN_APPEND1: "ALL (l2::'a::type list) n::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5369
   length l2 <= n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5370
   (ALL l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5371
       LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5372
  by (import rich_list LASTN_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5373
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5374
lemma LASTN_MAP: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5375
   n <= length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5376
   (ALL f::'a::type => 'b::type. LASTN n (map f l) = map f (LASTN n l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5377
  by (import rich_list LASTN_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5378
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5379
lemma BUTLASTN_MAP: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5380
   n <= length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5381
   (ALL f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5382
       BUTLASTN n (map f l) = map f (BUTLASTN n l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5383
  by (import rich_list BUTLASTN_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5384
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5385
lemma ALL_EL_LASTN: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5386
 (%P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5387
     (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5388
      (%l::'a::type list.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5389
          (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5390
           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5391
           ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5392
             (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5393
                 (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5394
                  ((op <=::nat => nat => bool) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5395
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5396
                  ((list_all::('a::type => bool) => 'a::type list => bool) P
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5397
                    ((LASTN::nat => 'a::type list => 'a::type list) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5398
                      l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5399
  by (import rich_list ALL_EL_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5400
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5401
lemma ALL_EL_BUTLASTN: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5402
 (%P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5403
     (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5404
      (%l::'a::type list.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5405
          (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5406
           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5407
           ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5408
             (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5409
                 (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5410
                  ((op <=::nat => nat => bool) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5411
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5412
                  ((list_all::('a::type => bool) => 'a::type list => bool) P
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5413
                    ((BUTLASTN::nat => 'a::type list => 'a::type list) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5414
                      l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5415
  by (import rich_list ALL_EL_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5416
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5417
lemma LENGTH_FIRSTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (FIRSTN n l) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5418
  by (import rich_list LENGTH_FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5419
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5420
lemma FIRSTN_FIRSTN: "(All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5421
 (%m::nat.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5422
     (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5423
      (%l::'a::type list.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5424
          (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5425
           ((op <=::nat => nat => bool) m ((size::'a::type list => nat) l))
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5426
           ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5427
             (%n::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5428
                 (op -->::bool => bool => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5429
                  ((op <=::nat => nat => bool) n m)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5430
                  ((op =::'a::type list => 'a::type list => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5431
                    ((FIRSTN::nat => 'a::type list => 'a::type list) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5432
                      ((FIRSTN::nat => 'a::type list => 'a::type list) m l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5433
                    ((FIRSTN::nat => 'a::type list => 'a::type list) n
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5434
                      l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5435
  by (import rich_list FIRSTN_FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5436
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5437
lemma LENGTH_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5438
   n <= length l --> length (BUTFIRSTN n l) = length l - n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5439
  by (import rich_list LENGTH_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5440
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5441
lemma BUTFIRSTN_LENGTH_NIL: "ALL l::'a::type list. BUTFIRSTN (length l) l = []"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5442
  by (import rich_list BUTFIRSTN_LENGTH_NIL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5443
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5444
lemma BUTFIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5445
   n <= length l1 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5446
   (ALL l2::'a::type list. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5447
  by (import rich_list BUTFIRSTN_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5448
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5449
lemma BUTFIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5450
   length l1 <= n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5451
   (ALL l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5452
       BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5453
  by (import rich_list BUTFIRSTN_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5454
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5455
lemma BUTFIRSTN_BUTFIRSTN: "ALL (n::nat) (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5456
   n + m <= length l --> BUTFIRSTN n (BUTFIRSTN m l) = BUTFIRSTN (n + m) l"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5457
  by (import rich_list BUTFIRSTN_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5458
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5459
lemma APPEND_FIRSTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5460
   n <= length l --> FIRSTN n l @ BUTFIRSTN n l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5461
  by (import rich_list APPEND_FIRSTN_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5462
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5463
lemma LASTN_SEG: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5464
   n <= length l --> LASTN n l = SEG n (length l - n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5465
  by (import rich_list LASTN_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5466
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5467
lemma FIRSTN_SEG: "ALL (n::nat) l::'a::type list. n <= length l --> FIRSTN n l = SEG n 0 l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5468
  by (import rich_list FIRSTN_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5469
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5470
lemma BUTFIRSTN_SEG: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5471
   n <= length l --> BUTFIRSTN n l = SEG (length l - n) n l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5472
  by (import rich_list BUTFIRSTN_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5473
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5474
lemma BUTFIRSTN_SNOC: "ALL (n::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5475
   n <= length l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5476
   (ALL x::'a::type. BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5477
  by (import rich_list BUTFIRSTN_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5478
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5479
lemma APPEND_BUTLASTN_BUTFIRSTN: "ALL (m::nat) (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5480
   m + n = length l --> BUTLASTN m l @ BUTFIRSTN n l = l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5481
  by (import rich_list APPEND_BUTLASTN_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5482
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5483
lemma SEG_SEG: "ALL (n1::nat) (m1::nat) (n2::nat) (m2::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5484
   n1 + m1 <= length l & n2 + m2 <= n1 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5485
   SEG n2 m2 (SEG n1 m1 l) = SEG n2 (m1 + m2) l"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5486
  by (import rich_list SEG_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5487
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5488
lemma SEG_APPEND1: "ALL (n::nat) (m::nat) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5489
   n + m <= length l1 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5490
   (ALL l2::'a::type list. SEG n m (l1 @ l2) = SEG n m l1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5491
  by (import rich_list SEG_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5492
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5493
lemma SEG_APPEND2: "ALL (l1::'a::type list) (m::nat) (n::nat) l2::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5494
   length l1 <= m & n <= length l2 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5495
   SEG n m (l1 @ l2) = SEG n (m - length l1) l2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5496
  by (import rich_list SEG_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5497
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5498
lemma SEG_FIRSTN_BUTFISTN: "ALL (n::nat) (m::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5499
   n + m <= length l --> SEG n m l = FIRSTN n (BUTFIRSTN m l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5500
  by (import rich_list SEG_FIRSTN_BUTFISTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5501
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5502
lemma SEG_APPEND: "ALL (m::nat) (l1::'a::type list) (n::nat) l2::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5503
   m < length l1 & length l1 <= n + m & n + m <= length l1 + length l2 -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5504
   SEG n m (l1 @ l2) =
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5505
   SEG (length l1 - m) m l1 @ SEG (n + m - length l1) 0 l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5506
  by (import rich_list SEG_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5507
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5508
lemma SEG_LENGTH_SNOC: "ALL (x::'a::type list) xa::'a::type. SEG 1 (length x) (SNOC xa x) = [xa]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5509
  by (import rich_list SEG_LENGTH_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5510
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5511
lemma SEG_SNOC: "ALL (n::nat) (m::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5512
   n + m <= length l --> (ALL x::'a::type. SEG n m (SNOC x l) = SEG n m l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5513
  by (import rich_list SEG_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5514
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5515
lemma ELL_SEG: "ALL (n::nat) l::'a::type list.
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5516
   n < length l --> ELL n l = hd (SEG 1 (PRE (length l - n)) l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5517
  by (import rich_list ELL_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5518
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5519
lemma SNOC_FOLDR: "ALL (x::'a::type) l::'a::type list. SNOC x l = foldr op # l [x]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5520
  by (import rich_list SNOC_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5521
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5522
lemma IS_EL_FOLDR_MAP: "ALL (x::'a::type) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5523
   x mem xa = foldr op | (map (op = x) xa) False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5524
  by (import rich_list IS_EL_FOLDR_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5525
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5526
lemma IS_EL_FOLDL_MAP: "ALL (x::'a::type) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5527
   x mem xa = foldl op | False (map (op = x) xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5528
  by (import rich_list IS_EL_FOLDL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5529
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5530
lemma FILTER_FILTER: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list.
23290
c358025ad8db filter syntax change
nipkow
parents: 17727
diff changeset
  5531
   filter P (filter Q l) = [x::'a::type<-l. P x & Q x]"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5532
  by (import rich_list FILTER_FILTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5533
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5534
lemma FCOMM_FOLDR_FLAT: "ALL (g::'a::type => 'a::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5535
   f::'b::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5536
   FCOMM g f -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5537
   (ALL e::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5538
       LEFT_ID g e -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5539
       (ALL l::'b::type list list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5540
           foldr f (concat l) e = foldr g (map (FOLDR f e) l) e))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5541
  by (import rich_list FCOMM_FOLDR_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5542
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5543
lemma FCOMM_FOLDL_FLAT: "ALL (f::'a::type => 'b::type => 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5544
   g::'a::type => 'a::type => 'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5545
   FCOMM f g -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5546
   (ALL e::'a::type.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5547
       RIGHT_ID g e -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5548
       (ALL l::'b::type list list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5549
           foldl f e (concat l) = foldl g e (map (foldl f e) l)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5550
  by (import rich_list FCOMM_FOLDL_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5551
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5552
lemma FOLDR_MAP_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5553
   (ALL (a::'a::type) (b::'a::type) c::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5554
       f a (f b c) = f b (f a c)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5555
   (ALL (e::'a::type) (g::'b::type => 'a::type) l::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5556
       foldr f (map g (rev l)) e = foldr f (map g l) e)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5557
  by (import rich_list FOLDR_MAP_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5558
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5559
lemma FOLDR_FILTER_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5560
   (ALL (a::'a::type) (b::'a::type) c::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5561
       f a (f b c) = f b (f a c)) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5562
   (ALL (e::'a::type) (P::'a::type => bool) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5563
       foldr f (filter P (rev l)) e = foldr f (filter P l) e)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5564
  by (import rich_list FOLDR_FILTER_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5565
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5566
lemma COMM_ASSOC_FOLDR_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5567
   COMM f -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5568
   ASSOC f -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5569
   (ALL (e::'a::type) l::'a::type list. foldr f (rev l) e = foldr f l e)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5570
  by (import rich_list COMM_ASSOC_FOLDR_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5571
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5572
lemma COMM_ASSOC_FOLDL_REVERSE: "ALL f::'a::type => 'a::type => 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5573
   COMM f -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5574
   ASSOC f -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5575
   (ALL (e::'a::type) l::'a::type list. foldl f e (rev l) = foldl f e l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5576
  by (import rich_list COMM_ASSOC_FOLDL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5577
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5578
lemma ELL_LAST: "ALL l::'a::type list. ~ null l --> ELL 0 l = last l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5579
  by (import rich_list ELL_LAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5580
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5581
lemma ELL_0_SNOC: "ALL (l::'a::type list) x::'a::type. ELL 0 (SNOC x l) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5582
  by (import rich_list ELL_0_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5583
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5584
lemma ELL_SNOC: "ALL n>0.
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5585
   ALL (x::'a::type) l::'a::type list. ELL n (SNOC x l) = ELL (PRE n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5586
  by (import rich_list ELL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5587
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5588
lemma ELL_SUC_SNOC: "ALL (n::nat) (x::'a::type) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5589
   ELL (Suc n) (SNOC x xa) = ELL n xa"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5590
  by (import rich_list ELL_SUC_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5591
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5592
lemma ELL_CONS: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5593
   n < length l --> (ALL x::'a::type. ELL n (x # l) = ELL n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5594
  by (import rich_list ELL_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5595
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5596
lemma ELL_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. ELL (length l) (x # l) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5597
  by (import rich_list ELL_LENGTH_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5598
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5599
lemma ELL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5600
   ELL (length l) (SNOC x l) = (if null l then x else hd l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5601
  by (import rich_list ELL_LENGTH_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5602
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5603
lemma ELL_APPEND2: "ALL (n::nat) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5604
   n < length l2 --> (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL n l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5605
  by (import rich_list ELL_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5606
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5607
lemma ELL_APPEND1: "ALL (l2::'a::type list) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5608
   length l2 <= n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5609
   (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL (n - length l2) l1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5610
  by (import rich_list ELL_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5611
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5612
lemma ELL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> ELL (PRE (length l)) l = hd l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5613
  by (import rich_list ELL_PRE_LENGTH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5614
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5615
lemma EL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type. EL (length l) (SNOC x l) = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5616
  by (import rich_list EL_LENGTH_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5617
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5618
lemma EL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> EL (PRE (length l)) l = last l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5619
  by (import rich_list EL_PRE_LENGTH)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5620
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5621
lemma EL_SNOC: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5622
   n < length l --> (ALL x::'a::type. EL n (SNOC x l) = EL n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5623
  by (import rich_list EL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5624
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5625
lemma EL_ELL: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5626
   n < length l --> EL n l = ELL (PRE (length l - n)) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5627
  by (import rich_list EL_ELL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5628
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5629
lemma EL_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5630
   ~ null l2 --> EL (length l1) (l1 @ l2) = hd l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5631
  by (import rich_list EL_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5632
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5633
lemma ELL_EL: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5634
   n < length l --> ELL n l = EL (PRE (length l - n)) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5635
  by (import rich_list ELL_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5636
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5637
lemma ELL_MAP: "ALL (n::nat) (l::'a::type list) f::'a::type => 'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5638
   n < length l --> ELL n (map f l) = f (ELL n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5639
  by (import rich_list ELL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5640
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5641
lemma LENGTH_BUTLAST: "ALL l::'a::type list. l ~= [] --> length (butlast l) = PRE (length l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5642
  by (import rich_list LENGTH_BUTLAST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5643
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5644
lemma BUTFIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5645
   BUTFIRSTN (length l1) (l1 @ l2) = l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5646
  by (import rich_list BUTFIRSTN_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5647
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5648
lemma FIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5649
   n <= length l1 -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5650
   (ALL l2::'a::type list. FIRSTN n (l1 @ l2) = FIRSTN n l1)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5651
  by (import rich_list FIRSTN_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5652
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5653
lemma FIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5654
   length l1 <= n -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5655
   (ALL l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5656
       FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5657
  by (import rich_list FIRSTN_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5658
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5659
lemma FIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list. FIRSTN (length l1) (l1 @ l2) = l1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5660
  by (import rich_list FIRSTN_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5661
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5662
lemma REVERSE_FLAT: "ALL l::'a::type list list. rev (concat l) = concat (rev (map rev l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5663
  by (import rich_list REVERSE_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5664
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5665
lemma MAP_FILTER: "ALL (f::'a::type => 'a::type) (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5666
   (ALL x::'a::type. P (f x) = P x) -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5667
   map f (filter P l) = filter P (map f l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5668
  by (import rich_list MAP_FILTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5669
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5670
lemma FLAT_REVERSE: "ALL l::'a::type list list. concat (rev l) = rev (concat (map rev l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5671
  by (import rich_list FLAT_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5672
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5673
lemma FLAT_FLAT: "ALL l::'a::type list list list. concat (concat l) = concat (map concat l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5674
  by (import rich_list FLAT_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5675
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5676
lemma SOME_EL_REVERSE: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5677
   list_exists P (rev l) = list_exists P l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5678
  by (import rich_list SOME_EL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5679
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5680
lemma ALL_EL_SEG: "ALL (P::'a::type => bool) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5681
   list_all P l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5682
   (ALL (m::nat) k::nat. m + k <= length l --> list_all P (SEG m k l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5683
  by (import rich_list ALL_EL_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5684
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5685
lemma ALL_EL_FIRSTN: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5686
 (%P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5687
     (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5688
      (%l::'a::type list.
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5689
          (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5690
           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
14847
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5691
           ((All::(nat => bool) => bool)
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5692
             (%m::nat.
44d92c12b255 updated;
wenzelm
parents: 14684
diff changeset
  5693
                 (op -->::bool => bool => bool)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5694
                  ((op <=::nat => nat => bool) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5695
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5696
                  ((list_all::('a::type => bool) => 'a::type list => bool) P
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5697
                    ((FIRSTN::nat => 'a::type list => 'a::type list) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5698
                      l))))))"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5699
  by (import rich_list ALL_EL_FIRSTN)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5700
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5701
lemma ALL_EL_BUTFIRSTN: "(All::(('a::type => bool) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5702
 (%P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5703
     (All::('a::type list => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5704
      (%l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5705
          (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5706
           ((list_all::('a::type => bool) => 'a::type list => bool) P l)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5707
           ((All::(nat => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5708
             (%m::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5709
                 (op -->::bool => bool => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5710
                  ((op <=::nat => nat => bool) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5711
                    ((size::'a::type list => nat) l))
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5712
                  ((list_all::('a::type => bool) => 'a::type list => bool) P
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5713
                    ((BUTFIRSTN::nat => 'a::type list => 'a::type list) m
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5714
                      l))))))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5715
  by (import rich_list ALL_EL_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5716
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5717
lemma SOME_EL_SEG: "ALL (m::nat) (k::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5718
   m + k <= length l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5719
   (ALL P::'a::type => bool. list_exists P (SEG m k l) --> list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5720
  by (import rich_list SOME_EL_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5721
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5722
lemma SOME_EL_FIRSTN: "ALL (m::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5723
   m <= length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5724
   (ALL P::'a::type => bool. list_exists P (FIRSTN m l) --> list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5725
  by (import rich_list SOME_EL_FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5726
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5727
lemma SOME_EL_BUTFIRSTN: "ALL (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5728
   m <= length l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5729
   (ALL P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5730
       list_exists P (BUTFIRSTN m l) --> list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5731
  by (import rich_list SOME_EL_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5732
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5733
lemma SOME_EL_LASTN: "ALL (m::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5734
   m <= length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5735
   (ALL P::'a::type => bool. list_exists P (LASTN m l) --> list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5736
  by (import rich_list SOME_EL_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5737
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5738
lemma SOME_EL_BUTLASTN: "ALL (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5739
   m <= length l -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5740
   (ALL P::'a::type => bool.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5741
       list_exists P (BUTLASTN m l) --> list_exists P l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5742
  by (import rich_list SOME_EL_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5743
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5744
lemma IS_EL_REVERSE: "ALL (x::'a::type) l::'a::type list. x mem rev l = x mem l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5745
  by (import rich_list IS_EL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5746
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5747
lemma IS_EL_FILTER: "ALL (P::'a::type => bool) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5748
   P x --> (ALL l::'a::type list. x mem filter P l = x mem l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5749
  by (import rich_list IS_EL_FILTER)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5750
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5751
lemma IS_EL_SEG: "ALL (n::nat) (m::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5752
   n + m <= length l --> (ALL x::'a::type. x mem SEG n m l --> x mem l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5753
  by (import rich_list IS_EL_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5754
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5755
lemma IS_EL_SOME_EL: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5756
  by (import rich_list IS_EL_SOME_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5757
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5758
lemma IS_EL_FIRSTN: "ALL (x::nat) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5759
   x <= length xa --> (ALL xb::'a::type. xb mem FIRSTN x xa --> xb mem xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5760
  by (import rich_list IS_EL_FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5761
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5762
lemma IS_EL_BUTFIRSTN: "ALL (x::nat) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5763
   x <= length xa -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5764
   (ALL xb::'a::type. xb mem BUTFIRSTN x xa --> xb mem xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5765
  by (import rich_list IS_EL_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5766
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5767
lemma IS_EL_BUTLASTN: "ALL (x::nat) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5768
   x <= length xa --> (ALL xb::'a::type. xb mem BUTLASTN x xa --> xb mem xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5769
  by (import rich_list IS_EL_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5770
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5771
lemma IS_EL_LASTN: "ALL (x::nat) xa::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5772
   x <= length xa --> (ALL xb::'a::type. xb mem LASTN x xa --> xb mem xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5773
  by (import rich_list IS_EL_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5774
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5775
lemma ZIP_SNOC: "ALL (l1::'a::type list) l2::'b::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5776
   length l1 = length l2 -->
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5777
   (ALL (x1::'a::type) x2::'b::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5778
       zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5779
  by (import rich_list ZIP_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5780
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5781
lemma UNZIP_SNOC: "ALL (x::'a::type * 'b::type) l::('a::type * 'b::type) list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5782
   unzip (SNOC x l) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5783
   (SNOC (fst x) (fst (unzip l)), SNOC (snd x) (snd (unzip l)))"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5784
  by (import rich_list UNZIP_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5785
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5786
lemma LENGTH_UNZIP_FST: "ALL x::('a::type * 'b::type) list. length (UNZIP_FST x) = length x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5787
  by (import rich_list LENGTH_UNZIP_FST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5788
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5789
lemma LENGTH_UNZIP_SND: "ALL x::('a::type * 'b::type) list. length (UNZIP_SND x) = length x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5790
  by (import rich_list LENGTH_UNZIP_SND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5791
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5792
lemma SUM_APPEND: "ALL (l1::nat list) l2::nat list. sum (l1 @ l2) = sum l1 + sum l2"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5793
  by (import rich_list SUM_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5794
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5795
lemma SUM_REVERSE: "ALL l::nat list. sum (rev l) = sum l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5796
  by (import rich_list SUM_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5797
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5798
lemma SUM_FLAT: "ALL l::nat list list. sum (concat l) = sum (map sum l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5799
  by (import rich_list SUM_FLAT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5800
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5801
lemma EL_APPEND1: "ALL (n::nat) (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5802
   n < length l1 --> EL n (l1 @ l2) = EL n l1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5803
  by (import rich_list EL_APPEND1)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5804
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5805
lemma EL_APPEND2: "ALL (l1::'a::type list) n::nat.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5806
   length l1 <= n -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5807
   (ALL l2::'a::type list. EL n (l1 @ l2) = EL (n - length l1) l2)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5808
  by (import rich_list EL_APPEND2)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5809
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5810
lemma EL_MAP: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5811
   n < length l -->
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5812
   (ALL f::'a::type => 'b::type. EL n (map f l) = f (EL n l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5813
  by (import rich_list EL_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5814
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5815
lemma EL_CONS: "ALL n>0. ALL (x::'a::type) l::'a::type list. EL n (x # l) = EL (PRE n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5816
  by (import rich_list EL_CONS)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5817
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5818
lemma EL_SEG: "ALL (n::nat) l::'a::type list. n < length l --> EL n l = hd (SEG 1 n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5819
  by (import rich_list EL_SEG)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5820
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5821
lemma EL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> EL n l mem l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5822
  by (import rich_list EL_IS_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5823
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5824
lemma TL_SNOC: "ALL (x::'a::type) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5825
   tl (SNOC x l) = (if null l then [] else SNOC x (tl l))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5826
  by (import rich_list TL_SNOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5827
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5828
lemma EL_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5829
   n < length l --> EL n (rev l) = EL (PRE (length l - n)) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5830
  by (import rich_list EL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5831
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5832
lemma EL_REVERSE_ELL: "ALL (n::nat) l::'a::type list. n < length l --> EL n (rev l) = ELL n l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5833
  by (import rich_list EL_REVERSE_ELL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5834
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5835
lemma ELL_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5836
   ~ null l1 --> ELL (length l2) (l1 @ l2) = last l1"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5837
  by (import rich_list ELL_LENGTH_APPEND)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5838
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5839
lemma ELL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n l mem l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5840
  by (import rich_list ELL_IS_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5841
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5842
lemma ELL_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5843
   n < length l --> ELL n (rev l) = ELL (PRE (length l - n)) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5844
  by (import rich_list ELL_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5845
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5846
lemma ELL_REVERSE_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n (rev l) = EL n l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5847
  by (import rich_list ELL_REVERSE_EL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5848
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5849
lemma FIRSTN_BUTLASTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5850
   n <= length l --> FIRSTN n l = BUTLASTN (length l - n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5851
  by (import rich_list FIRSTN_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5852
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5853
lemma BUTLASTN_FIRSTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5854
   n <= length l --> BUTLASTN n l = FIRSTN (length l - n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5855
  by (import rich_list BUTLASTN_FIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5856
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5857
lemma LASTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5858
   n <= length l --> LASTN n l = BUTFIRSTN (length l - n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5859
  by (import rich_list LASTN_BUTFIRSTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5860
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5861
lemma BUTFIRSTN_LASTN: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5862
   n <= length l --> BUTFIRSTN n l = LASTN (length l - n) l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5863
  by (import rich_list BUTFIRSTN_LASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5864
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5865
lemma SEG_LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5866
   n + m <= length l -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5867
   SEG n m l = LASTN n (BUTLASTN (length l - (n + m)) l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5868
  by (import rich_list SEG_LASTN_BUTLASTN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5869
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5870
lemma BUTFIRSTN_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5871
   n <= length l --> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5872
  by (import rich_list BUTFIRSTN_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5873
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5874
lemma BUTLASTN_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5875
   n <= length l --> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5876
  by (import rich_list BUTLASTN_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5877
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5878
lemma LASTN_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5879
   n <= length l --> LASTN n (rev l) = rev (FIRSTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5880
  by (import rich_list LASTN_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5881
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5882
lemma FIRSTN_REVERSE: "ALL (n::nat) l::'a::type list.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5883
   n <= length l --> FIRSTN n (rev l) = rev (LASTN n l)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5884
  by (import rich_list FIRSTN_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5885
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5886
lemma SEG_REVERSE: "ALL (n::nat) (m::nat) l::'a::type list.
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5887
   n + m <= length l -->
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5888
   SEG n m (rev l) = rev (SEG n (length l - (n + m)) l)"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5889
  by (import rich_list SEG_REVERSE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5890
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5891
lemma LENGTH_GENLIST: "ALL (f::nat => 'a::type) n::nat. length (GENLIST f n) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5892
  by (import rich_list LENGTH_GENLIST)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5893
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5894
lemma LENGTH_REPLICATE: "ALL (n::nat) x::'a::type. length (REPLICATE n x) = n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5895
  by (import rich_list LENGTH_REPLICATE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5896
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5897
lemma IS_EL_REPLICATE: "ALL n>0. ALL x::'a::type. x mem REPLICATE n x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5898
  by (import rich_list IS_EL_REPLICATE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5899
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5900
lemma ALL_EL_REPLICATE: "ALL (x::'a::type) n::nat. list_all (op = x) (REPLICATE n x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5901
  by (import rich_list ALL_EL_REPLICATE)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5902
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5903
lemma AND_EL_FOLDL: "ALL l::bool list. AND_EL l = foldl op & True l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5904
  by (import rich_list AND_EL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5905
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5906
lemma AND_EL_FOLDR: "ALL l::bool list. AND_EL l = foldr op & l True"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5907
  by (import rich_list AND_EL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5908
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5909
lemma OR_EL_FOLDL: "ALL l::bool list. OR_EL l = foldl op | False l"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5910
  by (import rich_list OR_EL_FOLDL)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5911
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5912
lemma OR_EL_FOLDR: "ALL l::bool list. OR_EL l = foldr op | l False"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5913
  by (import rich_list OR_EL_FOLDR)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5914
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5915
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5916
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5917
;setup_theory state_transformer
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5918
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5919
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5920
  UNIT :: "'b => 'a => 'b * 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5921
  "(op ==::('b::type => 'a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5922
        => ('b::type => 'a::type => 'b::type * 'a::type) => prop)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5923
 (UNIT::'b::type => 'a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5924
 (Pair::'b::type => 'a::type => 'b::type * 'a::type)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5925
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5926
lemma UNIT_DEF: "ALL x::'b::type. UNIT x = Pair x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5927
  by (import state_transformer UNIT_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5928
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5929
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5930
  BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5931
  "(op ==::(('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5932
         => ('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5933
            => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5934
        => (('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5935
            => ('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5936
               => 'a::type => 'c::type * 'a::type)
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
  5937
           => prop)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5938
 (BIND::('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5939
        => ('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5940
           => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5941
 (%(g::'a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5942
     f::'b::type => 'a::type => 'c::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5943
     (op o::('b::type * 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5944
            => ('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5945
               => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5946
      ((split::('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5947
               => 'b::type * 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5948
        f)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5949
      g)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5950
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5951
lemma BIND_DEF: "(All::(('a::type => 'b::type * 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5952
 (%g::'a::type => 'b::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5953
     (All::(('b::type => 'a::type => 'c::type * 'a::type) => bool) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5954
      (%f::'b::type => 'a::type => 'c::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5955
          (op =::('a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5956
                 => ('a::type => 'c::type * 'a::type) => bool)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5957
           ((BIND::('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5958
                   => ('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5959
                      => 'a::type => 'c::type * 'a::type)
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 14847
diff changeset
  5960
             g f)
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5961
           ((op o::('b::type * 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5962
                   => ('a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5963
                      => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5964
             ((split::('b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5965
                      => 'b::type * 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5966
               f)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5967
             g)))"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5968
  by (import state_transformer BIND_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5969
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5970
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5971
  MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5972
  "MMAP ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5973
%(f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5974
   BIND m (UNIT o f)"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5975
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5976
lemma MMAP_DEF: "ALL (f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5977
   MMAP f m = BIND m (UNIT o f)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5978
  by (import state_transformer MMAP_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5979
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5980
constdefs
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
  5981
  JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" 
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5982
  "JOIN ==
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5983
%z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. BIND z I"
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5984
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5985
lemma JOIN_DEF: "ALL z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5986
   JOIN z = BIND z I"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5987
  by (import state_transformer JOIN_DEF)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5988
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5989
lemma BIND_LEFT_UNIT: "ALL (k::'a::type => 'b::type => 'c::type * 'b::type) x::'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5990
   BIND (UNIT x) k = k x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5991
  by (import state_transformer BIND_LEFT_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5992
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5993
lemma UNIT_UNCURRY: "ALL x::'a::type * 'b::type. split UNIT x = x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5994
  by (import state_transformer UNIT_UNCURRY)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5995
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5996
lemma BIND_RIGHT_UNIT: "ALL k::'a::type => 'b::type * 'a::type. BIND k UNIT = k"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5997
  by (import state_transformer BIND_RIGHT_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  5998
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  5999
lemma BIND_ASSOC: "ALL (x::'a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6000
   (xa::'b::type => 'a::type => 'c::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6001
   xb::'c::type => 'a::type => 'd::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6002
   BIND x (%a::'b::type. BIND (xa a) xb) = BIND (BIND x xa) xb"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6003
  by (import state_transformer BIND_ASSOC)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6004
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6005
lemma MMAP_ID: "MMAP I = I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6006
  by (import state_transformer MMAP_ID)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6007
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6008
lemma MMAP_COMP: "ALL (f::'c::type => 'd::type) g::'b::type => 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6009
   MMAP (f o g) = MMAP f o MMAP g"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6010
  by (import state_transformer MMAP_COMP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6011
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6012
lemma MMAP_UNIT: "ALL f::'b::type => 'c::type. MMAP f o UNIT = UNIT o f"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6013
  by (import state_transformer MMAP_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6014
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6015
lemma MMAP_JOIN: "ALL f::'b::type => 'c::type. MMAP f o JOIN = JOIN o MMAP (MMAP f)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6016
  by (import state_transformer MMAP_JOIN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6017
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6018
lemma JOIN_UNIT: "JOIN o UNIT = I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6019
  by (import state_transformer JOIN_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6020
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6021
lemma JOIN_MMAP_UNIT: "JOIN o MMAP UNIT = I"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6022
  by (import state_transformer JOIN_MMAP_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6023
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6024
lemma JOIN_MAP_JOIN: "JOIN o MMAP JOIN = JOIN o JOIN"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6025
  by (import state_transformer JOIN_MAP_JOIN)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6026
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6027
lemma JOIN_MAP: "ALL (x::'a::type => 'b::type * 'a::type)
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6028
   xa::'b::type => 'a::type => 'c::type * 'a::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6029
   BIND x xa = JOIN (MMAP xa x)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6030
  by (import state_transformer JOIN_MAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6031
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6032
lemma FST_o_UNIT: "ALL x::'a::type. fst o UNIT x = K x"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6033
  by (import state_transformer FST_o_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6034
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6035
lemma SND_o_UNIT: "ALL x::'a::type. snd o UNIT x = I"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6036
  by (import state_transformer SND_o_UNIT)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6037
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6038
lemma FST_o_MMAP: "ALL (x::'a::type => 'b::type) xa::'c::type => 'a::type * 'c::type.
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17566
diff changeset
  6039
   fst o MMAP x xa = x o (fst o xa)"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6040
  by (import state_transformer FST_o_MMAP)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6041
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6042
;end_setup
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6043
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6044
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  6045